Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 030107f

Browse files
awainversejcommelin
andcommitted
feat(order/compactly_generated): A compactly-generated modular lattice is complemented iff atomistic (#6071)
Shows that a compactly-generated modular lattice is complemented iff it is atomistic Proves extra lemmas about atomistic or compactly-generated lattices Proves extra lemmas about `complete_lattice.independent` Fix the name of `is_modular_lattice.sup_inf_sup_assoc` Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent 7fb7fb3 commit 030107f

File tree

5 files changed

+214
-8
lines changed

5 files changed

+214
-8
lines changed

docs/references.bib

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -593,3 +593,12 @@ @article {MR317916
593593
DOI = {10.2307/2318447},
594594
URL = {https://doi.org/10.2307/2318447},
595595
}
596+
597+
@book{calugareanu,
598+
author = {C\v{a}lug\v{a}reanu, Grigore},
599+
year = {2000},
600+
month = {01},
601+
pages = {},
602+
title = {Lattice Concepts of Module Theory},
603+
doi = {10.1007/978-94-015-9588-9}
604+
}

src/order/atoms.lean

Lines changed: 26 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -219,6 +219,32 @@ instance : is_atomic α :=
219219

220220
end is_atomistic
221221

222+
section is_atomistic
223+
variables [is_atomistic α]
224+
225+
@[simp]
226+
theorem Sup_atoms_le_eq (b : α) : Sup {a : α | is_atom a ∧ a ≤ b} = b :=
227+
begin
228+
rcases eq_Sup_atoms b with ⟨s, rfl, hs⟩,
229+
exact le_antisymm (Sup_le (λ _, and.right)) (Sup_le_Sup (λ a ha, ⟨hs a ha, le_Sup ha⟩)),
230+
end
231+
232+
@[simp]
233+
theorem Sup_atoms_eq_top : Sup {a : α | is_atom a} = ⊤ :=
234+
begin
235+
refine eq.trans (congr rfl (set.ext (λ x, _))) (Sup_atoms_le_eq ⊤),
236+
exact (and_iff_left le_top).symm,
237+
end
238+
239+
theorem le_iff_atom_le_imp {a b : α} :
240+
a ≤ b ↔ ∀ c : α, is_atom c → c ≤ a → c ≤ b :=
241+
⟨λ ab c hc ca, le_trans ca ab, λ h, begin
242+
rw [← Sup_atoms_le_eq a, ← Sup_atoms_le_eq b],
243+
exact Sup_le_Sup (λ c hc, ⟨hc.1, h c hc.1 hc.2⟩),
244+
end
245+
246+
end is_atomistic
247+
222248
namespace is_coatomistic
223249

224250
instance is_atomistic_dual [h : is_coatomistic α] : is_atomistic (order_dual α) :=

src/order/compactly_generated.lean

Lines changed: 143 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,8 @@ Copyright (c) 2021 Oliver Nash. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Oliver Nash
55
-/
6+
import data.set.finite
7+
import data.finset.order
68
import order.well_founded
79
import order.order_iso_nat
810
import order.atoms
@@ -38,14 +40,19 @@ This is demonstrated by means of the following four lemmas:
3840
We also show well-founded lattices are compactly generated
3941
(`complete_lattice.compactly_generated_of_well_founded`).
4042
43+
## References
44+
- [G. Călugăreanu, *Lattice Concepts of Module Theory*][calugareanu]
45+
4146
## Tags
4247
4348
complete lattice, well-founded, compact
4449
-/
4550

51+
variables {α : Type*} [complete_lattice α]
52+
4653
namespace complete_lattice
4754

48-
variables : Type*) [complete_lattice α]
55+
variables)
4956

5057
/-- A compactness property for a complete lattice is that any `sup`-closed non-empty subset
5158
contains its `Sup`. -/
@@ -243,7 +250,7 @@ class is_compactly_generated (α : Type*) [complete_lattice α] : Prop :=
243250
∀ (x : α), ∃ (s : set α), (∀ x ∈ s, complete_lattice.is_compact_element x) ∧ Sup s = x)
244251

