Brendan Hart 2019-2020
Addition by Tom de Jong in July 2024.
\begin{code}
{-# OPTIONS --safe --without-K #-}
open import MLTT.Spartan
open import UF.PropTrunc
open import UF.FunExt
module DomainTheory.Basics.Products
(pt : propositional-truncations-exist)
(fe : โ {๐ค ๐ฅ} โ funext ๐ค ๐ฅ)
where
open PropositionalTruncation pt
open import OrderedTypes.Poset fe
open import UF.Base
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import UF.Sets
open PosetAxioms
\end{code}
First, let's define the product of two DCPOs.
\begin{code}
module DcpoProductsGeneral
(๐ฅ : Universe)
where
open import DomainTheory.Basics.Dcpo pt fe ๐ฅ
open import DomainTheory.Basics.Miscelanea pt fe ๐ฅ
open import DomainTheory.Basics.Pointed pt fe ๐ฅ
module _ {D : ๐ค ฬ } {E : ๐ค' ฬ } where
_โ-ร_ : (D โ D โ ๐ฃ ฬ )
โ (E โ E โ ๐ฃ' ฬ )
โ (D ร E โ D ร E โ ๐ฃ โ ๐ฃ' ฬ )
_โ-ร_ _โโ_ _โโ_ (a , b) (c , d) = (a โโ c) ร (b โโ d)
prโโฮฑ-is-directed : {I : ๐ฅ ฬ }
โ {ฮฑ : I โ D ร E}
โ (_โโ_ : D โ D โ ๐ฃ ฬ )
โ (_โโ_ : E โ E โ ๐ฃ' ฬ )
โ is-directed (_โโ_ โ-ร _โโ_) ฮฑ
โ is-directed _โโ_ (prโ โ ฮฑ)
prโโฮฑ-is-directed {_} {_} {I} {ฮฑ} _โโ_ _โโ_ ฮด =
inhabited-if-directed (_โโ_ โ-ร _โโ_) ฮฑ ฮด , o
where
o : (i j : I)
โ โ k ๊ I , ((prโ โ ฮฑ) i โโ (prโ โ ฮฑ) k ร
(prโ โ ฮฑ) j โโ (prโ โ ฮฑ) k)
o i j = โฅโฅ-functor
(ฮป (a , (b , _) , c , _) โ a , b , c)
(semidirected-if-directed (_โโ_ โ-ร _โโ_) ฮฑ ฮด i j)
prโโฮฑ-is-directed : {I : ๐ฅ ฬ }
โ {ฮฑ : I โ D ร E}
โ (_โโ_ : D โ D โ ๐ฃ ฬ )
โ (_โโ_ : E โ E โ ๐ฃ' ฬ )
โ is-directed (_โโ_ โ-ร _โโ_) ฮฑ
โ is-directed _โโ_ (prโ โ ฮฑ)
prโโฮฑ-is-directed {_} {_} {I} {ฮฑ} _โโ_ _โโ_ ฮด =
inhabited-if-directed (_โโ_ โ-ร _โโ_) ฮฑ ฮด , o
where
o : (i j : I)
โ โ k ๊ I , ((prโ โ ฮฑ) i โโ (prโ โ ฮฑ) k ร
(prโ โ ฮฑ) j โโ (prโ โ ฮฑ) k)
o i j = โฅโฅ-functor
(ฮป (a , (_ , b) , _ , c) โ a , b , c)
(semidirected-if-directed (_โโ_ โ-ร _โโ_) ฮฑ ฮด i j)
infixr 30 _รแตแถแตแต_
_รแตแถแตแต_ : DCPO {๐ค} {๐ฃ} โ DCPO {๐ค'} {๐ฃ'} โ DCPO {๐ค โ ๐ค'} {๐ฃ โ ๐ฃ'}
๐ รแตแถแตแต ๐ = (โจ ๐ โฉ ร โจ ๐ โฉ) ,
(underlying-order ๐) โ-ร (underlying-order ๐) ,
axioms
where
axioms : dcpo-axioms (underlying-order ๐ โ-ร underlying-order ๐)
axioms = (ร-is-set (sethood ๐) (sethood ๐) , prop , r , t , a) , c
where
_โ๐_ = underlying-order ๐
_โ๐_ = underlying-order ๐
_โ_ = _โ๐_ โ-ร _โ๐_
prop : is-prop-valued _โ_
prop (xโ , xโ) (yโ , yโ) (aโ , aโ) (bโ , bโ) =
to-ร-๏ผ
(prop-valuedness ๐ xโ yโ aโ bโ)
(prop-valuedness ๐ xโ yโ aโ bโ)
r : is-reflexive _โ_
r a = reflexivity ๐ (prโ a) ,
reflexivity ๐ (prโ a)
t : is-transitive _โ_
t (xโ , xโ) (yโ , yโ) (zโ , zโ) (uโ , uโ) (vโ , vโ) = wโ , wโ
where
wโ = xโ โโจ ๐ โฉ[ uโ ]
yโ โโจ ๐ โฉ[ vโ ]
zโ โโจ ๐ โฉ
wโ = xโ โโจ ๐ โฉ[ uโ ]
yโ โโจ ๐ โฉ[ vโ ]
zโ โโจ ๐ โฉ
a : is-antisymmetric _โ_
a (a , b) (c , d) (a-โ-c , b-โ-d) (c-โ-a , d-โ-b) =
to-ร-๏ผ
(antisymmetry ๐ a c a-โ-c c-โ-a)
(antisymmetry ๐ b d b-โ-d d-โ-b)
c : is-directed-complete _โ_
c I ฮฑ ฮด = (โ ๐ prโโฮฑ-is-dir , โ ๐ prโโฮฑ-is-dir) , s
where
prโโฮฑ-is-dir : is-Directed ๐ (prโ โ ฮฑ)
prโโฮฑ-is-dir = prโโฮฑ-is-directed _โ๐_ _โ๐_ ฮด
prโโฮฑ-is-dir : is-Directed ๐ (prโ โ ฮฑ)
prโโฮฑ-is-dir = prโโฮฑ-is-directed _โ๐_ _โ๐_ ฮด
s : is-sup _โ_ (โ ๐ prโโฮฑ-is-dir , โ ๐ prโโฮฑ-is-dir) ฮฑ
s = is-upper , is-least-upper
where
is-upper : is-upperbound _โ_ (โ ๐ prโโฮฑ-is-dir , โ ๐ prโโฮฑ-is-dir) ฮฑ
is-upper i = (โ-is-upperbound ๐ prโโฮฑ-is-dir i ,
โ-is-upperbound ๐ prโโฮฑ-is-dir i)
is-least-upper : (t : โจ ๐ โฉ ร โจ ๐ โฉ)
โ is-upperbound _โ_ t ฮฑ
โ _โ_ (โ ๐ prโโฮฑ-is-dir , โ ๐ prโโฮฑ-is-dir) t
is-least-upper t@(x , y) u = lub-of-x , lub-of-y
where
x-is-upperbound : is-upperbound _โ๐_ x (prโ โ ฮฑ)
x-is-upperbound i = prโ (u i)
y-is-upperbound : is-upperbound _โ๐_ y (prโ โ ฮฑ)
y-is-upperbound i = prโ (u i)
lub-of-x = โ-is-lowerbound-of-upperbounds ๐
prโโฮฑ-is-dir
x
x-is-upperbound
lub-of-y = โ-is-lowerbound-of-upperbounds
๐
prโโฮฑ-is-dir
y
y-is-upperbound
\end{code}
Some useful proofs on products.
\begin{code}
module _ (๐ : DCPO {๐ค} {๐ค'}) where
constant-function-is-directed : { I : ๐ฅ ฬ } (h : โฅ I โฅ) (d : โจ ๐ โฉ)
โ is-Directed ๐ (ฮป (i : I) โ d)
constant-function-is-directed h d =
h , ฮป i j โ โฃ i , (reflexivity ๐ d , reflexivity ๐ d) โฃ
constant-is-โ-of-constant-function : {I : ๐ฅ ฬ }
{d : โจ ๐ โฉ}
(ฮด : is-Directed ๐ (ฮป (i : I) โ d))
โ d ๏ผ โ ๐ ฮด
constant-is-โ-of-constant-function {I} {d} ฮด = antisymmetry ๐ d (โ ๐ ฮด) lโ lโ
where
lโ : d โโจ ๐ โฉ โ ๐ ฮด
lโ = โฅโฅ-rec
(prop-valuedness ๐ d (โ ๐ ฮด))
(ฮป (i : I) โ โ-is-upperbound ๐ ฮด i) (prโ ฮด)
lโ : โ ๐ ฮด โโจ ๐ โฉ d
lโ = โ-is-lowerbound-of-upperbounds ๐
ฮด
d
(ฮป i โ reflexivity ๐ d)
module _ (๐ : DCPO {๐ค} {๐ค'})
(๐ : DCPO {๐ฃ} {๐ฃ'})
where
prโโฮฑ-is-Directed : {I : ๐ฅ ฬ }
{ฮฑ : I โ โจ ๐ รแตแถแตแต ๐ โฉ}
โ is-Directed (๐ รแตแถแตแต ๐) ฮฑ
โ is-Directed ๐ (prโ โ ฮฑ)
prโโฮฑ-is-Directed {I} {ฮฑ} ฮด =
prโโฮฑ-is-directed (underlying-order ๐) (underlying-order ๐) ฮด
prโโฮฑ-is-Directed : {I : ๐ฅ ฬ }
{ฮฑ : I โ โจ ๐ รแตแถแตแต ๐ โฉ}
โ is-Directed (๐ รแตแถแตแต ๐) ฮฑ
โ is-Directed ๐ (prโ โ ฮฑ)
prโโฮฑ-is-Directed = prโโฮฑ-is-directed (underlying-order ๐) (underlying-order ๐)
โจprโ,prโโฉ-is-directed : {I : ๐ฅ ฬ }
โ {ฮฑโ : I โ โจ ๐ โฉ}
โ {ฮฑโ : I โ โจ ๐ โฉ}
โ is-Directed ๐ ฮฑโ
โ is-Directed ๐ ฮฑโ
โ is-Directed (๐ รแตแถแตแต ๐)
(ฮป ((iโ , iโ) : I ร I) โ (ฮฑโ iโ , ฮฑโ iโ))
โจprโ,prโโฉ-is-directed ฮดโ@(hโ , sโ) ฮดโ@(hโ , sโ) =
(binary-choice hโ hโ) ,
ฮป (iโ , iโ) (jโ , jโ)
โ โฅโฅ-functor
(ฮป ((aโ , bโ , cโ) , (aโ , bโ , cโ)) โ (aโ , aโ) , (bโ , bโ) , (cโ , cโ))
(binary-choice (sโ iโ jโ) (sโ iโ jโ))
โโจ,โฉ๏ผโจโ,โโฉ : {I : ๐ฅ ฬ } {ฮฑ : I โ โจ ๐ รแตแถแตแต ๐ โฉ}
โ (ฮด : is-Directed (๐ รแตแถแตแต ๐) ฮฑ)
โ โ (๐ รแตแถแตแต ๐) ฮด
๏ผ (โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด))
โโจ,โฉ๏ผโจโ,โโฉ {I} {ฮฑ} ฮด =
antisymmetry (๐ รแตแถแตแต ๐)
(โ (๐ รแตแถแตแต ๐) ฮด)
(โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด))
โโจ,โฉโโจโ,โโฉ
โจโ,โโฉโโโจ,โฉ
where
โโจ,โฉโโจโ,โโฉ : โ (๐ รแตแถแตแต ๐) ฮด
โโจ ๐ รแตแถแตแต ๐ โฉ
(โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด))
โโจ,โฉโโจโ,โโฉ = โ-is-lowerbound-of-upperbounds (๐ รแตแถแตแต ๐)
ฮด
(โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด))
โ,โ-is-upperbound
where
โ,โ-is-upperbound : (i : I)
โ (ฮฑ i)
โโจ ๐ รแตแถแตแต ๐ โฉ (โ ๐ (prโโฮฑ-is-Directed ฮด) ,
โ ๐ (prโโฮฑ-is-Directed ฮด))
โ,โ-is-upperbound i = (โ-is-upperbound ๐ (prโโฮฑ-is-Directed ฮด) i) ,
(โ-is-upperbound ๐ (prโโฮฑ-is-Directed ฮด) i)
โจโ,โโฉโโโจ,โฉ : (โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด))
โโจ ๐ รแตแถแตแต ๐ โฉ (โ (๐ รแตแถแตแต ๐) ฮด)
โจโ,โโฉโโโจ,โฉ = โโ , โโ
where
โโ : (โ ๐ (prโโฮฑ-is-Directed ฮด)) โโจ ๐ โฉ (prโ (โ (๐ รแตแถแตแต ๐) ฮด))
โโ = โ-is-lowerbound-of-upperbounds ๐
(prโโฮฑ-is-Directed ฮด)
(prโ (โ (๐ รแตแถแตแต ๐) ฮด))
prโ-โโจ,โฉ-is-upperbound
where
prโ-โโจ,โฉ-is-upperbound : (i : I)
โ ((prโ โ ฮฑ) i) โโจ ๐ โฉ (prโ (โ (๐ รแตแถแตแต ๐) ฮด))
prโ-โโจ,โฉ-is-upperbound i = prโ (โ-is-upperbound (๐ รแตแถแตแต ๐) ฮด i)
โโ : โ ๐ (prโโฮฑ-is-Directed ฮด) โโจ ๐ โฉ prโ (โ (๐ รแตแถแตแต ๐) ฮด)
โโ = โ-is-lowerbound-of-upperbounds ๐
(prโโฮฑ-is-Directed ฮด)
(prโ (โ (๐ รแตแถแตแต ๐) ฮด))
prโ-โโจ,โฉ-is-upperbound
where
prโ-โโจ,โฉ-is-upperbound : (i : I)
โ ((prโ โ ฮฑ) i) โโจ ๐ โฉ (prโ (โ (๐ รแตแถแตแต ๐) ฮด))
prโ-โโจ,โฉ-is-upperbound i = prโ (โ-is-upperbound (๐ รแตแถแตแต ๐) ฮด i)
โจโ,โโฉ๏ผโโจ,โฉ : {I : ๐ฅ ฬ }
โ {ฮฑโ : I โ โจ ๐ โฉ}
โ {ฮฑโ : I โ โจ ๐ โฉ}
โ (ฮดโ : is-Directed ๐ ฮฑโ)
โ (ฮดโ : is-Directed ๐ ฮฑโ)
โ (โ ๐ ฮดโ , โ ๐ ฮดโ) ๏ผ โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
โจโ,โโฉ๏ผโโจ,โฉ {I} {ฮฑโ} {ฮฑโ} ฮดโ ฮดโ = antisymmetry (๐ รแตแถแตแต ๐)
(โ ๐ ฮดโ , โ ๐ ฮดโ)
(โ (๐ รแตแถแตแต ๐)
(โจprโ,prโโฉ-is-directed ฮดโ ฮดโ))
โจโ,โโฉโโโจ,โฉ โโจ,โฉโโจโ,โโฉ
where
โจโ,โโฉโโโจ,โฉ : (โ ๐ ฮดโ , โ ๐ ฮดโ)
โโจ ๐ รแตแถแตแต ๐ โฉ โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
โจโ,โโฉโโโจ,โฉ = โโ , โโ
where
โโ : โ ๐ ฮดโ โโจ ๐ โฉ prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ))
โโ = โ-is-lowerbound-of-upperbounds ๐
ฮดโ
(prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)))
p
where
p : (i : I)
โ (ฮฑโ i) โโจ ๐ โฉ (prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)))
p i = prโ (โ-is-upperbound ((๐ รแตแถแตแต ๐))
(โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
(i , i))
โโ : โ ๐ ฮดโ โโจ ๐ โฉ prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ))
โโ = โ-is-lowerbound-of-upperbounds ๐
ฮดโ
(prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)))
p
where
p : (i : I)
โ (ฮฑโ i) โโจ ๐ โฉ (prโ (โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)))
p i = prโ (โ-is-upperbound (๐ รแตแถแตแต ๐)
(โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
(i , i))
โโจ,โฉโโจโ,โโฉ : โ (๐ รแตแถแตแต ๐) (โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
โโจ ๐ รแตแถแตแต ๐ โฉ (โ ๐ ฮดโ , โ ๐ ฮดโ)
โโจ,โฉโโจโ,โโฉ = โ-is-lowerbound-of-upperbounds (๐ รแตแถแตแต ๐)
(โจprโ,prโโฉ-is-directed ฮดโ ฮดโ)
(โ ๐ ฮดโ , โ ๐ ฮดโ)
upperbound
where
upperbound : (i : I ร I)
โ (ฮฑโ (prโ i) , ฮฑโ (prโ i)) โโจ ๐ รแตแถแตแต ๐ โฉ (โ ๐ ฮดโ , โ ๐ ฮดโ)
upperbound i = โ-is-upperbound ๐ ฮดโ (prโ i) ,
โ-is-upperbound ๐ ฮดโ (prโ i)
prโ-is-continuous : DCPO[ ๐ รแตแถแตแต ๐ , ๐ ]
prโ-is-continuous = prโ , c
where
c : is-continuous (๐ รแตแถแตแต ๐) ๐ prโ
c I ฮฑ ฮด = u , v
where
u : is-upperbound (underlying-order ๐) (prโ (โ (๐ รแตแถแตแต ๐) ฮด)) (prโ โ ฮฑ)
u i = prโ (โ-is-upperbound (๐ รแตแถแตแต ๐) ฮด i)
v : (x : โจ ๐ โฉ)
โ ((i : I) โ (prโ (ฮฑ i)) โโจ ๐ โฉ x)
โ (prโ (โ (๐ รแตแถแตแต ๐) ฮด)) โโจ ๐ โฉ x
v x p = transport (ฮป - โ prโ - โโจ ๐ โฉ x) (โโจ,โฉ๏ผโจโ,โโฉ ฮด) d
where
d : prโ (โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด)) โโจ ๐ โฉ x
d = โ-is-lowerbound-of-upperbounds ๐ (prโโฮฑ-is-Directed ฮด) x p
prโ-is-continuous : DCPO[ ๐ รแตแถแตแต ๐ , ๐ ]
prโ-is-continuous = prโ , c
where
c : is-continuous (๐ รแตแถแตแต ๐) ๐ prโ
c I ฮฑ ฮด = u , v
where
u : is-upperbound (underlying-order ๐) (prโ (โ (๐ รแตแถแตแต ๐) ฮด)) (prโ โ ฮฑ)
u i = prโ (โ-is-upperbound (๐ รแตแถแตแต ๐) ฮด i)
v : (y : โจ ๐ โฉ)
โ ((i : I) โ (prโ (ฮฑ i)) โโจ ๐ โฉ y)
โ (prโ (โ (๐ รแตแถแตแต ๐) ฮด)) โโจ ๐ โฉ y
v y p = transport (ฮป - โ prโ - โโจ ๐ โฉ y) (โโจ,โฉ๏ผโจโ,โโฉ ฮด) e
where
e : prโ (โ ๐ (prโโฮฑ-is-Directed ฮด) , โ ๐ (prโโฮฑ-is-Directed ฮด)) โโจ ๐ โฉ y
e = โ-is-lowerbound-of-upperbounds ๐ (prโโฮฑ-is-Directed ฮด) y p
infixr 30 _รแตแถแตแตโฅ_
_รแตแถแตแตโฅ_ : DCPOโฅ {๐ค} {๐ฃ}
โ DCPOโฅ {๐ค'} {๐ฃ'}
โ DCPOโฅ {๐ค โ ๐ค'} {๐ฃ โ ๐ฃ'}
๐ รแตแถแตแตโฅ ๐ = ((๐ โป) รแตแถแตแต (๐ โป)) , least , p
where
least : โจ (๐ โป) รแตแถแตแต (๐ โป) โฉ
least = โฅ ๐ , โฅ ๐
p : is-least (underlying-order ((๐ โป) รแตแถแตแต (๐ โป))) least
p x = (โฅ-is-least ๐ (prโ x)) , (โฅ-is-least ๐ (prโ x))
module _ (๐ : DCPO {๐ค} {๐ค'})
(๐ : DCPO {๐ฃ} {๐ฃ'})
(๐ : DCPO {๐ฆ} {๐ฆ'})
where
รแตแถแตแต-assocโ : DCPO[ ๐ รแตแถแตแต ๐ รแตแถแตแต ๐ , (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ ]
รแตแถแตแต-assocโ = f , c
where
f : โจ ๐ รแตแถแตแต ๐ รแตแถแตแต ๐ โฉ โ โจ (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ โฉ
f (d , e , f) = (d , e) , f
f-is-monotone : is-monotone (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ((๐ รแตแถแตแต ๐) รแตแถแตแต ๐) f
f-is-monotone x y p = ((prโ p) , (prโ (prโ p))) , (prโ (prโ p))
c : is-continuous (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ((๐ รแตแถแตแต ๐) รแตแถแตแต ๐) f
c I ฮฑ ฮด = u , v
where
assoc-โ : โจ (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ โฉ
assoc-โ = (prโ (โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด) ,
(prโ (prโ (โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด)))) ,
prโ (prโ (โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด))
u : is-upperbound
(underlying-order ((๐ รแตแถแตแต ๐) รแตแถแตแต ๐))
assoc-โ
(ฮป i โ ((prโ (ฮฑ i)) , (prโ (prโ (ฮฑ i)))) , (prโ (prโ (ฮฑ i))))
u i = (prโ p , prโ (prโ p)) , prโ (prโ p)
where
p = โ-is-upperbound (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด i
v : (w : โจ (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ โฉ)
โ ((i : I) โ f (ฮฑ i) โโจ (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ โฉ w)
โ f (โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด) โโจ (๐ รแตแถแตแต ๐) รแตแถแตแต ๐ โฉ w
v w@((x , y) , z) p = f-is-monotone (โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด) w' l
where
w' = x , (y , z)
w'-is-upperbound : is-upperbound
(underlying-order (๐ รแตแถแตแต ๐ รแตแถแตแต ๐))
w'
ฮฑ
w'-is-upperbound i = (prโ (prโ (p i))) , (prโ (prโ (p i))) , (prโ (p i))
l : โ (๐ รแตแถแตแต ๐ รแตแถแตแต ๐) ฮด โโจ ๐ รแตแถแตแต ๐ รแตแถแตแต ๐ โฉ w'
l = โ-is-lowerbound-of-upperbounds (๐ รแตแถแตแต ๐ รแตแถแตแต ๐)
ฮด
w'
w'-is-upperbound
to-ร-DCPO : DCPO[ ๐ , ๐ ] โ DCPO[ ๐ , ๐ ] โ DCPO[ ๐ , ๐ รแตแถแตแต ๐ ]
to-ร-DCPO ๐@(f , fc) ๐@(g , gc) = h , hc
where
h : โจ ๐ โฉ โ โจ ๐ รแตแถแตแต ๐ โฉ
h d = f d , g d
h-is-monotone : is-monotone ๐ (๐ รแตแถแตแต ๐) h
h-is-monotone x y p = monotone-if-continuous ๐ ๐ ๐ x y p ,
monotone-if-continuous ๐ ๐ ๐ x y p
hc : is-continuous ๐ (๐ รแตแถแตแต ๐) h
hc I ฮฑ ฮด = u , v
where
u : is-upperbound (underlying-order (๐ รแตแถแตแต ๐)) (h (โ ๐ ฮด)) (h โ ฮฑ)
u i = h-is-monotone (ฮฑ i) (โ ๐ ฮด) (โ-is-upperbound ๐ ฮด i)
v : (t : โจ ๐ รแตแถแตแต ๐ โฉ)
โ ((i : I) โ (h (ฮฑ i)) โโจ ๐ รแตแถแตแต ๐ โฉ t)
โ (h (โ ๐ ฮด)) โโจ ๐ รแตแถแตแต ๐ โฉ t
v t@(y , z) p = lf , lg
where
lf : f (โ ๐ ฮด) โโจ ๐ โฉ y
lf = transport (ฮป - โ - โโจ ๐ โฉ y) (q โปยน) l
where
ฮต : is-Directed ๐ (f โ ฮฑ)
ฮต = image-is-directed ๐ ๐ (monotone-if-continuous ๐ ๐ ๐) ฮด
q : f (โ ๐ ฮด) ๏ผ โ ๐ ฮต
q = continuous-โ-๏ผ ๐ ๐ ๐ ฮด
l : โ ๐ ฮต โโจ ๐ โฉ y
l = โ-is-lowerbound-of-upperbounds ๐ ฮต y (ฮป i โ prโ (p i))
lg : g (โ ๐ ฮด) โโจ ๐ โฉ z
lg = transport (ฮป - โ - โโจ ๐ โฉ z) (q โปยน) l
where
ฯ : is-Directed ๐ (g โ ฮฑ)
ฯ = image-is-directed ๐ ๐ (monotone-if-continuous ๐ ๐ ๐) ฮด
q : g (โ ๐ ฮด) ๏ผ โ ๐ ฯ
q = continuous-โ-๏ผ ๐ ๐ ๐ ฮด
l : โ ๐ ฯ โโจ ๐ โฉ z
l = โ-is-lowerbound-of-upperbounds ๐ ฯ z (ฮป i โ prโ (p i))
module _ (๐ : DCPOโฅ {๐ค} {๐ค'})
(๐ : DCPOโฅ {๐ฃ} {๐ฃ'})
(๐ : DCPOโฅ {๐ฆ} {๐ฆ'})
where
รแตแถแตแตโฅ-assocโ : DCPOโฅ[ ๐ รแตแถแตแตโฅ ๐ รแตแถแตแตโฅ ๐ , (๐ รแตแถแตแตโฅ ๐) รแตแถแตแตโฅ ๐ ]
รแตแถแตแตโฅ-assocโ = รแตแถแตแต-assocโ (๐ โป) (๐ โป) (๐ โป)
to-ร-DCPOโฅ : DCPOโฅ[ ๐ , ๐ ] โ DCPOโฅ[ ๐ , ๐ ] โ DCPOโฅ[ ๐ , ๐ รแตแถแตแตโฅ ๐ ]
to-ร-DCPOโฅ f g = to-ร-DCPO (๐ โป) (๐ โป) (๐ โป) f g
\end{code}
Added 3 July 2024 by Tom de Jong.
\begin{code}
รแตแถแตแต-is-product : (๐โ : DCPO {๐ค} {๐ฃ}) (๐โ : DCPO {๐ค'} {๐ฃ'})
(๐ : DCPO {๐ฆ} {๐ฆ'})
(f : โจ ๐ โฉ โ โจ ๐โ โฉ) (g : โจ ๐ โฉ โ โจ ๐โ โฉ)
โ is-continuous ๐ ๐โ f
โ is-continuous ๐ ๐โ g
โ โ! k ๊ (โจ ๐ โฉ โ โจ ๐โ รแตแถแตแต ๐โ โฉ) ,
is-continuous ๐ (๐โ รแตแถแตแต ๐โ) k
ร prโ โ k โผ f
ร prโ โ k โผ g
รแตแถแตแต-is-product ๐โ ๐โ ๐ f g cf cg =
(k , k-is-continuous , โผ-refl , โผ-refl) , k-is-unique
where
k-bundled : DCPO[ ๐ , ๐โ รแตแถแตแต ๐โ ]
k-bundled = to-ร-DCPO ๐ ๐โ ๐โ (f , cf) (g , cg)
k : โจ ๐ โฉ โ โจ ๐โ รแตแถแตแต ๐โ โฉ
k = prโ k-bundled
k-is-continuous : is-continuous ๐ (๐โ รแตแถแตแต ๐โ) k
k-is-continuous = prโ k-bundled
k-is-unique : is-central _ (k , k-is-continuous ,
(ฮป x โ refl) , (ฮป x โ refl))
k-is-unique (h , h-cont , h-eqf , h-eqg) =
to-subtype-๏ผ (ฮป j โ ร-is-prop
(being-continuous-is-prop ๐ (๐โ รแตแถแตแต ๐โ) j)
(ร-is-prop (ฮ -is-prop fe (ฮป e โ sethood ๐โ))
(ฮ -is-prop fe (ฮป e โ sethood ๐โ))))
(dfunext fe (ฮป e โ k e ๏ผโจ refl โฉ
f e , g e ๏ผโจ (eq e ) โปยน โฉ
h e โ))
where
eq : (e : โจ ๐ โฉ) โ h e ๏ผ f e , g e
eq e = apโ _,_ (h-eqf e) (h-eqg e)
\end{code}