2
2
Copyright (c) 2017 Johannes Hölzl. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Johannes Hölzl
5
-
6
- Theory of complete Boolean algebras.
7
5
-/
8
6
import order.complete_lattice
9
7
8
+ /-!
9
+ # Completely distributive lattices and Boolean algebras
10
+
11
+ In this file there are definitions and an API for completely distributive lattices and completely
12
+ distributive Boolean algebras.
13
+
14
+ ## Typeclasses
15
+
16
+ * `complete_distrib_lattice`: Completely distributive lattices: A complete lattice whose `⊓` and `⊔`
17
+ distribute over `⨆` and `⨅` respectively.
18
+ * `complete_boolean_algebra`: Completely distributive Boolean algebra: A Boolean algebra whose `⊓`
19
+ and `⊔` distribute over `⨆` and `⨅` respectively.
20
+ -/
21
+
10
22
set_option old_structure_cmd true
11
23
12
24
universes u v w
@@ -17,8 +29,8 @@ variables {α : Type u} {β : Type v} {ι : Sort w}
17
29
as this class includes a requirement that the lattice join
18
30
distribute over *arbitrary* infima, and similarly for the dual. -/
19
31
class complete_distrib_lattice α extends complete_lattice α :=
20
- (infi_sup_le_sup_Inf : ∀a s, (⨅ b ∈ s, a ⊔ b) ≤ a ⊔ Inf s)
21
- (inf_Sup_le_supr_inf : ∀a s, a ⊓ Sup s ≤ (⨆ b ∈ s, a ⊓ b))
32
+ (infi_sup_le_sup_Inf : ∀ a s, (⨅ b ∈ s, a ⊔ b) ≤ a ⊔ Inf s)
33
+ (inf_Sup_le_supr_inf : ∀ a s, a ⊓ Sup s ≤ (⨆ b ∈ s, a ⊓ b))
22
34
23
35
section complete_distrib_lattice
24
36
variables [complete_distrib_lattice α] {a b : α} {s t : set α}
@@ -62,31 +74,31 @@ instance pi.complete_distrib_lattice {ι : Type*} {π : ι → Type*}
62
74
← supr_subtype''],
63
75
.. pi.complete_lattice }
64
76
65
- theorem Inf_sup_Inf : Inf s ⊔ Inf t = (⨅p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) :=
77
+ theorem Inf_sup_Inf : Inf s ⊔ Inf t = (⨅ p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) :=
66
78
begin
67
79
apply le_antisymm,
68
80
{ simp only [and_imp, prod.forall, le_infi_iff, set.mem_prod],
69
81
intros a b ha hb,
70
82
exact sup_le_sup (Inf_le ha) (Inf_le hb) },
71
- { have : ∀ a ∈ s, (⨅p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) ≤ a ⊔ Inf t,
72
- { assume a ha,
73
- have : (⨅p ∈ set.prod s t, ((p : α × α).1 : α) ⊔ p.2 ) ≤
74
- (⨅p ∈ prod.mk a '' t, (p : α × α).1 ⊔ p.2 ),
83
+ { have : ∀ a ∈ s, (⨅ p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) ≤ a ⊔ Inf t,
84
+ { rintro a ha,
85
+ have : (⨅ p ∈ set.prod s t, ((p : α × α).1 : α) ⊔ p.2 ) ≤
86
+ (⨅ p ∈ prod.mk a '' t, (p : α × α).1 ⊔ p.2 ),
75
87
{ apply infi_le_infi_of_subset,
76
- rintros ⟨x, y⟩,
88
+ rintro ⟨x, y⟩,
77
89
simp only [and_imp, set.mem_image, prod.mk.inj_iff, set.prod_mk_mem_set_prod_eq,
78
90
exists_imp_distrib],
79
- assume x' x't ax x'y,
91
+ rintro x' x't ax x'y,
80
92
rw [← x'y, ← ax],
81
93
simp [ha, x't] },
82
94
rw [infi_image] at this ,
83
95
simp only at this ,
84
96
rwa ← sup_Inf_eq at this },
85
- calc (⨅p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) ≤ (⨅a∈ s, a ⊔ Inf t) : by simp; exact this
97
+ calc (⨅ p ∈ set.prod s t, (p : α × α).1 ⊔ p.2 ) ≤ (⨅ a ∈ s, a ⊔ Inf t) : by simp; exact this
86
98
... = Inf s ⊔ Inf t : Inf_sup_eq.symm }
87
99
end
88
100
89
- theorem Sup_inf_Sup : Sup s ⊓ Sup t = (⨆p ∈ set.prod s t, (p : α × α).1 ⊓ p.2 ) :=
101
+ theorem Sup_inf_Sup : Sup s ⊓ Sup t = (⨆ p ∈ set.prod s t, (p : α × α).1 ⊓ p.2 ) :=
90
102
@Inf_sup_Inf (order_dual α) _ _ _
91
103
92
104
lemma supr_disjoint_iff {f : ι → α} : disjoint (⨆ i, f i) a ↔ ∀ i, disjoint (f i) a :=
@@ -103,7 +115,7 @@ instance complete_distrib_lattice.bounded_distrib_lattice [d : complete_distrib_
103
115
{ le_sup_inf := λ x y z, by rw [← Inf_pair, ← Inf_pair, sup_Inf_eq, ← Inf_image, set.image_pair],
104
116
..d }
105
117
106
- /-- A complete boolean algebra is a completely distributive boolean algebra. -/
118
+ /-- A complete Boolean algebra is a completely distributive Boolean algebra. -/
107
119
class complete_boolean_algebra α extends boolean_algebra α, complete_distrib_lattice α
108
120
109
121
instance pi.complete_boolean_algebra {ι : Type *} {π : ι → Type *}
@@ -120,18 +132,18 @@ instance Prop.complete_boolean_algebra : complete_boolean_algebra Prop :=
120
132
section complete_boolean_algebra
121
133
variables [complete_boolean_algebra α] {a b : α} {s : set α} {f : ι → α}
122
134
123
- theorem compl_infi : (infi f)ᶜ = (⨆i, (f i)ᶜ) :=
135
+ theorem compl_infi : (infi f)ᶜ = (⨆ i, (f i)ᶜ) :=
124
136
le_antisymm
125
- (compl_le_of_compl_le $ le_infi $ assume i, compl_le_of_compl_le $ le_supr (compl ∘ f) i)
126
- (supr_le $ assume i, compl_le_compl $ infi_le _ _)
137
+ (compl_le_of_compl_le $ le_infi $ λ i, compl_le_of_compl_le $ le_supr (compl ∘ f) i)
138
+ (supr_le $ λ i, compl_le_compl $ infi_le _ _)
127
139
128
- theorem compl_supr : (supr f)ᶜ = (⨅i, (f i)ᶜ) :=
140
+ theorem compl_supr : (supr f)ᶜ = (⨅ i, (f i)ᶜ) :=
129
141
compl_injective (by simp [compl_infi])
130
142
131
- theorem compl_Inf : (Inf s)ᶜ = (⨆i∈ s, iᶜ) :=
143
+ theorem compl_Inf : (Inf s)ᶜ = (⨆ i ∈ s, iᶜ) :=
132
144
by simp only [Inf_eq_infi, compl_infi]
133
145
134
- theorem compl_Sup : (Sup s)ᶜ = (⨅i∈ s, iᶜ) :=
146
+ theorem compl_Sup : (Sup s)ᶜ = (⨅ i ∈ s, iᶜ) :=
135
147
by simp only [Sup_eq_supr, compl_supr]
136
148
137
149
end complete_boolean_algebra
0 commit comments