245252
section
246-
variables : Type*} [complete_lattice α] [is_compactly_generated α] {a b : α} {s : set α}
253+
variables} [is_compactly_generated α] {a b : α} {s : set α}
247254

248255
@[simp]
249256
lemma Sup_compact_le_eq (b) : Sup {c : α | complete_lattice.is_compact_element c ∧ c ≤ b} = b :=
@@ -252,6 +259,14 @@ begin
252259
exact le_antisymm (Sup_le (λ c hc, hc.2)) (Sup_le_Sup (λ c cs, ⟨hs c cs, le_Sup cs⟩)),
253260
end
254261

262+
@[simp]
263+
theorem Sup_compact_eq_top :
264+
Sup {a : α | complete_lattice.is_compact_element a} = ⊤ :=
265+
begin
266+
refine eq.trans (congr rfl (set.ext (λ x, _))) (Sup_compact_le_eq ⊤),
267+
exact (and_iff_left le_top).symm,
268+
end
269+
255270
theorem le_iff_compact_le_imp {a b : α} :
256271
a ≤ b ↔ ∀ c : α, complete_lattice.is_compact_element c → c ≤ a → c ≤ b :=
257272
⟨λ ab c hc ca, le_trans ca ab, λ h, begin
@@ -301,10 +316,34 @@ theorem complete_lattice.independent_iff_finite {s : set α} :
301316
exact ⟨ha, set.subset.trans ht (set.diff_subset _ _)⟩ }
302317
end
303318

319+
lemma complete_lattice.independent_Union_of_directed {η : Type*}
320+
{s : η → set α} (hs : directed (⊆) s)
321+
(h : ∀ i, complete_lattice.independent (s i)) :
322+
complete_lattice.independent (⋃ i, s i) :=
323+
begin
324+
by_cases hη : nonempty η,
325+
{ resetI,
326+
rw complete_lattice.independent_iff_finite,
327+
intros t ht,
328+
obtain ⟨I, fi, hI⟩ := set.finite_subset_Union t.finite_to_set ht,
329+
obtain ⟨i, hi⟩ := hs.finset_le fi.to_finset,
330+
exact (h i).mono (set.subset.trans hI $ set.bUnion_subset $
331+
λ j hj, hi j (set.finite.mem_to_finset.2 hj)) },
332+
{ rintros a ⟨_, ⟨i, _⟩, _⟩,
333+
exfalso, exact hη ⟨i⟩, },
334+
end
335+
336+
lemma complete_lattice.independent_sUnion_of_directed {s : set (set α)}
337+
(hs : directed_on (⊆) s)
338+
(h : ∀ a ∈ s, complete_lattice.independent a) :
339+
complete_lattice.independent (⋃₀ s) :=
340+
by rw set.sUnion_eq_Union; exact
341+
complete_lattice.independent_Union_of_directed hs.directed_coe (by simpa using h)
342+
343+
304344
end
305345

306346
namespace complete_lattice
307-
variables {α : Type*} [complete_lattice α]
308347

309348
lemma compactly_generated_of_well_founded (h : well_founded ((>) : α → α → Prop)) :
310349
is_compactly_generated α :=
@@ -334,3 +373,104 @@ theorem Iic_coatomic_of_compact_element {k : α} (h : is_compact_element k) :
334373
end
335374

