module Order.Frame where
Frames🔗
A frame is a lattice with finite meets1 and arbitrary joins satisfying the infinite distributive law
In the study of frames, for simplicity, we assume propositional resizing: that way, it suffices for a frame to have joins of families, for an arbitrary type in the same universe as to have joins for arbitrary subsets of
record is-frame {o ℓ} (P : Poset o ℓ) : Type (lsuc o ⊔ ℓ) where
no-eta-equality
open Poset P
field
_∩_ : Ob → Ob → Ob
: ∀ x y → is-meet P x y (x ∩ y)
∩-meets
: Top P
has-top
: ∀ {I : Type o} (k : I → Ob) → Ob
⋃ : ∀ {I : Type o} (k : I → Ob) → is-lub P k (⋃ k)
⋃-lubs
: ∀ {I} x (f : I → Ob) → x ∩ ⋃ f ≡ ⋃ λ i → x ∩ f i ⋃-distribl
We have explicitly required that a frame be a meet-semilattice, but it’s worth explicitly pointing out that the infinitary join operation can also be used for more mundane purposes: By taking a join over the type of booleans (and over the empty type), we can show that all frames are also join-semilattices.
infixr 25 _∩_
module is-lubs {I} {k : I → Ob} = is-lub (⋃-lubs k)
open Meets ∩-meets public
open Top has-top using (top; !) public
open Lubs P ⋃-lubs public
module _ {o ℓ} {P : Poset o ℓ} (f : is-frame P) where
: is-meet-semilattice P
is-frame→is-meet-slat = record { is-frame f }
is-frame→is-meet-slat
: is-join-semilattice P
is-frame→is-join-slat = record { is-frame f }
is-frame→is-join-slat
: is-lattice P
is-frame→is-lattice = record { is-frame f }
is-frame→is-lattice
private variable
: Level
o ℓ o' ℓ' : Poset o ℓ
P Q R
abstract
: is-prop (is-frame P)
is-frame-is-prop {P = P} p q = path where
is-frame-is-prop open Order.Diagram.Top P using (H-Level-Top)
module p = is-frame p
module q = is-frame q
open is-frame
: ∀ x y → x p.∩ y ≡ x q.∩ y
meetp = meet-unique (p.∩-meets x y) (q.∩-meets x y)
meetp x y
: ∀ {I} (k : I → ⌞ P ⌟) → p.⋃ k ≡ q.⋃ k
lubp = lub-unique (p.⋃-lubs k) (q.⋃-lubs k)
lubp k
: p ≡ q
path ._∩_ x y = meetp x y i
path i .∩-meets x y = is-prop→pathp (λ i → hlevel {T = is-meet P x y (meetp x y i)} 1) (p.∩-meets x y) (q.∩-meets x y) i
path i .has-top = hlevel {T = Top P} 1 p.has-top q.has-top i
path i .⋃ k = lubp k i
path i .⋃-lubs k = is-prop→pathp (λ i → hlevel {T = is-lub P k (lubp k i)} 1) (p.⋃-lubs k) (q.⋃-lubs k) i
path i .⋃-distribl x f j = is-set→squarep (λ _ _ → Poset.Ob-is-set P)
path i (λ i → meetp x (lubp f i) i)
(p.⋃-distribl x f) (q.⋃-distribl x f)
(λ i → lubp (λ e → meetp x (f e) i) i)
i j
instance
: ∀ {n} → H-Level (is-frame P) (suc n)
H-Level-is-frame = prop-instance is-frame-is-prop H-Level-is-frame
Of course, a frame is not just a lattice, but a complete lattice. Since the infinite distributive law says exactly that “meet with ” preserves joins, this implies that it has a right adjoint, so frames are also complete Heyting algebras. Once again, the difference in naming reflects the morphisms we will consider these structures under: A frame homomorphism is a monotone map which preserves the finite meets and the infinitary joins, but not necessarily the infinitary meets (or the Heyting implication).
Since meets and joins are defined by a universal property, and we have assumed that homomorphisms are a priori monotone, it suffices to show the following inequalities:
- For every we have
- and finally, for every family we have
record
is-frame-hom{P : Poset o ℓ} {Q : Poset o ℓ'}
(f : Monotone P Q) (P-frame : is-frame P) (Q-frame : is-frame Q)
: Type (lsuc o ⊔ ℓ') where
private
module P = Poset P
module Pᶠ = is-frame P-frame
module Q = Order.Reasoning Q
module Qᶠ = is-frame Q-frame
open is-lub
field
: ∀ x y → f · x Qᶠ.∩ f · y Q.≤ f · (x Pᶠ.∩ y)
∩-≤ : Qᶠ.top Q.≤ f · Pᶠ.top
top-≤ : ∀ {I : Type o} (k : I → ⌞ P ⌟) → (f · Pᶠ.⋃ k) Q.≤ Qᶠ.⋃ (apply f ⊙ k) ⋃-≤
If is a frame homomorphism, then it is also a homomorphism of meet semilattices.
: is-meet-slat-hom f record { Pᶠ } record { Qᶠ }
has-meet-slat-hom .is-meet-slat-hom.∩-≤ = ∩-≤
has-meet-slat-hom .is-meet-slat-hom.top-≤ = top-≤
has-meet-slat-hom
open is-meet-slat-hom has-meet-slat-hom hiding (∩-≤; top-≤) public
Furthermore, we can actually show from the inequality required above that preserves all joins up to equality.
: ∀ {I : Type o} (k : I → ⌞ P ⌟) → (f · Pᶠ.⋃ k) ≡ Qᶠ.⋃ (apply f ⊙ k)
pres-⋃ =
pres-⋃ k .≤-antisym
Q(⋃-≤ k)
(Qᶠ.⋃-universal _ (λ i → f .pres-≤ (Pᶠ.⋃-inj i)))
pres-lubs: ∀ {I : Type o} {k : I → ⌞ P ⌟} {l}
→ is-lub P k l
→ is-lub Q (apply f ⊙ k) (f · l)
{k = k} {l = l} l-lub .fam≤lub i = f .pres-≤ (l-lub .fam≤lub i)
pres-lubs {k = k} {l = l} l-lub .least ub p =
pres-lubs .≤⟨ f .pres-≤ (l-lub .least _ Pᶠ.⋃-inj) ⟩
f · l Q.⋃ k Q.≤⟨ ⋃-≤ k ⟩
f · Pᶠ.⋃ (apply f ⊙ k) Q.≤⟨ Qᶠ.⋃-universal ub p ⟩
Qᶠ.≤∎ ub Q
As a corollary, is also a homomorphism of the underlying join semilattices.
opaque.∪-joins Lubs.has-bottom
unfolding Lubs
: is-join-slat-hom f record { Pᶠ } record { Qᶠ }
has-join-slat-hom .is-join-slat-hom.∪-≤ x y =
has-join-slat-hom .≤-trans (⋃-≤ _) $ Qᶠ.⋃-universal _ λ where
Q(lift true) → Qᶠ.⋃-inj (lift true)
(lift false) → Qᶠ.⋃-inj (lift false)
.is-join-slat-hom.bot-≤ =
has-join-slat-hom .≤-trans (⋃-≤ _) (Qᶠ.⋃-universal _ (λ ()))
Q
open is-join-slat-hom has-join-slat-hom public
open is-frame-hom
unquoteDecl H-Level-is-frame-hom = declare-record-hlevel 1 H-Level-is-frame-hom (quote is-frame-hom)
Clearly, the identity function is a frame homomorphism.
id-frame-hom: ∀ (Pᶠ : is-frame P)
→ is-frame-hom idₘ Pᶠ Pᶠ
{P = P} Pᶠ .∩-≤ x y =
id-frame-hom .≤-refl P
Poset{P = P} Pᶠ .top-≤ =
id-frame-hom .≤-refl P
Poset{P = P} Pᶠ .⋃-≤ k =
id-frame-hom .≤-refl P Poset
Furthermore, frame homomorphisms are closed under composition.
∘-frame-hom: ∀ {Pₗ Qₗ Rₗ} {f : Monotone Q R} {g : Monotone P Q}
→ is-frame-hom f Qₗ Rₗ
→ is-frame-hom g Pₗ Qₗ
→ is-frame-hom (f ∘ₘ g) Pₗ Rₗ
{R = R} {f = f} {g = g} f-pres g-pres .∩-≤ x y =
∘-frame-hom .Poset.≤-trans (f-pres .∩-≤ (g · x) (g · y)) (f .pres-≤ (g-pres .∩-≤ x y))
R {R = R} {f = f} {g = g} f-pres g-pres .top-≤ =
∘-frame-hom .Poset.≤-trans (f-pres .top-≤) (f .pres-≤ (g-pres .top-≤))
R {R = R} {f = f} {g = g} f-pres g-pres .⋃-≤ k =
∘-frame-hom .Poset.≤-trans (f .pres-≤ (g-pres .⋃-≤ k)) (f-pres .⋃-≤ (λ i → g · k i)) R
This means that we can define the category of frames as a subcategory of the category of posets.
: ∀ o ℓ → Subcat (Posets o ℓ) _ _
Frame-subcat .Subcat.is-ob = is-frame
Frame-subcat o ℓ .Subcat.is-hom = is-frame-hom
Frame-subcat o ℓ .Subcat.is-hom-prop _ _ _ = hlevel 1
Frame-subcat o ℓ .Subcat.is-hom-id = id-frame-hom
Frame-subcat o ℓ .Subcat.is-hom-∘ = ∘-frame-hom
Frame-subcat o ℓ
: ∀ o ℓ → Precategory _ _
Frames = Subcategory (Frame-subcat o ℓ)
Frames o ℓ
module Frames {o} {ℓ} = Cat.Reasoning (Frames o ℓ)
: ∀ o ℓ → Type _
Frame = Frames.Ob {o} {ℓ} Frame o ℓ
Power sets as frames🔗
A canonical source of frames are power sets: The power set of any type is a frame, because it is a complete lattice satisfying the infinite distributive law; This can be seen by some computation, as is done below.
open is-frame
open is-meet-semilattice
: ∀ {ℓ} (A : Type ℓ) → Frame ℓ ℓ
Power-frame {ℓ = ℓ} A .fst = Subsets A
Power-frame .snd ._∩_ P Q i = P i ∧Ω Q i
Power-frame A .snd .∩-meets P Q =
Power-frame A λ _ → Props-has-meets (P _) (Q _)
is-meet-pointwise .snd .has-top =
Power-frame A λ _ → Props-has-top
has-top-pointwise .snd .⋃ k i = ∃Ω _ (λ j → k j i)
Power-frame A .snd .⋃-lubs k = is-lub-pointwise _ _ λ _ →
Power-frame A λ i → k i _
Props-has-lubs .snd .⋃-distribl x f = ext λ i → Ω-ua
Power-frame A (rec! λ xi j j~i → inc (j , xi , j~i))
(rec! λ j xi j~i → xi , inc (j , j~i))
So, in addition to the operation, we have a top element↩︎