├── .gitignore ├── README.org ├── SDT.agda └── synthetic-domain-theory.agda-lib /.gitignore: -------------------------------------------------------------------------------- 1 | _build/ 2 | -------------------------------------------------------------------------------- /README.org: -------------------------------------------------------------------------------- 1 | * Synthetic Domain Theory 2 | 3 | An experiment in doing synthetic domain theory in cubical agda. 4 | 5 | ** References 6 | 7 | We are mostly following Fiore and Plotkin's [[https://homepages.inf.ed.ac.uk/gdp/publications/ADT_and_SDT.pdf][An Extension of Models of 8 | Axiomatic Domain Theory to Models of Synthetic Domain Theory]]. 9 | -------------------------------------------------------------------------------- /SDT.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --guardedness #-} 2 | module SDT where 3 | 4 | open import Cubical.Foundations.Prelude 5 | open import Cubical.Foundations.HLevels 6 | open import Cubical.Foundations.Isomorphism 7 | open import Cubical.Foundations.Structure 8 | open import Cubical.Foundations.Univalence 9 | 10 | open import Cubical.Data.Empty 11 | open import Cubical.Data.Nat 12 | open import Cubical.Data.Sigma 13 | 14 | open import Cubical.Categories.Category 15 | open import Cubical.Categories.Constructions.FullSubcategory 16 | open import Cubical.Categories.Functor 17 | open import Cubical.Categories.Instances.Sets 18 | 19 | private 20 | variable 21 | ℓ : Level 22 | 23 | record Dominance {ℓ} : Type (ℓ-suc ℓ) where 24 | field 25 | isSemiDecidableProp : Type ℓ → Type ℓ 26 | isSemiDecidableProp→isProp : ∀ X → isSemiDecidableProp X → isProp X 27 | isPropisSemiDecidableProp : ∀ X → isProp (isSemiDecidableProp X) 28 | 29 | hSDProp : Type (ℓ-suc ℓ) 30 | hSDProp = TypeWithStr ℓ isSemiDecidableProp 31 | 32 | isSetSDProp : isSet hSDProp 33 | isSetSDProp = λ (X , a) (Y , b) → 34 | isPropRetract (cong fst) 35 | (Σ≡Prop isPropisSemiDecidableProp) 36 | (section-Σ≡Prop isPropisSemiDecidableProp) 37 | (isOfHLevel≡ 1 (isSemiDecidableProp→isProp X a) (isSemiDecidableProp→isProp Y b)) 38 | 39 | -- hSDPropExt : ∀ {A B} → isSemiDecidableProp A → isSemiDecidableProp B → (A → B) → (B → A) → A ≡ B 40 | -- hSDPropExt = {!!} 41 | 42 | field 43 | isSemiDecidable⊤ : isSemiDecidableProp Unit* 44 | isSemiDecidableΣ : ∀ {A B} → isSemiDecidableProp A → ((x : A) → isSemiDecidableProp (B x)) → isSemiDecidableProp (Σ A B) 45 | 46 | -- Domains wrt the dominance must live in the next universe in a 47 | -- predicative theory to accommoddate this notion of lifting. 48 | -- related: https://arxiv.org/abs/2008.01422 49 | module LiftMonad (ΣΣ : Dominance {ℓ}) where 50 | module ΣΣ = Dominance ΣΣ 51 | open ΣΣ 52 | L : Type (ℓ-suc ℓ) → Type (ℓ-suc ℓ) 53 | L X = Σ[ ϕ ∈ hSDProp ] (⟨ ϕ ⟩ → X) 54 | 55 | when_,_ : ∀ {X} → (ϕ : hSDProp) → (⟨ ϕ ⟩ → X) → L X 56 | when ϕ , x~ = ϕ , x~ 57 | 58 | supp : ∀ {X} → L X → hSDProp 59 | supp l = fst l 60 | 61 | elt : ∀ {X} → (l : L X) → ⟨ supp l ⟩ → X 62 | elt l p = snd l p 63 | 64 | isSetL : ∀ {A} → isSet A → isSet (L A) 65 | isSetL {A} a = isSetΣ ΣΣ.isSetSDProp (λ _ → isSet→ a) 66 | 67 | η : ∀ {X} → X → L X 68 | η x = when (Unit* , isSemiDecidable⊤) , λ tt → x 69 | 70 | ext : ∀ {X Y} → (X → L Y) → L X → L Y 71 | ext k l = when ((Σ ⟨ supp l ⟩ (λ p → ⟨ supp (k (elt l p)) ⟩)) , isSemiDecidableΣ (str (supp l)) λ p → str (supp (k (elt l p)))) , 72 | λ p → elt (k (elt l (fst p))) (snd p) 73 | 74 | map : ∀ {X Y} → (X → Y) → L X → L Y 75 | map f x~ = (supp x~) , (λ p → f (elt x~ p)) 76 | 77 | data ω : Type (ℓ-suc ℓ) where 78 | think : L ω → ω 79 | 80 | foldω : ∀ {X} → (L X → X) → ω → X 81 | foldω f (think x) = f ((fst x) , (λ p → foldω f (snd x p))) 82 | 83 | ω-is-initial : ∀ {X} → isSet X → (f : L X → X) → ∃![ h ∈ (ω → X) ] (∀ (l : L ω) → h (think l) ≡ f (map h l)) 84 | ω-is-initial {X} a f = 85 | uniqueExists (foldω f) 86 | (λ l → refl) 87 | (λ a' → isPropΠ λ x → a (a' (think x)) (f (map a' x))) 88 | (λ h hyp i o → lem h hyp o i) 89 | where lem : (h : ω → X) → ((x : L ω) → h (think x) ≡ f (map h x)) → ∀ o → foldω f o ≡ h o 90 | lem h hyp (think x) = 91 | f (fst x , (λ p → foldω f (snd x p))) ≡⟨ (λ i → f ((fst x) , (λ p → lem h hyp (snd x p) i))) ⟩ 92 | f (map h x) ≡⟨ sym (hyp x) ⟩ 93 | h (think x) ∎ 94 | 95 | -- final algebra 96 | record ω+ : Type (ℓ-suc ℓ) where 97 | coinductive 98 | field 99 | prj : L ω+ 100 | 101 | -- TODO 102 | -- isSetω+ : isSet ω+ 103 | -- isSetω+ x y = λ x₁ y₁ → {!!} 104 | 105 | open ω+ 106 | unfoldω+ : ∀ {X} → (X → L X) → X → ω+ 107 | unfoldω+ g x .prj = supp (g x) , λ p → unfoldω+ g (elt (g x) p) 108 | 109 | Ω : ω+ 110 | Ω = unfoldω+ (λ x → (Unit* , isSemiDecidable⊤) , λ _ → tt*) (lift tt) 111 | 112 | -- TODO 113 | -- ω+-is-final : ∀ {X} → isSet X → (g : X → L X) → ∃![ h ∈ (X → ω+) ] (∀ x → ω+.prj (h x) ≡ map h (g x)) 114 | -- ω+-is-final = {!!} 115 | 116 | ω→ω+ : ω → ω+ 117 | ω→ω+ = unfoldω+ (foldω (map think)) 118 | 119 | hasLimit : ∀ {X : Type (ℓ-suc ℓ)} → (ω → X) → Type (ℓ-suc ℓ) 120 | hasLimit {X} chainX = ∃![ limChain ∈ (ω+ → X) ] chainX ≡ (λ x → limChain (ω→ω+ x)) 121 | 122 | limitPt : ∀ {X : Type (ℓ-suc ℓ)} → (ω+ → X) → X 123 | limitPt limChain = limChain Ω 124 | 125 | isPropHasLimit : ∀ {X} → (chain : ω → X) → isProp (hasLimit chain) 126 | isPropHasLimit chain = isProp∃! 127 | 128 | isComplete : Type (ℓ-suc ℓ) → Type (ℓ-suc ℓ) 129 | isComplete X = ∀ (chain : ω → X) → hasLimit chain 130 | 131 | isPropIsComplete : ∀ {X} → isProp (isComplete X) 132 | isPropIsComplete = isPropΠ isPropHasLimit 133 | 134 | isWellComplete : Type (ℓ-suc ℓ) → Type (ℓ-suc ℓ) 135 | isWellComplete X = isComplete (L X) 136 | 137 | 138 | 139 | isPredomain : Type (ℓ-suc ℓ) → Type (ℓ-suc ℓ) 140 | isPredomain X = isSet X × isWellComplete X 141 | 142 | Predomain : Type (ℓ-suc (ℓ-suc ℓ)) 143 | Predomain = TypeWithStr (ℓ-suc ℓ) isPredomain 144 | 145 | PREDOMAIN : Category _ _ 146 | PREDOMAIN = FullSubcategory (SET (ℓ-suc ℓ)) λ x → isComplete ⟨ x ⟩ 147 | 148 | module SDT (ΣΣ : Dominance {ℓ}) 149 | (⊥-isSemiDecidable : Dominance.isSemiDecidableProp ΣΣ ⊥*) 150 | (Σ-is-complete : LiftMonad.isComplete ΣΣ (Dominance.hSDProp ΣΣ)) 151 | where 152 | -- Do we need another axiom? 153 | 154 | -- The last axiom in Fiore-Rosolini is that the lifting functor L 155 | -- has "rank", meaning it preserves κ-filtered colimits for some 156 | -- regular cardinal κ. 157 | 158 | -- but they say it is also sufficient that L preserve reflexive 159 | -- coequalizers which seems maybe possible to prove? 160 | 161 | module ΣΣ = Dominance ΣΣ 162 | open ΣΣ 163 | open LiftMonad ΣΣ 164 | 165 | -- | TODO: 166 | -- | 1. Prove Σ is *well*-complete, i.e., a predomain 167 | -- | 2. Show L is a monad on PREDOMAIN 168 | -- | 3. Prove Well-complete objects are complete 169 | -- | 4. Construct DOMAIN (w/ strict maps) as the EM category of L 170 | -- | 5. Construct a fixed point combinator for domains 171 | 172 | L-unit-R : ∀ {X} → (l : L X) → ext {X} η l ≡ l 173 | L-unit-R l = {!!} 174 | 175 | L-unit-L : ∀ {X Y} → (x : X) → (k : X → L Y) → ext k (η x) ≡ k x 176 | L-unit-L x k = {!!} 177 | 178 | L-assoc : ∀ {X Y Z} (l : L X) (k : X → L Y) (h : Y → L Z) → ext (λ x → ext h (k x)) l ≡ ext h (ext k l) 179 | L-assoc = {!!} 180 | 181 | L0≡1 : L ⊥* ≡ Unit* 182 | L0≡1 = ua (isoToEquiv (iso (λ x → lift tt) (λ x → (⊥* , ⊥-isSemiDecidable) , λ lifted → Cubical.Data.Empty.elim (lower lifted)) (λ b → refl) λ a → Σ≡Prop (λ x → isProp→ isProp⊥*) (Σ≡Prop isPropisSemiDecidableProp (ua (uninhabEquiv lower λ x → lower (snd a x)))))) 183 | 184 | -- The information ordering 185 | _⊑_ : ∀ {X : Type (ℓ-suc ℓ)} → X → X → Type (ℓ-suc ℓ) 186 | _⊑_ {X} x y = (∀ (ϕ : X → hSDProp) → ⟨ ϕ x ⟩ → ⟨ ϕ y ⟩) 187 | 188 | isProp⊑ : ∀ {X} (x y : X) → isProp (x ⊑ y) 189 | isProp⊑ x y = isPropΠ (λ ϕ → isProp→ (isSemiDecidableProp→isProp _ (str (ϕ y)))) 190 | 191 | refl⊑ : ∀ {X} (x : X) → x ⊑ x 192 | refl⊑ x = λ ϕ x₁ → x₁ 193 | 194 | trans⊑ : ∀ {X}{x₁ x₂ x₃ : X} → x₁ ⊑ x₂ → x₂ ⊑ x₃ → x₁ ⊑ x₃ 195 | trans⊑ x12 x23 = λ ϕ p → x23 ϕ (x12 ϕ p) 196 | 197 | all-functions-are-monotone : ∀ {X Y} (f : X → Y) → ∀ {x}{x'} → x ⊑ x' → f x ⊑ f x' 198 | all-functions-are-monotone f {x}{x'} x⊑x' = λ ϕ p → x⊑x' (λ x₁ → ϕ (f x₁)) p 199 | 200 | _⊑→_ : ∀ {X Y : Type (ℓ-suc ℓ)} (f g : X → Y) → Type (ℓ-suc ℓ) 201 | _⊑→_ {X}{Y} f g = ∀ x → f x ⊑ g x 202 | 203 | record _◃_ (X Y : Predomain) : Type (ℓ-suc ℓ) where 204 | field 205 | embed : ⟨ X ⟩ → ⟨ Y ⟩ 206 | project : ⟨ Y ⟩ → ⟨ X ⟩ 207 | retraction : ∀ x → project (embed x) ≡ x 208 | projection : ∀ y → y ⊑ embed (project y) 209 | -------------------------------------------------------------------------------- /synthetic-domain-theory.agda-lib: -------------------------------------------------------------------------------- 1 | name: synthetic-domain-theory 2 | include: . 3 | depend: cubical 4 | flags: --cubical --no-import-sorts --------------------------------------------------------------------------------