336375
end complete_lattice
376+
377+
section
378+
variables [is_modular_lattice α] [is_compactly_generated α]
379+
380+
@[priority 100]
381+
instance is_atomic_of_is_complemented [is_complemented α] : is_atomic α :=
382+
⟨λ b, begin
383+
by_cases h : {c : α | complete_lattice.is_compact_element c ∧ c ≤ b} ⊆ {⊥},
384+
{ left,
385+
rw [← Sup_compact_le_eq b, Sup_eq_bot],
386+
exact h },
387+
{ rcases set.not_subset.1 h with ⟨c, ⟨hc, hcb⟩, hcbot⟩,
388+
right,
389+
have hc' := complete_lattice.Iic_coatomic_of_compact_element hc,
390+
rw ← is_atomic_iff_is_coatomic at hc',
391+
haveI := hc',
392+
obtain con | ⟨a, ha, hac⟩ := eq_bot_or_exists_atom_le (⟨c, le_refl c⟩ : set.Iic c),
393+
{ exfalso,
394+
apply hcbot,
395+
simp only [subtype.ext_iff, set.Iic.coe_bot, subtype.coe_mk] at con,
396+
exact con },
397+
rw [← subtype.coe_le_coe, subtype.coe_mk] at hac,
398+
exact ⟨a, ha.of_is_atom_coe_Iic, hac.trans hcb⟩ },
399+
end
400+
401+
/-- See Lemma 5.1, Călugăreanu -/
402+
@[priority 100]
403+
instance is_atomistic_of_is_complemented [is_complemented α] : is_atomistic α :=
404+
⟨λ b, ⟨{a | is_atom a ∧ a ≤ b}, begin
405+
symmetry,
406+
have hle : Sup {a : α | is_atom a ∧ a ≤ b} ≤ b := (Sup_le $ λ _, and.right),
407+
apply (lt_or_eq_of_le hle).resolve_left (λ con, _),
408+
obtain ⟨c, hc⟩ := exists_is_compl (⟨Sup {a : α | is_atom a ∧ a ≤ b}, hle⟩ : set.Iic b),
409+
obtain rfl | ⟨a, ha, hac⟩ := eq_bot_or_exists_atom_le c,
410+
{ exact ne_of_lt con (subtype.ext_iff.1 (eq_top_of_is_compl_bot hc)) },
411+
{ apply ha.1,
412+
rw eq_bot_iff,
413+
apply le_trans (le_inf _ hac) hc.1,
414+
rw [← subtype.coe_le_coe, subtype.coe_mk],
415+
exact le_Sup ⟨ha.of_is_atom_coe_Iic, a.2⟩ }
416+
end, λ _, and.left⟩⟩
417+
418+
/-- See Theorem 6.6, Călugăreanu -/
419+
theorem is_complemented_of_is_atomistic [is_atomistic α] : is_complemented α :=
420+
⟨λ b, begin
421+
rcases zorn.zorn_subset
422+
{s : set α | complete_lattice.independent s ∧ b ⊓ Sup s = ⊥ ∧ ∀ a ∈ s, is_atom a} _ with
423+
⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩,
424+
{ refine ⟨Sup s, le_of_eq b_inf_Sup_s, le_iff_atom_le_imp.2 (λ a ha _, _)⟩,
425+
rw ← inf_eq_left,
426+
refine (eq_bot_or_eq_of_le_atom ha inf_le_left).resolve_left (λ con, ha.1 _),
427+
rw [eq_bot_iff, ← con],
428+
refine le_inf (le_refl a) ((le_Sup _).trans le_sup_right),
429+
rw ← disjoint_iff at *,
430+
have a_dis_Sup_s : disjoint a (Sup s) := con.mono_right le_sup_right,
431+
rw ← s_max (s ∪ {a}) ⟨λ x hx, _, ⟨_, λ x hx, _⟩⟩ (set.subset_union_left _ _),
432+
{ exact set.mem_union_right _ (set.mem_singleton _) },
433+
{ rw [set.mem_union, set.mem_singleton_iff] at hx,
434+
by_cases xa : x = a,
435+
{ simp only [xa, set.mem_singleton, set.insert_diff_of_mem, set.union_singleton],
436+
exact con.mono_right (le_trans (Sup_le_Sup (set.diff_subset s {a})) le_sup_right) },
437+
{ have h : (s ∪ {a}) \ {x} = (s \ {x}) ∪ {a},
438+
{ simp only [set.union_singleton],
439+
rw set.insert_diff_of_not_mem,
440+
rw set.mem_singleton_iff,
441+
exact ne.symm xa },
442+
rw [h, Sup_union, Sup_singleton],
443+
apply (s_ind x (hx.resolve_right xa)).disjoint_sup_right_of_disjoint_sup_left
444+
(a_dis_Sup_s.mono_right _).symm,
445+
rw [← Sup_insert, set.insert_diff_singleton,
446+
set.insert_eq_of_mem (hx.resolve_right xa)] } },
447+
{ rw [Sup_union, Sup_singleton, ← disjoint_iff],
448+
exact b_inf_Sup_s.disjoint_sup_right_of_disjoint_sup_left con.symm },
449+
{ rw [set.mem_union, set.mem_singleton_iff] at hx,
450+
cases hx,
451+
{ exact s_atoms x hx },
452+
{ rw hx,
453+
exact ha } } },
454+
{ intros c hc1 hc2,
455+
refine ⟨⋃₀ c, ⟨complete_lattice.independent_sUnion_of_directed hc2.directed_on
456+
(λ s hs, (hc1 hs).1), _, λ a ha, _⟩, λ _, set.subset_sUnion_of_mem⟩,
457+
{ rw [Sup_sUnion, ← Sup_image, inf_Sup_eq_of_directed_on, supr_eq_bot],
458+
{ intro i,
459+
rw supr_eq_bot,
460+
intro hi,
461+
obtain ⟨x, xc, rfl⟩ := (set.mem_image _ _ _).1 hi,
462+
exact (hc1 xc).2.1 },
463+
{ rw directed_on_image,
464+
refine hc2.directed_on.mono (λ s t, Sup_le_Sup) } },
465+
{ rcases set.mem_sUnion.1 ha with ⟨s, sc, as⟩,
466+
exact (hc1 sc).2.2 a as } }
467+
end
468+
469+
theorem is_complemented_iff_is_atomistic : is_complemented α ↔ is_atomistic α :=
470+
begin
471+
split; introsI,
472+
{ exact is_atomistic_of_is_complemented },
473+
{ exact is_complemented_of_is_atomistic }
474+
end
475+
476+
end

