@@ -3,6 +3,8 @@ Copyright (c) 2021 Oliver Nash. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Oliver Nash
5
5
-/
6
+ import data.set.finite
7
+ import data.finset.order
6
8
import order.well_founded
7
9
import order.order_iso_nat
8
10
import order.atoms
@@ -38,14 +40,19 @@ This is demonstrated by means of the following four lemmas:
38
40
We also show well-founded lattices are compactly generated
39
41
(`complete_lattice.compactly_generated_of_well_founded`).
40
42
43
+ ## References
44
+ - [ G. Călugăreanu, *Lattice Concepts of Module Theory* ] [calugareanu ]
45
+
41
46
## Tags
42
47
43
48
complete lattice, well-founded, compact
44
49
-/
45
50
51
+ variables {α : Type *} [complete_lattice α]
52
+
46
53
namespace complete_lattice
47
54
48
- variables (α : Type *) [complete_lattice α]
55
+ variables (α)
49
56
50
57
/-- A compactness property for a complete lattice is that any `sup`-closed non-empty subset
51
58
contains its `Sup`. -/
@@ -243,7 +250,7 @@ class is_compactly_generated (α : Type*) [complete_lattice α] : Prop :=
243
250
∀ (x : α), ∃ (s : set α), (∀ x ∈ s, complete_lattice.is_compact_element x) ∧ Sup s = x)
244
251
245
252
section
246
- variables {α : Type *} [complete_lattice α] [is_compactly_generated α] {a b : α} {s : set α}
253
+ variables {α} [is_compactly_generated α] {a b : α} {s : set α}
247
254
248
255
@[simp]
249
256
lemma Sup_compact_le_eq (b) : Sup {c : α | complete_lattice.is_compact_element c ∧ c ≤ b} = b :=
@@ -252,6 +259,14 @@ begin
252
259
exact le_antisymm (Sup_le (λ c hc, hc.2 )) (Sup_le_Sup (λ c cs, ⟨hs c cs, le_Sup cs⟩)),
253
260
end
254
261
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
+
255
270
theorem le_iff_compact_le_imp {a b : α} :
256
271
a ≤ b ↔ ∀ c : α, complete_lattice.is_compact_element c → c ≤ a → c ≤ b :=
257
272
⟨λ ab c hc ca, le_trans ca ab, λ h, begin
@@ -301,10 +316,34 @@ theorem complete_lattice.independent_iff_finite {s : set α} :
301
316
exact ⟨ha, set.subset.trans ht (set.diff_subset _ _)⟩ }
302
317
end ⟩
303
318
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
+
304
344
end
305
345
306
346
namespace complete_lattice
307
- variables {α : Type *} [complete_lattice α]
308
347
309
348
lemma compactly_generated_of_well_founded (h : well_founded ((>) : α → α → Prop )) :
310
349
is_compactly_generated α :=
@@ -334,3 +373,104 @@ theorem Iic_coatomic_of_compact_element {k : α} (h : is_compact_element k) :
334
373
end ⟩
335
374
336
375
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
0 commit comments