src/order/complete_lattice.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -380,6 +380,18 @@ le_antisymm
380380
(Sup_le $ assume b h, le_supr_of_le b $ le_supr _ h)
381381
(supr_le $ assume b, supr_le $ assume h, le_Sup h)
382382

383+
lemma Sup_sUnion {s : set (set α)} :
384+
Sup (⋃₀ s) = ⨆ (t ∈ s), Sup t :=
385+
begin
386+
apply le_antisymm,
387+
{ apply Sup_le (λ b hb, _),
388+
rcases hb with ⟨t, ts, bt⟩,
389+
apply le_trans _ (le_supr _ t),
390+
exact le_trans (le_Sup bt) (le_supr _ ts), },
391+
{ apply supr_le (λ t, _),
392+
exact supr_le (λ ts, Sup_le_Sup (λ x xt, ⟨t, ts, xt⟩)) }
393+
end
394+
383395
lemma le_supr_iff : (a ≤ supr s) ↔ (∀ b, (∀ i, s i ≤ b) → a ≤ b) :=
384396
⟨λ h b hb, le_trans h (supr_le hb), λ h, h _ $ λ i, le_supr s i⟩
385397

@@ -1027,6 +1039,10 @@ variables [complete_lattice α]
10271039
from the `Sup` of the rest. -/
10281040
def complete_lattice.independent (s : set α) : Prop := ∀ a ∈ s, disjoint a (Sup (s \ {a}))
10291041

1042+
@[simp]
1043+
lemma complete_lattice.independent_empty : complete_lattice.independent (∅ : set α) :=
1044+
λ x hx, (set.not_mem_empty x hx).elim
1045+
10301046
theorem complete_lattice.independent.mono {s t : set α}
10311047
(ht : complete_lattice.independent t) (hst : s ⊆ t) :
10321048
complete_lattice.independent s :=

src/order/modular_lattice.lean

Lines changed: 20 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -21,8 +21,8 @@ any distributive lattice.
2121
This corresponds to the diamond (or second) isomorphism theorems of algebra.
2222
2323
## Main Results
24-
- `is_modular_lattice_iff_sup_inf_sup_assoc`:
25-
Modularity is equivalent to the `sup_inf_sup_assoc`: `(x ⊓ z) ⊔ (y ⊓ z) = ((x ⊓ z) ⊔ y) ⊓ z`
24+
- `is_modular_lattice_iff_inf_sup_inf_assoc`:
25+
Modularity is equivalent to the `inf_sup_inf_assoc`: `(x ⊓ z) ⊔ (y ⊓ z) = ((x ⊓ z) ⊔ y) ⊓ z`
2626
- `distrib_lattice.is_modular_lattice`: Distributive lattices are modular.
2727
2828
## To do
@@ -44,7 +44,7 @@ theorem sup_inf_assoc_of_le {x : α} (y : α) {z : α} (h : x ≤ z) :
4444
le_antisymm (is_modular_lattice.sup_inf_le_assoc_of_le y h)
4545
(le_inf (sup_le_sup_left inf_le_left _) (sup_le h inf_le_right))
4646

47-
theorem is_modular_lattice.sup_inf_sup_assoc {x y z : α} :
47+
theorem is_modular_lattice.inf_sup_inf_assoc {x y z : α} :
4848
(x ⊓ z) ⊔ (y ⊓ z) = ((x ⊓ z) ⊔ y) ⊓ z :=
4949
(sup_inf_assoc_of_le y inf_le_right).symm
5050

@@ -56,6 +56,10 @@ instance : is_modular_lattice (order_dual α) :=
5656
⟨λ x y z xz, le_of_eq (by { rw [inf_comm, sup_comm, eq_comm, inf_comm, sup_comm],
5757
convert sup_inf_assoc_of_le (order_dual.of_dual y) (order_dual.dual_le.2 xz) })⟩
5858

59+
theorem is_modular_lattice.sup_inf_sup_assoc {x y z : α} :
60+
(x ⊔ z) ⊓ (y ⊔ z) = ((x ⊔ z) ⊓ y) ⊔ z :=
61+
@is_modular_lattice.inf_sup_inf_assoc (order_dual α) _ _ _ _ _
62+
5963
/-- The diamond isomorphism between the intervals `[a ⊓ b, a]` and `[b, a ⊔ b]` -/
6064
def inf_Icc_order_iso_Icc_sup (a b : α) : set.Icc (a ⊓ b) a ≃o set.Icc b (a ⊔ b) :=
6165
{ to_fun := λ x, ⟨x ⊔ b, ⟨le_sup_right, sup_le_sup_right x.prop.2 b⟩⟩,
@@ -85,9 +89,9 @@ def Iic_order_iso_Ici {a b : α} (h : is_compl a b) : set.Iic a ≃o set.Ici b :
8589

8690
end is_compl
8791

88-
theorem is_modular_lattice_iff_sup_inf_sup_assoc [lattice α] :
92+
theorem is_modular_lattice_iff_inf_sup_inf_assoc [lattice α] :
8993
is_modular_lattice α ↔ ∀ (x y z : α), (x ⊓ z) ⊔ (y ⊓ z) = ((x ⊓ z) ⊔ y) ⊓ z :=
90-
⟨λ h, @is_modular_lattice.sup_inf_sup_assoc _ _ h, λ h, ⟨λ x y z xz, by rw [← inf_eq_left.2 xz, h]⟩⟩
94+
⟨λ h, @is_modular_lattice.inf_sup_inf_assoc _ _ h, λ h, ⟨λ x y z xz, by rw [← inf_eq_left.2 xz, h]⟩⟩
9195

9296
namespace distrib_lattice
9397

@@ -97,6 +101,17 @@ instance [distrib_lattice α] : is_modular_lattice α :=
97101

98102
end distrib_lattice
99103

104+
theorem disjoint.disjoint_sup_right_of_disjoint_sup_left
105+
[bounded_lattice α] [is_modular_lattice α] {a b c : α}
106+
(h : disjoint a b) (hsup : disjoint (a ⊔ b) c) :
107+
disjoint a (b ⊔ c) :=
108+
begin
109+
rw [disjoint, ← h.eq_bot, sup_comm],
110+
apply le_inf inf_le_left,
111+
apply (inf_le_inf_right (c ⊔ b) le_sup_right).trans,
112+
rw [sup_comm, is_modular_lattice.sup_inf_sup_assoc, hsup.eq_bot, bot_sup_eq]
113+
end
114+
100115
namespace is_modular_lattice
101116

102117
variables [bounded_lattice α] [is_modular_lattice α] {a : α}

0 commit comments

Comments
 (0)