├── .gitignore ├── FormalTopology ├── Makefile ├── generate_html.sh ├── Main.lagda.md ├── NegriFormalTopology.lagda.md ├── Prenucleus.lagda.md ├── ProductTopology.lagda.md ├── Cofinality.lagda.md ├── Compactness.lagda.md ├── Ideal.lagda.md ├── BaireSpace.lagda.md ├── InducedNucleus.lagda.md ├── FormalTopology.lagda.md ├── GaloisConnection.lagda.md ├── Automata.lagda.md ├── WayBelow.lagda.md ├── Cover.lagda.md ├── CoverFormsNucleus.lagda.md ├── AdditionalFrameTheorems.lagda.md ├── RightAdjoint.lagda.md ├── SnocList.lagda.md ├── Omega.lagda.md ├── Stone.lagda.md ├── HeytingImplication.lagda.md ├── Base.lagda.md ├── Basis.lagda.md ├── Index.lagda.md ├── Regular.lagda.md ├── CantorSpace.lagda.md ├── Nucleus.lagda.md ├── UniversalProperty.lagda.md ├── Spectral.lagda.md └── Sierpinski.lagda.md ├── .github └── workflows │ └── main.yml ├── README.md └── resources └── Agda.css /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | FormalTopology/html 3 | -------------------------------------------------------------------------------- /FormalTopology/Makefile: -------------------------------------------------------------------------------- 1 | all: html 2 | 3 | html: 4 | ./generate_html.sh 5 | 6 | .PHONY: 7 | clean: 8 | rm -rf html 9 | -------------------------------------------------------------------------------- /FormalTopology/generate_html.sh: -------------------------------------------------------------------------------- 1 | #!/bin/bash 2 | 3 | echo "Running Agda..." 4 | 5 | agda --safe --html --html-highlight=auto Main.lagda.md 6 | 7 | cp -f ../resources/Agda.css html/Agda.css 8 | 9 | cd html 10 | 11 | for f in `ls *.md`; do 12 | echo "Compiling Markdown: $f..." 13 | if [ $f == "FrameOfNuclei.md" ]; then 14 | echo "Handling the FrameOfNuclei module..." 15 | pandoc $f --css Agda.css --toc -o "$(basename --suffix='.md' $f).html" 16 | else 17 | pandoc $f -s --css Agda.css --toc -o "$(basename --suffix='.md' $f).html" 18 | fi 19 | done 20 | -------------------------------------------------------------------------------- /FormalTopology/Main.lagda.md: -------------------------------------------------------------------------------- 1 | ``` 2 | {-# OPTIONS --without-K --cubical --safe #-} 3 | 4 | module Main where 5 | 6 | open import Index 7 | open import Basis 8 | open import Poset 9 | open import Frame 10 | open import Nucleus 11 | open import Cover 12 | open import CoverFormsNucleus 13 | open import BaireSpace 14 | open import UniversalProperty 15 | open import ProductTopology 16 | open import Compactness 17 | open import CantorSpace 18 | open import Sierpinski 19 | open import GaloisConnection 20 | open import Regular 21 | open import PatchFrame 22 | open import Omega 23 | open import ClosedNuclei 24 | open import PatchFrameNucleusLemma 25 | ``` 26 | -------------------------------------------------------------------------------- /.github/workflows/main.yml: -------------------------------------------------------------------------------- 1 | on: [push] 2 | 3 | jobs: 4 | typechecking: 5 | runs-on: ubuntu-latest 6 | name: A job to typecheck Agda code 7 | steps: 8 | - name: "Clone repository" 9 | uses: actions/checkout@v2 10 | - name: Typecheck 11 | id: typecheck 12 | uses: ayberkt/agda-github-action@master 13 | with: 14 | main-file: Main.lagda.md 15 | source-dir: FormalTopology 16 | - run: pwd 17 | - run: ls 18 | - run: sudo cp -f resources/Agda.css FormalTopology/html/Agda.css 19 | - name: Upload HTML 20 | id: html-upload 21 | if: github.ref == 'refs/heads/master' 22 | uses: actions/upload-artifact@v1 23 | with: 24 | name: html 25 | path: FormalTopology/html 26 | - name: Deploy to GitHub Pages 27 | id: deploy 28 | if: github.ref == 'refs/heads/master' 29 | uses: JamesIves/github-pages-deploy-action@releases/v3 30 | with: 31 | ACCESS_TOKEN: ${{ secrets.ACCESS_TOKEN }} 32 | BRANCH: gh-pages 33 | FOLDER: FormalTopology/html -------------------------------------------------------------------------------- /FormalTopology/NegriFormalTopology.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Heyting Implication in a Frame 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 16 | 17 | ## Definition 18 | 19 | ```agda 20 | record FormalTopologyStr₀ (P : Poset 𝓤 𝓥) : (𝓤 ⁺ ∨ 𝓥) ̇ where 21 | field 22 | _◁_ : ∣ P ∣ₚ → (∣ P ∣ₚ → hProp 𝓤) → hProp 𝓤 23 | 24 | _◁◁_ : 𝒫 ∣ P ∣ₚ → 𝒫 ∣ P ∣ₚ → hProp 𝓤 25 | U ◁◁ V = ∀[ u ∶ ∣ P ∣ₚ ] u ∈ U ⇒ u ◁ V 26 | 27 | _∧_ : (∣ P ∣ₚ → hProp 𝓥) → (∣ P ∣ₚ → hProp 𝓥) → ∣ P ∣ₚ → hProp (𝓤 ∨ 𝓥) 28 | (U ∧ V) x = 29 | ∥ Σ[ (u , v) ∈ ∣ P ∣ₚ × ∣ P ∣ₚ ] 30 | [ U u ] × [ V v ] × [ x ⊑[ P ] u ] × [ x ⊑[ P ] v ] ∥Ω 31 | 32 | field 33 | ◁-refl : (U : 𝒫 ∣ P ∣ₚ) → [ U ⊆ (λ - → - ◁ U) ] 34 | ◁-trans : (a : ∣ P ∣ₚ) (U V : 𝒫 ∣ P ∣ₚ) 35 | → [ a ◁ U ] → [ U ◁◁ V ] → [ a ◁ V ] 36 | ◁-left : (a b : ∣ P ∣ₚ) 37 | → [ b ⊑[ P ] a ] → [ a ◁ U ] → [ b ◁ U ] 38 | -- ◁-right : (U : 𝒫 ∣ P ∣ₚ) (a : ∣ P ∣ₚ) 39 | -- → [ a ◁ U ] → [ a ◁ V ] → [ a ◁ (U ∧ V) ] 40 | ``` 41 | -------------------------------------------------------------------------------- /FormalTopology/Prenucleus.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module Prenucleus where 5 | 6 | open import Cubical.Core.Everything hiding (_∧_) 7 | 8 | open import Poset 9 | open import Frame 10 | open import Cubical.Functions.Logic hiding (_⊓_) 11 | open import Cubical.Foundations.Prelude using (refl; sym; cong; _≡⟨_⟩_; _∎) 12 | open import Cubical.Data.Sigma using (Σ≡Prop; _×_) 13 | open import Cubical.Foundations.Function using (const; _∘_) 14 | open import Basis renaming (_⊓_ to _∧_) 15 | ``` 16 | 17 | ```agda 18 | isPrenuclear : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → ∣ L ∣F) → Type (ℓ-max ℓ₀ ℓ₁) 19 | isPrenuclear L j = N₀ × N₁ 20 | where 21 | N₀ = (x y : ∣ L ∣F) → j (x ⊓[ L ] y) ≡ (j x) ⊓[ L ] (j y) 22 | N₁ = (x : ∣ L ∣F) → [ x ⊑[ pos L ] (j x) ] 23 | ``` 24 | 25 | ```agda 26 | Prenucleus : Frame ℓ₀ ℓ₁ ℓ₂ → Type (ℓ-max ℓ₀ ℓ₁) 27 | Prenucleus L = Σ (∣ L ∣F → ∣ L ∣F) (isPrenuclear L) 28 | ``` 29 | 30 | Every prenucleus is monotonic. 31 | 32 | ```agda 33 | monop : (L : Frame ℓ₀ ℓ₁ ℓ₂) ((j , _) : Prenucleus L) 34 | → (x y : ∣ L ∣F) → [ x ⊑[ pos L ] y ] → [ (j x) ⊑[ pos L ] (j y) ] 35 | monop L (j , N₀ , _) x y x⊑y = 36 | j x ⊑⟨ ≡⇒⊑ (pos L) (cong j (x⊑y⇒x=x∧y L x⊑y)) ⟩ 37 | j (x ⊓[ L ] y) ⊑⟨ ≡⇒⊑ (pos L) (N₀ x y) ⟩ 38 | j x ⊓[ L ] j y ⊑⟨ ⊓[ L ]-lower₁ (j x) (j y) ⟩ 39 | j y ■ 40 | where 41 | open PosetReasoning (pos L) 42 | ``` 43 | -------------------------------------------------------------------------------- /FormalTopology/ProductTopology.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Product Topology 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | module ProductTopology where 9 | 10 | open import Basis 11 | open import Poset 12 | open import FormalTopology 13 | 14 | module _ (𝔉 𝔊 : FormalTopology ℓ₀ ℓ₀) where 15 | P = pos 𝔉 16 | Q = pos 𝔊 17 | 18 | 𝔉×𝔊 : FormalTopology ℓ₀ ℓ₀ 19 | 𝔉×𝔊 = P ×ₚ Q , ×-IS , ×-mono , ×-sim 20 | where 21 | ×-exp : ∣ P ×ₚ Q ∣ₚ → Type ℓ₀ 22 | ×-exp (a₀ , a₁) = exp 𝔉 a₀ ⊎ exp 𝔊 a₁ 23 | 24 | ×-out : {a : ∣ P ×ₚ Q ∣ₚ} → ×-exp a → Type ℓ₀ 25 | ×-out (inl b) = outcome 𝔉 b 26 | ×-out (inr b) = outcome 𝔊 b 27 | 28 | ×-next : {a : ∣ P ×ₚ Q ∣ₚ} {b : ×-exp a} → ×-out {a = a} b → ∣ P ×ₚ Q ∣ₚ 29 | ×-next {a = (a₀ , a₁)} {b = inl b} c = (next 𝔉 c) , a₁ 30 | ×-next {a = (a₀ , a₁)} {b = inr b} c = a₀ , (next 𝔊 c) 31 | 32 | ×-IS : InteractionStr ∣ P ×ₚ Q ∣ₚ 33 | ×-IS = ×-exp , ×-out , λ {a} {b} c → ×-next {b = b} c 34 | 35 | ×-mono : hasMono (P ×ₚ Q) ×-IS 36 | ×-mono (a₀ , a₁) (inl b) c = (mono 𝔉 a₀ b c) , ⊑[ Q ]-refl a₁ 37 | ×-mono (a₀ , a₁) (inr b) c = (⊑[ P ]-refl a₀) , mono 𝔊 a₁ b c 38 | 39 | ×-sim : hasSimulation (P ×ₚ Q) ×-IS 40 | ×-sim (a₀ , a₁) (a , a′) (a₀⊑a , a₁⊑a′) b with b 41 | ... | inl b₀ = let (b₀′ , p) = sim 𝔉 _ _ a₀⊑a b₀ 42 | in inl b₀′ , λ c₀ → π₀ (p c₀) , π₁ (p c₀) , a₁⊑a′ 43 | ... | inr b₁ = let (b₁′ , p) = sim 𝔊 _ _ a₁⊑a′ b₁ 44 | in inr b₁′ , λ c₁ → π₀ (p c₁) , a₀⊑a , π₁ (p c₁) 45 | -------------------------------------------------------------------------------- /FormalTopology/Cofinality.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | open import Basis 5 | open import Poset 6 | open import Frame 7 | 8 | module Cofinality (F : Frame ℓ₀ ℓ₁ ℓ₂) where 9 | ``` 10 | 11 | Definition 12 | ========== 13 | 14 | ```agda 15 | _cofinal-in_ : Fam ℓ₂ ∣ F ∣F → Fam ℓ₂ ∣ F ∣F → Type (ℓ-max ℓ₁ ℓ₂) 16 | _cofinal-in_ U V = 17 | (i : index U) → Σ[ j ∈ index V ] [ (U $ i) ⊑[ pos F ] (V $ j) ] 18 | ``` 19 | 20 | Bi-cofinality 21 | ============= 22 | 23 | ```agda 24 | bicofinal : Fam ℓ₂ ∣ F ∣F → Fam ℓ₂ ∣ F ∣F → Type (ℓ-max ℓ₁ ℓ₂) 25 | bicofinal U V = (U cofinal-in V) × (V cofinal-in U) 26 | ``` 27 | 28 | Joins of bi-cofinal families are equal 29 | ====================================== 30 | 31 | ```agda 32 | bicofinal→same-join-lemma : (U V : Fam ℓ₂ ∣ F ∣F) 33 | → U cofinal-in V → [ ⋁[ F ] U ⊑[ pos F ] ⋁[ F ] V ] 34 | bicofinal→same-join-lemma U V U-cf-V = ⋁[ F ]-least U (⋁[ F ] V) nts where 35 | 36 | open PosetReasoning (pos F) 37 | 38 | nts : [ ∀[ u ε U ] (u ⊑[ pos F ] (⋁[ F ] V)) ] 39 | nts u (i , eq) = subst (λ - → [ - ⊑[ pos F ] _ ]) eq nts₀ where 40 | 41 | j = π₀ (U-cf-V i) 42 | 43 | nts₀ : [ (U $ i) ⊑[ pos F ] (⋁[ F ] V) ] 44 | nts₀ = U $ i ⊑⟨ π₁ (U-cf-V i) ⟩ 45 | V $ j ⊑⟨ ⋁[ F ]-upper _ _ (j , refl) ⟩ 46 | ⋁[ F ] V ■ 47 | 48 | bicofinal→same-join : (U V : Fam ℓ₂ ∣ F ∣F) 49 | → bicofinal U V → ⋁[ F ] U ≡ ⋁[ F ] V 50 | bicofinal→same-join U V (φ , ψ) = ⊑[ pos F ]-antisym _ _ ⋁U⊑⋁V ⋁V⊑⋁U where 51 | 52 | ⋁U⊑⋁V : [ (⋁[ F ] U) ⊑[ pos F ] (⋁[ F ] V) ] 53 | ⋁U⊑⋁V = bicofinal→same-join-lemma U V φ 54 | 55 | ⋁V⊑⋁U : [ (⋁[ F ] V) ⊑[ pos F ] (⋁[ F ] U) ] 56 | ⋁V⊑⋁U = bicofinal→same-join-lemma V U ψ 57 | ``` 58 | -------------------------------------------------------------------------------- /FormalTopology/Compactness.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module Compactness where 5 | 6 | open import Basis hiding (A; B) 7 | open import Cubical.Data.List hiding ([_]) 8 | open import Cubical.Data.Nat using (ℕ) 9 | open import Poset 10 | open import FormalTopology 11 | open import Cover 12 | open import Frame renaming (pos to posf) 13 | ``` 14 | 15 | # Compactness for formal topologies 16 | 17 | ```agda 18 | module _ (F : FormalTopology ℓ₀ ℓ₀) where 19 | 20 | open CoverFromFormalTopology F using (_◁_) 21 | 22 | private 23 | A = stage F 24 | B = exp F 25 | C = outcome F 26 | d = next F 27 | 28 | down : List A → 𝒫 A 29 | down [] = λ _ → bot ℓ₀ 30 | down (x ∷ xs) = λ y → ∥ [ y ⊑[ pos F ] x ] ⊎ [ y ∈ down xs ] ∥ , ∥∥-prop _ 31 | 32 | isCompact : Type (ℓ-suc ℓ₀) 33 | isCompact = (a : A) (U : 𝒫 A) (U-dc : [ isDownwardsClosed (pos F) U ]) → 34 | a ◁ U → ∥ Σ[ as ∈ List A ] (a ◁ down as) × [ down as ⊆ U ] ∥ 35 | ``` 36 | 37 | ## Compactness for frames 38 | 39 | Johnstone's definition of a compact frame is simply a frame whose top element is 40 | finite. Therefore, let us start by writing down Johnstone's notion of 41 | finiteness (Lemma 3.1.ii, pg. 63). 42 | 43 | ```agda 44 | isFinite₂ : (F : Frame ℓ₀ ℓ₁ ℓ₂) 45 | → ∣ F ∣F → Type (ℓ-max (ℓ-max ℓ₀ ℓ₁) (ℓ-suc ℓ₂)) 46 | isFinite₂ {ℓ₂ = ℓ₂} F x = 47 | (U : Fam ℓ₂ ∣ F ∣F) → 48 | [ isDirected (Frame.pos F) U ] → 49 | x ≡ ⋁[ F ] U → ∥ Σ[ i ∈ index U ] [ x ⊑[ Frame.pos F ] (U $ i) ] ∥ 50 | ``` 51 | 52 | A frame is compact iff its top element is finite. 53 | 54 | ```agda 55 | isCompactFrame : (F : Frame ℓ₀ ℓ₁ ℓ₂) → hProp (ℓ-max (ℓ-max ℓ₀ ℓ₁) (ℓ-suc ℓ₂)) 56 | isCompactFrame F = isFinite₂ F ⊤[ F ] , isPropΠ3 (λ x y z → ∥∥-prop _) 57 | ``` 58 | -------------------------------------------------------------------------------- /FormalTopology/Ideal.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module Ideal where 5 | 6 | open import Basis 7 | open import Poset hiding (isDownwardsClosed) 8 | open import Frame 9 | ``` 10 | 11 | ```agda 12 | ↓ : (L : Frame 𝓤 𝓥 𝓦) → ∣ L ∣F → ∣ L ∣F → hProp 𝓥 13 | ↓ L x y = y ⊑[ pos L ] x 14 | 15 | isDownwardsClosed : (L : Frame 𝓤 𝓥 𝓦) → (∣ L ∣F → hProp 𝓥) → hProp (𝓤 ∨ 𝓥) 16 | isDownwardsClosed L U = 17 | ∀[ x ∶ ∣ L ∣F ] U x ⇒ (∀[ y ∶ ∣ L ∣F ] y ⊑[ pos L ] x ⇒ U y) 18 | 19 | isUpwardsDirected : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → hProp ℓ₁) → hProp _ 20 | isUpwardsDirected L U = 21 | U ⊥[ L ] ⊓ (∀[ x ∶ ∣ L ∣F ] ∀[ y ∶ ∣ L ∣F ] U x ⇒ U y ⇒ U (x ∨[ L ] y)) 22 | 23 | isIdeal : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → hProp ℓ₁) → hProp _ 24 | isIdeal L U = isDownwardsClosed L U ⊓ isUpwardsDirected L U 25 | 26 | isIdeal′ : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → hProp ℓ₁) → hProp _ 27 | isIdeal′ L U = isDownwardsClosed L U ⊓ ϕ 28 | where 29 | ϕ : hProp _ 30 | ϕ = ∀[ x ∶ ∣ L ∣F ] ∀[ y ∶ ∣ L ∣F ] U x ⇒ U y ⇒ U (x ∨[ L ] y) 31 | 32 | ideal→ideal′ : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (U : ∣ L ∣F → hProp ℓ₁) 33 | → [ isIdeal L U ] → [ isIdeal′ L U ] 34 | ideal→ideal′ L U (dc , ud) = dc , † 35 | where 36 | † : [ ∀[ x ∶ ∣ L ∣F ] ∀[ y ∶ ∣ L ∣F ] U x ⇒ U y ⇒ U (x ∨[ L ] y) ] 37 | † x y x∈U y∈U = ? 38 | ``` 39 | 40 | ```agda 41 | ↓-ideal : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (x : ∣ L ∣F) → [ isIdeal L (↓ L x) ] 42 | ↓-ideal L x = dc , ud 43 | where 44 | open PosetReasoning (pos L) 45 | 46 | dc : [ isDownwardsClosed L (↓ L x) ] 47 | dc y y⊑x z z⊑y = z ⊑⟨ z⊑y ⟩ y ⊑⟨ y⊑x ⟩ x ■ 48 | 49 | ud : [ isUpwardsDirected L (↓ L x) ] 50 | ud = ⊥[ L ]-bottom x , λ y z y∈U z∈U → ⊔[ L ]-least y z x y∈U z∈U 51 | ``` 52 | 53 | ```agda 54 | isAPrincipalIdeal : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → hProp ℓ₁) → Type _ 55 | isAPrincipalIdeal L U = Σ[ x ∈ ∣ L ∣F ] U ≡ ↓ L x 56 | ``` 57 | -------------------------------------------------------------------------------- /FormalTopology/BaireSpace.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Baire Space 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | open import Cubical.Core.Everything 9 | open import Cubical.Data.Nat using (ℕ) 10 | open import Cubical.Foundations.Prelude using (isProp) 11 | open import Cubical.Data.Sigma using (_×_; _,_) 12 | 13 | data 𝔻 : Type₀ where 14 | [] : 𝔻 15 | _⌢_ : 𝔻 → ℕ → 𝔻 16 | 17 | IsDC : (𝔻 → Type₀) → Type₀ 18 | IsDC P = (σ : 𝔻) (n : ℕ) → P σ → P (σ ⌢ n) 19 | 20 | data _◀_ (σ : 𝔻) (P : 𝔻 → Type₀) : Type₀ where 21 | dir : P σ → σ ◀ P 22 | branch : ((n : ℕ) → (σ ⌢ n) ◀ P) → σ ◀ P 23 | squash : (φ ψ : σ ◀ P) → φ ≡ ψ 24 | 25 | variable 26 | m n : ℕ; σ τ : 𝔻; P Q : 𝔻 → Type₀ 27 | 28 | ◀-prop : isProp (σ ◀ P) 29 | ◀-prop = squash 30 | 31 | δ : σ ◀ P → ((v : 𝔻) → P v → v ◀ Q) → σ ◀ Q 32 | δ (dir uεP) φ = φ _ uεP 33 | δ (branch f) φ = branch (λ n → δ (f n) φ) 34 | δ (squash u◀P₀ u◀P₁ i) φ = squash (δ u◀P₀ φ) (δ u◀P₁ φ) i 35 | 36 | δ-corollary : σ ◀ (λ - → - ◀ P) → σ ◀ P 37 | δ-corollary u◀◀P = δ u◀◀P (λ _ v◀P → v◀P) 38 | 39 | ζ : (n : ℕ) → IsDC P → σ ◀ P → (σ ⌢ n) ◀ P 40 | ζ n dc (dir σεP) = dir (dc _ n σεP) 41 | ζ n dc (branch f) = branch λ m → ζ m dc (f n) 42 | ζ n dc (squash σ◀P σ◀P′ i) = squash (ζ n dc σ◀P) (ζ n dc σ◀P′) i 43 | 44 | ζ′ : IsDC P → IsDC (λ - → - ◀ P) 45 | ζ′ P-dc σ n σ◀P = ζ n P-dc σ◀P 46 | 47 | lemma : IsDC P → P σ → σ ◀ Q → σ ◀ (λ - → P - × Q -) 48 | lemma P-dc σεP (dir σεQ) = dir (σεP , σεQ) 49 | lemma P-dc σεP (branch f) = branch (λ n → lemma P-dc (P-dc _ n σεP) (f n)) 50 | lemma P-dc σεP (squash σ◀Q σ◀Q′ i) = squash (lemma P-dc σεP σ◀Q) (lemma P-dc σεP σ◀Q′) i 51 | 52 | -- Meet preservation. 53 | mp : IsDC P → IsDC Q → σ ◀ P → σ ◀ Q → σ ◀ (λ - → P - × Q -) 54 | mp P-dc Q-dc (dir σεP) σ◀Q = lemma P-dc σεP σ◀Q 55 | mp P-dc Q-dc (branch f) σ◀Q = branch (λ n → mp P-dc Q-dc (f n) (ζ n Q-dc σ◀Q)) 56 | mp P-dc Q-dc (squash σ◀P σ◀P′ i) σ◀Q = squash (mp P-dc Q-dc σ◀P σ◀Q) (mp P-dc Q-dc σ◀P′ σ◀Q) i 57 | ``` 58 | -------------------------------------------------------------------------------- /FormalTopology/InducedNucleus.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Induced Nucleus 3 | --- 4 | 5 | 20 | 21 | ```agda 22 | induced₀ : (A : Frame 𝓤 𝓥 𝓥) (B : Frame 𝓤′ 𝓥 𝓥) 23 | → (ℬ : Fam 𝓥 ∣ A ∣F) 24 | → isDirBasisFor A ℬ 25 | → (f : A ─f→ B) → ∣ A ∣F → ∣ A ∣F 26 | induced₀ A B ℬ dir 𝒻@((f , _) , _) = 𝒻 ⋆ ∘ f 27 | where 28 | open AdjointFunctorTheorem A B (ℬ , dir) renaming (RA-of-homo to _⋆) 29 | ``` 30 | 31 | ```agda 32 | induced₀-is-nuclear : (A : Frame 𝓤 𝓥 𝓥) (B : Frame 𝓤′ 𝓥 𝓥) 33 | → (ℬ : Fam 𝓥 ∣ A ∣F) 34 | → (b : isDirBasisFor A ℬ) 35 | → (h : A ─f→ B) 36 | → isNuclear A (induced₀ A B ℬ b h) 37 | induced₀-is-nuclear A B ℬ b 𝒻@((f , _) , _ , f∧ , f⋁) = 38 | ind-𝓃₀ , ind-𝓃₁ , ind-𝓃₂ 39 | where 40 | open AdjointFunctorTheorem A B (ℬ , b) renaming (RA-of-homo to _⋆) 41 | open PosetReasoning (pos A) 42 | 43 | f⋆ : ∣ B ∣F → ∣ A ∣F 44 | f⋆ = _⋆ 𝒻 45 | 46 | f⋆∘f : ∣ A ∣F → ∣ A ∣F 47 | f⋆∘f = induced₀ A B ℬ b 𝒻 48 | 49 | open GaloisConnection (pos A) (pos B) 50 | 51 | ind-𝓃₀ : (x y : ∣ A ∣F) → f⋆∘f (x ⊓[ A ] y) ≡ (f⋆∘f x) ⊓[ A ] (f⋆∘f y) 52 | ind-𝓃₀ x y = f⋆ (f (x ⊓[ A ] y)) ≡⟨ cong f⋆ (f∧ x y) ⟩ 53 | f⋆ (f x ⊓[ B ] f y) ≡⟨ RAPL (π₀ 𝒻) (sym ∘ f⋁) (f x) (f y) ⟩ 54 | f⋆ (f x) ⊓[ A ] f⋆ (f y) ∎ 55 | 56 | ind-𝓃₁ : (x : ∣ A ∣F) → [ x ⊑[ pos A ] f⋆∘f x ] 57 | ind-𝓃₁ x = π₀ (^*-RA (π₀ 𝒻) (sym ∘ f⋁) x (f x)) (⊑[ pos B ]-refl (f x)) 58 | 59 | ind-𝓃₂ : (x : ∣ A ∣F) → [ f⋆∘f (f⋆∘f x) ⊑[ pos A ] f⋆∘f x ] 60 | ind-𝓃₂ x = f⋆∘f (f⋆∘f x) ⊑⟨ ⊑[ pos A ]-refl _ ⟩ 61 | f⋆ (f (f⋆ (f x))) ⊑⟨ † ⟩ 62 | f⋆∘f x ■ 63 | where 64 | † = f⋆∘f-idempotent (π₀ 𝒻) (_^*ᴹ (π₀ 𝒻) (sym ∘ f⋁)) (^*-RA (π₀ 𝒻) (sym ∘ f⋁)) x 65 | ``` 66 | -------------------------------------------------------------------------------- /FormalTopology/FormalTopology.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Formal Topology 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | module FormalTopology where 9 | 10 | open import Basis 11 | open import Poset 12 | 13 | InteractionStr : (A : Type ℓ) → Type (ℓ-suc ℓ) 14 | InteractionStr {ℓ = ℓ} A = 15 | Σ[ B ∈ (A → Type ℓ) ] Σ[ C ∈ ({x : A} → B x → Type ℓ) ]({x : A} → {y : B x} → C y → A) 16 | 17 | InteractionSys : (ℓ : Level) → Type (ℓ-suc ℓ) 18 | InteractionSys ℓ = Σ (Type ℓ) InteractionStr 19 | 20 | state : InteractionSys ℓ → Type ℓ 21 | action : (D : InteractionSys ℓ) → state D → Type ℓ 22 | reaction : (D : InteractionSys ℓ) → {x : state D} → action D x → Type ℓ 23 | δ : (D : InteractionSys ℓ) {x : state D} {y : action D x} 24 | → reaction D y → state D 25 | 26 | state (A , _ , _ , _) = A 27 | action (_ , B , _ , _) = B 28 | reaction (_ , _ , C , _) = C 29 | δ (_ , _ , _ , d) = d 30 | 31 | hasMono : (P : Poset ℓ₀ ℓ₁) → InteractionStr ∣ P ∣ₚ → Type (ℓ-max ℓ₀ ℓ₁) 32 | hasMono P i = 33 | (a : state IS) (b : action IS a) (c : reaction IS b) → [ δ IS c ⊑[ P ] a ] 34 | where 35 | IS : InteractionSys _ 36 | IS = ∣ P ∣ₚ , i 37 | 38 | module _ (P : Poset ℓ₀ ℓ₁) (ℐ-str : InteractionStr ∣ P ∣ₚ) where 39 | ℐ : InteractionSys ℓ₀ 40 | ℐ = (∣ P ∣ₚ , ℐ-str) 41 | 42 | hasSimulation : Type (ℓ-max ℓ₀ ℓ₁) 43 | hasSimulation = 44 | (a′ a : ∣ P ∣ₚ) → [ a′ ⊑[ P ] a ] → 45 | (b : action ℐ a) → Σ[ b′ ∈ action ℐ a′ ] 46 | ((c′ : reaction ℐ b′) → Σ[ c ∈ reaction ℐ b ] [ δ ℐ c′ ⊑[ P ] δ ℐ c ]) 47 | 48 | FormalTopology : (ℓ₀ ℓ₁ : Level) → Type (ℓ-max (ℓ-suc ℓ₀) (ℓ-suc ℓ₁)) 49 | FormalTopology ℓ₀ ℓ₁ = 50 | Σ[ P ∈ Poset ℓ₀ ℓ₁ ] Σ[ ℐ ∈ InteractionStr ∣ P ∣ₚ ] hasMono P ℐ × hasSimulation P ℐ 51 | 52 | pos : FormalTopology ℓ₀ ℓ₁ → Poset ℓ₀ ℓ₁ 53 | pos (P , _) = P 54 | 55 | IS : (ℱ : FormalTopology ℓ₀ ℓ₁) → InteractionStr ∣ pos ℱ ∣ₚ 56 | IS (_ , is , _) = is 57 | 58 | stage : FormalTopology ℓ₀ ℓ₁ → Type ℓ₀ 59 | stage (P , ℐ-str , _) = state (∣ P ∣ₚ , ℐ-str) 60 | 61 | exp : (ℱ : FormalTopology ℓ₀ ℓ₁) → stage ℱ → Type ℓ₀ 62 | exp (P , ℐ-str , _) = action (∣ P ∣ₚ , ℐ-str) 63 | 64 | outcome : (ℱ : FormalTopology ℓ₀ ℓ₁) → {a : stage ℱ} → exp ℱ a → Type ℓ₀ 65 | outcome (P , ℐ-str , _) = reaction (∣ P ∣ₚ , ℐ-str) 66 | 67 | next : (ℱ : FormalTopology ℓ₀ ℓ₁) {a : stage ℱ} {b : exp ℱ a} → outcome ℱ b → stage ℱ 68 | next (P , ℐ-str , _) = δ (∣ P ∣ₚ , ℐ-str) 69 | 70 | mono : (ℱ : FormalTopology ℓ₀ ℓ₁) → hasMono (pos ℱ) (IS ℱ) 71 | mono (_ , _ , φ , _) = φ 72 | 73 | sim : (ℱ : FormalTopology ℓ₀ ℓ₁) → hasSimulation (pos ℱ) (IS ℱ) 74 | sim (_ , _ , _ , φ) = φ 75 | ``` 76 | -------------------------------------------------------------------------------- /FormalTopology/GaloisConnection.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Galois Connection 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | module GaloisConnection where 9 | 10 | open import Cubical.Core.Everything 11 | open import Basis 12 | open import Poset 13 | open import Frame 14 | ``` 15 | 16 | The monotonic map that is right adjoint to the diagonal monotonic map. 17 | 18 | ```agda 19 | hasBinMeets : (P : Poset ℓ₀ ℓ₁) → Type (ℓ-max ℓ₀ ℓ₁) 20 | hasBinMeets P = Σ[ _⊓_ ∈ (∣ P ∣ₚ → ∣ P ∣ₚ → ∣ P ∣ₚ) ] [ isGLB P _⊓_ ] 21 | 22 | hasBinMeets′ : (P : Poset ℓ₀ ℓ₁) → Type (ℓ-max ℓ₀ ℓ₁) 23 | hasBinMeets′ P = Σ[ g ∈ (P ×ₚ P) ─m→ P ] [ (Δ P) ⊣ g ] 24 | where 25 | open GaloisConnection P (P ×ₚ P) 26 | 27 | hasBinMeets′-prop : (P : Poset ℓ₀ ℓ₁) → isProp (hasBinMeets′ P) 28 | hasBinMeets′-prop P (_⊓₀_ , Δ⊣⊓₀) (_⊓₁_ , Δ⊣⊓₁) = 29 | Σ≡Prop (λ g → isProp[] (Δ P ⊣ g)) (⊣-unique-right (Δ P) _⊓₀_ _⊓₁_ Δ⊣⊓₀ Δ⊣⊓₁) 30 | where 31 | open GaloisConnection P (P ×ₚ P) 32 | ``` 33 | 34 | ```agda 35 | bin-meets→bin-meets′ : (P : Poset ℓ₀ ℓ₁) → hasBinMeets P → hasBinMeets′ P 36 | bin-meets→bin-meets′ P meet = (f , f-mono) , f⊣Δ 37 | where 38 | open GaloisConnection P (P ×ₚ P) 39 | open PosetReasoning P 40 | 41 | f : ∣ P ×ₚ P ∣ₚ → ∣ P ∣ₚ 42 | f (x , y) = fst meet x y 43 | 44 | f-mono : isMonotonic (P ×ₚ P) P f 45 | f-mono (x₀ , x₁) (y₀ , y₁) (x₀⊑y₀ , x₁⊑y₁) = 46 | π₁ (π₁ meet) y₀ y₁ (f (x₀ , x₁)) (φ , ψ) 47 | where 48 | φ : [ f (x₀ , x₁) ⊑[ P ] y₀ ] 49 | φ = f (x₀ , x₁) ⊑⟨ fst (fst (snd meet) x₀ x₁) ⟩ x₀ ⊑⟨ x₀⊑y₀ ⟩ y₀ ■ 50 | 51 | ψ : [ f (x₀ , x₁) ⊑[ P ] y₁ ] 52 | ψ = f (x₀ , x₁) ⊑⟨ snd (fst (snd meet) x₀ x₁) ⟩ x₁ ⊑⟨ x₁⊑y₁ ⟩ y₁ ■ 53 | 54 | f⊣Δ : [ Δ P ⊣ (f , f-mono) ] 55 | f⊣Δ x (y₀ , y₁) = NTS₀ , NTS₁ 56 | where 57 | NTS₀ : [ (x , x) ⊑[ P ×ₚ P ] (y₀ , y₁) ⇒ x ⊑[ P ] f (y₀ , y₁) ] 58 | NTS₀ (x⊑y₀ , x⊑y₁) = π₁ (π₁ meet) y₀ y₁ x (x⊑y₀ , x⊑y₁) 59 | 60 | NTS₁ : [ x ⊑[ P ] f (y₀ , y₁) ] → [ (x , x) ⊑[ P ×ₚ P ] (y₀ , y₁) ] 61 | NTS₁ x⊑fy₀y₁ = x⊑y₀ , x⊑y₁ 62 | where 63 | x⊑y₀ : [ x ⊑[ P ] y₀ ] 64 | x⊑y₀ = x ⊑⟨ x⊑fy₀y₁ ⟩ f (y₀ , y₁) ⊑⟨ fst (fst (snd meet) y₀ y₁) ⟩ y₀ ■ 65 | 66 | x⊑y₁ : [ x ⊑[ P ] y₁ ] 67 | x⊑y₁ = x ⊑⟨ x⊑fy₀y₁ ⟩ f (y₀ , y₁) ⊑⟨ snd (fst (snd meet) y₀ y₁) ⟩ y₁ ■ 68 | ``` 69 | 70 | ```agda 71 | bin-meets′→bin-meets : (P : Poset ℓ₀ ℓ₁) 72 | → hasBinMeets′ P 73 | → hasBinMeets P 74 | bin-meets′→bin-meets P (g , Δ⊣g) = inf , inf-is-glb 75 | where 76 | inf : ∣ P ∣ₚ → ∣ P ∣ₚ → ∣ P ∣ₚ 77 | inf x y = g $ₘ (x , y) 78 | 79 | inf-is-glb : [ isGLB P inf ] 80 | inf-is-glb = NTS₀ , NTS₁ 81 | where 82 | NTS₀ : (x y : ∣ P ∣ₚ) → [ (inf x y) ⊑[ P ] x ⊓ (inf x y) ⊑[ P ] y ] 83 | NTS₀ x y = snd (Δ⊣g (inf x y) (x , y)) (⊑[ P ]-refl _) 84 | 85 | NTS₁ : (x y z : ∣ P ∣ₚ) → [ z ⊑[ P ] x ⊓ z ⊑[ P ] y ⇒ z ⊑[ P ] inf x y ] 86 | NTS₁ x y z (z⊑x , z⊑y) = π₀ (Δ⊣g z (x , y)) (z⊑x , z⊑y) 87 | ``` 88 | 89 | ```agda 90 | meets⇔meets′ : (P : Poset ℓ₀ ℓ₁) → hasBinMeets P ↔ hasBinMeets′ P 91 | meets⇔meets′ P = bin-meets→bin-meets′ P , bin-meets′→bin-meets P 92 | ``` 93 | 94 | -------------------------------------------------------------------------------- /FormalTopology/Automata.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Some Experiments on Automata 3 | --- 4 | 5 | 23 | 24 | ```agda 25 | isFinite : 𝓤 ̇ → 𝓤 ̇ 26 | isFinite A = Σ[ n ∈ ℕ ] Fin n ≃ A 27 | ``` 28 | 29 | 30 | ## Definition of DFA 31 | 32 | ```agda 33 | record DFA : 𝓤₁ ̇ where 34 | field 35 | Q : 𝓤₀ ̇ 36 | S : 𝓤₀ ̇ 37 | δ : Q → S → Q 38 | q₀ : Q 39 | F : Q → hProp 𝓤₀ 40 | 41 | field 42 | ∣Q∣ : ℕ 43 | Q-fin : Fin ∣Q∣ ≃ Q 44 | S-fin : isFinite S 45 | 46 | Q-set : isSet Q 47 | Q-set = subst isSet (ua Q-fin) isSetFin 48 | 49 | open DFA 50 | ``` 51 | 52 | ## Acceptance 53 | 54 | ```agda 55 | _accepts_at_ : (M : DFA) → List (M .S) → M .Q → hProp 𝓤₀ 56 | _accepts_at_ M [] x = M .F x 57 | _accepts_at_ M (a ∷ as) x = M accepts as at (M .δ x a) 58 | ``` 59 | 60 | ```agda 61 | _accepts_ : (M : DFA) → List (M .S) → hProp 𝓤₀ 62 | M accepts as = M accepts as at M .q₀ 63 | ``` 64 | 65 | ## The specialisation preorder 66 | 67 | ```agda 68 | _<<<_ : (M : DFA) → M .Q → M .Q → hProp 𝓤₀ 69 | _<<<_ M x y = ∀[ s ∶ List (M .S) ] M accepts s at y ⇒ M accepts s at x 70 | ``` 71 | 72 | ```agda 73 | <<<-refl : (M : DFA) → [ ∀[ x ∶ M .Q ] (_<<<_ M x x) ] 74 | <<<-refl M x s = idfun _ 75 | ``` 76 | 77 | ```agda 78 | <<<-trans : (M : DFA) 79 | → (x y z : M .Q) → [ _<<<_ M x y ⇒ _<<<_ M y z ⇒ _<<<_ M x z ] 80 | <<<-trans M x y z p q s = p s ∘ q s 81 | ``` 82 | 83 | ```agda 84 | isAntisymmetric : (M : DFA) → 𝓤₀ ̇ 85 | isAntisymmetric M = (x y : M .Q) → [ _<<<_ M x y ] → [ _<<<_ M y x ] → x ≡ y 86 | ``` 87 | 88 | ```agda 89 | dfa-poset : (M : DFA) → isAntisymmetric M → Poset 𝓤₀ 𝓤₀ 90 | dfa-poset M p = M .Q , _<<<_ M , Q-set M , <<<-refl M , <<<-trans M , p 91 | ``` 92 | 93 | ## Interaction system 94 | 95 | ```agda 96 | to-is : (M : DFA) → InteractionSys 𝓤₀ 97 | to-is M = M .Q , BB , CC , (λ {x} {y} → dd {x} {y}) 98 | where 99 | BB : M .Q → 𝓤₀ ̇ 100 | BB = const (Unit 𝓤₀) 101 | 102 | CC : Unit 𝓤₀ → Type 𝓤₀ 103 | CC _ = M .S 104 | 105 | dd : {x : M .Q} {y : BB x} → CC y → M .Q 106 | dd {x = x} = M .δ x 107 | ``` 108 | 109 | ## Formal topology 110 | 111 | ```agda 112 | to-ft : (M : DFA) → isAntisymmetric M → FormalTopology 𝓤₀ 𝓤₀ 113 | to-ft M p = dfa-poset M p , π₁ (to-is M) , m , s 114 | where 115 | -- Kind of like a safety property. 116 | -- For the language of bad prefixes: If you accept the word, then you accept any 117 | -- extension of the word. 118 | m′ : (x : M .Q) (a : M .S) → [ _<<<_ M (M .δ x a) x ] 119 | m′ = {!!} 120 | 121 | m : hasMono (dfa-poset M p) (π₁ (to-is M)) 122 | m a tt = m′ a 123 | 124 | s′ : (x′ x : M .Q) 125 | → [ _<<<_ M x′ x ] 126 | → (a′ : M .S) 127 | → Σ[ a ∈ M .S ] [ _<<<_ M (M .δ x′ a′) (M .δ x a) ] 128 | s′ = {!!} 129 | 130 | -- ∀ x′, x : Q. x′ ≤ x → ∀ a′ : Σ. ∃ a : Σ. δ(x′, a′) ≤ δ(x, a) 131 | 132 | s : hasSimulation (dfa-poset M p) (π₁ (to-is M)) 133 | s a′ a x b = tt , s′ a′ a x 134 | 135 | 136 | -- --} 137 | ``` 138 | -------------------------------------------------------------------------------- /FormalTopology/WayBelow.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: The Way Below Relation 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 22 | 23 | ## Notation 24 | 25 | ```agda 26 | infix 7 _≤_ 27 | 28 | _≤_ : ∣ F ∣F → ∣ F ∣F → hProp 𝓥 29 | x ≤ y = x ⊑[ pos F ] y 30 | ``` 31 | 32 | ```agda 33 | infix 8 ⋁_ 34 | 35 | ⋁_ : Fam 𝓦 ∣ F ∣F → ∣ F ∣F 36 | ⋁ U = ⋁[ F ] U 37 | ``` 38 | 39 | ## Definition of way below 40 | 41 | Definition copied from Escardó and de Jong. 42 | 43 | ```agda 44 | _≪_ : ∣ F ∣F → ∣ F ∣F → hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 45 | x ≪ y = 46 | ∀[ S ∶ Fam 𝓦 ∣ F ∣F ] 47 | isDirected (pos F) S ⇒ y ≤ ⋁ S ⇒ ∥ Σ[ i ∈ index S ] [ x ≤ (S $ i) ] ∥Ω 48 | ``` 49 | 50 | ## Definition of a compact element 51 | 52 | ```agda 53 | isCompactOpen : ∣ F ∣F → hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 54 | isCompactOpen x = x ≪ x 55 | ``` 56 | 57 | ```agda 58 | compacts-of : (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 59 | compacts-of = Σ[ x ∈ ∣ F ∣F ] [ isCompactOpen x ] 60 | ``` 61 | 62 | ## Definition of a compact frame 63 | 64 | ```agda 65 | isCompact : hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 66 | isCompact = isCompactOpen ⊤[ F ] 67 | ``` 68 | 69 | ## Continuity 70 | 71 | ```agda 72 | isContinuous : hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 73 | isContinuous = 74 | ∀[ x ∶ ∣ F ∣F ] isSup (pos F) ((Σ[ y ∈ ∣ F ∣F ] [ y ≪ x ]) , fst) x 75 | ``` 76 | 77 | ## Properties of the way below relation 78 | 79 | ```agda 80 | compactness-closed-under-joins : (x y : ∣ F ∣F) 81 | → [ x ≪ x ] → [ y ≪ y ] → [ (x ∨[ F ] y) ≪ (x ∨[ F ] y) ] 82 | compactness-closed-under-joins x y x≪x y≪y S S-dir p = 83 | ∥∥-rec (∥∥-prop _) (uncurry nts) (∥∥-× s-x s-y) 84 | where 85 | open PosetReasoning (pos F) 86 | 87 | x≤⋁S : [ x ⊑[ pos F ] ⋁[ F ] S ] 88 | x≤⋁S = x ⊑⟨ ⊔[ F ]-upper₀ x y ⟩ 89 | x ∨[ F ] y ⊑⟨ p ⟩ 90 | ⋁[ F ] S ■ 91 | 92 | y≤⋁S : [ y ⊑[ pos F ] ⋁[ F ] S ] 93 | y≤⋁S = y ⊑⟨ ⊔[ F ]-upper₁ x y ⟩ 94 | x ∨[ F ] y ⊑⟨ p ⟩ 95 | ⋁[ F ] S ■ 96 | 97 | s-x : ∥ Σ[ i₀ ∈ index S ] [ x ⊑[ pos F ] (S $ i₀) ] ∥ 98 | s-x = x≪x S S-dir x≤⋁S 99 | 100 | s-y : ∥ Σ[ i₁ ∈ index S ] [ y ⊑[ pos F ] (S $ i₁) ] ∥ 101 | s-y = y≪y S S-dir y≤⋁S 102 | 103 | nts : Σ[ i₀ ∈ index S ] [ x ⊑[ pos F ] (S $ i₀) ] 104 | → Σ[ i₁ ∈ index S ] [ y ⊑[ pos F ] (S $ i₁) ] 105 | → ∥ Σ[ i ∈ index S ] [ (x ∨[ F ] y) ⊑[ pos F ] (S $ i) ] ∥ 106 | nts (i₀ , x≤s₀) (i₁ , y≤s₁) = 107 | ∥∥-rec 108 | (∥∥-prop (Σ[ i ∈ index S ] [ (x ∨[ F ] y) ⊑[ pos F ] (S $ i) ])) 109 | rem 110 | (π₁ S-dir i₀ i₁) 111 | where 112 | rem : Σ[ k ∈ index S ] [ (S $ i₀) ⊑[ pos F ] (S $ k) ] × [ (S $ i₁) ⊑[ pos F ] (S $ k) ] 113 | → ∥ Σ[ i ∈ index S ] [ (x ∨[ F ] y) ⊑[ pos F ] (S $ i) ] ∥ 114 | rem (i , Si₀≤Sk , Si₁≤Sk) = ∣ i , (⋁[ F ]-least _ _ goal) ∣ 115 | where 116 | goal : [ ∀[ z ε ⁅ x , y ⁆ ] (z ⊑[ pos F ] (S $ i)) ] 117 | goal z (true , eq) = subst (λ - → [ - ⊑[ pos F ] (S $ i) ]) eq 118 | (x ⊑⟨ x≤s₀ ⟩ S $ i₀ ⊑⟨ Si₀≤Sk ⟩ S $ i ■) 119 | goal z (false , eq) = subst (λ - → [ - ⊑[ pos F ] (S $ i) ]) eq 120 | (y ⊑⟨ y≤s₁ ⟩ S $ i₁ ⊑⟨ Si₁≤Sk ⟩ S $ i ■) 121 | ``` 122 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Formal Topology in Univalent Foundations 2 | 3 | [**Link to thesis**](https://hdl.handle.net/20.500.12380/301098). 4 | 5 | This is the Agda development accompanying my (upcoming) master's thesis at Chalmers 6 | University of Technology to be titled _Formal Topology in Univalent Foundations_. 7 | 8 | 9 | **Note**: This library is not actively maintained. The dependency `cubical` 10 | needs to be checked out to commit `09cc7134082573cf82436f5d317405812856f7f6`. 11 | For an actively developed version of the locale theory development here, see 12 | the 13 | [`Locales`](https://github.com/martinescardo/TypeTopology/tree/master/source/Locales) 14 | in [`TypeTopology`](https://github.com/martinescardo/TypeTopology). 15 | 16 | The approach to formal topology implemented here follows an idea of Thierry Coquand [0] to 17 | define formal topologies as posets endowed with [interaction systems][2]. The main novelty 18 | in this development is the definition of covering as an HIT. This seems to be necessary in 19 | the context of univalent type theory to avoid using a form of the axiom of choice. 20 | 21 | The version of the code presented in the thesis will be archived whereas this repository 22 | (which is, as of now, mostly the same) will be maintained and developed further. 23 | 24 | ## Question: what is formal topology? 25 | 26 | Here is an answer by Giovanni Sambin [1]: 27 | 28 | > What is formal topology? A good approximation to the correct answer is: formal topology 29 | > is topology as developed in (Martin-Löf's) type theory. 30 | 31 | Also, [this blog post][6] by Mike Shulman contains interesting remarks about predicative 32 | mathematics and formal topology. 33 | 34 | ## Overview 35 | 36 | The main development comprises nine modules. If you are interested in reading the code, I 37 | suggest the following order: 38 | 39 | 1. `Basis`. Basic definitions of univalent type theory, many of which are imported from 40 | `agda/cubical` and some of which are adapted from Martín Escardó's 41 | [introduction to HoTT/UF][4]. 42 | 2. `Poset`. 43 | 3. `Frame`. A rudimentary development of [frames][5]. 44 | 4. `Nucleus`. The notion of a [nucleus][3] on a frame. 45 | 5. `FormalTopology`. Definition of a formal topology as an interaction system. 46 | 6. `Cover`. The cover relation induced by the structure of a formal topology. 47 | 7. `CoverFormsNucleus`. Contains the proof that the cover relation of a formal topology is 48 | a nucleus on the frame of downwards-closed subsets of its underlying poset. 49 | 8. `UniversalProperty`. Contains the proof that formal topologies present frames as 50 | expected. 51 | 9. `CantorSpace`. The definition of the formal Cantor topology along with a proof that it 52 | is compact. 53 | 10. `Sierpinski`. The formal topology of the Sierpinski space. 54 | 55 | ## Clickable HTML 56 | 57 | A rendering of the code in glorious clickable HTML can be [accessed here][7]. 58 | 59 | ## Credits 60 | 61 | This work was carried out under the supervision of Thierry Coquand. 62 | 63 | ## References 64 | 65 | 0. Coquand, T. 1996. Formal Topology with Posets. http://www.cse.chalmers.se/~coquand/formal.html 66 | 1. Sambin, G. 2000. Formal Topology and Domains. Electronic Notes in Theoretical Computer Science. 35, (Jan. 2000), 177–190. DOI:https://doi.org/10.1016/S1571-0661(05)80742-X. 67 | 68 | [2]: http://www.dcs.ed.ac.uk/home/pgh/interactive_systems.html 69 | [3]: https://ncatlab.org/nlab/show/nucleus 70 | [4]: https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html 71 | [5]: https://ncatlab.org/nlab/show/frame 72 | [6]: https://golem.ph.utexas.edu/category/2011/01/topos_theory_can_make_you_a_pr.html 73 | [7]: https://ayberkt.github.io/formal-topology-in-UF/Main.html 74 | -------------------------------------------------------------------------------- /FormalTopology/Cover.lagda.md: -------------------------------------------------------------------------------- 1 | # Some lemmas about the cover relation 2 | 3 | ```agda 4 | {-# OPTIONS --cubical --safe #-} 5 | 6 | module Cover where 7 | 8 | open import FormalTopology 9 | open import Poset 10 | open import Basis 11 | ``` 12 | 13 | We define a submodule `CoverFromFormalTopology` parameterised by a formal topology `ℱ` 14 | since all of the functions in this module take as argument a certain formal topology. 15 | 16 | ```agda 17 | module CoverFromFormalTopology (ℱ : FormalTopology ℓ ℓ′) where 18 | ``` 19 | 20 | We refer to the underlying poset of the formal topology `ℱ` as `P`, and its outcome 21 | function as `out`. 22 | 23 | ```agda 24 | private 25 | P = pos ℱ 26 | out = outcome 27 | ``` 28 | 29 | ## Definition of the covering relation 30 | 31 | The covering relation is defined as follows: 32 | 33 | ```agda 34 | data _◁_ (a : ∣ P ∣ₚ) (U : ∣ P ∣ₚ → hProp ℓ) : Type ℓ where 35 | dir : [ U a ] → a ◁ U 36 | branch : (b : exp ℱ a) → (f : (c : out ℱ b) → next ℱ c ◁ U) → a ◁ U 37 | squash : (p q : a ◁ U) → p ≡ q 38 | 39 | ◁-prop : (a : ∣ P ∣ₚ) (U : 𝒫 ∣ P ∣ₚ) → isProp (a ◁ U) 40 | ◁-prop a U = squash 41 | ``` 42 | 43 | ## Lemmas about the covering relation 44 | 45 | We now prove four crucial lemmas about the cover. 46 | 47 | ### Lemma 1 48 | 49 | ```agda 50 | module _ {U : ∣ P ∣ₚ → hProp ℓ} (U-down : [ isDownwardsClosed P U ]) where 51 | 52 | ◁-lem₁ : {a a′ : ∣ P ∣ₚ} → [ a′ ⊑[ P ] a ] → a ◁ U → a′ ◁ U 53 | ◁-lem₁ {_} {_} h (squash p₀ p₁ i) = squash (◁-lem₁ h p₀) (◁-lem₁ h p₁) i 54 | ◁-lem₁ {_} {_} h (dir q) = dir (U-down _ _ q h) 55 | ◁-lem₁ {a = a} {a′} h (branch b f) = branch b′ g 56 | where 57 | b′ : exp ℱ a′ 58 | b′ = π₀ (sim ℱ a′ a h b) 59 | 60 | g : (c′ : out ℱ b′) → next ℱ c′ ◁ U 61 | g c′ = ◁-lem₁ δc′⊑δc (f c) 62 | where 63 | c : out ℱ b 64 | c = π₀ (π₁ (sim ℱ a′ a h b) c′) 65 | 66 | δc′⊑δc : [ next ℱ c′ ⊑[ P ] next ℱ c ] 67 | δc′⊑δc = π₁ (π₁ (sim ℱ a′ a h b) c′) 68 | ``` 69 | 70 | ### Lemma 2 71 | 72 | ```agda 73 | module _ (U : 𝒫 ∣ P ∣ₚ) (V : 𝒫 ∣ P ∣ₚ) (V-dc : [ isDownwardsClosed P V ]) where 74 | 75 | ◁-lem₂ : {a : ∣ P ∣ₚ} → a ◁ U → [ a ∈ V ] → a ◁ (U ∩ V) 76 | ◁-lem₂ (squash p₀ p₁ i) h = squash (◁-lem₂ p₀ h) (◁-lem₂ p₁ h) i 77 | ◁-lem₂ (dir q) h = dir (q , h) 78 | ◁-lem₂ (branch b f) h = branch b λ c → ◁-lem₂ (f c) (V-dc _ _ h (mono ℱ _ b c)) 79 | ``` 80 | 81 | ### Lemma 3 82 | 83 | ``` 84 | module _ (U : 𝒫 ∣ P ∣ₚ) (V : 𝒫 ∣ P ∣ₚ) 85 | (U-dc : [ isDownwardsClosed P U ]) 86 | (V-dc : [ isDownwardsClosed P V ]) where 87 | 88 | ◁-lem₃ : {a a′ : ∣ P ∣ₚ} → [ a′ ⊑[ P ] a ] → a ◁ U → a′ ◁ V → a′ ◁ (V ∩ U) 89 | ◁-lem₃ {a} {a′} a′⊑a (squash p q i) r = squash (◁-lem₃ a′⊑a p r) (◁-lem₃ a′⊑a q r) i 90 | ◁-lem₃ {a} {a′} a′⊑a (dir a∈U) r = ◁-lem₂ V U U-dc r (U-dc a a′ a∈U a′⊑a) 91 | ◁-lem₃ {a} {a′} a′⊑a (branch b f) r = branch b′ g 92 | where 93 | b′ : exp ℱ a′ 94 | b′ = π₀ (sim ℱ a′ a a′⊑a b) 95 | 96 | g : (c′ : out ℱ b′) → next ℱ c′ ◁ (V ∩ U) 97 | g c′ = ◁-lem₃ NTS (f c) (◁-lem₁ V-dc (mono ℱ a′ b′ c′) r) 98 | where 99 | c : out ℱ b 100 | c = π₀ (π₁ (sim ℱ a′ a a′⊑a b) c′) 101 | 102 | NTS : [ next ℱ c′ ⊑[ P ] next ℱ c ] 103 | NTS = π₁ (π₁ (sim ℱ a′ a a′⊑a b) c′) 104 | ``` 105 | 106 | ### Lemma 4 107 | 108 | ```agda 109 | ◁-lem₄ : (U : 𝒫 ∣ P ∣ₚ) (V : 𝒫 ∣ P ∣ₚ) 110 | → ((u : ∣ P ∣ₚ) → [ u ∈ U ] → u ◁ V) → (a : ∣ P ∣ₚ) → a ◁ U → a ◁ V 111 | ◁-lem₄ U V h a (squash p₀ p₁ i) = squash (◁-lem₄ U V h a p₀) (◁-lem₄ U V h a p₁) i 112 | ◁-lem₄ U V h a (dir p) = h a p 113 | ◁-lem₄ U V h a (branch b f) = branch b λ c → ◁-lem₄ U V h (next ℱ c) (f c) 114 | ``` 115 | -------------------------------------------------------------------------------- /FormalTopology/CoverFormsNucleus.lagda.md: -------------------------------------------------------------------------------- 1 | ``` 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module CoverFormsNucleus where 5 | 6 | open import Basis hiding (A) renaming (squash to squash′) 7 | open import Poset 8 | open import Frame 9 | open import Cover 10 | open import Nucleus using (isNuclear; Nucleus; 𝔣𝔦𝔵; idem) 11 | open import Cubical.Data.Bool using (Bool; true; false) 12 | open import FormalTopology renaming (pos to pos′) 13 | ``` 14 | 15 | We use a module that takes some formal topology `F` as a parameter to reduce 16 | parameter-passing. 17 | 18 | ``` 19 | module NucleusFrom (ℱ : FormalTopology ℓ₀ ℓ₀) where 20 | ``` 21 | 22 | We refer to the underlying poset of `F` as `P` and the frame of downwards-closed subsets 23 | of `P` as `P↓`. 24 | 25 | ``` 26 | private 27 | P = pos′ ℱ 28 | P↓ = DCFrame P 29 | _⊑_ = λ (x y : stage ℱ) → x ⊑[ P ] y 30 | 31 | open CoverFromFormalTopology ℱ public 32 | ``` 33 | 34 | Now, we define the *covering nucleus* which we denote by `𝕛`. At its heart, this is 35 | nothing but the map `U ↦ - ◁ U`. 36 | 37 | ``` 38 | 𝕛 : ∣ P↓ ∣F → ∣ P↓ ∣F 39 | 𝕛 (U , U-down) = (λ - → U ▷ -) , U▷-dc 40 | where 41 | -- This is not propositional unless we force it to be using the HIT definition! 42 | _▷_ : 𝒫 ∣ P ∣ₚ → 𝒫 ∣ P ∣ₚ 43 | U ▷ a = a ◁ U , squash 44 | 45 | U▷-dc : [ isDownwardsClosed P (λ - → (- ◁ U) , squash) ] 46 | U▷-dc a a₀ aεU₁ a₀⊑a = ◁-lem₁ U-down a₀⊑a aεU₁ 47 | 48 | 𝕛-nuclear : isNuclear P↓ 𝕛 49 | 𝕛-nuclear = N₀ , N₁ , N₂ 50 | where 51 | -- We reason by antisymmetry and prove in (d) 𝕛 (a₀ ⊓ a₁) ⊑ (𝕛 a₀) ⊓ (𝕛 a₁) and 52 | -- in (u) (𝕛 a₀) ⊓ (𝕛 a₁) ⊑ 𝕛 (a₀ ⊓ a₁). 53 | N₀ : (𝔘 𝔙 : ∣ P↓ ∣F) → 𝕛 (𝔘 ⊓[ P↓ ] 𝔙) ≡ (𝕛 𝔘) ⊓[ P↓ ] (𝕛 𝔙) 54 | N₀ 𝕌@(U , U-down) 𝕍@(V , V-down) = 55 | ⊑[ pos P↓ ]-antisym (𝕛 (𝕌 ⊓[ P↓ ] 𝕍)) (𝕛 𝕌 ⊓[ P↓ ] 𝕛 𝕍) down up 56 | where 57 | down : [ (𝕛 (𝕌 ⊓[ P↓ ] 𝕍)) ⊑[ pos P↓ ] (𝕛 𝕌 ⊓[ P↓ ] 𝕛 𝕍) ] 58 | down a p = (◁-lem₄ (U ∩ V) U (λ { u (u∈U , u∈V) → dir u∈U }) a p) 59 | , (◁-lem₄ (U ∩ V) V (λ { u (u∈U , u∈V) → dir u∈V }) a p) 60 | 61 | up : [ (𝕛 𝕌 ⊓[ P↓ ] 𝕛 𝕍) ⊑[ pos P↓ ] 𝕛 (𝕌 ⊓[ P↓ ] 𝕍) ] 62 | up a (a◁U , a◁V) = ◁-lem₃ V U V-down U-down (⊑[ P ]-refl a) a◁V a◁U 63 | 64 | N₁ : (𝔘 : ∣ P↓ ∣F) → [ 𝔘 ⊑[ pos P↓ ] (𝕛 𝔘) ] 65 | N₁ _ a₀ a∈U = dir a∈U 66 | 67 | N₂ : (𝔘 : ∣ P↓ ∣F) → [ π₀ (𝕛 (𝕛 𝔘)) ⊆ π₀ (𝕛 𝔘) ] 68 | N₂ 𝔘@(U , _) = ◁-lem₄ (π₀ (𝕛 𝔘)) U (λ _ q → q) 69 | ``` 70 | 71 | We denote by `L` the frame of fixed points for `𝕛`. 72 | 73 | ``` 74 | L : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀ 75 | L = 𝔣𝔦𝔵 P↓ (𝕛 , 𝕛-nuclear) 76 | ``` 77 | 78 | The following is a just a piece of convenient notation for projecting out the underlying 79 | set of a downwards-closed subset equipped with the information that it is a fixed point 80 | for `𝕛`. 81 | 82 | ``` 83 | ⦅_⦆ : ∣ L ∣F → 𝒫 ∣ P ∣ₚ 84 | ⦅ ((U , _) , _) ⦆ = U 85 | ``` 86 | 87 | Given some `x` in `F`, we define a map taking `x` to its *downwards-closure*. 88 | 89 | ``` 90 | ↓-clos : stage ℱ → ∣ P↓ ∣F 91 | ↓-clos x = x↓ , down-DC 92 | where 93 | x↓ = λ y → y ⊑[ P ] x 94 | down-DC : [ isDownwardsClosed P x↓ ] 95 | down-DC z y z⊑x y⊑z = ⊑[ P ]-trans y z x y⊑z z⊑x 96 | 97 | x◁x↓ : (x : stage ℱ) → x ◁ (λ - → - ⊑[ P ] x) 98 | x◁x↓ x = dir (⊑[ P ]-refl x) 99 | ``` 100 | 101 | By composing this with the covering nucleus, we define a map `e` from `F` to `P↓`. 102 | 103 | ``` 104 | e : ∣ P ∣ₚ → ∣ P↓ ∣F 105 | e z = (λ a → (a ◁ (π₀ (↓-clos z))) , squash) , NTS 106 | where 107 | NTS : [ isDownwardsClosed P (λ a → (a ◁ (λ - → - ⊑[ P ] z)) , squash) ] 108 | NTS _ _ x y = ◁-lem₁ (λ _ _ x⊑y y⊑z → ⊑[ P ]-trans _ _ z y⊑z x⊑y) y x 109 | ``` 110 | 111 | We can further refine the codomain of `e` to `L`. In other words, we can prove that `j (e 112 | x) = e x` for every `x`. We call the version `e` with the refined codomain `η`. 113 | 114 | ``` 115 | fixing : (x : ∣ P ∣ₚ) → 𝕛 (e x) ≡ e x 116 | fixing x = ⊑[ pos P↓ ]-antisym (𝕛 (e x)) (e x) down up 117 | where 118 | down : ∀ y → [ π₀ (𝕛 (e x)) y ] → [ π₀ (e x) y ] 119 | down = ◁-lem₄ (π₀ (e x)) (π₀ (↓-clos x)) (λ _ q → q) 120 | 121 | up : [ e x ⊑[ pos P↓ ] 𝕛 (e x) ] 122 | up = π₀ (π₁ 𝕛-nuclear) (e x) 123 | 124 | η : stage ℱ → ∣ L ∣F 125 | η x = e x , fixing x 126 | ``` 127 | 128 | Furthermore, `η` is a monotonic map. 129 | 130 | ``` 131 | ηm : P ─m→ pos L 132 | ηm = η , η-mono 133 | where 134 | η-mono : isMonotonic P (pos L) η 135 | η-mono x y x⊑y = ◁-lem₄ (π₀ (↓-clos x)) (π₀ (↓-clos y)) NTS 136 | where 137 | NTS : (u : ∣ P ∣ₚ) → [ u ∈ π₀ (↓-clos x) ] → u ◁ π₀ (↓-clos y) 138 | NTS _ p = ◁-lem₁ (π₁ (↓-clos y)) p (dir x⊑y) 139 | ``` 140 | -------------------------------------------------------------------------------- /resources/Agda.css: -------------------------------------------------------------------------------- 1 | /* Copied from the Agda tutorial by Péter Diviánszky, Ambrus Kaposi, and */ 2 | /* Gábor Páli. Contains modifications. */ 3 | 4 | /* See https://hub.darcs.net/divip/AgdaTutorial/browse/LICENSE for their */ 5 | /* license. */ 6 | 7 | @import url('https://fonts.googleapis.com/css2?family=Fira+Mono&display=swap'); 8 | @import url('https://fonts.googleapis.com/css2?family=Share+Tech+Mono&display=swap'); 9 | 10 | @font-face { 11 | font-family: "PragmataPro Mono Liga"; 12 | src: url("fonts/PragmataPro_Mono_R_liga_0828.woff2") format('woff'); 13 | } 14 | 15 | /* Aspects. */ 16 | .Comment { color: #B22222 } 17 | .Keyword { color: #CD6600 } 18 | .String { color: #B22222 } 19 | .Number { color: #A020F0 } 20 | .Symbol { color: #404040 } 21 | .PrimitiveType { color: #0000CD } 22 | .Operator {} 23 | 24 | /* NameKinds. */ 25 | .Bound { color: black } 26 | .InductiveConstructor { color: #008B00 } 27 | .CoinductiveConstructor { color: #8B7500 } 28 | .Datatype { color: #0000CD } 29 | .Field { color: #EE1289 } 30 | .Function { color: #0000CD } 31 | .Module { color: #A020F0 } 32 | .Postulate { color: #0000CD } 33 | .Primitive { color: #0000CD } 34 | .Record { color: #0000CD } 35 | 36 | /* OtherAspects. */ 37 | .DottedPattern {} 38 | .UnsolvedMeta { color: black; background: yellow } 39 | .UnsolvedConstraint { color: black; background: yellow } 40 | .TerminationProblem { color: black; background: #FFA07A } 41 | .IncompletePattern { color: black; background: #F5DEB3 } 42 | .Error { color: red; text-decoration: underline } 43 | .TypeChecks { color: black; background: #ADD8E6 } 44 | 45 | /* Standard attributes. */ 46 | a { text-decoration: none } 47 | a[href]:hover { background-color: #B4EEB4 } 48 | 49 | /* padding */ 50 | 51 | body { 52 | margin: auto; 53 | padding: 1em 1em 1em 1em; 54 | max-width: 50em; 55 | line-height: 130%; 56 | background-color: white; 57 | color: black; 58 | } 59 | 60 | h1 { 61 | padding-top: 1.5em; 62 | } 63 | 64 | h2 { 65 | padding-top: 1em; 66 | } 67 | 68 | h3 { 69 | padding-top: 0.5em; 70 | } 71 | 72 | h1.title { 73 | padding-top: 1em; 74 | padding-bottom: 0.5em; 75 | } 76 | h1.title, h2.author, h3.date { 77 | text-align: center; 78 | line-height: normal; 79 | } 80 | 81 | .author { 82 | text-align: center; 83 | padding-bottom: 3em; 84 | font-size: 22px; 85 | } 86 | 87 | h2.author, h3.date { visibility: hidden } 88 | h2.author, h3.date { display: none } 89 | 90 | 91 | hr { 92 | height: 10px; 93 | } 94 | 95 | 96 | /* margin */ 97 | 98 | div.indent, pre, table { 99 | margin-left: 5%; 100 | width: 95%; 101 | } 102 | 103 | div.indent, pre { 104 | margin-top: 0.5em; 105 | margin-bottom: 0.5em; 106 | } 107 | 108 | pre.normal { 109 | margin: 0; 110 | width: auto; 111 | } 112 | 113 | 114 | div.indent input { 115 | margin-bottom: 0.5em; 116 | } 117 | 118 | input, textarea { 119 | margin: 0; 120 | } 121 | 122 | 123 | /* border */ 124 | 125 | h1, h2, h3 { 126 | border-bottom: 1px dotted black; 127 | } 128 | 129 | h1.title, h2.author, h3.date, input, textarea, hr { 130 | border-style: none; 131 | } 132 | 133 | 134 | /* font sizes */ 135 | 136 | h1 { font-size: 130%; } 137 | h1.title { font-size: 200%; } 138 | h2 { font-size: 110%; } 139 | input, textarea { font-size: 100%; } /* monospace correction */ 140 | 141 | 142 | /* font family, style, weight, text decoration */ 143 | 144 | h1, h2, h3 { 145 | font-weight: bold; 146 | } 147 | 148 | pre code, textarea, div.string { 149 | font-family: 'PragmataPro Mono Liga', monospace; 150 | font-size: 100; 151 | } 152 | 153 | pre.Agda { 154 | font-family: 'PragmataPro Mono Liga', monospace; 155 | line-height: 120%; 156 | } 157 | 158 | pre.normal, pre.normal code, code, .error, input { 159 | font-family: 'PragmataPro Mono Liga', monospace; 160 | } 161 | 162 | 163 | a { 164 | text-decoration: none; 165 | } 166 | 167 | div.serious {font-weight: bold;} 168 | div.info {color: gray;} 169 | 170 | 171 | /* colors other than highlighting */ 172 | 173 | input { background-color: #eeeeee; } 174 | textarea { background-color: #ffffba; } 175 | hr { background-color: #f5f0ee; } 176 | div.string span { background-color: yellow; } 177 | 178 | a { color: #2a207a; } 179 | h1 a, h2 a, h3 a { color: black; } 180 | code { color: red; } 181 | .type { color: #0080ff; } 182 | .error { color: #a030a0; } 183 | .comment { color: gray; float: right; } 184 | 185 | 186 | /* misc */ 187 | 188 | p.caption { 189 | display: none; 190 | } 191 | 192 | div#lang { 193 | position:fixed; 194 | z-index:1000; 195 | right: 1%; 196 | top: 1%; 197 | } 198 | 199 | div#info { 200 | display: none; 201 | position:fixed; 202 | z-index:1000; 203 | left: 40%; 204 | width: 20%; 205 | top: 40%; 206 | height: 20%; 207 | } 208 | 209 | .wait { 210 | padding: 10px; 211 | width: 8em; 212 | border: 1px solid grey; 213 | } 214 | 215 | /* highlighting css */ 216 | 217 | table.sourceCode, tr.sourceCode, td.lineNumbers, td.sourceCode, table.sourceCode 218 | { margin: 0; padding: 0; border: 0; vertical-align: baseline; border: none; } 219 | td.lineNumbers { border-right: 1px solid #AAAAAA; text-align: right; color: #AAAAAA; padding-right: 5px; padding-left: 5px; } 220 | td.sourceCode { padding-left: 5px; } 221 | pre.sourceCode span.kw { color: #007020; /* font-weight: bold; */ } 222 | pre.sourceCode span.dt { color: #902000; } 223 | pre.sourceCode span.dv { color: #40a070; } 224 | pre.sourceCode span.bn { color: #40a070; } 225 | pre.sourceCode span.fl { color: #40a070; } 226 | pre.sourceCode span.ch { color: #4070a0; } 227 | pre.sourceCode span.st { color: #4070a0; } 228 | pre.sourceCode span.co { color: #60a0b0; /* font-style: italic; */ } 229 | pre.sourceCode span.ot { color: #007020; } 230 | pre.sourceCode span.al { color: red; /* font-weight: bold; */ } 231 | pre.sourceCode span.fu { color: #06287e; } 232 | pre.sourceCode span.re { } 233 | pre.sourceCode span.er { color: red; /* font-weight: bold; */ } 234 | -------------------------------------------------------------------------------- /FormalTopology/AdditionalFrameTheorems.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Properties of Frames 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 21 | 22 | Bunch of theorems about frames that should really go in the `Frame` module. That module 23 | has gotten too big which is slowing down the typechecking which is why they are in this 24 | module. 25 | 26 | ```agda 27 | resp-∨ : (F : Frame 𝓤 𝓥 𝓦) (G : Frame 𝓤′ 𝓥′ 𝓦) (f : F ─f→ G) 28 | → (x y : ∣ F ∣F) 29 | → π₀ (π₀ f) (x ∨[ F ] y) ≡ π₀ (π₀ f) x ∨[ G ] π₀ (π₀ f) y 30 | resp-∨ F G ((f , f-mono) , _ , _ , r) x y = 31 | f (x ∨[ F ] y) ≡⟨ r ⁅ x , y ⁆ ⟩ 32 | ⋁[ G ] (f ⟨$⟩ ⁅ x , y ⁆) ≡⟨ nts ⟩ 33 | f x ∨[ G ] f y ∎ 34 | where 35 | open PosetReasoning (pos G) 36 | 37 | nts : _ 38 | nts = ⋁-unique G _ _ G𝟏 G𝟐 39 | where 40 | G𝟏 : _ 41 | G𝟏 z (true , eq) = z ⊑⟨ ≡⇒⊑ (pos G) (sym eq) ⟩ 42 | f x ⊑⟨ ⋁[ G ]-upper _ _ (true , refl) ⟩ 43 | ⋁[ G ] (f ⟨$⟩ ⁅ x , y ⁆) ■ 44 | G𝟏 z (false , eq) = z ⊑⟨ ≡⇒⊑ (pos G) (sym eq) ⟩ 45 | f y ⊑⟨ ⋁[ G ]-upper _ _ (false , refl) ⟩ 46 | ⋁[ G ] (f ⟨$⟩ ⁅ x , y ⁆) ■ 47 | 48 | G𝟐 : _ 49 | G𝟐 z ϕ = ⋁[ G ]-least _ _ G𝟐a 50 | where 51 | G𝟐a : [ ∀[ w ε (f ⟨$⟩ ⁅ x , y ⁆) ] (w ⊑[ pos G ] z) ] 52 | G𝟐a w (true , eq) = w ⊑⟨ ≡⇒⊑ (pos G) (sym eq) ⟩ f x ⊑⟨ ϕ (f x) (true , refl) ⟩ z ■ 53 | G𝟐a w (false , eq) = w ⊑⟨ ≡⇒⊑ (pos G) (sym eq) ⟩ f y ⊑⟨ ϕ (f y) (false , refl) ⟩ z ■ 54 | 55 | resp-⊥ : (A : Frame 𝓤 𝓥 𝓦) (B : Frame 𝓤′ 𝓥′ 𝓦) 56 | → (f⋆ : A ─f→ B) 57 | → f⋆ $f ⊥[ A ] ≡ ⊥[ B ] 58 | resp-⊥ A B ((f⋆ , _) , (_ , _ , f⋁)) = 59 | f⋆ ⊥[ A ] ≡⟨ f⋁ (𝟘 _ , λ ()) ⟩ 60 | ⋁[ B ] (f⋆ ⟨$⟩ (𝟘 _ , λ ())) ≡⟨ † ⟩ 61 | ⊥[ B ] ∎ 62 | where 63 | † = cong (λ - → ⋁[ B ] (𝟘 _ , -)) (funExt λ ()) 64 | 65 | complement-preservation : (F : Frame 𝓤 𝓥 𝓦) (G : Frame 𝓤′ 𝓥′ 𝓦) (f : F ─f→ G) 66 | → (x x′ : ∣ F ∣F) 67 | → complements F x x′ 68 | → complements G (π₀ (π₀ f) x) (π₀ (π₀ f) x′) 69 | complement-preservation F G 𝒻@((f , f-mono) , p , q , r) x x′ (x∧x′=⊥ , x∨x′=⊤) = 70 | G𝟏 , G𝟐 71 | where 72 | open PosetReasoning (pos G) 73 | 74 | abstract 75 | G𝟏 : f x ⊓[ G ] f x′ ≡ ⊥[ G ] 76 | G𝟏 = ⊑[ pos G ]-antisym _ _ G𝟏a (⊥[ G ]-bottom (f x ⊓[ G ] f x′)) 77 | where 78 | G𝟏a : [ (f x ⊓[ G ] f x′) ⊑[ pos G ] ⊥[ G ] ] 79 | G𝟏a = (f x) ⊓[ G ] (f x′) ⊑⟨ ≡⇒⊑ (pos G) (sym (q x x′)) ⟩ 80 | f (x ⊓[ F ] x′) ⊑⟨ f-mono _ _ (≡⇒⊑ (pos F) x∧x′=⊥) ⟩ 81 | f ⊥[ F ] ⊑⟨ ⊑[ pos G ]-refl _ ⟩ 82 | f (⋁[ F ] (𝟘 _ , λ ())) ⊑⟨ ≡⇒⊑ (pos G) (r (𝟘 _ , λ ())) ⟩ 83 | ⋁[ G ] (f ⟨$⟩ (𝟘 _ , λ ())) ⊑⟨ ≡⇒⊑ (pos G) (cong (λ - → ⋁[ G ] (𝟘 _ , -)) (funExt λ ())) ⟩ 84 | ⋁[ G ] (𝟘 _ , λ ()) ⊑⟨ ⊑[ pos G ]-refl _ ⟩ 85 | ⊥[ G ] ■ 86 | 87 | G𝟐 : f x ∨[ G ] f x′ ≡ ⊤[ G ] 88 | G𝟐 = ⊑[ pos G ]-antisym _ _ (⊤[ G ]-top _) G𝟐a 89 | where 90 | G𝟐a : [ ⊤[ G ] ⊑[ pos G ] (f x ∨[ G ] f x′) ] 91 | G𝟐a = ⊤[ G ] ⊑⟨ ≡⇒⊑ (pos G) (sym p) ⟩ 92 | f ⊤[ F ] ⊑⟨ f-mono ⊤[ F ] (x ∨[ F ] x′) (≡⇒⊑ (pos F) (sym x∨x′=⊤)) ⟩ 93 | f (x ∨[ F ] x′) ⊑⟨ ≡⇒⊑ (pos G) (resp-∨ F G 𝒻 x x′) ⟩ 94 | f x ∨[ G ] f x′ ■ 95 | 96 | iso-inj-surj : (F : Frame 𝓤 𝓥 𝓦) (G : Frame 𝓤′ 𝓥′ 𝓦) 97 | → (𝒻@((f , _) , _) : F ─f→ G) 98 | → isEmbedding {A = ∣ F ∣F} {B = ∣ G ∣F} f 99 | → isSurjection {A = ∣ F ∣F} {B = ∣ G ∣F} f 100 | → isFrameIso {F = F} {G} 𝒻 101 | iso-inj-surj F G 𝒻@((f , f-mono) , f-resp-⊤ , f-resp-∧ , f-resp-⋁) f-emb f-surj = 102 | ((g , g-mono) , resp-⊤ , resp-∧ , resp-⋁) , sec , ret 103 | where 104 | ι : Iso ∣ F ∣F ∣ G ∣F 105 | ι = equivToIso (f , isEmbedding×isSurjection→isEquiv (f-emb , f-surj)) 106 | 107 | g : ∣ G ∣F → ∣ F ∣F 108 | g = Iso.inv ι 109 | 110 | sec : section f g 111 | sec = Iso.rightInv ι 112 | 113 | ret : retract f g 114 | ret = Iso.leftInv ι 115 | 116 | resp-⊤ : g ⊤[ G ] ≡ ⊤[ F ] 117 | resp-⊤ = g ⊤[ G ] ≡⟨ cong g (sym f-resp-⊤) ⟩ 118 | g (f ⊤[ F ]) ≡⟨ ret ⊤[ F ] ⟩ 119 | ⊤[ F ] ∎ 120 | 121 | f-inj : (x y : ∣ F ∣F) → f x ≡ f y → x ≡ y 122 | f-inj x y = Iso.inv (equivToIso ((λ p i → f (p i)) , f-emb x y)) 123 | 124 | resp-∧ : (x y : ∣ G ∣F) 125 | → g (x ⊓[ G ] y) ≡ g x ⊓[ F ] g y 126 | resp-∧ x y = f-inj (g (x ⊓[ G ] y)) (g x ⊓[ F ] g y) † 127 | where 128 | † : f (g (x ⊓[ G ] y)) ≡ f (g x ⊓[ F ] g y) 129 | † = f (g (x ⊓[ G ] y)) ≡⟨ sec (x ⊓[ G ] y) ⟩ 130 | x ⊓[ G ] y ≡⟨ cong (λ - → - ⊓[ G ] y) (sym (sec x)) ⟩ 131 | f (g x) ⊓[ G ] y ≡⟨ cong (λ - → _ ⊓[ G ] -) (sym (sec y)) ⟩ 132 | f (g x) ⊓[ G ] f (g y) ≡⟨ sym (f-resp-∧ (g x) (g y)) ⟩ 133 | f (g x ⊓[ F ] g y) ∎ 134 | 135 | g-mono : isMonotonic (pos G) (pos F) g 136 | g-mono y₀ y₁ y₀≤y₁ = x=x∧y⇒x⊑y F † 137 | where 138 | † : g y₀ ≡ g y₀ ⊓[ F ] g y₁ 139 | † = g y₀ ≡⟨ cong g (x⊑y⇒x=x∧y G y₀≤y₁) ⟩ 140 | g (y₀ ⊓[ G ] y₁) ≡⟨ resp-∧ y₀ y₁ ⟩ 141 | g y₀ ⊓[ F ] g y₁ ∎ 142 | 143 | resp-⋁ : (S : Fam _ ∣ G ∣F) → g (⋁[ G ] S) ≡ ⋁[ F ] ⁅ g s ∣ s ε S ⁆ 144 | resp-⋁ S = f-inj (g (⋁[ G ] S)) (⋁[ F ] ⁅ g s ∣ s ε S ⁆) † 145 | where 146 | † : f (g (⋁[ G ] S)) ≡ f (⋁[ F ] ⁅ g s ∣ s ε S ⁆) 147 | † = f (g (⋁[ G ] S)) ≡⟨ ⦅𝟏⦆ ⟩ 148 | ⋁[ G ] S ≡⟨ ⦅𝟐⦆ ⟩ 149 | ⋁[ G ] ⁅ f (g s) ∣ s ε S ⁆ ≡⟨ ⦅𝟑⦆ ⟩ 150 | f (⋁[ F ] ⁅ g s ∣ s ε S ⁆) ∎ 151 | where 152 | ⦅𝟏⦆ = sec (⋁[ G ] S) 153 | ⦅𝟐⦆ = cong (λ - → ⋁[ G ] (index S , -)) (funExt λ i → sym (sec (S $ i))) 154 | ⦅𝟑⦆ = sym (f-resp-⋁ ⁅ g s ∣ s ε S ⁆) 155 | ``` 156 | 157 | ```agda 158 | ⊥-compact : (A : Frame 𝓤 𝓥 𝓦) → [ _≪_ A ⊥[ A ] ⊥[ A ] ] 159 | ⊥-compact A S dir p = ∥∥-rec (∥∥-prop _) γ (π₀ dir) 160 | where 161 | γ : index S → ∥ Σ[ i ∈ index S ] [ ⊥[ A ] ⊑[ pos A ] (S $ i) ] ∥ 162 | γ i = ∣ i , (⊥[ A ]-bottom (S $ i)) ∣ 163 | ``` 164 | -------------------------------------------------------------------------------- /FormalTopology/RightAdjoint.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Right Adjoints for Frame Homomorphisms 3 | author: Ayberk Tosun (j.w.w. Martín Escardó) 4 | --- 5 | 6 | 18 | 19 | ```agda 20 | hasBasis : (F : Frame 𝓤 𝓥 𝓦) → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 21 | hasBasis {𝓦 = 𝓦} F = Σ[ B ∈ Fam 𝓦 ∣ F ∣F ] isDirBasisFor F B 22 | ``` 23 | 24 | ```agda 25 | module AdjointFunctorTheorem (F : Frame 𝓤 𝓥 𝓥) (G : Frame 𝓤′ 𝓥 𝓥) (basis : hasBasis F) where 26 | 27 | open GaloisConnection (pos F) (pos G) 28 | 29 | open PosetReasoning (pos G) 30 | open PosetReasoning (pos F) using () renaming (_⊑⟨_⟩_ to _⊑F⟨_⟩_; _■ to _■F) 31 | 32 | aft-1 : (f : pos F ─m→ pos G) 33 | → [ f hasRightAdjoint ] 34 | → ((S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ π₀ f s ∣ s ε S ⁆) ≡ (π₀ f (⋁[ F ] S))) 35 | aft-1 (f , f-mono) ((g , g-mono) , f⊣g) S = sym (⋁-unique G _ _ G𝟏 G𝟐) 36 | where 37 | G𝟏 : (x : ∣ G ∣F) → x ε (f ⟨$⟩ S) → [ x ⊑[ pos G ] (f (⋁[ F ] S)) ] 38 | G𝟏 x (i , eq) = subst (λ - → [ - ⊑[ pos G ] f (⋁[ F ] S) ]) eq G𝟏a 39 | where 40 | G𝟏a : [ f (S $ i) ⊑[ pos G ] f (⋁[ F ] S) ] 41 | G𝟏a = f-mono _ _ (⋁[ F ]-upper _ _ (i , refl)) 42 | 43 | G𝟐 : (z : ∣ G ∣F) 44 | → ((x : ∣ G ∣F) → x ε (f ⟨$⟩ S) → [ x ⊑[ pos G ] z ]) 45 | → [ f (⋁[ F ] S) ⊑[ pos G ] z ] 46 | G𝟐 z ϕ = π₁ (f⊣g (⋁[ F ] S) z) (⋁[ F ]-least _ _ G𝟐a) 47 | where 48 | G𝟐a : [ ∀[ s ε S ] (s ⊑[ pos F ] g z) ] 49 | G𝟐a s (i , eq) = 50 | subst (λ - → [ - ⊑[ pos F ] g z ]) eq (π₀ (f⊣g (π₁ S i) z) (ϕ (f (S $ i)) (i , refl))) 51 | 52 | aft-2 : (f : pos F ─m→ pos G) 53 | → ((S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ π₀ f s ∣ s ε S ⁆) ≡ (π₀ f (⋁[ F ] S))) 54 | → [ f hasRightAdjoint ] 55 | aft-2 f eq = (g , g-mono) , f⊣g 56 | where 57 | ℬ-F = π₀ basis 58 | 59 | g : ∣ G ∣F → ∣ F ∣F 60 | g y = 61 | ⋁[ F ] ⁅ π₁ ℬ-F i ∣ (i , _) ∶ Σ[ i ∈ π₀ ℬ-F ] [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] ⁆ 62 | 63 | g-mono : isMonotonic (pos G) (pos F) g 64 | g-mono x y x⊑y = ⋁[ F ]-least _ _ γ 65 | where 66 | γ : _ 67 | γ z ((i , p) , eq) = subst (λ - → [ - ⊑[ pos F ] (g y) ]) eq δ 68 | where 69 | nts : [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] 70 | nts = f $ₘ (π₁ ℬ-F i) ⊑⟨ p ⟩ x ⊑⟨ x⊑y ⟩ y ■ 71 | 72 | δ : [ _ ⊑[ pos F ] g y ] 73 | δ = ⋁[ F ]-upper _ _ ((i , nts) , refl) 74 | 75 | gm : pos G ─m→ pos F 76 | gm = g , g-mono 77 | 78 | f⊣g : [ f ⊣ gm ] 79 | f⊣g x y = nts₀ , nts₁ 80 | where 81 | nts₀ : [ f $ₘ x ⊑[ pos G ] y ⇒ x ⊑[ pos F ] g y ] 82 | nts₀ fx≤y = 83 | x ⊑F⟨ ≡⇒⊑ (pos F) (sym x-eq) ⟩ 84 | ⋁[ F ] ⁅ π₁ ℬ-F j ∣ j ε 𝒥 ⁆ ⊑F⟨ ⋁[ F ]-least _ _ rem ⟩ 85 | g y ■F 86 | where 87 | 𝒥 = π₀ (π₁ basis x) 88 | 89 | ϕ : _ 90 | ϕ = π₀ (π₁ (π₁ (π₁ basis x))) 91 | 92 | ψ : _ 93 | ψ = π₁ (π₁ (π₁ (π₁ basis x))) 94 | 95 | x-eq : (⋁[ F ] fmap (λ j → π₁ ℬ-F j) 𝒥) ≡ x 96 | x-eq = sym (⋁-unique F _ _ ϕ ψ) 97 | 98 | rem : [ ∀[ z ε ⁅ π₁ ℬ-F j ∣ j ε 𝒥 ⁆ ] (z ⊑[ pos F ] g y) ] 99 | rem z (j , eq) = subst (λ - → [ - ⊑[ pos F ] g y ]) eq rem′ 100 | where 101 | goal : [ f $ₘ π₁ ℬ-F (𝒥 $ j) ⊑[ pos G ] y ] 102 | goal = f $ₘ π₁ ℬ-F (𝒥 $ j) ⊑⟨ π₁ f _ _ (⋁[ F ]-upper _ _ (j , refl)) ⟩ 103 | f $ₘ (⋁[ F ] ⁅ π₁ ℬ-F j ∣ j ε 𝒥 ⁆) ⊑⟨ ≡⇒⊑ (pos G) (cong (f $ₘ_) x-eq) ⟩ 104 | f $ₘ x ⊑⟨ fx≤y ⟩ 105 | y ■ 106 | 107 | rem′ : [ (π₁ ((π₁ ℬ-F) ⟨$⟩ 𝒥) j) ⊑[ pos F ] (g y) ] 108 | rem′ = ⋁[ F ]-upper _ _ ((𝒥 $ j , goal) , refl) 109 | 110 | nts₁ : [ x ⊑[ pos F ] g y ⇒ f $ₘ x ⊑[ pos G ] y ] 111 | nts₁ x≤gy = 112 | f $ₘ x ⊑⟨ π₁ f _ _ x≤gy ⟩ 113 | f $ₘ (⋁[ F ] ⁅ π₁ ℬ-F i ∣ (i , _) ∶ Σ[ i ∈ π₀ ℬ-F ] [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] ⁆) ⊑⟨ ≡⇒⊑ (pos G) (sym (eq _)) ⟩ 114 | (⋁[ G ] ⁅ f $ₘ (π₁ ℬ-F i) ∣ (i , _) ∶ Σ[ i ∈ π₀ ℬ-F ] [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] ⁆ ) ⊑⟨ rem ⟩ 115 | y ■ 116 | where 117 | 𝒥 = π₀ (π₁ basis x) 118 | 119 | rem : [ ⋁[ G ] (⁅ f $ₘ (π₁ ℬ-F i) ∣ (i , _) ∶ Σ[ i ∈ π₀ ℬ-F ] [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] ⁆) ⊑[ pos G ] y ] 120 | rem = ⋁[ G ]-least _ _ goal 121 | where 122 | goal : [ ∀[ z ε ⁅ f $ₘ (π₁ ℬ-F i) ∣ (i , _) ∶ Σ[ i ∈ π₀ ℬ-F ] [ f $ₘ (π₁ ℬ-F i) ⊑[ pos G ] y ] ⁆ ] (z ⊑[ pos G ] y) ] 123 | goal z ((i , p) , eq) = subst (λ - → [ - ⊑[ pos G ] y ]) eq p 124 | 125 | _^*ᴹ : (f : pos F ─m→ pos G) → ((S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ π₀ f s ∣ s ε S ⁆) ≡ f $ₘ (⋁[ F ] S)) → pos G ─m→ pos F 126 | _^*ᴹ f rem = (π₀ (aft-2 f rem)) 127 | 128 | _^* : (f : pos F ─m→ pos G) → ((S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ π₀ f s ∣ s ε S ⁆) ≡ f $ₘ (⋁[ F ] S)) → ∣ G ∣F → ∣ F ∣F 129 | _^* f rem = π₀ (π₀ (aft-2 f rem)) 130 | 131 | ^*-RA : (f : pos F ─m→ pos G) → (rem : (S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ π₀ f s ∣ s ε S ⁆) ≡ f $ₘ (⋁[ F ] S)) → [ f ⊣ (_^*ᴹ f rem) ] 132 | ^*-RA f = π₁ ∘ aft-2 f 133 | 134 | RA-of-homo : (F ─f→ G) → ∣ G ∣F → ∣ F ∣F 135 | RA-of-homo (f , _ , _ , p) = π₀ (_^*ᴹ f λ S → sym (p S)) 136 | 137 | closure : (f : pos F ─m→ pos G) 138 | → (f⋆ : pos G ─m→ pos F) 139 | → [ f ⊣ f⋆ ] 140 | → (x : ∣ G ∣F) → [ f $ₘ (f⋆ $ₘ x) ⊑[ pos G ] x ] 141 | closure f f⋆ f⊣f⋆ x = π₁ (f⊣f⋆ (f⋆ $ₘ x) x) (⊑[ pos F ]-refl (f⋆ $ₘ x)) 142 | 143 | RAPL : (f : pos F ─m→ pos G) 144 | → (f⋁ : ((S : Fam 𝓥 ∣ F ∣F) → (⋁[ G ] ⁅ f $ₘ s ∣ s ε S ⁆) ≡ f $ₘ (⋁[ F ] S))) 145 | → (x y : ∣ G ∣F) → _^* f f⋁ (x ⊓[ G ] y) ≡ (_^* f f⋁ x) ⊓[ F ] (_^* f f⋁ y) 146 | RAPL 𝒻@(f , _) f⋁ x y = ⊑[ pos F ]-antisym _ _ G𝟏 G𝟐 147 | where 148 | f⊣f⋆ : [ 𝒻 ⊣ _^*ᴹ 𝒻 f⋁ ] 149 | f⊣f⋆ = ^*-RA 𝒻 f⋁ 150 | 151 | 𝒻⋆ : pos G ─m→ pos F 152 | 𝒻⋆ = _^*ᴹ 𝒻 f⋁ 153 | 154 | f⋆ : ∣ G ∣F → ∣ F ∣F 155 | f⋆ = _^* 𝒻 f⋁ 156 | 157 | G𝟏 : [ (𝒻 ^*) f⋁ (x ⊓[ G ] y) ⊑[ pos F ] (𝒻 ^*) f⋁ x ⊓[ F ] (𝒻 ^*) f⋁ y ] 158 | G𝟏 = ⊓[ F ]-greatest _ _ _ G𝟏a G𝟏b 159 | where 160 | G𝟏a : [ (𝒻 ^*) f⋁ (x ⊓[ G ] y) ⊑[ pos F ] (𝒻 ^*) f⋁ x ] 161 | G𝟏a = π₁ ((𝒻 ^*ᴹ) f⋁) _ _ (⊓[ G ]-lower₀ x y) 162 | 163 | G𝟏b : [ (𝒻 ^*) f⋁ (x ⊓[ G ] y) ⊑[ pos F ] (𝒻 ^*) f⋁ y ] 164 | G𝟏b = (π₁ ((𝒻 ^*ᴹ) f⋁) _ _ (⊓[ G ]-lower₁ x y)) 165 | 166 | G𝟐 : [ (𝒻 ^*) f⋁ x ⊓[ F ] (𝒻 ^*) f⋁ y ⊑[ pos F ] (𝒻 ^*) f⋁ (x ⊓[ G ] y) ] 167 | G𝟐 = π₀ (f⊣f⋆ (glb-of F ((𝒻 ^*) f⋁ x) ((𝒻 ^*) f⋁ y)) (x ⊓[ G ] y)) G𝟐a 168 | where 169 | G𝟐a : [ f (f⋆ x ⊓[ F ] f⋆ y) ⊑[ pos G ] x ⊓[ G ] y ] 170 | G𝟐a = ⊓[ G ]-greatest x y (f (f⋆ x ⊓[ F ] f⋆ y)) G𝟐b G𝟐c 171 | where 172 | G𝟐b : [ f (f⋆ x ⊓[ F ] f⋆ y) ⊑[ pos G ] x ] 173 | G𝟐b = f (f⋆ x ⊓[ F ] f⋆ y) ⊑⟨ ⦅𝟏⦆ ⟩ f (f⋆ x) ⊑⟨ ⦅𝟐⦆ ⟩ x ■ 174 | where 175 | ⦅𝟏⦆ = π₁ 𝒻 _ _ (⊓[ F ]-lower₀ _ _) 176 | ⦅𝟐⦆ = closure 𝒻 𝒻⋆ (^*-RA 𝒻 f⋁) x 177 | 178 | G𝟐c : [ f (f⋆ x ⊓[ F ] f⋆ y) ⊑[ pos G ] y ] 179 | G𝟐c = f (f⋆ x ⊓[ F ] f⋆ y) ⊑⟨ ⦅𝟏⦆ ⟩ f (f⋆ y) ⊑⟨ ⦅𝟐⦆ ⟩ y ■ 180 | where 181 | ⦅𝟏⦆ = π₁ 𝒻 _ _ (⊓[ F ]-lower₁ _ _) 182 | ⦅𝟐⦆ = closure 𝒻 𝒻⋆ (^*-RA 𝒻 f⋁) y 183 | ``` 184 | -------------------------------------------------------------------------------- /FormalTopology/SnocList.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Snoc List 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | open import Basis hiding (¬_) 9 | open import Cubical.Data.Empty.Base using (⊥; rec) 10 | open import Cubical.Data.Empty.Properties using (isProp⊥) 11 | open import Cubical.Relation.Nullary.DecidableEq using (Discrete→isSet) 12 | open import Cubical.Relation.Nullary using (Discrete; yes; no; Dec; ¬_) 13 | 14 | module SnocList (Z : Type ℓ-zero) (_≟_ : Discrete Z) where 15 | 16 | data SnocList : Type ℓ-zero where 17 | [] : SnocList 18 | _⌢_ : SnocList → Z → SnocList 19 | 20 | infixl 5 _⌢_ 21 | 22 | _elem_ : (z : Z) → SnocList → hProp ℓ-zero 23 | z elem [] = ⊥ , isProp⊥ 24 | z elem (zs ⌢ z′) with z ≟ z′ 25 | z elem (zs ⌢ z′) | yes p = (Unit ℓ-zero) , Unit-prop 26 | z elem (zs ⌢ z′) | no ¬p = z elem zs 27 | 28 | ⌢-eq-left : {xs ys : SnocList} {x y : Z} → xs ⌢ x ≡ ys ⌢ y → xs ≡ ys 29 | ⌢-eq-left {ys = ys} p = subst (λ { (zs ⌢ _) → zs ≡ ys ; [] → Z }) (sym p) refl 30 | 31 | ⌢-eq-right : {xs ys : SnocList} {x y : Z} → xs ⌢ x ≡ ys ⌢ y → x ≡ y 32 | ⌢-eq-right {x = x} {y = y} p with x ≟ y 33 | ⌢-eq-right {x = x} {y = y} p | yes q = q 34 | ⌢-eq-right {x = x} {y = y} p | no ¬q = sym (subst P p refl) 35 | where 36 | P : SnocList → Type ℓ-zero 37 | P [] = ⊥ 38 | P (_ ⌢ z) = z ≡ x 39 | 40 | ⌢≠[] : {xs : SnocList} {x : Z} → ¬ (xs ⌢ x ≡ []) 41 | ⌢≠[] p = subst P p tt 42 | where 43 | P : SnocList → Type ℓ-zero 44 | P [] = ⊥ 45 | P (_ ⌢ _) = Unit ℓ-zero 46 | 47 | 48 | SnocList-discrete : Discrete SnocList 49 | SnocList-discrete [] [] = yes refl 50 | SnocList-discrete [] (ys ⌢ y) = no λ p → ⌢≠[] (sym p) 51 | SnocList-discrete (xs ⌢ x) [] = no λ p → ⌢≠[] p 52 | SnocList-discrete (xs ⌢ x) (ys ⌢ y) 53 | with x ≟ y | SnocList-discrete xs ys 54 | ... | _ | no ¬q = no λ q → ¬q (⌢-eq-left q) 55 | ... | no ¬p | _ = no λ p → ¬p (⌢-eq-right p) 56 | ... | yes p | yes q = subst P (sym q) (subst Q (sym p) (yes refl)) 57 | where 58 | P = λ xs′ → Dec (xs′ ⌢ x ≡ ys ⌢ y) 59 | Q = λ x′ → Dec (ys ⌢ x′ ≡ ys ⌢ y) 60 | 61 | SnocList-set : isSet SnocList 62 | SnocList-set = Discrete→isSet SnocList-discrete 63 | 64 | _++_ : SnocList → SnocList → SnocList 65 | xs ++ [] = xs 66 | xs ++ (ys ⌢ y) = (xs ++ ys) ⌢ y 67 | 68 | elem-lemma : (xs ys : SnocList) (z : Z) → [ z elem xs ] → [ z elem (xs ++ ys) ] 69 | elem-lemma xs [] z z∈xs = z∈xs 70 | elem-lemma xs (ys ⌢ y) z z∈xs with z ≟ y 71 | elem-lemma xs (ys ⌢ y) z z∈xs | yes p = tt 72 | elem-lemma xs (ys ⌢ y) z z∈xs | no ¬p = elem-lemma xs ys z z∈xs 73 | 74 | []≠⌢++ : (xs ys : SnocList) (x : Z) → ¬ ([] ≡ ((xs ⌢ x) ++ ys)) 75 | []≠⌢++ xs [] x p = ⌢≠[] (sym p) 76 | []≠⌢++ xs (ys ⌢ y) x p = ⌢≠[] (sym p) 77 | 78 | snoc-lemma : (xs ys : SnocList) (x : Z) → (xs ⌢ x) ++ ys ≡ xs ++ (([] ⌢ x) ++ ys) 79 | snoc-lemma xs [] x = refl 80 | snoc-lemma xs (ys ⌢ y) x = cong (λ - → - ⌢ y) (snoc-lemma xs ys x) 81 | 82 | ++-left-id : (xs : SnocList) → [] ++ xs ≡ xs 83 | ++-left-id [] = refl 84 | ++-left-id (xs ⌢ x) = cong (λ - → - ⌢ x) (++-left-id xs) 85 | 86 | assoc : (xs ys zs : SnocList) → xs ++ (ys ++ zs) ≡ (xs ++ ys) ++ zs 87 | assoc xs ys [] = refl 88 | assoc xs ys (zs ⌢ z) = cong (λ - → - ⌢ z) (assoc xs ys zs) 89 | 90 | xs≠xs⌢y : {xs : SnocList} {y : Z} → ¬ (xs ≡ xs ⌢ y) 91 | xs≠xs⌢y {[]} p = subst (λ { [] → Unit ℓ-zero ; (_ ⌢ _) → ⊥ }) p tt 92 | xs≠xs⌢y {xs ⌢ x} p = rec (xs≠xs⌢y (⌢-eq-left p)) 93 | 94 | xs≰xs⌢y : {xs ys : SnocList} {x : Z} → ¬ (xs ≡ (xs ⌢ x) ++ ys) 95 | xs≰xs⌢y {[]} {[]} {y} p = subst (λ { (_ ⌢ _) → ⊥ ; [] → Unit ℓ-zero }) p tt 96 | xs≰xs⌢y {[]} {ys ⌢ _} {y} p = subst (λ { (_ ⌢ _) → ⊥ ; [] → Unit ℓ-zero }) p tt 97 | xs≰xs⌢y {xs ⌢ x} {[]} {y} p = rec (xs≠xs⌢y (⌢-eq-left p)) 98 | xs≰xs⌢y {xs ⌢ x} {ys ⌢ y} {y′} p = rec (xs≰xs⌢y NTS) 99 | where 100 | NTS : xs ≡ (xs ⌢ x) ++ (([] ⌢ y′) ++ ys) 101 | NTS = xs ≡⟨ ⌢-eq-left p ⟩ 102 | (xs ⌢ x ⌢ y′) ++ ys ≡⟨ snoc-lemma ((xs ⌢ x) ++ []) ys y′ ⟩ 103 | (xs ⌢ x) ++ (([] ⌢ y′) ++ ys) ∎ 104 | 105 | lemma3 : {xs : SnocList} {ys : SnocList} {y : Z} → ¬ (xs ≡ xs ++ (ys ⌢ y)) 106 | lemma3 {[]} {ys} {y} p = rec (⌢≠[] NTS) 107 | where 108 | NTS : ys ⌢ y ≡ [] 109 | NTS = (ys ⌢ y) ≡⟨ sym (cong (λ - → - ⌢ y) (++-left-id ys)) ⟩ 110 | ([] ++ ys) ⌢ y ≡⟨ sym p ⟩ 111 | [] ∎ 112 | lemma3 {xs ⌢ x} {[]} {y} p = rec (lemma3 (⌢-eq-left p)) 113 | lemma3 {xs ⌢ x} {ys ⌢ y} {y′} p = rec (lemma3 NTS) 114 | where 115 | NTS : xs ≡ (xs ++ (([] ⌢ x) ++ ys)) ⌢ y 116 | NTS = xs ≡⟨ ⌢-eq-left p ⟩ 117 | ((xs ⌢ x) ++ ys) ⌢ y ≡⟨ snoc-lemma xs (ys ⌢ y) x ⟩ 118 | xs ++ (([] ⌢ x) ++ ys) ⌢ y ∎ 119 | 120 | 121 | ++-lemma : {xs ys zs₀ zs₁ : SnocList} → xs ≡ ys ++ zs₀ → xs ≡ ys ++ zs₁ → zs₀ ≡ zs₁ 122 | ++-lemma {[]} {[]} {zs₀} {zs₁} p q = zs₀ ≡⟨ sym (++-left-id zs₀) ⟩ 123 | [] ++ zs₀ ≡⟨ sym p ⟩ 124 | [] ≡⟨ q ⟩ 125 | [] ++ zs₁ ≡⟨ ++-left-id zs₁ ⟩ 126 | zs₁ ∎ 127 | ++-lemma {[]} {ys ⌢ y} {[]} {[]} p q = refl 128 | ++-lemma {[]} {ys ⌢ y} {[]} {zs₁ ⌢ x} p q = rec (⌢≠[] (sym p)) 129 | ++-lemma {[]} {ys ⌢ y} {zs₀ ⌢ x} {[]} p q = rec (⌢≠[] (sym q)) 130 | ++-lemma {[]} {ys ⌢ y} {zs₀ ⌢ x} {zs₁ ⌢ x₁} p q = rec (⌢≠[] (sym p)) 131 | ++-lemma {xs ⌢ x} {[]} {[]} {zs₁} p q = rec (⌢≠[] p) 132 | ++-lemma {xs ⌢ x} {[]} {zs₀ ⌢ z₀} {[]} p q = rec (⌢≠[] q) 133 | ++-lemma {xs ⌢ x} {[]} {zs₀ ⌢ z₀} {zs₁ ⌢ z₁} p q = 134 | zs₀ ⌢ z₀ ≡⟨ cong (λ - → - ⌢ z₀) (sym (++-left-id zs₀)) ⟩ 135 | ([] ++ zs₀) ⌢ z₀ ≡⟨ sym p ⟩ 136 | xs ⌢ x ≡⟨ q ⟩ 137 | ([] ++ zs₁) ⌢ z₁ ≡⟨ cong (λ - → - ⌢ z₁) (++-left-id zs₁) ⟩ 138 | zs₁ ⌢ z₁ ∎ 139 | ++-lemma {xs ⌢ x} {ys ⌢ y} {[]} {[]} p q = refl 140 | ++-lemma {xs ⌢ x} {ys ⌢ y} {[]} {zs₁ ⌢ z₁} p q = rec (xs≰xs⌢y (⌢-eq-left NTS)) 141 | where 142 | NTS : ys ⌢ y ≡ ((ys ⌢ y) ++ zs₁) ⌢ z₁ 143 | NTS = ys ⌢ y ≡⟨ sym p ⟩ xs ⌢ x ≡⟨ q ⟩ ((ys ⌢ y) ++ zs₁) ⌢ z₁ ∎ 144 | ++-lemma {xs ⌢ x} {ys ⌢ y} {zs₀ ⌢ z₀} {[]} p q = rec (xs≰xs⌢y (⌢-eq-left NTS)) 145 | where 146 | NTS : ys ⌢ y ≡ ((ys ⌢ y) ++ zs₀) ⌢ z₀ 147 | NTS = ys ⌢ y ≡⟨ sym q ⟩ xs ⌢ x ≡⟨ p ⟩ ((ys ⌢ y) ++ (zs₀ ⌢ z₀)) ∎ 148 | ++-lemma {xs ⌢ x} {ys ⌢ y} {zs₀ ⌢ z₀} {zs₁ ⌢ z₁} p q with z₀ ≟ z₁ 149 | ... | yes z₀=z₁ = zs₀ ⌢ z₀ ≡⟨ cong (λ - → - ⌢ z₀) IH ⟩ 150 | zs₁ ⌢ z₀ ≡⟨ cong (_⌢_ zs₁) z₀=z₁ ⟩ 151 | zs₁ ⌢ z₁ ∎ 152 | where 153 | IH : zs₀ ≡ zs₁ 154 | IH = ++-lemma (⌢-eq-left p) (⌢-eq-left q) 155 | ... | no ¬p = rec (¬p (⌢-eq-right NTS)) 156 | where 157 | NTS : ((ys ⌢ y) ++ zs₀) ⌢ z₀ ≡ ((ys ⌢ y) ++ zs₁) ⌢ z₁ 158 | NTS = (ys ⌢ y) ++ zs₀ ⌢ z₀ ≡⟨ sym p ⟩ xs ⌢ x ≡⟨ q ⟩ (ys ⌢ y) ++ zs₁ ⌢ z₁ ∎ 159 | 160 | nonempty : SnocList → Type ℓ-zero 161 | nonempty [] = ⊥ 162 | nonempty (xs ⌢ x) = Unit ℓ-zero 163 | 164 | head : (xs : SnocList) → nonempty xs → Z 165 | head ([] ⌢ x) p = x 166 | head ((xs ⌢ x′) ⌢ x) p = head (xs ⌢ x′) tt 167 | 168 | tail : (xs : SnocList) → nonempty xs → SnocList 169 | tail ([] ⌢ x) tt = [] 170 | tail (xs ⌢ x₀ ⌢ x₁) p = (tail (xs ⌢ x₀) tt) ⌢ x₁ 171 | 172 | hd-tl-lemma : (xs : SnocList) (ne : nonempty xs) → ([] ⌢ head xs ne) ++ tail xs ne ≡ xs 173 | hd-tl-lemma ([] ⌢ x) tt = refl 174 | hd-tl-lemma (xs ⌢ x₀ ⌢ x₁) tt = cong (λ - → - ⌢ x₁) (hd-tl-lemma (xs ⌢ x₀) tt) 175 | -------------------------------------------------------------------------------- /FormalTopology/Omega.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: The Initial Frame 3 | author: Ayberk Tosun
(j.w.w. with Martín Escardó) 4 | --- 5 | 6 | 20 | 21 | This module contains a construction of the initial object in the category 22 | **Frm** of frames. 23 | 24 | ## The underlying poset of the frame 25 | 26 | Let us first write the down the partial order underlying the frame: 27 | 28 | ```agda 29 | _≤_ : hProp 𝓤 → hProp 𝓤 → hProp 𝓤 30 | A ≤ B = A ⇒ B 31 | ``` 32 | 33 | This gives us a poset structure at every universe 𝓤: 34 | 35 | ```agda 36 | ΩP-str : (𝓤 : Universe) → PosetStr 𝓤 (hProp 𝓤) 37 | ΩP-str 𝓤 = _≤_ , isSetHProp , ≤-reflexive , ≤-trans , ≤-antisym 38 | where 39 | ≤-reflexive : [ isReflexive _≤_ ] 40 | ≤-reflexive x = idfun _ 41 | 42 | ≤-trans : [ isTransitive _≤_ ] 43 | ≤-trans _ _ _ p q = q ∘ p 44 | 45 | ≤-antisym : [ isAntisym isSetHProp _≤_ ] 46 | ≤-antisym _ _ = ⇔toPath 47 | ``` 48 | 49 | We will denote this poset by ΩP: 50 | 51 | ```agda 52 | ΩP : (𝓤 : Universe) → Poset (𝓤 ⁺) 𝓤 53 | ΩP 𝓤 = hProp 𝓤 , ΩP-str 𝓤 54 | ``` 55 | 56 | ## Definition of the initial frame 57 | 58 | ```agda 59 | Ω : (𝓤 : Universe) → Frame (𝓤 ⁺) 𝓤 𝓤 60 | Ω 𝓤 = 61 | hProp 𝓤 , (ΩP-str 𝓤 , top 𝓤 , _⊓_ , ⋁_) , is-top , is-glb , is-lub , distr 62 | where 63 | ⋁_ : Fam 𝓤 (hProp 𝓤) → hProp 𝓤 64 | ⋁ (I , f) = ∥ (Σ[ i ∈ I ] [ f i ]) ∥Ω 65 | 66 | is-top : [ isTop (ΩP 𝓤) (top 𝓤) ] 67 | is-top _ _ = tt 68 | 69 | is-glb : [ isGLB (ΩP 𝓤) _⊓_ ] 70 | is-glb = (λ _ _ → fst , snd) , λ x y z p q → fst p q , snd p q 71 | 72 | is-lub : [ isLUB (ΩP 𝓤) ⋁_ ] 73 | is-lub = (λ { U _ (i , eq) q → ∣ i , subst [_] (sym eq) q ∣ }) , nts 74 | where 75 | nts : (U : Fam 𝓤 (hProp 𝓤)) (A : hProp 𝓤) 76 | → [ ∀[ B ε U ] (B ≤ A) ⇒ (⋁ U) ≤ A ] 77 | nts U x p q = ∥∥-rec (isProp[] x) rem q 78 | where 79 | rem : Σ[ i ∈ index U ] [ U $ i ] → [ x ] 80 | rem (i , eq) = p (U $ i) (i , refl) eq 81 | 82 | distr : [ isDist (ΩP 𝓤) _⊓_ ⋁_ ] 83 | distr A U = ⇔toPath nts₀ nts₁ 84 | where 85 | open JoinSyntax (hProp 𝓤) ⋁_ 86 | 87 | nts₀ : _ 88 | nts₀ (x , y) = ∥∥-rec (∥∥-prop _) (λ { (i , Uᵢ) → ∣ i , x , Uᵢ ∣ }) y 89 | 90 | nts₁ : [ ⋁⟨ i ⟩ (A ⊓ (U $ i)) ⇒ A ⊓ (⋁ U) ] 91 | nts₁ y = ∥∥-rec (isProp[] (A ⊓ (⋁ U))) (λ { (i , x , y) → x , ∣ i , y ∣ }) y 92 | ``` 93 | 94 | ## Initiality 95 | 96 | We now construct, for each frame $A$, a frame homomorphism ‼ : Ω → A. Let us 97 | start by writing down the underlying function: 98 | 99 | ```agda 100 | ∣‼∣ : (A : Frame 𝓦 𝓥 𝓤) → hProp 𝓤 → ∣ A ∣F 101 | ∣‼∣ A P = ⋁[ A ] ⁅ ⊤[ A ] ∣ _ ∶ [ P ] ⁆ 102 | ``` 103 | 104 | ### ‼ is monotonic 105 | 106 | This is a monotonic map between the underlying frames: 107 | 108 | ```agda 109 | ∣‼∣-mono : (A : Frame 𝓦 𝓥 𝓤) → isMonotonic (ΩP 𝓤) (pos A) (∣‼∣ A) 110 | ∣‼∣-mono A P Q P≤Q = 111 | ⋁[ A ]-least _ (∣‼∣ A Q) nts 112 | where 113 | nts : _ 114 | nts x (p , eq) = ⋁[ A ]-upper _ _ (P≤Q p , eq) 115 | ``` 116 | 117 | ```agda 118 | ‼m : (A : Frame 𝓦 𝓥 𝓤) → ΩP 𝓤 ─m→ pos A 119 | ‼m A = ∣‼∣ A , ∣‼∣-mono A 120 | ``` 121 | 122 | ### ‼ is a frame homomorphism 123 | 124 | ```agda 125 | ∣‼∣-resp-⊤ : (A : Frame 𝓦 𝓥 𝓤) → ∣‼∣ A (top 𝓤) ≡ ⊤[ A ] 126 | ∣‼∣-resp-⊤ A = 127 | ⊑[ pos A ]-antisym _ _ (⊤[ A ]-top _) (⋁[ A ]-upper _ _ (tt , refl)) 128 | ``` 129 | 130 | ```agda 131 | ∣‼∣-resp-∧ : (A : Frame 𝓦 𝓥 𝓤) (P Q : hProp 𝓤) 132 | → ∣‼∣ A (P ⊓ Q) ≡ (∣‼∣ A P) ⊓[ A ] (∣‼∣ A Q) 133 | ∣‼∣-resp-∧ {𝓤 = 𝓤} A P Q = 134 | ∣‼∣ A (P ⊓ Q) ≡⟨ refl ⟩ 135 | ⋁ ⁅ 𝟏 ∣ _ ∶ [ P ⊓ Q ] ⁆ ≡⟨ cong (λ - → ⋁ ⁅ - ∣ _ ∶ _ ⁆) p ⟩ 136 | ⋁ ⁅ 𝟏 ⊓[ A ] 𝟏 ∣ _ ∶ [ P ⊓ Q ] ⁆ ≡⟨ sym (sym-distr A _ _ ) ⟩ 137 | ∣‼∣ A P ⊓[ A ] ∣‼∣ A Q ∎ 138 | where 139 | 𝟏 = ⊤[ A ] 140 | 141 | ⋁_ : Fam 𝓤 ∣ A ∣F → ∣ A ∣F 142 | ⋁ U = ⋁[ A ] U 143 | 144 | p : 𝟏 ≡ 𝟏 ⊓[ A ] 𝟏 145 | p = sym (x∧x=x A 𝟏) 146 | ``` 147 | 148 | ```agda 149 | ∣‼∣-resp-⋁ : (A : Frame 𝓦 𝓥 𝓤) (U : Fam 𝓤 (hProp 𝓤)) 150 | → ∣‼∣ A (⋁[ Ω 𝓤 ] U) ≡ ⋁[ A ] ⁅ ∣‼∣ A x ∣ x ε U ⁆ 151 | ∣‼∣-resp-⋁ A U = ⊑[ pos A ]-antisym _ _ below above 152 | where 153 | open PosetNotation (pos A) using () renaming (_≤_ to _⊑_) 154 | 155 | below : [ ∣‼∣ A (⋁[ Ω _ ] U) ⊑ (⋁[ A ] ⁅ ∣‼∣ A P ∣ P ε U ⁆) ] 156 | below = ⋁[ A ]-least _ _ goal 157 | where 158 | goal : _ 159 | goal x (q , eq) = ∥∥-rec (isProp[] (_ ⊑[ pos A ] _)) rem q 160 | where 161 | rem : Σ[ i ∈ index U ] [ U $ i ] → [ x ⊑ (⋁[ A ] ⁅ ∣‼∣ A P ∣ P ε U ⁆) ] 162 | rem (i , p) = ⋁[ A ]-upper _ _ (i , sym (⋁-unique A _ x γ δ)) 163 | where 164 | γ : _ 165 | γ y (j , q) = subst (λ - → [ rel (pos A) y - ]) eq (⊤[ A ]-top y) 166 | 167 | δ : _ 168 | δ w q = q x (p , eq) 169 | 170 | above : [ (⋁[ A ] ⁅ ∣‼∣ A P ∣ P ε U ⁆) ⊑ ∣‼∣ A (⋁[ Ω _ ] U) ] 171 | above = ⋁[ A ]-least _ _ goal 172 | where 173 | goal : _ 174 | goal x (i , eq) = subst (λ - → [ - ⊑[ pos A ] _ ]) eq nts 175 | where 176 | rem : _ 177 | rem x (p , eq-j) = ⋁[ A ]-upper _ _ (∣ i , p ∣ , eq-j) 178 | 179 | nts : [ ∣‼∣ A (U $ i) ⊑ ∣‼∣ A (⋁[ Ω _ ] U) ] 180 | nts = ⋁[ A ]-least _ _ rem 181 | ``` 182 | 183 | ## Definition of ‼ 184 | 185 | ```agda 186 | ‼ : (𝓤 : Universe) (A : Frame 𝓦 𝓥 𝓤) → Ω 𝓤 ─f→ A 187 | ‼ 𝓤 A = ‼m A , ∣‼∣-resp-⊤ A , ∣‼∣-resp-∧ A , ∣‼∣-resp-⋁ A 188 | ``` 189 | 190 | ## Uniqueness of ‼ 191 | 192 | To prove uniqueness, we will need the following easy lemma: 193 | 194 | ```agda 195 | main-lemma : (P : hProp 𝓤) → [ P ⊑[ ΩP 𝓤 ] ⋁[ Ω 𝓤 ] ⁅ top 𝓤 ∣ _ ∶ [ P ] ⁆ ] 196 | main-lemma {𝓤 = 𝓤} P p = 197 | ⋁[ Ω 𝓤 ]-upper (⁅ top 𝓤 ∣ _ ∶ [ P ] ⁆) _ (p , refl) tt 198 | ``` 199 | 200 | We now prove uniqueness: given any other frame homomorphism ⁇ out of Ω, ‼ is 201 | equal to ⁇: 202 | 203 | ```agda 204 | ‼-is-unique : (A : Frame 𝓦 𝓥 𝓤) → (⁇ : (Ω 𝓤) ─f→ A) → ‼ 𝓤 A ≡ ⁇ 205 | ‼-is-unique {𝓤 = 𝓤} A ⁇@((∣⁇∣ , ⁇-mono) , (⁇-resp-⊤ , ⁇-resp-∧ , ⁇-resp-⋁)) = 206 | forget-homo (Ω 𝓤) A (‼ 𝓤 A) ⁇ (sym ∘ goal) 207 | where 208 | open PosetNotation (pos A) using () renaming (_≤_ to _⊑_) 209 | open PosetReasoning (pos A) using (_⊑⟨_⟩_; _■) 210 | 𝟏 = ⊤[ A ] 211 | 212 | goal : (P : hProp 𝓤) → ∣⁇∣ P ≡ ∣‼∣ A P 213 | goal P = ⋁-unique A _ _ upper-bound least 214 | where 215 | upper-bound : (x : ∣ A ∣F) → x ε ⁅ 𝟏 ∣ _ ∶ [ P ] ⁆ → [ x ⊑ ∣⁇∣ P ] 216 | upper-bound x (p , eq) = 217 | x ⊑⟨ ≡⇒⊑ (pos A) (sym eq) ⟩ 218 | 𝟏 ⊑⟨ ≡⇒⊑ (pos A) (sym ⁇-resp-⊤) ⟩ 219 | ∣⁇∣ (top 𝓤) ⊑⟨ ⁇-mono (top 𝓤) P (const p) ⟩ 220 | ∣⁇∣ P ■ 221 | 222 | least : (u : ∣ A ∣F) 223 | → ((x : ∣ A ∣F) → x ε ⁅ 𝟏 ∣ _ ∶ [ P ] ⁆ → [ x ⊑ u ]) → [ ∣⁇∣ P ⊑ u ] 224 | least u u-upper-bound = 225 | ∣⁇∣ P ⊑⟨ ⁇-mono P _ (main-lemma P) ⟩ 226 | ∣⁇∣ (⋁[ Ω 𝓤 ] ⁅ top 𝓤 ∣ p ∶ [ P ] ⁆) ⊑⟨ ≡⇒⊑ (pos A) (⁇-resp-⋁ _) ⟩ 227 | ⋁[ A ] ⁅ ∣⁇∣ (top 𝓤) ∣ p ∶ [ P ] ⁆ ⊑⟨ † ⟩ 228 | ⋁[ A ] ⁅ ⊤[ A ] ∣ p ∶ [ P ] ⁆ ⊑⟨ ⋁[ A ]-least _ _ u-upper-bound ⟩ 229 | u ■ 230 | where 231 | † : [ (⋁[ A ] ⁅ ∣⁇∣ (top 𝓤) ∣ p ∶ [ P ] ⁆) ⊑ (⋁[ A ] ⁅ 𝟏 ∣ p ∶ [ P ] ⁆) ] 232 | † = ≡⇒⊑ (pos A) (cong (λ - → (⋁[ A ] ⁅ - ∣ p ∶ [ P ] ⁆)) ⁇-resp-⊤) 233 | ``` 234 | 235 | ```agda 236 | ΩF-initial : (A : Frame 𝓦 𝓥 𝓤) → isContr (Ω 𝓤 ─f→ A) 237 | ΩF-initial {𝓤 = 𝓤} A = ‼ 𝓤 A , ‼-is-unique A 238 | ``` 239 | -------------------------------------------------------------------------------- /FormalTopology/Stone.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Stone Locales 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 24 | 25 | ## Definition of Stone 26 | 27 | ```agda 28 | isComplemented : Fam 𝓦 ∣ F ∣F → (𝓤 ∨ 𝓦) ̇ 29 | isComplemented S = (x : ∣ F ∣F) → x ε S → hasComplement F x 30 | ``` 31 | 32 | ```agda 33 | isZeroDimensionalₛ : (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 34 | isZeroDimensionalₛ = Σ[ ℬ ∈ Fam 𝓦 ∣ F ∣F ] isBasisFor F ℬ × isComplemented ℬ 35 | ``` 36 | 37 | ```agda 38 | isZeroDimensional : hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 39 | isZeroDimensional = ∥ isZeroDimensionalₛ ∥Ω 40 | ``` 41 | 42 | ```agda 43 | isStone′ : hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 44 | isStone′ = isCompact F ⊓ isZeroDimensional 45 | ``` 46 | 47 | ```agda 48 | isStoneₛ : (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 49 | isStoneₛ = [ isCompact F ] × isZeroDimensionalₛ 50 | ``` 51 | 52 | ```agda 53 | ⋜→≪ : [ isCompact F ] 54 | → (x y : ∣ F ∣F) → x ⋜[ F ] y → [ _≪_ F x y ] 55 | ⋜→≪ F-comp x y (z , comp₀ , comp₁) S S-dir p = 56 | ∥∥-rec (∥∥-prop _) main rem 57 | where 58 | open PosetReasoning (pos F) 59 | 60 | nts : [ ⊤[ F ] ⊑[ pos F ] ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆ ] 61 | nts = ⊤[ F ] ⊑⟨ ≡⇒⊑ (pos F) (sym comp₁) ⟩ 62 | y ∨[ F ] z ⊑⟨ ∨-cleft F _ _ _ p ⟩ 63 | (⋁[ F ] S) ∨[ F ] z ⊑⟨ G𝟏 ⟩ 64 | (⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆) ■ 65 | where 66 | G𝟏 : [ (⋁[ F ] S) ∨[ F ] z ⊑[ pos F ] ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆ ] 67 | G𝟏 = ⋁[ F ]-least _ _ G𝟐 68 | where 69 | G𝟐 : [ ∀[ w ε ⁅ (⋁[ F ] S) , z ⁆ ] (w ⊑[ pos F ] ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆) ] 70 | G𝟐 = ∥∥-rec (isProp[] (∀[ w ε ⁅ (⋁[ F ] S) , z ⁆ ] (_ ⊑[ pos F ] _))) G𝟑 (π₀ S-dir) 71 | where 72 | G𝟑 : index S 73 | → [ ∀[ w ε ⁅ (⋁[ F ] S) , z ⁆ ] (w ⊑[ pos F ] ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆) ] 74 | G𝟑 i w (true , eq) = w ⊑⟨ ≡⇒⊑ (pos F) (sym eq) ⟩ 75 | ⋁[ F ] S ⊑⟨ ⋁[ F ]-least _ _ G𝟒 ⟩ 76 | ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆ ■ 77 | where 78 | G𝟒 : _ 79 | G𝟒 v (k , eq′) = 80 | v ⊑⟨ ≡⇒⊑ (pos F) (sym eq′) ⟩ 81 | S $ k ⊑⟨ ⋁[ F ]-upper _ _ (true , refl) ⟩ 82 | (S $ k) ∨[ F ] z ⊑⟨ ⋁[ F ]-upper _ _ (k , refl) ⟩ 83 | ⋁[ F ] ⁅ s ∨[ F ] z ∣ s ε S ⁆ ■ 84 | G𝟑 i w (false , eq) = 85 | w ⊑⟨ ≡⇒⊑ (pos F) (sym eq) ⟩ 86 | z ⊑⟨ ⋁[ F ]-upper _ _ (false , refl) ⟩ 87 | (S $ i) ∨[ F ] z ⊑⟨ ⋁[ F ]-upper _ _ (i , refl) ⟩ 88 | _ ■ 89 | 90 | dir : [ isDirected (pos F) (⁅ s ∨[ F ] z ∣ s ε S ⁆) ] 91 | π₀ dir = π₀ S-dir 92 | π₁ dir i j = ∥∥-rec (∥∥-prop _) G𝟏 (π₁ S-dir i j) 93 | where 94 | G𝟏 : _ 95 | G𝟏 (k , p , q) = ∣ k , (∨-cleft F _ _ _ p , ∨-cleft F _ _ _ q) ∣ 96 | 97 | rem : ∥ Σ[ i ∈ index S ] (S $ i) ∨[ F ] z ≡ ⊤[ F ] ∥ 98 | rem = ∥∥-rec (∥∥-prop _) goal (F-comp (⁅ s ∨[ F ] z ∣ s ε S ⁆) dir nts) 99 | where 100 | goal : _ 101 | goal (i , ϕ) = ∣ i , ⊑[ pos F ]-antisym _ _ (⊤[ F ]-top _) ϕ ∣ 102 | 103 | main : Σ[ i ∈ index S ] (S $ i) ∨[ F ] z ≡ ⊤[ F ] 104 | → ∥ Σ[ i ∈ index S ] [ x ⊑[ pos F ] (S $ i) ] ∥ 105 | main (i , ϕ) = ∣ i , G𝟏 ∣ 106 | where 107 | open SomePropertiesOf⋜ F 108 | 109 | G𝟏 : [ x ⊑[ pos F ] (S $ i) ] 110 | G𝟏 = a⋜b→a⊑b x (S $ i) (z , (comp₀ , ϕ)) 111 | ``` 112 | 113 | ```agda 114 | clopen→compact-in-compact-locale : [ isCompact F ] 115 | → (x : ∣ F ∣F) 116 | → hasComplement F x 117 | → [ _≪_ F x x ] 118 | clopen→compact-in-compact-locale ⊤≪⊤ x x-clop = ⋜→≪ ⊤≪⊤ x x x-clop 119 | ``` 120 | 121 | ```agda 122 | directify-clopen : (ℬ : Fam 𝓦 ∣ F ∣F) 123 | → isComplemented ℬ 124 | → isComplemented (directify F ℬ) 125 | directify-clopen ℬ@(I , β) ψ b (is , eq) = subst (hasComplement F) eq (nts is) 126 | where 127 | nts : (is : List I) → hasComplement F (directify F ℬ $ is) 128 | nts [] = subst (hasComplement F) refl (⊤[ F ] , G𝟏 , G𝟐) 129 | where 130 | G𝟏 : ⊥[ F ] ⊓[ F ] ⊤[ F ] ≡ ⊥[ F ] 131 | G𝟏 = x∧⊤=x F ⊥[ F ] 132 | 133 | G𝟐 : ⊥[ F ] ∨[ F ] ⊤[ F ] ≡ ⊤[ F ] 134 | G𝟐 = x∨⊥=x F ⊤[ F ] 135 | nts (i ∷ is) = subst (hasComplement F) refl goal 136 | where 137 | ¬βᵢ : ∣ F ∣F 138 | ¬βᵢ = π₀ (ψ (β i) (i , refl)) 139 | 140 | ¬dir : ∣ F ∣F 141 | ¬dir = π₀ (nts is) 142 | 143 | goal : hasComplement F (β i ∨[ F ] (directify F ℬ $ is)) 144 | goal = (¬βᵢ ⊓[ F ] ¬dir) 145 | , (complements-sym F (∧-complement F _ _ _ _ (complements-sym F (π₁ (ψ (β i) (i , refl)))) ((complements-sym F (π₁ (nts is)))))) 146 | ``` 147 | 148 | ```agda 149 | clopen-basis-∧-complement : (ℬ : Fam 𝓦 ∣ F ∣F) 150 | → isComplemented ℬ 151 | → (i j : index ℬ) 152 | → hasComplement F ((ℬ $ i) ⊓[ F ] (ℬ $ j)) 153 | clopen-basis-∧-complement ℬ κ i j = 154 | (¬ℬᵢ ∨[ F ] ¬ℬⱼ) 155 | , ∧-complement F (ℬ $ i) (ℬ $ j) ¬ℬᵢ ¬ℬⱼ ¬ℬᵢ-complements-ℬᵢ ¬ℬⱼ-complements-ℬⱼ 156 | where 157 | ¬ℬᵢ : ∣ F ∣F 158 | ¬ℬᵢ = π₀ (κ (ℬ $ i) (i , refl)) 159 | 160 | ¬ℬⱼ : ∣ F ∣F 161 | ¬ℬⱼ = π₀ (κ (ℬ $ j) (j , refl)) 162 | 163 | ¬ℬᵢ-complements-ℬᵢ : complements F (ℬ $ i) ¬ℬᵢ 164 | ¬ℬᵢ-complements-ℬᵢ = π₁ (κ (ℬ $ i) (i , refl)) 165 | 166 | ¬ℬⱼ-complements-ℬⱼ : complements F (ℬ $ j) ¬ℬⱼ 167 | ¬ℬⱼ-complements-ℬⱼ = π₁ (κ (ℬ $ j) (j , refl)) 168 | ``` 169 | 170 | ```agda 171 | compact→basic : (ℬ : Fam 𝓦 ∣ F ∣F) → isDirBasisFor F ℬ 172 | → (x : ∣ F ∣F) → [ _≪_ F x x ] → ∥ x ε ℬ ∥ 173 | compact→basic ℬ basis x x≪x = ∥∥-rec (∥∥-prop (x ε ℬ)) goal (x≪x C C-dir C-covers-x₀) 174 | where 175 | open PosetReasoning (pos F) 176 | 177 | 𝒥 : Fam 𝓦 (index ℬ) 178 | 𝒥 = π₀ (basis x) 179 | 180 | C : Fam 𝓦 ∣ F ∣F 181 | C = ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ 182 | 183 | C-dir : [ isDirected (pos F) C ] 184 | C-dir = π₀ (π₁ (basis x)) 185 | 186 | C-covers-x : x ≡ (⋁[ F ] C) 187 | C-covers-x = ⋁-unique F _ _ (π₀ (π₁ (π₁ (basis x)))) (π₁ (π₁ (π₁ (basis x)))) 188 | 189 | C-covers-x₀ : [ x ⊑[ pos F ] ⋁[ F ] C ] 190 | C-covers-x₀ = ≡⇒⊑ (pos F) C-covers-x 191 | 192 | goal : Σ[ i ∈ index C ] [ x ⊑[ pos F ] (C $ i) ] → ∥ x ε ℬ ∥ 193 | goal (i , x≤cᵢ) = ∣ 𝒥 $ i , ⊑[ pos F ]-antisym _ _ down up ∣ 194 | where 195 | down : [ (ℬ $ (𝒥 $ i)) ⊑[ pos F ] x ] 196 | down = ℬ $ 𝒥 $ i ⊑⟨ ⋁[ F ]-upper _ _ (i , refl) ⟩ 197 | ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ⊑⟨ ≡⇒⊑ (pos F) (sym C-covers-x) ⟩ 198 | x ■ 199 | 200 | up : [ x ⊑[ pos F ] (ℬ $ (𝒥 $ i)) ] 201 | up = x≤cᵢ 202 | ``` 203 | 204 | ```agda 205 | clopen→basic-in-stone-locale : isStoneₛ 206 | → (ℬ : Fam 𝓦 ∣ F ∣F) 207 | → isDirBasisFor F ℬ 208 | → (x : ∣ F ∣F) → hasComplement F x → ∥ x ε ℬ ∥ 209 | clopen→basic-in-stone-locale (⊤≪⊤ , _) ℬ basis x x-clop = 210 | compact→basic ℬ basis x (clopen→compact-in-compact-locale ⊤≪⊤ x x-clop) 211 | where 212 | 𝒥 = π₀ (basis x) 213 | 214 | C : Fam 𝓦 ∣ F ∣F 215 | C = ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ 216 | ``` 217 | 218 | ```agda 219 | stone→spectral : isStoneₛ → isSpectralₛ F 220 | stone→spectral stone@(⊤≪⊤ , ℬ , b , κ) = ℬ↑ , κ↑ , dir-basis , ∧-closure 221 | where 222 | ℬ↑ : Fam 𝓦 ∣ F ∣F 223 | ℬ↑ = directify F ℬ 224 | 225 | ζ : (is : index ℬ↑) → hasComplement F (ℬ↑ $ is) 226 | ζ is = directify-clopen ℬ κ (ℬ↑ $ is) (is , refl) 227 | 228 | κ↑ : (is : List (index ℬ)) → [ isCompactOpen F (ℬ↑ $ is) ] 229 | κ↑ is = clopen→compact-in-compact-locale ⊤≪⊤ (ℬ↑ $ is) (ζ is) 230 | 231 | dir-basis : isDirBasisFor F ℬ↑ 232 | dir-basis = directified-basis-gives-basis F ℬ b 233 | 234 | ∧-closure : closedUnderFinMeets F ℬ↑ 235 | ∧-closure = G𝟏 , G𝟐 236 | where 237 | G𝟏 : ∥ Σ[ t ∈ index ℬ↑ ] [ isTop (pos F) (ℬ↑ $ t) ] ∥ 238 | G𝟏 = ∥∥-rec (∥∥-prop _) G𝟏a ⊤-basic 239 | where 240 | G𝟏a : ⊤[ F ] ε ℬ↑ → ∥ Σ[ t ∈ index ℬ↑ ] [ isTop (pos F) (ℬ↑ $ t) ] ∥ 241 | G𝟏a (t , eq) = ∣ t , subst ([_] ∘ isTop (pos F)) (sym eq) ⊤[ F ]-top ∣ 242 | 243 | ⊤-basic : ∥ ⊤[ F ] ε ℬ↑ ∥ 244 | ⊤-basic = compact→basic ℬ↑ dir-basis ⊤[ F ] ⊤≪⊤ 245 | 246 | G𝟐 : (is js : index ℬ↑) 247 | → ∥ Σ[ ks ∈ index ℬ↑ ] [ isInf (pos F) (ℬ↑ $ ks) (ℬ↑ $ is) (ℬ↑ $ js) ] ∥ 248 | G𝟐 is js = ∥∥-rec (∥∥-prop _) goal foo 249 | where 250 | ∧↑-clop : hasComplement F ((ℬ↑ $ is) ⊓[ F ] (ℬ↑ $ js)) 251 | ∧↑-clop = ∧-has-complement F (ℬ↑ $ is) (ℬ↑ $ js) (ζ is) (ζ js) 252 | 253 | foo : ∥ ((ℬ↑ $ is) ⊓[ F ] (ℬ↑ $ js)) ε ℬ↑ ∥ 254 | foo = clopen→basic-in-stone-locale stone ℬ↑ dir-basis ((ℬ↑ $ is) ⊓[ F ] (ℬ↑ $ js)) ∧↑-clop 255 | 256 | goal : ((ℬ↑ $ is) ⊓[ F ] (ℬ↑ $ js)) ε ℬ↑ 257 | → ∥ Σ[ ks ∈ index ℬ↑ ] [ isInf (pos F) (ℬ↑ $ ks) (ℬ↑ $ is) (ℬ↑ $ js) ] ∥ 258 | goal (ks , eq) = 259 | ∣ ks , subst (λ - → [ isInf (pos F) - (ℬ↑ $ is) (ℬ↑ $ js) ]) (sym eq) inf ∣ 260 | where 261 | inf : [ isInf (pos F) ((ℬ↑ $ is) ⊓[ F ] (ℬ↑ $ js)) (ℬ↑ $ is) (ℬ↑ $ js) ] 262 | inf = (⊓[ F ]-lower₀ _ _ , ⊓[ F ]-lower₁ _ _) 263 | , λ z (p , q) → ⊓[ F ]-greatest (ℬ↑ $ is) (ℬ↑ $ js) z p q 264 | ``` 265 | -------------------------------------------------------------------------------- /FormalTopology/HeytingImplication.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Heyting Implication in a Frame 3 | author: Ayberk Tosun (j.w.w. Martín Escardó) 4 | --- 5 | 6 | 19 | 20 | ## Basis 21 | 22 | Given a (𝓤, 𝓥, 𝓦)-frame F, we say that some 𝓦-family on ∣ F ∣ forms a basis for 23 | it iff for every x : ∣ F ∣, there exists a family U of basic elements such that 24 | ⋁ U is the join of x. 25 | 26 | ```agda 27 | formsBasis : (F : Frame 𝓤 𝓥 𝓦) → Fam 𝓦 ∣ F ∣F → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 28 | formsBasis {𝓥 = 𝓥} {𝓦} F B = 29 | ((x : ∣ F ∣F) → 30 | Σ[ U ∈ Fam 𝓦 (index B) ] 31 | [ isDirected (pos F) ⁅ B $ u ∣ u ε U ⁆ ⊓ isSup (pos F) ⁅ B $ u ∣ u ε U ⁆ x ]) 32 | 33 | -- F has a compact basis iff there is some 𝓦-family B s.t. 34 | -- 35 | -- 1. what I now have in `formsBasis`, 36 | -- 2. the compact elements form a meet-sublattice. 37 | 38 | hasBasis : (F : Frame 𝓤 𝓥 𝓦) → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 39 | hasBasis {𝓦 = 𝓦} F = Σ[ B ∈ Fam 𝓦 ∣ F ∣F ] formsBasis F B 40 | ``` 41 | 42 | ## Definition of Heyting implication 43 | 44 | ```agda 45 | module DefnOfHeytingImplication (F : Frame 𝓤 𝓥 𝓦) where 46 | ``` 47 | 48 | We say that some z is the Heyting implication for (x, y) iff for every w, w ≤ z 49 | iff (w ∧ x) ≤ y. 50 | 51 | ```agda 52 | _isHeytingImplicationFor_ : ∣ F ∣F → ∣ F ∣F × ∣ F ∣F → hProp (𝓤 ∨ 𝓥) 53 | z isHeytingImplicationFor (x , y) = 54 | ∀[ w ∶ ∣ F ∣F ] (w ⊑[ pos F ] z) ⇔ ((x ⊓[ F ] w) ⊑[ pos F ] y) 55 | ``` 56 | 57 | Given a binary operation on the frame, we say that it is a Heyting implication 58 | iff it gives the Heyting implication for any two opens. 59 | 60 | ```agda 61 | isHeytingImplication : (∣ F ∣F → ∣ F ∣F → ∣ F ∣F) → hProp (𝓤 ∨ 𝓥) 62 | isHeytingImplication _⇒_ = 63 | ∀[ x ∶ ∣ F ∣F ] ∀[ y ∶ ∣ F ∣F ] 64 | ((x ⇒ y) isHeytingImplicationFor (x , y)) 65 | ``` 66 | 67 | ```agda 68 | module HeytingImplicationProperties (F : Frame 𝓤 𝓥 𝓦) where 69 | 70 | open DefnOfHeytingImplication F 71 | 72 | lemma₀ : {x y z u v w : ∣ F ∣F} 73 | → [ u isHeytingImplicationFor (x , (y ⊓[ F ] z)) ] 74 | → [ v isHeytingImplicationFor (x , y) ] 75 | → [ w isHeytingImplicationFor (x , z) ] 76 | → u ≡ v ⊓[ F ] w 77 | lemma₀ {x} {y} {z} {u} {v} {w} p q r = ⊓-unique F v w u β₀ β₁ γ 78 | where 79 | open PosetReasoning (pos F) 80 | 81 | δ₀ : [ (x ⊓[ F ] u) ⊑[ pos F ] y ] 82 | δ₀ = x ⊓[ F ] u ⊑⟨ π₀ (p u) (⊑[ pos F ]-refl u) ⟩ 83 | y ⊓[ F ] z ⊑⟨ ⊓[ F ]-lower₀ y z ⟩ 84 | y ■ 85 | 86 | δ₁ : [ (x ⊓[ F ] u) ⊑[ pos F ] z ] 87 | δ₁ = x ⊓[ F ] u ⊑⟨ π₀ (p u) (⊑[ pos F ]-refl u) ⟩ 88 | y ⊓[ F ] z ⊑⟨ ⊓[ F ]-lower₁ y z ⟩ 89 | z ■ 90 | 91 | β₀ : [ u ⊑[ pos F ] v ] 92 | β₀ = π₁ (q u) δ₀ 93 | 94 | β₁ : [ u ⊑[ pos F ] w ] 95 | β₁ = π₁ (r u) δ₁ 96 | 97 | γ : (t : ∣ F ∣F) 98 | → [ t ⊑[ pos F ] v ] → [ t ⊑[ pos F ] w ] → [ t ⊑[ pos F ] u ] 99 | γ t ϕ ψ = π₁ (p t) ((⊓[ F ]-greatest y z (x ⊓[ F ] t) ε) ζ) 100 | where 101 | ε : [ x ⊓[ F ] t ⊑[ pos F ] y ] 102 | ε = x ⊓[ F ] t ⊑⟨ cright F x ϕ ⟩ 103 | x ⊓[ F ] v ⊑⟨ π₀ (q v) (⊑[ pos F ]-refl v) ⟩ 104 | y ■ 105 | 106 | ζ : [ (x ⊓[ F ] t) ⊑[ pos F ] z ] 107 | ζ = x ⊓[ F ] t ⊑⟨ cright F x ψ ⟩ 108 | x ⊓[ F ] w ⊑⟨ π₀ (r w) (⊑[ pos F ]-refl w) ⟩ 109 | z ■ 110 | ``` 111 | 112 | ## Construction of the Heyting implication 113 | 114 | ```agda 115 | module Definition (F : Frame 𝓤 𝓥 𝓥) (basis : hasBasis F) where 116 | ``` 117 | 118 | ```agda 119 | open DefnOfHeytingImplication F 120 | 121 | _≤_ : ∣ F ∣F → ∣ F ∣F → hProp 𝓥 122 | x ≤ y = x ⊑[ pos F ] y 123 | 124 | infix 3 _≤_ 125 | 126 | _∧_ : ∣ F ∣F → ∣ F ∣F → ∣ F ∣F 127 | x ∧ y = x ⊓[ F ] y 128 | 129 | ℬ : 𝓥 ̇ 130 | ℬ = π₀ (π₀ basis) 131 | 132 | β : ℬ → ∣ F ∣F 133 | β = π₁ (π₀ basis) 134 | ``` 135 | 136 | We define the implication as follows: 137 | 138 | ```agda 139 | _==>_ : ∣ F ∣F → ∣ F ∣F → ∣ F ∣F 140 | x ==> y = 141 | ⋁[ F ] ⁅ β b ∣ (b , _) ∶ Σ[ b ∈ ℬ ] [ (x ∧ β b) ≤ y ] ⁆ 142 | ``` 143 | 144 | Of course, the first thing we prove is modus ponens: 145 | 146 | ```agda 147 | mp : (x y : ∣ F ∣F) → [ x ⊓[ F ] (x ==> y) ⊑[ pos F ] y ] 148 | mp x y = 149 | x ⊓[ F ] (x ==> y) ⊑⟨ ≡⇒⊑ (pos F ) (dist F _ _) ⟩ 150 | ⋁[ F ] ⁅ x ⊓[ F ] β b ∣ (b , _) ∶ _ ⁆ ⊑⟨ ⋁[ F ]-least _ _ γ ⟩ 151 | y ■ 152 | where 153 | open PosetReasoning (pos F) 154 | 155 | γ : _ 156 | γ z ((_ , q) , eq) = subst (λ - → [ - ⊑[ pos F ] _ ]) eq q 157 | 158 | mp-op : (x y : ∣ F ∣F) → [ (x ==> y) ⊓[ F ] x ⊑[ pos F ] y ] 159 | mp-op x y = (x ==> y) ⊓[ F ] x ⊑⟨ ≡⇒⊑ (pos F) (comm F (x ==> y) x) ⟩ 160 | x ⊓[ F ] (x ==> y) ⊑⟨ mp x y ⟩ 161 | y ■ 162 | where 163 | open PosetReasoning (pos F) 164 | ``` 165 | 166 | We now proceed to prove that this is the Heyting implication: 167 | 168 | ```agda 169 | lemma : (x y : ∣ F ∣F) (c : ℬ) 170 | → [ (x ∧ β c ≤ y) ⇒ (β c ≤ x ==> y) ] 171 | lemma x y c p = ⋁[ F ]-upper _ (β c) ((c , p) , refl) 172 | 173 | ==>-is-HI : [ isHeytingImplication _==>_ ] 174 | ==>-is-HI x y z = forward , backward 175 | where 176 | open PosetReasoning (pos F) 177 | 178 | forward : [ (z ≤ x ==> y) ⇒ (x ∧ z ≤ y) ] 179 | forward p = x ⊓[ F ] z ⊑⟨ cright F _ p ⟩ 180 | x ⊓[ F ] (x ==> y) ⊑⟨ mp x y ⟩ 181 | y ■ 182 | 183 | backward : [ ((x ∧ z) ≤ y) ⇒ (z ≤ x ==> y) ] 184 | backward p = z ⊑⟨ ≡⇒⊑ (pos F) ε ⟩ 185 | ⋁[ F ] ⁅ β v ∣ v ε VV ⁆ ⊑⟨ ⋁[ F ]-least _ _ nts ⟩ 186 | x ==> y ■ 187 | where 188 | VV : Fam 𝓥 ℬ 189 | VV = π₀ (π₁ basis z) 190 | 191 | ε : z ≡ ⋁[ F ] ⁅ β v ∣ v ε VV ⁆ 192 | ε = ⋁-unique F ⁅ β v ∣ v ε VV ⁆ _ (π₀ lub) (π₁ lub) 193 | where 194 | lub = π₁ (π₁ (π₁ basis z)) 195 | 196 | nts : _ 197 | nts w (i , eq) = w ⊑⟨ ≡⇒⊑ (pos F) (sym eq) ⟩ 198 | β (VV $ i) ⊑⟨ ⋁[ F ]-upper _ (β (VV $ i)) ((VV $ i , δ) , refl) ⟩ 199 | x ==> y ■ 200 | where 201 | δ : [ x ⊓[ F ] β (VV $ i) ≤ y ] 202 | δ = x ∧ β (VV $ i) ⊑⟨ cright F _ (⋁[ F ]-upper _ _ (i , refl)) ⟩ 203 | x ∧ (⋁[ F ] ⁅ β v ∣ v ε VV ⁆) ⊑⟨ cright F _ (≡⇒⊑ (pos F) (sym ε)) ⟩ 204 | x ∧ z ⊑⟨ p ⟩ 205 | y ■ 206 | 207 | ==>-id : (x : ∣ F ∣F) → x ==> x ≡ ⊤[ F ] 208 | ==>-id x = ⊑[ pos F ]-antisym (x ==> x) ⊤[ F ] (⊤[ F ]-top (x ==> x)) G𝟐 209 | where 210 | G𝟐 : [ ⊤[ F ] ⊑[ pos F ] (x ==> x) ] 211 | G𝟐 = π₁ (==>-is-HI x x ⊤[ F ]) (⊓[ F ]-lower₀ x ⊤[ F ]) 212 | 213 | ex-falso-quodlibet : (x : ∣ F ∣F) → ⊥[ F ] ==> x ≡ ⊤[ F ] 214 | ex-falso-quodlibet x = 215 | ⊑[ pos F ]-antisym (⊥[ F ] ==> x) ⊤[ F ] (⊤[ F ]-top (⊥[ F ] ==> x)) γ 216 | where 217 | open PosetReasoning (pos F) 218 | 219 | γ : [ ⊤[ F ] ⊑[ pos F ] ⊥[ F ] ==> x ] 220 | γ = π₁ (==>-is-HI ⊥[ F ] x ⊤[ F ]) (⊥[ F ] ⊓[ F ] ⊤[ F ] ⊑⟨ ⊓[ F ]-lower₀ _ _ ⟩ 221 | ⊥[ F ] ⊑⟨ ⊥[ F ]-bottom x ⟩ 222 | x ■) 223 | 224 | ==>-nucleus-lemma : (x y : ∣ F ∣F) (j : Nucleus F) 225 | → [ (x ==> y) ⊑[ pos F ] (π₀ j x ==> π₀ j y) ] 226 | ==>-nucleus-lemma x y 𝒿@(j , 𝓃₀ , 𝓃₁ , 𝓃₂) = 227 | (x ==> y) ⊑⟨ ⦅𝟏⦆ ⟩ 228 | (x ==> j y) ⊑⟨ ⦅𝟐⦆ ⟩ 229 | (j x ==> j y) ■ 230 | where 231 | open PosetReasoning (pos F) 232 | 233 | ⦅𝟏⦆ : [ (x ==> y) ⊑[ pos F ] (x ==> j y) ] 234 | ⦅𝟏⦆ = π₁ (==>-is-HI x (j y) (x ==> y)) (_ ⊑⟨ mp x y ⟩ y ⊑⟨ 𝓃₁ y ⟩ j y ■) 235 | 236 | ⦅𝟐⦆ : [ (x ==> j y) ⊑[ pos F ] (j x ==> j y) ] 237 | ⦅𝟐⦆ = π₁ (==>-is-HI (j x) (j y) (x ==> j y)) ⦅𝟐a⦆ 238 | where 239 | ⦅𝟐a⦆ : [ j x ⊓[ F ] (x ==> j y) ⊑[ pos F ] (j y) ] 240 | ⦅𝟐a⦆ = 241 | j x ⊓[ F ] (x ==> j y) ⊑⟨ cright F (j x) (𝓃₁ (x ==> j y)) ⟩ 242 | j x ⊓[ F ] j (x ==> j y) ⊑⟨ ≡⇒⊑ (pos F) (sym (𝓃₀ x (x ==> j y))) ⟩ 243 | j (x ⊓[ F ] (x ==> j y)) ⊑⟨ mono F 𝒿 _ _ (mp x (j y)) ⟩ 244 | j (j y) ⊑⟨ 𝓃₂ y ⟩ 245 | j y ■ 246 | 247 | ==>-∨-lemma : (x y z : ∣ F ∣F) 248 | → [ ((x ==> z) ⊓[ F ] (y ==> z)) ⊑[ pos F ] ((x ∨[ F ] y) ==> z) ] 249 | ==>-∨-lemma x y z = π₁ (==>-is-HI _ _ _) G𝟏 250 | where 251 | open PosetReasoning (pos F) 252 | 253 | G𝟏 : [ (x ∨[ F ] y) ⊓[ F ] ((x ==> z) ⊓[ F ] (y ==> z)) ⊑[ pos F ] z ] 254 | G𝟏 = (x ∨[ F ] y) ⊓[ F ] ((x ==> z) ⊓[ F ] (y ==> z)) ⊑⟨ ≡⇒⊑ (pos F) (comm F _ _) ⟩ 255 | ((x ==> z) ⊓[ F ] (y ==> z)) ⊓[ F ](x ∨[ F ] y) ⊑⟨ ≡⇒⊑ (pos F) (bin-dist F ((x ==> z) ⊓[ F ] (y ==> z)) x y) ⟩ 256 | (((x ==> z) ⊓[ F ] (y ==> z)) ⊓[ F ] x) ∨[ F ] (((x ==> z) ⊓[ F ] (y ==> z)) ⊓[ F ] y) ⊑⟨ ⦅𝟏⦆ ⟩ 257 | ((x ==> z) ⊓[ F ] x) ∨[ F ] (((x ==> z) ⊓[ F ] (y ==> z)) ⊓[ F ] y) ⊑⟨ ⦅𝟐⦆ ⟩ 258 | ((x ==> z) ⊓[ F ] x) ∨[ F ] ((y ==> z) ⊓[ F ] y) ⊑⟨ ∨-cright F _ _ _ (≡⇒⊑ (pos F) (comm F (y ==> z) y)) ⟩ 259 | ((x ==> z) ⊓[ F ] x) ∨[ F ] (y ⊓[ F ] (y ==> z)) ⊑⟨ ∨-cright F _ _ _ (mp y z) ⟩ 260 | ((x ==> z) ⊓[ F ] x) ∨[ F ] z ⊑⟨ ∨-cleft F _ _ _ (≡⇒⊑ (pos F) (comm F _ _)) ⟩ 261 | (x ⊓[ F ] (x ==> z)) ∨[ F ] z ⊑⟨ ∨-cleft F _ _ _ (mp x z) ⟩ 262 | z ∨[ F ] z ⊑⟨ ≡⇒⊑ (pos F) (x∨x=x F z) ⟩ 263 | z ■ 264 | where 265 | G𝟐 = ∨-cright F z _ _ (mp y z) 266 | 267 | ⦅𝟏⦆ = ∨-cleft F _ _ _ (cleft F _ (⊓[ F ]-lower₀ _ _)) 268 | 269 | ⦅𝟐⦆ : _ 270 | ⦅𝟐⦆ = ∨-cright F _ _ _ (cleft F _ (⊓[ F ]-lower₁ _ _)) 271 | ``` 272 | -------------------------------------------------------------------------------- /FormalTopology/Base.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Base for a Frame 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 20 | 21 | ```agda 22 | isBasisFor : (F : Frame 𝓤 𝓥 𝓦) → Fam 𝓦 ∣ F ∣F → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 23 | isBasisFor {𝓦 = 𝓦} F ℬ = 24 | (x : ∣ F ∣F) → 25 | Σ[ J ∈ Fam 𝓦 (index ℬ) ] [ isSup (pos F) ⁅ ℬ $ j ∣ j ε J ⁆ x ] 26 | 27 | isDirBasisFor : (F : Frame 𝓤 𝓥 𝓦) → Fam 𝓦 ∣ F ∣F → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 28 | isDirBasisFor {𝓥 = 𝓥} {𝓦} F B = 29 | (x : ∣ F ∣F) → 30 | Σ[ U ∈ Fam 𝓦 (index B) ] 31 | [ isDirected (pos F) ⁅ B $ u ∣ u ε U ⁆ ⊓ isSup (pos F) ⁅ B $ u ∣ u ε U ⁆ x ] 32 | ``` 33 | 34 | ```agda 35 | dir-basis→basis : (F : Frame 𝓤 𝓥 𝓦) 36 | → (ℬ : Fam 𝓦 ∣ F ∣F) → isDirBasisFor F ℬ → isBasisFor F ℬ 37 | dir-basis→basis F ℬ ψ x with ψ x 38 | dir-basis→basis F ℬ ψ x | (U , _ , sup) = U , sup 39 | ``` 40 | 41 | ```agda 42 | hasBasis : (F : Frame 𝓤 𝓥 𝓦) → hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 43 | hasBasis {𝓦 = 𝓦} F = ∥ Σ[ ℬ ∈ Fam 𝓦 ∣ F ∣F ] isBasisFor F ℬ ∥Ω 44 | ``` 45 | 46 | ```agda 47 | hasDirBasis : (F : Frame 𝓤 𝓥 𝓦) → hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) 48 | hasDirBasis {𝓦 = 𝓦} F = ∥ Σ[ ℬ ∈ Fam 𝓦 ∣ F ∣F ] isDirBasisFor F ℬ ∥Ω 49 | ``` 50 | 51 | ```agda 52 | directify : (F : Frame 𝓤 𝓥 𝓦) (S : Fam 𝓦 ∣ F ∣F) → Fam 𝓦 ∣ F ∣F 53 | directify F (I , f) = (List I , g) 54 | where 55 | open PosetReasoning (pos F) 56 | 57 | g : List I → ∣ F ∣F 58 | g = foldr (λ i - → f i ∨[ F ] -) ⊥[ F ] 59 | 60 | directify-directed : (F : Frame 𝓤 𝓥 𝓦) (S : Fam 𝓦 ∣ F ∣F) 61 | → [ isDirected (pos F) (directify F S) ] 62 | directify-directed F S@(I , f) = directed 63 | where 64 | open PosetReasoning (pos F) 65 | 66 | g = (directify F S) $_ 67 | 68 | g-functorial : (is js : List I) → g (is ++ js) ≡ g is ∨[ F ] g js 69 | g-functorial [] js = g ([] ++ js) ≡⟨ refl ⟩ 70 | g js ≡⟨ sym (x∨⊥=x F (g js)) ⟩ 71 | ⊥[ F ] ∨[ F ] g js ≡⟨ refl ⟩ 72 | g [] ∨[ F ] g js ∎ 73 | g-functorial (i ∷ is) js = g (i ∷ (is ++ js)) ≡⟨ refl ⟩ 74 | f i ∨[ F ] g (is ++ js) ≡⟨ ⦅𝟏⦆ ⟩ 75 | f i ∨[ F ] (g is ∨[ F ] g js) ≡⟨ ⦅𝟐⦆ ⟩ 76 | (f i ∨[ F ] g is) ∨[ F ] g js ≡⟨ refl ⟩ 77 | g (i ∷ is) ∨[ F ] g js ∎ 78 | where 79 | ⦅𝟏⦆ = cong (λ - → f i ∨[ F ] -) (g-functorial is js) 80 | ⦅𝟐⦆ = sym (∨[ F ]-assoc (f i) (g is) (g js)) 81 | 82 | upwards : [ ∀[ is ∶ List I ] ∀[ js ∶ List I ] 83 | ∥ Σ[ ks ∈ List I ] [ g is ⊑[ pos F ] g ks ] × [ g js ⊑[ pos F ] g ks ] ∥Ω ] 84 | upwards is js = ∣ is ++ js , G𝟏 , G𝟐 ∣ 85 | where 86 | G𝟏 : [ g is ⊑[ pos F ] g (is ++ js) ] 87 | G𝟏 = g is ⊑⟨ ⊔[ F ]-upper₀ (g is) (g js) ⟩ 88 | g is ∨[ F ] g js ⊑⟨ ≡⇒⊑ (pos F) (sym (g-functorial is js)) ⟩ 89 | g (is ++ js) ■ 90 | 91 | G𝟐 : [ g js ⊑[ pos F ] g (is ++ js) ] 92 | G𝟐 = g js ⊑⟨ ⊔[ F ]-upper₁ (g is) (g js) ⟩ 93 | g is ∨[ F ] g js ⊑⟨ ≡⇒⊑ (pos F) (sym (g-functorial is js)) ⟩ 94 | g (is ++ js) ■ 95 | 96 | directed : [ isDirected (pos F) (List I , g) ] 97 | directed = ∣ [] ∣ , upwards 98 | 99 | directification-resp-join : (F : Frame 𝓤 𝓥 𝓦) (S : Fam 𝓦 ∣ F ∣F) 100 | → ⋁[ F ] S ≡ ⋁[ F ] (directify F S) 101 | directification-resp-join F S = ⊑[ pos F ]-antisym _ _ G𝟏 G𝟐 102 | where 103 | open PosetReasoning (pos F) 104 | 105 | G𝟏 : [ ⋁[ F ] S ⊑[ pos F ] ⋁[ F ] (directify F S) ] 106 | G𝟏 = ⋁[ F ]-least _ _ nts 107 | where 108 | nts : [ ∀[ s ε S ] (s ⊑[ pos F ] ⋁[ F ] (directify F S)) ] 109 | nts s (i , eq) = ⋁[ F ]-upper (directify F S) s (i ∷ [] , goal) 110 | where 111 | goal : (S $ i) ∨[ F ] ⊥[ F ] ≡ s 112 | goal = (S $ i) ∨[ F ] ⊥[ F ] ≡⟨ ∨-comm F (S $ i) ⊥[ F ] ⟩ 113 | ⊥[ F ] ∨[ F ] (S $ i) ≡⟨ x∨⊥=x F (S $ i) ⟩ 114 | S $ i ≡⟨ eq ⟩ 115 | s ∎ 116 | 117 | G𝟐-lemma : (is : List (index S)) → [ (directify F S $ is) ⊑[ pos F ] ⋁[ F ] S ] 118 | G𝟐-lemma [] = ⊥[ F ] ⊑⟨ ⊥[ F ]-bottom (⋁[ F ] S) ⟩ ⋁[ F ] S ■ 119 | G𝟐-lemma (i ∷ is) = ⊔[ F ]-least _ _ _ nts₀ nts₁ 120 | where 121 | nts₀ : [ (S $ i) ⊑[ pos F ] (⋁[ F ] S) ] 122 | nts₀ = ⋁[ F ]-upper _ _ (i , refl) 123 | 124 | nts₁ : [ (directify F S $ is) ⊑[ pos F ] (⋁[ F ] S) ] 125 | nts₁ = G𝟐-lemma is 126 | 127 | G𝟐 : [ ⋁[ F ] (directify F S) ⊑[ pos F ] ⋁[ F ] S ] 128 | G𝟐 = ⋁[ F ]-least (directify F S) (⋁[ F ] S) nts 129 | where 130 | nts : [ ∀[ s ε (directify F S) ] (s ⊑[ pos F ] ⋁[ F ] S) ] 131 | nts s (is , eq) = subst (λ - → [ rel (pos F) - (⋁[ F ] S) ]) eq (G𝟐-lemma is) 132 | ``` 133 | 134 | ```agda 135 | directified-basis-gives-basis : (F : Frame 𝓤 𝓥 𝓦) 136 | → (ℬ : Fam 𝓦 ∣ F ∣F) 137 | → isBasisFor F ℬ 138 | → isDirBasisFor F (directify F ℬ) 139 | directified-basis-gives-basis {𝓦 = 𝓦} F ℬ basis = goal 140 | where 141 | 𝒥 : ∣ F ∣F → Fam 𝓦 (index ℬ) 142 | 𝒥 = π₀ ∘ basis 143 | 144 | approx : ∣ F ∣F → Fam 𝓦 ∣ F ∣F 145 | approx x = ⁅ ℬ $ j ∣ j ε 𝒥 x ⁆ 146 | 147 | goal : isDirBasisFor F (directify F ℬ) 148 | goal x = (List (index (approx x)) , g) , dir , sup 149 | where 150 | g : List (index (approx x)) → index (directify F ℬ) 151 | g = map (𝒥 x $_) 152 | 153 | what-we-need : ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ ≡ directify F (approx x) 154 | what-we-need = ΣPathTransport→PathΣ _ _ (refl , (transport refl ((directify F ℬ $_) ∘ g) ≡⟨ transportRefl ((directify F ℬ $_) ∘ g) ⟩ (directify F ℬ $_) ∘ g ≡⟨ funExt final-goal ⟩ π₁ (directify F (approx x)) ∎)) 155 | where 156 | final-goal : (is : List (index (approx x))) → directify F ℬ $ g is ≡ (directify F (approx x)) $ is 157 | final-goal [] = refl 158 | final-goal (i ∷ is) = directify F ℬ $ g (i ∷ is) ≡⟨ refl ⟩ 159 | (ℬ $ 𝒥 x $ i) ∨[ F ] (directify F ℬ $ (g is)) ≡⟨ cong (λ - → (ℬ $ 𝒥 x $ i) ∨[ F ] -) (final-goal is) ⟩ 160 | (ℬ $ (𝒥 x $ i)) ∨[ F ] (directify F (approx x) $ is) ≡⟨ refl ⟩ 161 | directify F (approx x) $ (i ∷ is) ∎ 162 | 163 | dir : [ isDirected (pos F) ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ ] 164 | dir = subst (λ - → [ isDirected (pos F) - ]) (sym what-we-need) (directify-directed F (approx x)) 165 | 166 | sup : [ isSup (pos F) ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ x ] 167 | sup = subst (λ - → [ isSup (pos F) - x ]) (sym what-we-need) (subst (λ - → [ isSup (pos F) (directify F (approx x)) - ]) (sym foo) (sup-goal₁ , sup-goal₂)) 168 | where 169 | sup-goal₁ = ⋁[ F ]-upper (directify F (approx x)) 170 | 171 | sup-goal₂ = ⋁[ F ]-least (directify F (approx x)) 172 | 173 | foo : x ≡ ⋁[ F ] directify F (approx x) 174 | foo = x ≡⟨ uncurry (⋁-unique F (approx x) x) (π₁ (basis x)) ⟩ 175 | ⋁[ F ] (approx x) ≡⟨ directification-resp-join F (approx x) ⟩ 176 | ⋁[ F ] directify F (approx x) ∎ 177 | ``` 178 | 179 | ```agda 180 | basis→dir-basis : (F : Frame 𝓤 𝓥 𝓦) → [ hasBasis F ] → [ hasDirBasis F ] 181 | basis→dir-basis {𝓦 = 𝓦} F = ∥∥-rec (isProp[] (hasDirBasis F)) nts 182 | where 183 | nts : Σ[ ℬ ∈ Fam 𝓦 ∣ F ∣F ] (isBasisFor F ℬ) → [ hasDirBasis F ] 184 | nts (ℬ , basis) = ∣ directify F ℬ , goal ∣ 185 | where 186 | 𝒥 : ∣ F ∣F → Fam 𝓦 (index ℬ) 187 | 𝒥 = π₀ ∘ basis 188 | 189 | approx : ∣ F ∣F → Fam 𝓦 ∣ F ∣F 190 | approx x = ⁅ ℬ $ j ∣ j ε 𝒥 x ⁆ 191 | 192 | goal : isDirBasisFor F (directify F ℬ) 193 | goal x = (List (index (approx x)) , g) , dir , sup 194 | where 195 | g : List (index (approx x)) → index (directify F ℬ) 196 | g = map (𝒥 x $_) 197 | 198 | what-we-need : ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ ≡ directify F (approx x) 199 | what-we-need = ΣPathTransport→PathΣ _ _ (refl , (transport refl ((directify F ℬ $_) ∘ g) ≡⟨ transportRefl ((directify F ℬ $_) ∘ g) ⟩ (directify F ℬ $_) ∘ g ≡⟨ funExt final-goal ⟩ π₁ (directify F (approx x)) ∎)) 200 | where 201 | final-goal : (is : List (index (approx x))) → directify F ℬ $ g is ≡ (directify F (approx x)) $ is 202 | final-goal [] = refl 203 | final-goal (i ∷ is) = directify F ℬ $ g (i ∷ is) ≡⟨ refl ⟩ 204 | (ℬ $ 𝒥 x $ i) ∨[ F ] (directify F ℬ $ (g is)) ≡⟨ cong (λ - → (ℬ $ 𝒥 x $ i) ∨[ F ] -) (final-goal is) ⟩ 205 | (ℬ $ (𝒥 x $ i)) ∨[ F ] (directify F (approx x) $ is) ≡⟨ refl ⟩ 206 | directify F (approx x) $ (i ∷ is) ∎ 207 | 208 | dir : [ isDirected (pos F) ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ ] 209 | dir = subst (λ - → [ isDirected (pos F) - ]) (sym what-we-need) (directify-directed F (approx x)) 210 | 211 | sup : [ isSup (pos F) ⁅ directify F ℬ $ is ∣ is ε (List (index (approx x)) , g) ⁆ x ] 212 | sup = subst (λ - → [ isSup (pos F) - x ]) (sym what-we-need) (subst (λ - → [ isSup (pos F) (directify F (approx x)) - ]) (sym foo) (sup-goal₁ , sup-goal₂)) 213 | where 214 | sup-goal₁ = ⋁[ F ]-upper (directify F (approx x)) 215 | 216 | sup-goal₂ = ⋁[ F ]-least (directify F (approx x)) 217 | 218 | foo : x ≡ ⋁[ F ] directify F (approx x) 219 | foo = x ≡⟨ uncurry (⋁-unique F (approx x) x) (π₁ (basis x)) ⟩ 220 | ⋁[ F ] (approx x) ≡⟨ directification-resp-join F (approx x) ⟩ 221 | ⋁[ F ] directify F (approx x) ∎ 222 | ``` 223 | 224 | ```agda 225 | ``` 226 | -------------------------------------------------------------------------------- /FormalTopology/Basis.lagda.md: -------------------------------------------------------------------------------- 1 | ``` 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module Basis where 5 | 6 | import Cubical.Core.Everything as CE 7 | import Cubical.Data.Sigma as DΣ 8 | import Cubical.Data.Sum as DS 9 | import Cubical.Foundations.Prelude as FP 10 | import Cubical.Foundations.Equiv as FE 11 | import Cubical.Functions.Logic as FL 12 | import Cubical.Foundations.Structure as FS 13 | import Cubical.Foundations.HLevels as FH 14 | import Cubical.Foundations.Isomorphism as FI 15 | import Cubical.Foundations.Equiv.HalfAdjoint as HAE 16 | import Cubical.Functions.FunExtEquiv as FEE 17 | import Cubical.Foundations.Function as FF 18 | 19 | open import Cubical.Foundations.Univalence public using (ua) 20 | 21 | open CE public using (_≡_; Type; Σ; Σ-syntax; _,_; _≃_; equivFun; isEquiv; Level; 22 | ℓ-max; ℓ-zero; ℓ-suc) 23 | open DΣ public using (Σ≡Prop; ΣPathTransport→PathΣ; PathΣ→ΣPathTransport; _×_; _,_) 24 | renaming (fst to π₀; snd to π₁) 25 | open DS public using (inl; inr; _⊎_) 26 | open FP public using (funExt; funExt⁻; subst; isContr; isProp; isPropIsProp; isPropIsContr; isSet; 27 | isProp→isSet; cong; refl; sym; _≡⟨_⟩_; _∎; transport; 28 | transportRefl; J; JRefl) 29 | open FE public using (idEquiv; invEquiv; secEq; retEq; fiber; equivToIso; 30 | isPropIsEquiv) 31 | open FL public using ( _⇔_ ; _⇒_ ; ⇔toPath ; _⊓_ ; ∃[∶]-syntax; ∀[∶]-syntax; ∀[]-syntax; ¬_) 32 | renaming (isProp⟨⟩ to isProp[]) 33 | open FS public using () renaming (⟨_⟩ to [_]) 34 | open FH public using (hProp; isSetHProp; isPropIsSet; isPropΣ; isOfHLevel; 35 | isOfHLevelΠ; isOfHLevelΣ; isOfHLevelSuc; isSetΣ; 36 | isSetΠ; isSetΠ2; isPropΠ; isPropΠ2; isPropΠ3) 37 | open FI public using (isoToPath; isoToEquiv; iso; section; retract; Iso) 38 | open FF public using (_∘_) renaming (idfun to id) 39 | open FEE public using (funExtEquiv; funExt₂; funExt₂Equiv; funExt₂Path) 40 | open HAE public using (isHAEquiv; equiv→HAEquiv) 41 | 42 | open import Cubical.Data.Nat using (ℕ; _+_) renaming (suc to sucℕ; zero to zeroℕ) 43 | open import Cubical.Data.Nat.Properties using (injSuc; snotz; +-comm) 44 | open import Cubical.Data.Nat.Order 45 | open import Cubical.Data.Empty using (rec; ⊥) 46 | open import Cubical.Data.List using (List; length; _∷_; []) 47 | open import Cubical.Data.Fin using (Fin) 48 | open import Cubical.Relation.Nullary.DecidableEq using (Discrete→isSet) 49 | open import Cubical.Relation.Nullary using (Discrete; yes; no; Dec) 50 | ``` 51 | 52 | ``` 53 | variable 54 | ℓ ℓ′ ℓ₀ ℓ₁ ℓ₂ ℓ₃ ℓ₀′ ℓ₁′ ℓ₂′ ℓ₀′′ ℓ₁′′ ℓ₂′′ : Level 55 | 56 | variable 57 | A : Type ℓ₀ 58 | B : A → Type ℓ₀ 59 | A₀ : Type ℓ₁ 60 | 61 | _↔_ : (A : Type ℓ) (B : Type ℓ′) → Type _ 62 | A ↔ B = (A → B) × (B → A) 63 | 64 | ↔-to : {A : Type ℓ} {B : Type ℓ′} → A ↔ B → A → B 65 | ↔-to (to , _) = to 66 | 67 | ↔-from : {A : Type ℓ} {B : Type ℓ′} → A ↔ B → B → A 68 | ↔-from (_ , from) = from 69 | ``` 70 | 71 | ## Levels 72 | 73 | Escardó-style level notation. 74 | 75 | ```agda 76 | infixr 5 _∨_ 77 | 78 | Universe : Type₀ 79 | Universe = Level 80 | 81 | variable 82 | 𝒰 𝒱 𝒲 𝓤 𝓥 𝓦 𝓤′ 𝓥′ 𝓦′ 𝓤′′ 𝓥′′ 𝓦′′ : Universe 83 | 84 | _∨_ : Level → Level → Level 85 | ℓ₀ ∨ ℓ₁ = ℓ-max ℓ₀ ℓ₁ 86 | 87 | infix 6 _⁺ 88 | 89 | _⁺ : Level → Level 90 | ℓ ⁺ = ℓ-suc ℓ 91 | 92 | infix 6 _̇ 93 | 94 | _̇ : (ℓ : Level) → Type (ℓ ⁺) 95 | ℓ ̇ = Type ℓ 96 | 97 | 𝓤₀ : Level 98 | 𝓤₀ = ℓ-zero 99 | 100 | 𝓤₁ : Level 101 | 𝓤₁ = 𝓤₀ ⁺ 102 | ``` 103 | 104 | ## The unit type 105 | 106 | ``` 107 | data Unit (ℓ : Level) : Type ℓ where 108 | tt : Unit ℓ 109 | 110 | Unit-prop : {ℓ : Level} → isProp (Unit ℓ) 111 | Unit-prop tt tt = refl 112 | 113 | top : (𝓤 : Universe) → hProp 𝓤 114 | top 𝓤 = Unit 𝓤 , Unit-prop 115 | ``` 116 | 117 | ## Bottom 118 | 119 | ``` 120 | data 𝟘 (ℓ : Level) : Type ℓ where 121 | 122 | 𝟘-elim : {A : Type ℓ} → 𝟘 ℓ′ → A 123 | 𝟘-elim () 124 | 125 | bot : (ℓ : Level) → hProp ℓ 126 | bot ℓ = 𝟘 ℓ , λ () 127 | ``` 128 | 129 | ## Booleans 130 | 131 | ```agda 132 | data Bool (ℓ : Level) : Type ℓ where 133 | true : Bool ℓ 134 | false : Bool ℓ 135 | ``` 136 | 137 | ```agda 138 | true≠false : _≡_ {ℓ = 𝓤} true false → ⊥ 139 | true≠false p = subst (λ { true → Unit 𝓤₀ ; false → ⊥ }) p tt 140 | ``` 141 | 142 | ```agda 143 | _=b=_ : Discrete (Bool 𝓤) 144 | true =b= true = yes refl 145 | true =b= false = no true≠false 146 | false =b= true = no (true≠false ∘ sym) 147 | false =b= false = yes refl 148 | ``` 149 | 150 | ```agda 151 | Bool-set : isSet (Bool 𝓤) 152 | Bool-set = Discrete→isSet _=b=_ 153 | ``` 154 | 155 | ```agda 156 | if_then_else_ : {A : Type ℓ₀} → Bool ℓ₁ → A → A → A 157 | if true then x else y = x 158 | if false then x else y = y 159 | ``` 160 | 161 | ## Propositions 162 | 163 | ``` 164 | is-true-prop : (P : hProp ℓ) → isProp [ P ] 165 | is-true-prop (P , P-prop) = P-prop 166 | ``` 167 | 168 | ``` 169 | ∃_ : {A : Type ℓ₀} → (A → hProp ℓ₁) → Type _ 170 | ∃_ {A = A} P = Σ[ x ∈ A ] [ P x ] 171 | ``` 172 | 173 | ## Extensional equality 174 | 175 | ``` 176 | _~_ : (f g : (x : A) → B x) → Type _ 177 | _~_ {A = A} f g = (x : A) → f x ≡ g x 178 | ``` 179 | 180 | ## Powerset 181 | 182 | ``` 183 | 𝒫 : Type ℓ → Type _ 184 | 𝒫 {ℓ} A = A → hProp ℓ 185 | 186 | _∈_ : A → 𝒫 A → hProp _ 187 | x ∈ U = U x 188 | 189 | ∈-prop : {A : Type ℓ} {x : A} → (U : 𝒫 A) → isProp [ x ∈ U ] 190 | ∈-prop {x = x} U = is-true-prop (x ∈ U) 191 | 192 | 𝒫-set : (A : Type ℓ) → isSet (𝒫 A) 193 | 𝒫-set A = isSetΠ λ _ → isSetHProp 194 | 195 | _^c : {A : Type ℓ} → 𝒫 A → 𝒫 A 196 | U ^c = λ x → ¬ (x ∈ U) 197 | 198 | variable 199 | U V : 𝒫 A 200 | 201 | _⊆⊆_ : {A : Type ℓ} → (A → Type ℓ₀) → (A → Type ℓ₁) → Type _ 202 | _⊆⊆_ {A = A} U V = (x : A) → U x → V x 203 | 204 | _⊆_ : {A : Type ℓ} → 𝒫 A → 𝒫 A → hProp ℓ 205 | _⊆_ {A = A} U V = ((λ - → [ U - ]) ⊆⊆ (λ - → [ V - ])) , prop 206 | where 207 | prop : isProp ((x : A) → [ U x ] → [ V x ]) 208 | prop = isPropΠ λ x → isPropΠ λ _ → is-true-prop (V x) 209 | 210 | ⊆-antisym : [ U ⊆ V ] → [ V ⊆ U ] → U ≡ V 211 | ⊆-antisym {U = U} {V} U⊆V V⊆V = funExt (λ x → ⇔toPath (U⊆V x) (V⊆V x)) 212 | 213 | ∅ : 𝒫 A 214 | ∅ _ = bot _ 215 | 216 | entire : {A : Type ℓ} → 𝒫 A 217 | entire {ℓ = ℓ} _ = Unit ℓ , Unit-prop 218 | 219 | _∩_ : 𝒫 A → 𝒫 A → 𝒫 A 220 | _∩_ {A = A} U V = λ x → ([ U x ] × [ V x ]) , prop x 221 | where 222 | prop : (x : A) → isProp ([ U x ] × [ V x ]) 223 | prop x = isPropΣ (is-true-prop (x ∈ U)) λ _ → is-true-prop (V x) 224 | ``` 225 | 226 | ```agda 227 | U∩U^c=∅ : {A : Type ℓ} → (U : 𝒫 A) → Σ[ x ∈ A ] [ x ∈ (U ∩ (U ^c)) ] → Cubical.Data.Empty.⊥ 228 | U∩U^c=∅ U (x , (x∈U , x∈U^c)) = rec (x∈U^c x∈U) 229 | ``` 230 | 231 | ## Family 232 | 233 | ``` 234 | Fam : (ℓ₀ : Level) → Type ℓ₁ → Type _ 235 | Fam ℓ₀ A = Σ (Set ℓ₀) (λ I → I → A) 236 | 237 | index : Fam ℓ₁ A → Type ℓ₁ 238 | index (I , _) = I 239 | 240 | -- Application of a family over X to an index. 241 | _$_ : (ℱ : Fam ℓ₀ A) → index ℱ → A 242 | _$_ (_ , f) = f 243 | 244 | infixr 7 _$_ 245 | 246 | -- Membership for families. 247 | _ε_ : A → Fam ℓ₁ A → Type _ 248 | x ε (_ , f) = fiber f x 249 | 250 | _⊆fam_ : {A : Type ℓ} → Fam ℓ₁ A → Fam ℓ₁ A → Type (ℓ-max ℓ ℓ₁) 251 | _⊆fam_ {A = A} U V = (x : A) → x ε U → x ε V 252 | 253 | -- Composition of a family with a function. 254 | _⟨$⟩_ : {X : Type ℓ₀} {Y : Type ℓ₁} → (g : X → Y) → (ℱ : Fam ℓ₂ X) → Fam ℓ₂ Y 255 | g ⟨$⟩ ℱ = (index ℱ) , g ∘ (_$_ ℱ) 256 | 257 | fmap : {X : Type ℓ₀} {Y : Type ℓ₁} → (g : X → Y) → (ℱ : Fam ℓ₂ X) → Fam ℓ₂ Y 258 | fmap = _⟨$⟩_ 259 | 260 | syntax fmap (λ x → e) ℱ = ⁅ e ∣ x ε ℱ ⁆ 261 | 262 | compr-∶-syntax : {X : Type ℓ₀} → (I : Type ℓ₂) → (I → X) → Fam ℓ₂ X 263 | compr-∶-syntax I f = (I , f) 264 | 265 | syntax compr-∶-syntax I (λ i → e) = ⁅ e ∣ i ∶ I ⁆ 266 | 267 | -- Forall quantification for families. 268 | fam-forall : {X : Type ℓ₀} (ℱ : Fam ℓ₂ X) → (X → hProp ℓ₁) → hProp _ 269 | fam-forall {X = X} ℱ P = ((x : X) → x ε ℱ → [ P x ]) , prop 270 | where 271 | prop : isProp ((x : X) → x ε ℱ → [ P x ]) 272 | prop = isPropΠ λ x → isPropΠ λ _ → is-true-prop (P x) 273 | 274 | syntax fam-forall ℱ (λ x → P) = ∀[ x ε ℱ ] P 275 | 276 | -- Familification of a given powerset. 277 | ⟪_⟫ : {A : Type ℓ₀} → (A → hProp ℓ₁) → Fam _ A 278 | ⟪_⟫ {A = A} U = (Σ[ x ∈ A ] [ U x ]) , π₀ 279 | 280 | lookup : {A : Type ℓ₀} → (xs : List A) → Fin (length xs) → A 281 | lookup [] (_ , zeroℕ , p) = rec (snotz p) 282 | lookup [] (_ , sucℕ i , p) = rec (snotz p) 283 | lookup (x ∷ xs) (zeroℕ , _) = x 284 | lookup (x ∷ xs) (sucℕ i , p) = lookup xs (i , pred-≤-pred p) 285 | 286 | famFromList : {A : Type ℓ₀} → List A → Fam _ A 287 | famFromList xs = Fin (length xs) , lookup xs 288 | 289 | _×f_ : {A : Type ℓ₀} → Fam ℓ₂ A → Fam ℓ₂′ A → Fam (ℓ-max ℓ₂ ℓ₂′) (A × A) 290 | _×f_ (I , f) (J , g) = I × J , (λ { (i , j) → f i , g j }) 291 | 292 | _∪f_ : {A : 𝓤 ̇} → Fam 𝓦 A → Fam 𝓦 A → Fam 𝓦 A 293 | _∪f_ (I , f) (J , g) = I ⊎ J , λ { (inl i) → f i ; (inr j) → g j } 294 | 295 | ⁅_,_⁆ : {A : 𝓤 ̇} {𝓦 : Universe} → A → A → Fam 𝓦 A 296 | ⁅_,_⁆ {𝓦 = 𝓦} x y = Bool 𝓦 , λ b → if b then x else y 297 | 298 | ⁅⁆-distr : {A : 𝓤 ̇ } {B : 𝓥 ̇ } {𝓦 : Universe} (x y : A) (f : A → B) 299 | → fmap {X = A} {B} f (⁅_,_⁆ {𝓦 = 𝓦} x y) ≡ ⁅_,_⁆ {𝓦 = 𝓦} (f x) (f y) 300 | ⁅⁆-distr x y f = ΣPathTransport→PathΣ (f ⟨$⟩ ⁅ x , y ⁆) ⁅ f x , f y ⁆ (refl , γ) 301 | where 302 | β : π₁ (f ⟨$⟩ ⁅ x , y ⁆) ≡ π₁ ⁅ f x , f y ⁆ 303 | β = funExt λ { true → refl ; false → refl } 304 | 305 | γ : transport refl (π₁ (f ⟨$⟩ ⁅ x , y ⁆)) 306 | ≡ π₁ ⁅ f x , f y ⁆ 307 | γ = transport refl (π₁ (f ⟨$⟩ ⁅ x , y ⁆)) ≡⟨ transportRefl (π₁ (f ⟨$⟩ ⁅ x , y ⁆)) ⟩ 308 | π₁ (f ⟨$⟩ ⁅ x , y ⁆) ≡⟨ β ⟩ 309 | π₁ ⁅ f x , f y ⁆ ∎ 310 | ``` 311 | 312 | ## Truncation 313 | 314 | ``` 315 | data ∥_∥ (A : Type ℓ) : Type ℓ where 316 | ∣_∣ : A → ∥ A ∥ 317 | squash : (x y : ∥ A ∥) → x ≡ y 318 | 319 | ∥∥-prop : (A : Type ℓ) → isProp ∥ A ∥ 320 | ∥∥-prop _ = squash 321 | 322 | ∥_∥Ω : (A : 𝓤 ̇) → hProp 𝓤 323 | ∥ A ∥Ω = ∥ A ∥ , ∥∥-prop A 324 | 325 | ∥∥-rec : {X : Type ℓ} {Y : Type ℓ₀} → isProp Y → (X → Y) → ∥ X ∥ → Y 326 | ∥∥-rec Y-prop f ∣ x ∣ = f x 327 | ∥∥-rec Y-prop f (squash ∣x∣₀ ∣x∣₁ i) = 328 | Y-prop (∥∥-rec Y-prop f ∣x∣₀) (∥∥-rec Y-prop f ∣x∣₁) i 329 | 330 | ∥∥-× : {A : Type ℓ₀} {B : Type ℓ₁} → ∥ A ∥ → ∥ B ∥ → ∥ A × B ∥ 331 | ∥∥-× {A = A} {B} p q = ∥∥-rec (∥∥-prop (A × B)) NTS p 332 | where 333 | NTS′ : B → A → ∥ A × B ∥ 334 | NTS′ y x = ∣ x , y ∣ 335 | 336 | NTS : A → ∥ A × B ∥ 337 | NTS = ∥∥-rec (isPropΠ (λ _ → ∥∥-prop (A × B))) NTS′ q 338 | ``` 339 | -------------------------------------------------------------------------------- /FormalTopology/Index.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module Index where 5 | 6 | import Basis 7 | open import FormalTopology 8 | open import Cover 9 | 10 | import Poset 11 | import Frame 12 | 13 | import Nucleus 14 | import CoverFormsNucleus 15 | import UniversalProperty 16 | 17 | import Cubical.Relation.Nullary 18 | import Cubical.Relation.Nullary.DecidableEq 19 | import Cubical.Foundations.Univalence 20 | import Cubical.Foundations.SIP 21 | 22 | open import Cubical.Data.Bool using (Bool; _≟_) 23 | 24 | open import SnocList Bool _≟_ 25 | import Compactness 26 | import CantorSpace 27 | ``` 28 | 29 | ## Chapter 2: Foundations 30 | 31 | ### 2.3: Equivalence and univalence 32 | 33 | **Definition 2.1**. 34 | 35 | ```agda 36 | _ = Basis.isContr 37 | ``` 38 | 39 | **Definition 2.2**. 40 | 41 | ```agda 42 | _ = Basis.fiber 43 | ``` 44 | 45 | **Definition 2.3**. 46 | 47 | ```agda 48 | _ = Basis.isEquiv 49 | ``` 50 | 51 | **Definition 2.4**. 52 | 53 | ```agda 54 | _ = Basis._≃_ 55 | ``` 56 | 57 | **Definition 2.5**. 58 | 59 | ```agda 60 | _ = Basis.idEquiv 61 | ``` 62 | 63 | **Definition 2.6**. 64 | 65 | ```agda 66 | _ = Cubical.Foundations.Univalence.pathToEquiv 67 | ``` 68 | 69 | **Definition 2.7**. 70 | 71 | ```agda 72 | _ = Basis._~_ 73 | ``` 74 | 75 | **Proposition 2.8**. 76 | 77 | ```agda 78 | _ = Basis.funExt 79 | ``` 80 | 81 | ### 2.4: Homotopy levels 82 | 83 | **Definition 2.9**. 84 | 85 | ```agda 86 | _ = Basis.isOfHLevel 87 | ``` 88 | 89 | **Proposition 2.10**. 90 | 91 | ```agda 92 | _ = Basis.isOfHLevelSuc 93 | ``` 94 | 95 | **Proposition 2.11**. 96 | 97 | ```agda 98 | _ = Basis.isOfHLevelΠ 99 | ``` 100 | 101 | **Proposition 2.12**. 102 | 103 | ```agda 104 | _ = Basis.isOfHLevelΣ 105 | ``` 106 | 107 | Definition 2.13 is omitted. 108 | 109 | **Definition 2.14**. 110 | 111 | ```agda 112 | _ = Basis.isProp 113 | ``` 114 | 115 | Proposition 2.15 is omitted. 116 | 117 | **Definition 2.16**. 118 | 119 | ```agda 120 | _ = Basis.hProp 121 | ``` 122 | 123 | **Proposition 2.17**. 124 | 125 | ```agda 126 | _ = Basis.isPropΠ 127 | ``` 128 | 129 | **Proposition 2.18**. 130 | 131 | ```agda 132 | _ = Basis.isPropΣ 133 | ``` 134 | 135 | **Proposition 2.19**. 136 | 137 | ```agda 138 | _ = Basis.isPropIsProp 139 | ``` 140 | 141 | **Proposition 2.20**. 142 | 143 | ```agda 144 | _ = Basis.Σ≡Prop 145 | ``` 146 | 147 | **Definition 2.21**. This is defined directly for h-propositions in the 148 | `cubical` library. 149 | 150 | ```agda 151 | _ = Basis._⇔_ 152 | ``` 153 | 154 | **Proposition 2.22**. 155 | 156 | ```agda 157 | _ = Basis.⇔toPath 158 | ``` 159 | 160 | **Definition 2.23**. 161 | 162 | ```agda 163 | _ = Basis.Iso 164 | ``` 165 | 166 | **Proposition 2.24**. 167 | 168 | ```agda 169 | _ = Basis.isoToEquiv 170 | _ = Basis.equivToIso 171 | ``` 172 | 173 | **Definition 2.25**. 174 | 175 | ```agda 176 | _ = Basis.isSet 177 | ``` 178 | 179 | **Proposition 2.26**. 180 | 181 | ```agda 182 | _ = Basis.isProp→isSet 183 | ``` 184 | 185 | **Proposition 2.27**. 186 | 187 | ```agda 188 | _ = Basis.isSetHProp 189 | ``` 190 | 191 | **Proposition 2.28**. 192 | 193 | ```agda 194 | _ = Basis.isSetΠ 195 | ``` 196 | 197 | **Proposition 2.29**. 198 | 199 | ```agda 200 | _ = Basis.isSetΣ 201 | ``` 202 | 203 | **Proposition 2.30**. 204 | 205 | ```agda 206 | _ = Basis.isPropIsSet 207 | ``` 208 | 209 | **Definition 2.31**. 210 | 211 | ```agda 212 | _ = Cubical.Relation.Nullary.Dec 213 | ``` 214 | 215 | **Definition 2.32**. 216 | 217 | ```agda 218 | _ = Cubical.Relation.Nullary.Discrete 219 | ``` 220 | 221 | **Theorem 2.33**. 222 | 223 | ```agda 224 | _ = Cubical.Relation.Nullary.DecidableEq.Discrete→isSet 225 | ``` 226 | 227 | ### 2.5: Powersets 228 | 229 | **Definition 2.34**. 230 | 231 | ```agda 232 | _ = Basis.𝒫 233 | ``` 234 | 235 | **Proposition 2.35**. 236 | 237 | ```agda 238 | _ = Basis.𝒫-set 239 | ``` 240 | 241 | **Definition 2.36**. 242 | 243 | ```agda 244 | _ = Basis._⊆_ 245 | ``` 246 | 247 | **Definition 2.37**. 248 | 249 | ```agda 250 | _ = Basis.entire 251 | ``` 252 | 253 | **Definition 2.38**. 254 | 255 | ```agda 256 | _ = Basis._∩_ 257 | ``` 258 | 259 | ### 2.6: Families 260 | 261 | **Definition 2.39**. 262 | 263 | ```agda 264 | _ = Basis.Fam 265 | ``` 266 | 267 | **Definition 2.41**. 268 | 269 | ```agda 270 | _ = Basis._ε_ 271 | ``` 272 | 273 | **Definition 2.42**. 274 | 275 | ```agda 276 | _ = Basis._⟨$⟩_ 277 | ``` 278 | 279 | **Definition 2.43**. 280 | 281 | ```agda 282 | _ = Basis.⟪_⟫ 283 | ``` 284 | 285 | ### 2.8: Truncation 286 | 287 | **Definition 2.44**. 288 | 289 | ```agda 290 | _ = Basis.∥_∥ 291 | _ = Basis.∥∥-prop 292 | _ = Basis.∥∥-rec 293 | ``` 294 | 295 | ## Chapter 3: Frames 296 | 297 | ### 3.1: Partially ordered sets 298 | 299 | **Definition 3.1**. 300 | 301 | ```agda 302 | _ = Poset.Poset 303 | _ = Poset.PosetStr 304 | _ = Poset.PosetAx 305 | ``` 306 | 307 | **Definition 3.2**. 308 | 309 | ```agda 310 | _ = Poset.isDownwardsClosed 311 | _ = Poset.DCSubset 312 | ``` 313 | 314 | **Proposition 3.3**. 315 | 316 | ```agda 317 | _ = Poset.DCSubset-set 318 | ``` 319 | 320 | **Proposition 3.4**. 321 | 322 | ```agda 323 | _ = Frame.DCPoset 324 | ``` 325 | 326 | **Definition 3.5**. 327 | 328 | ```agda 329 | _ = Poset.isMonotonic 330 | ``` 331 | 332 | **Definition 3.6**. 333 | 334 | ```agda 335 | _ = Poset.isPosetIso 336 | ``` 337 | 338 | ### 3.2: Definition of a frame 339 | 340 | **Definition 3.7**. 341 | 342 | ```agda 343 | _ = Frame.RawFrameStr 344 | _ = Frame.FrameAx 345 | _ = Frame.FrameStr 346 | ``` 347 | 348 | **Proposition 3.8**. 349 | 350 | ```agda 351 | _ = Frame.FrameAx-props 352 | ``` 353 | 354 | **Definition 3.9**. 355 | 356 | ```agda 357 | _ = Frame.isFrameHomomorphism 358 | _ = Frame._─f→_ 359 | _ = Frame._≅f_ 360 | ``` 361 | 362 | **Definition 3.10**. 363 | 364 | ```agda 365 | _ = Frame.isFrameIso 366 | ``` 367 | 368 | **Definition 3.11** is not explicitly defined. We refer to it in an ad hoc way 369 | by referring to `_≅ₚ_` on the underlying posets. 370 | 371 | The equivalence of Defn. 3.10 and Defn. 3.11 is stated only in passing in the 372 | thesis, not as an explicit proposition but is witnessed in the Agda code 373 | by the following function: 374 | 375 | ```agda 376 | _ = Frame.≅ₚ≃≅f 377 | ``` 378 | 379 | ### 3.3: Some properties of frames 380 | 381 | **Proposition 3.12**. 382 | 383 | ```agda 384 | _ = Frame.comm 385 | ``` 386 | 387 | **Lemma 3.13**. 388 | 389 | ```agda 390 | _ = Frame.flatten 391 | ``` 392 | 393 | **Proposition 3.14**. 394 | 395 | ```agda 396 | _ = Frame.family-iff 397 | ``` 398 | 399 | **Proposition 3.15**. 400 | 401 | ```agda 402 | _ = Frame.sym-distr 403 | ``` 404 | 405 | ### 3.4: Univalence for frames 406 | 407 | **Definition 3.16**. 408 | 409 | ```agda 410 | _ = Poset.isAMonotonicEqv 411 | ``` 412 | 413 | **Definition 3.17**. 414 | 415 | ```agda 416 | _ = Cubical.Foundations.SIP.SIP 417 | ``` 418 | 419 | **Definition 3.18**. 420 | 421 | ```agda 422 | _ = Frame.isHomoEqv 423 | ``` 424 | 425 | Equation 3.19. 426 | 427 | ```agda 428 | _ = Frame.≃f≃≡ 429 | ``` 430 | 431 | Equation 3.20. 432 | 433 | ```agda 434 | _ = Frame.≃f≃≅ₚ 435 | ``` 436 | 437 | Equation 3.21. 438 | 439 | ```agda 440 | _ = Frame.≅ₚ≃≡ 441 | ``` 442 | 443 | ### 3.5: Frames of downwards-closed subsets 444 | 445 | **Theorem 3.22**. 446 | 447 | ```agda 448 | _ = Frame.DCFrame 449 | ``` 450 | 451 | ### 3.6: Nuclei and their fixed points 452 | 453 | **Definition 3.23**. 454 | 455 | ```agda 456 | _ = Nucleus.isNuclear 457 | _ = Nucleus.Nucleus 458 | ``` 459 | 460 | **Proposition 3.24**. 461 | 462 | ```agda 463 | _ = Nucleus.nuclei-resp-⊤ 464 | ``` 465 | 466 | **Lemma 3.25**. This is broken up into two functions in the Agda formalisatoin. 467 | 468 | ```agda 469 | _ = Frame.x⊑y⇒x=x∧y 470 | _ = Frame.x=x∧y⇒x⊑y 471 | ``` 472 | 473 | **Proposition 3.26**. 474 | 475 | ```agda 476 | _ = Nucleus.mono 477 | ``` 478 | 479 | **Proposition 3.27**. 480 | 481 | ```agda 482 | _ = Nucleus.𝔣𝔦𝔵-pos 483 | ``` 484 | 485 | **Theorem 3.28**. 486 | 487 | ```agda 488 | _ = Nucleus.𝔣𝔦𝔵 489 | ``` 490 | 491 | ## Chapter 4: Formal Topology 492 | 493 | ### 4.1: Interaction systems 494 | 495 | **Definition 4.1**. 496 | 497 | ```agda 498 | _ = InteractionStr 499 | _ = InteractionSys 500 | ``` 501 | 502 | **Definition 4.2**. 503 | 504 | ```agda 505 | _ = hasMono 506 | ``` 507 | 508 | **Definition 4.3**. 509 | 510 | ```agda 511 | _ = hasSimulation 512 | ``` 513 | 514 | **Definition 4.4**. 515 | 516 | ```agda 517 | _ = FormalTopology 518 | ``` 519 | ### 4.2: Cover relation 520 | 521 | Note that **Definition 4.5** is not formalised. 522 | 523 | **Definition 4.7**. 524 | 525 | ```agda 526 | _ = CoverFromFormalTopology._◁_ 527 | ``` 528 | 529 | **Proposition 4.8**. 530 | 531 | ```agda 532 | _ = CoverFromFormalTopology.◁-lem₁ 533 | ``` 534 | 535 | **Proposition 4.9**. 536 | 537 | ```agda 538 | _ = CoverFromFormalTopology.◁-lem₂ 539 | ``` 540 | 541 | **Proposition 4.10**. 542 | 543 | ```agda 544 | _ = CoverFromFormalTopology.◁-lem₃ 545 | ``` 546 | 547 | **Proposition 4.11**. 548 | 549 | ```agda 550 | _ = CoverFromFormalTopology.◁-lem₄ 551 | ``` 552 | 553 | ### 4.3: Covers are nuclei 554 | 555 | **Theorem 4.12**. 556 | 557 | ```agda 558 | _ = CoverFormsNucleus.NucleusFrom.𝕛-nuclear 559 | ``` 560 | 561 | ### 4.4: Lifting into the generated frame 562 | 563 | **Definition 4.13**. 564 | 565 | ```agda 566 | _ = CoverFormsNucleus.NucleusFrom.η 567 | ``` 568 | 569 | ### 4.5: Formal topologies present 570 | 571 | **Definition 4.14**. 572 | 573 | ```agda 574 | _ = UniversalProperty.represents 575 | ``` 576 | 577 | **Definition 4.15**. 578 | 579 | ```agda 580 | _ = UniversalProperty.isFlat 581 | ``` 582 | 583 | **Theorem 4.16**. 584 | 585 | The theorem statement is given in: 586 | 587 | ```agda 588 | _ = UniversalProperty.universal-prop 589 | ``` 590 | 591 | The proof is given in: 592 | 593 | ```agda 594 | _ = UniversalProperty.main 595 | ``` 596 | 597 | **Lemma 4.17**. 598 | 599 | ```agda 600 | _ = UniversalProperty.main-lemma 601 | ``` 602 | 603 | **Lemma 4.18**. 604 | 605 | ```agda 606 | _ = UniversalProperty.MainProof.resp-⋁-lem 607 | ``` 608 | 609 | ## Chapter 5: The Cantor space 610 | 611 | ### 5.1: The Cantor interaction system 612 | 613 | **Definition 5.1**. 614 | 615 | ```agda 616 | _ = SnocList._⌢_ 617 | _ = SnocList.[] 618 | ``` 619 | 620 | **Definition 5.2**. 621 | 622 | ```agda 623 | _ = _++_ 624 | ``` 625 | 626 | **Proposition 5.3**. 627 | 628 | ```agda 629 | _ = SnocList-discrete 630 | ``` 631 | 632 | **Definition 5.4**. 633 | 634 | ```agda 635 | _ = CantorSpace.ℂ-pos 636 | ``` 637 | 638 | **Definition 5.5**. 639 | 640 | ```agda 641 | _ = CantorSpace.ℂ-IS 642 | ``` 643 | 644 | **Theorem 5.6**. 645 | 646 | ```agda 647 | _ = CantorSpace.cantor 648 | ``` 649 | 650 | ### 5.2: The Cantor space is compact 651 | 652 | **Definition 5.7**. 653 | 654 | ```agda 655 | _ = Compactness.down 656 | ``` 657 | 658 | **Definition 5.8**. 659 | 660 | ```agda 661 | _ = Compactness.isCompact 662 | ``` 663 | 664 | **Lemma 5.9**. 665 | 666 | ```agda 667 | _ = CantorSpace.U⊆V⇒◁U⊆◁V 668 | ``` 669 | 670 | **Lemma 5.10**. In the Agda formalisation, this is broken up into two functions. 671 | 672 | ```agda 673 | _ = CantorSpace.↓-++-left 674 | _ = CantorSpace.↓-++-right 675 | ``` 676 | 677 | **Lemma 5.11**. 678 | 679 | ```agda 680 | _ = CantorSpace.◁^-decide 681 | ``` 682 | 683 | **Theorem 5.12**. 684 | 685 | ```agda 686 | _ = CantorSpace.compact 687 | ``` 688 | -------------------------------------------------------------------------------- /FormalTopology/Regular.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Regular Locales 3 | --- 4 | 5 | ```agda 6 | {-# OPTIONS --cubical --safe #-} 7 | 8 | module Regular where 9 | 10 | open import Cubical.Core.Everything hiding (_∧_; _∨_) 11 | open import Basis hiding (_∨_) 12 | open import Poset 13 | open import Nucleus 14 | open import Frame 15 | open import CoverFormsNucleus 16 | 17 | ``` 18 | 19 | ## Definition 20 | 21 | We define the well-inside relation exactly as in Johstone. This is a pointless 22 | characterisation of the relation that *U* ⋜ *V* := Clos(*U*) ⊆ *V*. Notice why this is 23 | called "well-inside": if Clos(*U*) ⊆ *V* it means *U* is inside *V* in such way that it 24 | doesn't touch its boundary. 25 | 26 | ```agda 27 | well-inside : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ∣ F ∣F → ∣ F ∣F → Type ℓ₀ 28 | well-inside F x y = 29 | Σ[ z ∈ ∣ F ∣F ] (x ⊓[ F ] z ≡ ⊥[ F ]) × (y ∨[ F ] z ≡ ⊤[ F ]) 30 | 31 | syntax well-inside F x y = x ⋜[ F ] y 32 | ``` 33 | 34 | We denote by ⇊ *x* the set of everything well-inside *x*. 35 | 36 | ```agda 37 | ⇊ : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ∣ F ∣F → Fam ℓ₀ ∣ F ∣F 38 | ⇊ F x = (Σ[ y ∈ ∣ F ∣F ] y ⋜[ F ] x) , π₀ 39 | ``` 40 | 41 | A locale *L* is then called **regular** iff every element in it is the join of all 42 | elements well-inside it. 43 | 44 | ```agda 45 | isRegular : (F : Frame ℓ₀ ℓ₁ ℓ₂) → hProp (ℓ-max ℓ₀ ℓ₁) 46 | isRegular F = ∀[ x ∶ ∣ F ∣F ] isSup (pos F) (⇊ F x) x 47 | ``` 48 | 49 | ## Some properties 50 | 51 | ```agda 52 | complements : (F : Frame 𝓤 𝓥 𝓦) → ∣ F ∣F → ∣ F ∣F → Type 𝓤 53 | complements F x y = (x ⊓[ F ] y ≡ ⊥[ F ]) × (x ∨[ F ] y ≡ ⊤[ F ]) 54 | ``` 55 | 56 | ```agda 57 | -- In other words, x is clopen. 58 | hasComplement : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ∣ F ∣F → Type ℓ₀ 59 | hasComplement F x = 60 | Σ[ y ∈ ∣ F ∣F ] (x ⊓[ F ] y ≡ ⊥[ F ]) × (x ∨[ F ] y ≡ ⊤[ F ]) 61 | 62 | hasComplement′ : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ∣ F ∣F → Type (ℓ-max (ℓ-max ℓ₀ ℓ₁) ℓ₂) 63 | hasComplement′ {ℓ₂ = ℓ₂} F x = 64 | Σ[ y ∈ ∣ F ∣F ] [ isInf (pos F) ⊥[ F ] x y ] × [ isSup {ℓ₂ = ℓ₂} (pos F) ⁅ x , y ⁆ ⊤[ F ] ] 65 | 66 | fclopens : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ∣ F ∣F → Fam ℓ₀ ∣ F ∣F 67 | fclopens F x = (Σ[ x ∈ ∣ F ∣F ] hasComplement F x) , π₀ 68 | 69 | complements-unique : (F : Frame 𝓤 𝓥 𝓦) 70 | → (x y₀ y₁ : ∣ F ∣F) 71 | → complements F x y₀ → complements F x y₁ → y₀ ≡ y₁ 72 | complements-unique F x y y′ (p , q) (p′ , q′) = nts 73 | where 74 | abstract 75 | nts : y ≡ y′ 76 | nts = 77 | y ≡⟨ sym (x∧⊤=x F y) ⟩ 78 | y ⊓[ F ] ⊤[ F ] ≡⟨ cong (λ - → y ⊓[ F ] -) (sym q′) ⟩ 79 | y ⊓[ F ] (x ∨[ F ] y′) ≡⟨ bin-dist F y x y′ ⟩ 80 | (y ⊓[ F ] x) ∨[ F ] (y ⊓[ F ] y′) ≡⟨ cong (λ - → - ∨[ F ] (y ⊓[ F ] y′)) (comm F y x) ⟩ 81 | (x ⊓[ F ] y) ∨[ F ] (y ⊓[ F ] y′) ≡⟨ cong (λ - → - ∨[ F ] (y ⊓[ F ] y′)) p ⟩ 82 | ⊥[ F ] ∨[ F ] (y ⊓[ F ] y′) ≡⟨ cong (λ - → - ∨[ F ] (y ⊓[ F ] y′)) (sym p′) ⟩ 83 | (x ⊓[ F ] y′) ∨[ F ] (y ⊓[ F ] y′) ≡⟨ cong (λ - → - ∨[ F ] (y ⊓[ F ] y′)) (comm F x y′) ⟩ 84 | (y′ ⊓[ F ] x) ∨[ F ] (y ⊓[ F ] y′) ≡⟨ sym (cong (λ - → (y′ ⊓[ F ] x) ∨[ F ] -) (comm F y′ y)) ⟩ 85 | (y′ ⊓[ F ] x) ∨[ F ] (y′ ⊓[ F ] y) ≡⟨ sym (bin-dist F y′ x y) ⟩ 86 | y′ ⊓[ F ] (x ∨[ F ] y) ≡⟨ comm F y′ (x ∨[ F ] y) ⟩ 87 | (x ∨[ F ] y) ⊓[ F ] y′ ≡⟨ cong (λ - → - ⊓[ F ] y′) q ⟩ 88 | ⊤[ F ] ⊓[ F ] y′ ≡⟨ comm F ⊤[ F ] y′ ⟩ 89 | y′ ⊓[ F ] ⊤[ F ] ≡⟨ x∧⊤=x F y′ ⟩ 90 | y′ ∎ 91 | 92 | hasComplement-prop : (F : Frame 𝓤 𝓥 𝓦) → (x : ∣ F ∣F) → isProp (hasComplement F x) 93 | hasComplement-prop F x (y , p , q) (y′ , p′ , q′) = 94 | Σ≡Prop 95 | (λ w → isPropΣ (F-set _ _) λ _ → F-set _ _) 96 | (complements-unique F x y y′ (p , q) (p′ , q′)) 97 | where 98 | F-set : isSet ∣ F ∣F 99 | F-set = carrier-is-set (pos F) 100 | 101 | module SomePropertiesOf⋜ (F : Frame ℓ₀ ℓ₁ ℓ₂) where 102 | 103 | private 104 | _⊑_ = λ (x y : ∣ F ∣F) → x ⊑[ pos F ] y 105 | 106 | ⋜-comp : (x : ∣ F ∣F) → (x ⋜[ F ] x) ↔ hasComplement F x 107 | ⋜-comp x = (λ x → x) , (λ x → x) 108 | 109 | a⋜b→a⊑b : (x y : ∣ F ∣F) → x ⋜[ F ] y → [ x ⊑[ pos F ] y ] 110 | a⋜b→a⊑b x y (z , p , q) = x=x∧y⇒x⊑y F NTS 111 | where 112 | open PosetReasoning (pos F) 113 | 114 | NTS : x ≡ x ⊓[ F ] y 115 | NTS = x ≡⟨ sym (x∧⊤=x F x) ⟩ 116 | x ⊓[ F ] ⊤[ F ] ≡⟨ cong (λ - → x ⊓[ F ] -) (sym q) ⟩ 117 | x ⊓[ F ] (y ∨[ F ] z) ≡⟨ bin-dist F x y z ⟩ 118 | (x ⊓[ F ] y) ∨[ F ] (x ⊓[ F ] z) ≡⟨ cong (λ - → _ ∨[ F ] -) p ⟩ 119 | (x ⊓[ F ] y) ∨[ F ] ⊥[ F ] ≡⟨ ∨-comm F (x ⊓[ F ] y) ⊥[ F ] ⟩ 120 | ⊥[ F ] ∨[ F ] (x ⊓[ F ] y) ≡⟨ x∨⊥=x F (x ⊓[ F ] y) ⟩ 121 | x ⊓[ F ] y ∎ 122 | 123 | a⋜c≤d : {x y z : ∣ F ∣F} → x ⋜[ F ] y → [ y ⊑[ pos F ] z ] → x ⋜[ F ] z 124 | a⋜c≤d {x} {y} {z} (c , p , q) y⊑z = 125 | c , (p , ⊑[ pos F ]-antisym _ _ (⊤[ F ]-top _) ⊤⊑z∨c) 126 | where 127 | open PosetReasoning (pos F) 128 | 129 | y⊑z∨c : [ y ⊑[ pos F ] (z ∨[ F ] c) ] 130 | y⊑z∨c = y ⊑⟨ y⊑z ⟩ z ⊑⟨ ⊔[ F ]-upper₀ z c ⟩ z ∨[ F ] c ■ 131 | 132 | ⊤⊑z∨c : [ ⊤[ F ] ⊑[ pos F ] (z ∨[ F ] c) ] 133 | ⊤⊑z∨c = 134 | subst (λ - → [ - ⊑[ pos F ] _ ]) q (⊔[ F ]-least _ _ _ y⊑z∨c (⊔[ F ]-upper₁ z c)) 135 | ``` 136 | 137 | ## Alternative characterisation 138 | 139 | Another way of characterising regularity is this: a locale *L* is called regular iff each 140 | of its elements can be written as the join of a _clopen_ family. Before looking at this 141 | though, let us first discuss how we can express clopen-ness. 142 | 143 | We say that some open *x* ∈ *L* is clopen iff it has a complement. This can be motivated 144 | by the fact that a set is clopen iff its boundary is empty i.e. it satisfies LEM. Now 145 | we can write down the alternative characterisation we mentioned. 146 | 147 | ```agda 148 | hasClopenBasis : (F : Frame ℓ₀ ℓ₁ ℓ₂) → Type (ℓ-max ℓ₀ (ℓ-suc ℓ₂)) 149 | hasClopenBasis {ℓ₂ = ℓ₂} F = 150 | (x : ∣ F ∣F) → 151 | Σ[ U ∈ Fam ℓ₂ ∣ F ∣F ] ((y : ∣ F ∣F) → y ε U → hasComplement F y) × (x ≡ ⋁[ F ] U) 152 | ``` 153 | 154 | ```agda 155 | regularity-lemma : (F : Frame ℓ₀ ℓ₁ ℓ₂) → hasClopenBasis F → [ isRegular F ] 156 | regularity-lemma F p x = upper , subst goal (sym x=⋁𝔘) ψ 157 | where 158 | open PosetReasoning (pos F) 159 | open SomePropertiesOf⋜ F 160 | 161 | goal = λ - → (y : ∣ F ∣F) → [ ∀[ k ε ⇊ F x ] (k ⊑[ pos F ] y) ] → [ - ⊑[ pos F ] y ] 162 | 163 | 𝔘 = π₀ (p x) 164 | 165 | x=⋁𝔘 : x ≡ ⋁[ F ] 𝔘 166 | x=⋁𝔘 = π₁ (π₁ (p x)) 167 | 168 | has-comp : (y : ∣ F ∣F) → y ε 𝔘 → hasComplement F y 169 | has-comp = π₀ (π₁ (p x)) 170 | 171 | upper : [ ∀[ y ε (⇊ F x) ] (y ⊑[ pos F ] x) ] 172 | upper y ((_ , wi) , eq) = subst (λ - → [ - ⊑[ pos F ] x ]) eq (a⋜b→a⊑b _ x wi) 173 | 174 | ψ : (y : ∣ F ∣F) → [ ∀[ k ε ⇊ F x ] (k ⊑[ pos F ] y) ] → [ (⋁[ F ] 𝔘) ⊑[ pos F ] y ] 175 | ψ y q = ⋁[ F ]-least 𝔘 y NTS 176 | where 177 | NTS : [ ∀[ k ε 𝔘 ] (k ⊑[ pos F ] y) ] 178 | NTS k (i , eq) = q k kε⇊Fx 179 | where 180 | 𝔘ᵢ-has-comp : hasComplement F (𝔘 $ i) 181 | 𝔘ᵢ-has-comp = has-comp (𝔘 $ i) (i , refl) 182 | 183 | 𝔘ᵢ⋜⋁𝔘 : (𝔘 $ i) ⋜[ F ] (⋁[ F ] 𝔘) 184 | 𝔘ᵢ⋜⋁𝔘 = a⋜c≤d 𝔘ᵢ-has-comp (⋁[ F ]-upper 𝔘 (𝔘 $ i) (i , refl)) 185 | 186 | kε⇊Fx : k ε ⇊ F x 187 | kε⇊Fx = (𝔘 $ i , subst (λ - → _ ⋜[ F ] -) (sym x=⋁𝔘) 𝔘ᵢ⋜⋁𝔘) , eq 188 | ``` 189 | 190 | ```agda 191 | ∧-complement : (F : Frame ℓ₀ ℓ₁ ℓ₂) 192 | → (x y x′ y′ : ∣ F ∣F) 193 | → complements F x x′ 194 | → complements F y y′ 195 | → complements F (x ⊓[ F ] y) (x′ ∨[ F ] y′) 196 | ∧-complement F x y x′ y′ (x∧x=⊥ , x∨x′=⊤) (y∧y′=⊥ , y∨y′=⊤) = G𝟏 , G𝟐 197 | where 198 | _∧_ : ∣ F ∣F → ∣ F ∣F → ∣ F ∣F 199 | x ∧ y = x ⊓[ F ] y 200 | 201 | _∨_ : ∣ F ∣F → ∣ F ∣F → ∣ F ∣F 202 | x ∨ y = x ∨[ F ] y 203 | 204 | G𝟏 : (x ∧ y) ∧ (x′ ∨ y′) ≡ ⊥[ F ] 205 | G𝟏 = (x ∧ y) ∧ (x′ ∨ y′) ≡⟨ ⦅𝟏⦆ ⟩ 206 | ((x ∧ y) ∧ x′) ∨ ((x ∧ y) ∧ y′) ≡⟨ ⦅𝟐⦆ ⟩ 207 | ((y ∧ x) ∧ x′) ∨ ((x ∧ y) ∧ y′) ≡⟨ ⦅𝟑⦆ ⟩ 208 | (y ∧ (x ∧ x′)) ∨ ((x ∧ y) ∧ y′) ≡⟨ ⦅𝟒⦆ ⟩ 209 | (y ∧ ⊥[ F ]) ∨ ((x ∧ y) ∧ y′) ≡⟨ ⦅𝟓⦆ ⟩ 210 | ⊥[ F ] ∨ ((x ∧ y) ∧ y′) ≡⟨ ⦅𝟔⦆ ⟩ 211 | (x ∧ y) ∧ y′ ≡⟨ ⦅𝟕⦆ ⟩ 212 | x ∧ (y ∧ y′) ≡⟨ ⦅𝟖⦆ ⟩ 213 | x ∧ ⊥[ F ] ≡⟨ ⦅𝟗⦆ ⟩ 214 | ⊥[ F ] ∎ 215 | where 216 | ⦅𝟏⦆ = bin-dist F (x ∧ y) x′ y′ 217 | ⦅𝟐⦆ = cong (λ - → (- ∧ x′) ∨ ((x ∧ y) ∧ y′)) (comm F x y) 218 | ⦅𝟑⦆ = cong (λ - → - ∨ ((x ∧ y) ∧ y′)) (⊓[ F ]-assoc y x x′) 219 | ⦅𝟒⦆ = cong (λ - → (y ∧ -) ∨ ((x ∧ y) ∧ y′)) x∧x=⊥ 220 | ⦅𝟓⦆ = cong (λ - → - ∨ ((x ∧ y) ∧ y′)) (x∧⊥=⊥ F y) 221 | ⦅𝟔⦆ = x∨⊥=x F ((x ∧ y) ∧ y′) 222 | ⦅𝟕⦆ = ⊓[ F ]-assoc x y y′ 223 | ⦅𝟖⦆ = cong (λ - → x ∧ -) y∧y′=⊥ 224 | ⦅𝟗⦆ = x∧⊥=⊥ F x 225 | 226 | G𝟐 : (x ∧ y) ∨ (x′ ∨ y′) ≡ ⊤[ F ] 227 | G𝟐 = (x ∧ y) ∨ (x′ ∨ y′) ≡⟨ ⦅𝟏⦆ ⟩ 228 | (x′ ∨ y′) ∨ (x ∧ y) ≡⟨ ⦅𝟐⦆ ⟩ 229 | (((x′ ∨ y′) ∨ x) ∧ ((x′ ∨ y′) ∨ y)) ≡⟨ ⦅𝟑⦆ ⟩ 230 | (((y′ ∨ x′) ∨ x) ∧ ((x′ ∨ y′) ∨ y)) ≡⟨ ⦅𝟒⦆ ⟩ 231 | ((y′ ∨ (x′ ∨ x)) ∧ ((x′ ∨ y′) ∨ y)) ≡⟨ ⦅𝟓⦆ ⟩ 232 | ((y′ ∨ ⊤[ F ]) ∧ ((x′ ∨ y′) ∨ y)) ≡⟨ ⦅𝟔⦆ ⟩ 233 | ⊤[ F ] ∧ ((x′ ∨ y′) ∨ y) ≡⟨ ⦅𝟕⦆ ⟩ 234 | ((x′ ∨ y′) ∨[ F ] y) ≡⟨ ⦅𝟖⦆ ⟩ 235 | x′ ∨ (y′ ∨[ F ] y) ≡⟨ ⦅𝟗⦆ ⟩ 236 | x′ ∨ ⊤[ F ] ≡⟨ ⦅𝟏𝟎⦆ ⟩ 237 | ⊤[ F ] ∎ 238 | where 239 | ⦅𝟏⦆ = ∨-comm F (x ∧ y) (x′ ∨ y′) 240 | ⦅𝟐⦆ = bin-dist-op F (x′ ∨ y′) x y 241 | ⦅𝟑⦆ = cong (λ - → (- ∨ x) ∧ ((x′ ∨ y′) ∨ y)) (∨-comm F x′ y′) 242 | ⦅𝟒⦆ = cong (λ - → - ∧ ((x′ ∨ y′) ∨ y)) (∨[ F ]-assoc y′ x′ x) 243 | ⦅𝟓⦆ = cong (λ - → (y′ ∨ -) ∧ ((x′ ∨ y′) ∨ y)) (_ ≡⟨ ∨-comm F x′ x ⟩ _ ≡⟨ x∨x′=⊤ ⟩ _ ∎) 244 | ⦅𝟔⦆ = cong (λ - → - ∧ ((x′ ∨ y′) ∨ y)) (x∨⊤=⊤ F y′ ) 245 | ⦅𝟕⦆ = ⊤[ F ] ∧ ((x′ ∨ y′) ∨ y) ≡⟨ comm F ⊤[ F ] ((x′ ∨ y′) ∨ y) ⟩ 246 | ((x′ ∨ y′) ∨ y) ∧ ⊤[ F ] ≡⟨ x∧⊤=x F ((x′ ∨ y′) ∨ y) ⟩ 247 | (x′ ∨ y′) ∨ y ∎ 248 | ⦅𝟖⦆ = ∨[ F ]-assoc x′ y′ y 249 | ⦅𝟗⦆ = cong (λ - → x′ ∨ -) (_ ≡⟨ ∨-comm F y′ y ⟩ y ∨[ F ] y′ ≡⟨ y∨y′=⊤ ⟩ _ ∎) 250 | ⦅𝟏𝟎⦆ = x∨⊤=⊤ F x′ 251 | 252 | complements-sym : (F : Frame 𝓤 𝓥 𝓦) 253 | → {x x′ : ∣ F ∣F} 254 | → complements F x x′ 255 | → complements F x′ x 256 | complements-sym F {x} {x′} (p , q) = G𝟏 , G𝟐 257 | where 258 | G𝟏 : x′ ⊓[ F ] x ≡ ⊥[ F ] 259 | G𝟏 = x′ ⊓[ F ] x ≡⟨ comm F x′ x ⟩ 260 | x ⊓[ F ] x′ ≡⟨ p ⟩ 261 | ⊥[ F ] ∎ 262 | 263 | G𝟐 : x′ ∨[ F ] x ≡ ⊤[ F ] 264 | G𝟐 = x′ ∨[ F ] x ≡⟨ ∨-comm F x′ x ⟩ 265 | x ∨[ F ] x′ ≡⟨ q ⟩ 266 | ⊤[ F ] ∎ 267 | 268 | ∧-has-complement : (F : Frame 𝓤 𝓥 𝓦) (x y : ∣ F ∣F) 269 | → hasComplement F x 270 | → hasComplement F y 271 | → hasComplement F (x ⊓[ F ] y) 272 | ∧-has-complement F x y (¬x , ¬x-complements-x) (¬y , ¬y-complements-y) = 273 | ¬x ∨[ F ] ¬y , ∧-complement F x y ¬x ¬y ¬x-complements-x ¬y-complements-y 274 | ``` 275 | -------------------------------------------------------------------------------- /FormalTopology/CantorSpace.lagda.md: -------------------------------------------------------------------------------- 1 | ``` 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module CantorSpace where 5 | 6 | open import Basis hiding (A; B) 7 | open import Cubical.Data.Empty.Base using (⊥; rec) 8 | open import Cubical.Data.Bool.Base using (true; false; _≟_) renaming (Bool to 𝔹) 9 | open import Cubical.Data.List using (List; _∷_; []) renaming (_++_ to _^_) 10 | open import Frame 11 | open import Nucleus 12 | open import CoverFormsNucleus 13 | open import Cover 14 | open import Poset 15 | open import FormalTopology 16 | open import Compactness 17 | ``` 18 | 19 | We open the `SnocList` module with the type `𝔹` of booleans. 20 | 21 | ``` 22 | open import SnocList 𝔹 _≟_ renaming (SnocList to ℂ; SnocList-set to ℂ-set) 23 | ``` 24 | 25 | The empty list and the snoc operator are called `[]` and `⌢` respectively. Concatenation 26 | operation on snoc lists is called `_++_`. Note that concatenation on lists is therefore 27 | renamed to `_^_` to prevent conflict. 28 | 29 | ## The Cantor poset 30 | 31 | `xs` is less than `ys` if there is some `zs` such that `xs = ys ++ zs`. 32 | 33 | ``` 34 | _≤_ : ℂ → ℂ → hProp ℓ-zero 35 | xs ≤ ys = (Σ[ zs ∈ ℂ ] xs ≡ ys ++ zs) , prop 36 | where 37 | prop : isProp (Σ[ zs ∈ ℂ ] xs ≡ ys ++ zs) 38 | prop (_ , p) (_ , q) = Σ≡Prop (λ ws → ℂ-set xs (ys ++ ws)) (++-lemma p q) 39 | ``` 40 | 41 | As `_≤_` is a partial order, we package it up as a poset. 42 | 43 | ``` 44 | ℂ-pos : Poset ℓ-zero ℓ-zero 45 | ℂ-pos = ℂ , _≤_ , ℂ-set , ≤-refl , ≤-trans , ≤-antisym 46 | where 47 | ≤-refl : (xs : ℂ) → [ xs ≤ xs ] 48 | ≤-refl xs = [] , refl 49 | 50 | ≤-trans : (xs ys zs : ℂ) → [ xs ≤ ys ] → [ ys ≤ zs ] → [ xs ≤ zs ] 51 | ≤-trans xs ys zs (as , p) (bs , q) = 52 | (bs ++ as) , NTS 53 | where 54 | NTS : xs ≡ zs ++ (bs ++ as) 55 | NTS = xs ≡⟨ p ⟩ 56 | ys ++ as ≡⟨ cong (λ - → - ++ as) q ⟩ 57 | (zs ++ bs) ++ as ≡⟨ sym (assoc zs bs as) ⟩ 58 | zs ++ (bs ++ as) ∎ 59 | 60 | ≤-antisym : (xs ys : ℂ) → [ xs ≤ ys ] → [ ys ≤ xs ] → xs ≡ ys 61 | ≤-antisym xs ys ([] , p) ([] , q) = p 62 | ≤-antisym xs ys ([] , p) (bs ⌢ x , q) = p 63 | ≤-antisym xs ys (as ⌢ x , p) ([] , q) = sym q 64 | ≤-antisym xs ys (as ⌢ a , p) (bs ⌢ b , q) = 65 | rec (lemma3 NTS) 66 | where 67 | NTS : xs ≡ xs ++ ((bs ⌢ b) ++ (as ⌢ a)) 68 | NTS = xs ≡⟨ p ⟩ 69 | ys ++ (as ⌢ a) ≡⟨ cong (λ - → - ++ as ⌢ a) q ⟩ 70 | (xs ++ (bs ⌢ b)) ++ (as ⌢ a) ≡⟨ sym (assoc xs (bs ⌢ b) (as ⌢ a)) ⟩ 71 | xs ++ ((bs ⌢ b) ++ (as ⌢ a)) ∎ 72 | ``` 73 | 74 | ## The Cantor formal topology 75 | 76 | We give the formal topology of the Cantor space as an 77 | [interaction system](http://www.dcs.ed.ac.uk/home/pgh/interactive_systems.html). 78 | 79 | 1. Each inhabitant of `ℂ` is like a stage of information. 80 | 81 | 1. At each stage of information we can perform a trivial experiment: querying the next 82 | bit. 83 | ``` 84 | ℂ-exp = λ (_ : ℂ) → Unit ℓ-zero 85 | ``` 86 | 87 | 1. Outcome of the trivial experiment is the delivery of the new bit. 88 | ``` 89 | ℂ-out = λ (_ : Unit ℓ-zero) → 𝔹 90 | ``` 91 | 92 | 1. This takes us to a new stage information, obtained by snoc'ing in the new bit to the 93 | current stage of information. 94 | ``` 95 | ℂ-rev : {_ : ℂ} → 𝔹 → ℂ 96 | ℂ-rev {xs} b = xs ⌢ b 97 | ``` 98 | 99 | These four components together form an interaction system that satiesfies the monotonicity 100 | and simulation properties (given in `ℂ-mono` and `ℂ-sim`). 101 | 102 | ``` 103 | ℂ-IS : InteractionStr ℂ 104 | ℂ-IS = ℂ-exp , ℂ-out , λ {xs} → ℂ-rev {xs} 105 | 106 | ℂ-mono : hasMono ℂ-pos ℂ-IS 107 | ℂ-mono _ _ c = [] ⌢ c , refl 108 | 109 | ℂ-sim : hasSimulation ℂ-pos ℂ-IS 110 | ℂ-sim xs ys xs≤ys@([] , p) tt = tt , NTS 111 | where 112 | NTS : (b₁ : 𝔹) → Σ[ b₀ ∈ 𝔹 ] [ (xs ⌢ b₁) ≤ (ys ⌢ b₀) ] 113 | NTS b₁ = b₁ , subst (λ - → [ (xs ⌢ b₁) ≤ (- ⌢ b₁) ]) p (⊑[ ℂ-pos ]-refl _) 114 | ℂ-sim xs ys xs≤ys@(zs ⌢ z , p) tt = tt , NTS 115 | where 116 | NTS : (c₀ : 𝔹) → Σ[ c ∈ 𝔹 ] [ (xs ⌢ c₀) ≤ (ys ⌢ c) ] 117 | NTS c₀ = 118 | head (zs ⌢ z) tt , subst (λ - → [ (- ⌢ c₀) ≤ _ ]) (sym p) NTS′ 119 | where 120 | φ = cong (λ - → ys ++ (- ⌢ c₀)) (sym (hd-tl-lemma (zs ⌢ z) tt)) 121 | ψ = cong (λ - → - ⌢ c₀) (sym (snoc-lemma ys _ _)) 122 | rem = (ys ++ zs) ⌢ z ⌢ c₀ ≡⟨ φ ⟩ 123 | (ys ++ (([] ⌢ head (zs ⌢ z) tt) ++ (tail (zs ⌢ z) tt))) ⌢ c₀ ≡⟨ ψ ⟩ 124 | ((ys ⌢ head (zs ⌢ z) tt) ++ tail (zs ⌢ z) tt) ⌢ c₀ ∎ 125 | NTS′ : [ ((ys ++ zs) ⌢ z ⌢ c₀) ≤ (ys ⌢ head (zs ⌢ z) tt) ] 126 | NTS′ = ((tail (zs ⌢ z) tt) ⌢ c₀) , rem 127 | ``` 128 | 129 | We finally package up all this as a formal topology 130 | 131 | ``` 132 | cantor : FormalTopology ℓ-zero ℓ-zero 133 | cantor = ℂ-pos , ℂ-IS , ℂ-mono , ℂ-sim 134 | 135 | open NucleusFrom cantor using () renaming (L to cantor-frame) 136 | 137 | _ : Frame (ℓ-suc ℓ-zero) ℓ-zero ℓ-zero 138 | _ = cantor-frame 139 | ``` 140 | 141 | from which we get a covering relation 142 | 143 | ``` 144 | open CoverFromFormalTopology cantor renaming (_◁_ to _<ℂ|_) 145 | 146 | _ : ℂ → (ℂ → hProp ℓ-zero) → Type ℓ-zero 147 | _ = _<ℂ|_ 148 | ``` 149 | 150 | ## The Cantor formal topology is compact 151 | 152 | We now want to view a list of `ℂ`s as a _finite cover_. We associate with some 153 | `xss : List ℂ` a subset, being covered by which corresponds to being covered by this list. 154 | 155 | ``` 156 | ℂ-down : List ℂ → 𝒫 ℂ 157 | ℂ-down = down cantor 158 | 159 | syntax ℂ-down xss xs = xs ↓ xss 160 | ``` 161 | 162 | This subset is downwards-closed. 163 | 164 | ``` 165 | ↓-dc : (xss : List ℂ) → [ isDownwardsClosed ℂ-pos (λ - → - ↓ xss) ] 166 | ↓-dc (xs ∷ xss) ys zs ys◁xs∷xss zs≤ys = 167 | ∥∥-rec (is-true-prop (zs ↓ (xs ∷ xss))) NTS ys◁xs∷xss 168 | where 169 | open PosetReasoning ℂ-pos using (_⊑⟨_⟩_; _■) 170 | 171 | NTS : [ ys ≤ xs ] ⊎ [ ys ↓ xss ] → [ zs ↓ (xs ∷ xss) ] 172 | NTS (inl ys≤xs) = ∣ inl (zs ⊑⟨ zs≤ys ⟩ ys ⊑⟨ ys≤xs ⟩ xs ■) ∣ 173 | NTS (inr ys◁xss) = ∣ inr (↓-dc xss ys zs ys◁xss zs≤ys) ∣ 174 | ``` 175 | 176 | We claim that the Cantor space is compact. 177 | 178 | ``` 179 | compact : isCompact cantor 180 | ``` 181 | 182 | ### Two little lemmas 183 | 184 | ``` 185 | U⊆V⇒◁U⊆◁V : (xs : ℂ) (U : 𝒫 ℂ) (V : 𝒫 ℂ) → [ U ⊆ V ] → xs <ℂ| U → xs <ℂ| V 186 | U⊆V⇒◁U⊆◁V xs U V U⊆V = ◁-lem₄ U V NTS xs 187 | where 188 | NTS : (u : ℂ) → [ u ∈ U ] → u <ℂ| V 189 | NTS u u∈U = dir (U⊆V u u∈U) 190 | 191 | ↓-++-left : (xss yss : List ℂ) → [ (λ - → - ↓ xss) ⊆ (λ - → - ↓ (xss ^ yss)) ] 192 | ↓-++-left [] yss _ () 193 | ↓-++-left (xs ∷ xss) yss ys ys∈down-xs-xss = 194 | ∥∥-rec (is-true-prop (ys ↓ ((xs ∷ xss) ^ yss))) NTS ys∈down-xs-xss 195 | where 196 | NTS : [ ys ≤ xs ] ⊎ [ ys ↓ xss ] → [ ys ↓ (xs ∷ xss ^ yss) ] 197 | NTS (inl ys≤xs) = ∣ inl ys≤xs ∣ 198 | NTS (inr ys∈down-xss) = ∣ inr (↓-++-left xss yss ys ys∈down-xss) ∣ 199 | 200 | ↓-++-right : (xss yss : List ℂ) → [ (λ - → - ↓ yss) ⊆ (λ - → - ↓ (xss ^ yss)) ] 201 | ↓-++-right xss [] _ () 202 | ↓-++-right [] (ys ∷ yss) zs zs∈◁ys∷yss = zs∈◁ys∷yss 203 | ↓-++-right (xs ∷ xss) (ys ∷ yss) zs zs∈◁ys∷yss = 204 | ∥∥-rec (is-true-prop (zs ↓ (xs ∷ xss ^ ys ∷ yss))) NTS zs∈◁ys∷yss 205 | where 206 | NTS : [ zs ≤ ys ] ⊎ [ zs ↓ yss ] → [ zs ↓ (xs ∷ xss ^ ys ∷ yss) ] 207 | NTS (inl zs≤ys) = let IH = ↓-++-right xss _ _ ∣ inl (⊑[ ℂ-pos ]-refl ys) ∣ 208 | in ∣ inr (↓-dc (xss ^ ys ∷ yss) ys zs IH zs≤ys) ∣ 209 | NTS (inr zs◁yss) = ∣ inr (↓-++-right xss _ zs ∣ inr zs◁yss ∣) ∣ 210 | 211 | ◁^-decide : (xs : ℂ) (yss zss : List ℂ) 212 | → [ xs ↓ (yss ^ zss) ] 213 | → ∥ [ xs ↓ yss ] ⊎ [ xs ↓ zss ] ∥ 214 | ◁^-decide xs [] zss k = ∣ inr k ∣ 215 | ◁^-decide xs (ys ∷ yss) zss k = ∥∥-rec (∥∥-prop _) NTS₀ k 216 | where 217 | NTS₀ : [ xs ≤ ys ] ⊎ [ xs ↓ (yss ^ zss) ] → ∥ [ xs ↓ (ys ∷ yss) ] ⊎ [ xs ↓ zss ] ∥ 218 | NTS₀ (inl xs≤ys) = ∣ inl ∣ inl xs≤ys ∣ ∣ 219 | NTS₀ (inr xs◁yss^zss) = ∥∥-rec (∥∥-prop _) NTS₁ (◁^-decide xs yss zss xs◁yss^zss) 220 | where 221 | NTS₁ : [ xs ↓ yss ] ⊎ [ xs ↓ zss ] → ∥ [ xs ↓ (ys ∷ yss) ] ⊎ [ xs ↓ zss ] ∥ 222 | NTS₁ (inl xs◁yss) = ∣ inl ∣ inr xs◁yss ∣ ∣ 223 | NTS₁ (inr xs◁zss) = ∣ inr xs◁zss ∣ 224 | ``` 225 | 226 | ### The proof 227 | 228 | The proof is by induction on the proof of `xs ◁ U`. 229 | 230 | ``` 231 | compact xs U U-dc (dir xs∈U) = ∣ xs ∷ [] , NTS₀ , NTS₁ ∣ 232 | where 233 | NTS₀ : xs <ℂ| (λ - → - ↓ (xs ∷ [])) 234 | NTS₀ = dir ∣ inl (⊑[ ℂ-pos ]-refl xs) ∣ 235 | 236 | NTS₁ : [ (λ - → - ↓ (xs ∷ [])) ⊆ U ] 237 | NTS₁ ys ∣ys◁[xs]∣ = ∥∥-rec (is-true-prop (ys ∈ U)) NTS₁′ ∣ys◁[xs]∣ 238 | where 239 | NTS₁′ : [ ys ≤ xs ] ⊎ [ ys ↓ [] ] → [ U ys ] 240 | NTS₁′ (inl ys≤xs) = U-dc xs ys xs∈U ys≤xs 241 | 242 | compact xs U U-dc (branch tt f) = 243 | let 244 | IH₀ : ∥ Σ[ yss₀ ∈ List ℂ ] 245 | ((xs ⌢ true) <ℂ| (λ - → - ↓ yss₀)) × [ ℂ-down yss₀ ⊆ U ] ∥ 246 | IH₀ = compact (xs ⌢ true) U U-dc (f true) 247 | IH₁ : ∥ Σ[ yss ∈ List ℂ ] 248 | ((xs ⌢ false) <ℂ| (λ - → - ↓ yss) × [ ℂ-down yss ⊆ U ]) ∥ 249 | IH₁ = compact (xs ⌢ false) U U-dc (f false) 250 | in 251 | ∥∥-rec (∥∥-prop _) (λ φ → ∥∥-rec (∥∥-prop _) (λ ψ → ∣ NTS φ ψ ∣) IH₁) IH₀ 252 | where 253 | NTS : Σ[ yss₀ ∈ _ ] ((xs ⌢ true) <ℂ| λ - → - ↓ yss₀) × [ ℂ-down yss₀ ⊆ U ] 254 | → Σ[ yss₁ ∈ _ ] ((xs ⌢ false) <ℂ| λ - → - ↓ yss₁) × [ ℂ-down yss₁ ⊆ U ] 255 | → Σ[ yss ∈ _ ] (xs <ℂ| λ - → - ↓ yss) × [ ℂ-down yss ⊆ U ] 256 | NTS (yss , φ , p) (zss , ψ , q) = yss ^ zss , branch tt g , NTS′ 257 | where 258 | g : (c : 𝔹) → (xs ⌢ c) <ℂ| (λ - → ℂ-down (yss ^ zss) -) 259 | g false = U⊆V⇒◁U⊆◁V _ (ℂ-down zss) (ℂ-down (yss ^ zss)) (↓-++-right yss zss) ψ 260 | g true = U⊆V⇒◁U⊆◁V _ (ℂ-down yss) (ℂ-down (yss ^ zss)) (↓-++-left yss zss) φ 261 | 262 | NTS′ : [ (λ - → - ↓ (yss ^ zss)) ⊆ U ] 263 | NTS′ ys ys◁yss₀^yss₁ = 264 | ∥∥-rec (is-true-prop (ys ∈ U)) NTS₂ (◁^-decide _ yss _ ys◁yss₀^yss₁) 265 | where 266 | NTS₂ : [ ys ↓ yss ] ⊎ [ ys ↓ zss ] → [ ys ∈ U ] 267 | NTS₂ (inl ys◁yss₀) = p ys ys◁yss₀ 268 | NTS₂ (inr ys◁yss₁) = q ys ys◁yss₁ 269 | 270 | compact xs U U-dc (squash xs◁U₀ xs◁U₁ i) = 271 | squash (compact xs U U-dc xs◁U₀) (compact xs U U-dc xs◁U₁) i 272 | ``` 273 | 274 | ## Some examples 275 | 276 | An example of an element of the Cantor frame is the set of opens that contain `true`. This 277 | is clearly downwards-closed and a fixed point for the covering relation. 278 | 279 | ```agda 280 | containing-true : ∣ cantor-frame ∣F 281 | containing-true = (W , W-dc) , fixing 282 | where 283 | W : 𝒫 ℂ 284 | W xs = true elem xs 285 | 286 | W-dc : [ isDownwardsClosed ℂ-pos W ] 287 | W-dc xs ys xs∈W ys≤xs@(zs , ys=xs++zs) = 288 | subst (λ - → [ - ∈ W ]) (sym ys=xs++zs) (elem-lemma xs zs true xs∈W) 289 | 290 | lemma : (xs : ℂ) → ((x : 𝔹) → [ true elem (xs ⌢ x) ]) → [ true elem xs ] 291 | lemma [] f with f false 292 | lemma [] f | () 293 | lemma (xs ⌢ x) f with x 294 | lemma (xs ⌢ x) f | false = lemma xs λ { false → f false ; true → tt } 295 | lemma (xs ⌢ x) f | true = tt 296 | 297 | fixing : NucleusFrom.𝕛 cantor (W , W-dc) ≡ (W , W-dc) 298 | fixing = 299 | Σ≡Prop 300 | (isProp[] ∘ isDownwardsClosed ℂ-pos) 301 | (funExt λ xs → ⇔toPath (fixing₀ xs) (fixing₁ xs)) 302 | where 303 | fixing₀ : (xs : ℂ) → [ xs ∈ (NucleusFrom.𝕛 cantor (W , W-dc) .π₀) ] → [ xs ∈ W ] 304 | fixing₀ xs (dir p) = p 305 | fixing₀ xs (branch b f) = lemma xs (λ x → fixing₀ (xs ⌢ x) (f x)) 306 | fixing₀ xs (squash p q i) = isProp[] (W xs) (fixing₀ xs p) (fixing₀ xs q) i 307 | 308 | fixing₁ : (xs : ℂ) → [ xs ∈ W ] → [ xs ∈ (NucleusFrom.𝕛 cantor (W , W-dc) .π₀) ] 309 | fixing₁ xs xs∈W = dir xs∈W 310 | ``` 311 | 312 | -------------------------------------------------------------------------------- /FormalTopology/Nucleus.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Nucleus 3 | --- 4 | 5 | 18 | 19 | ## Definition of nuclei 20 | 21 | We first write down what it means for a function to be “nuclear”. 22 | 23 | ```agda 24 | isNuclear : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (∣ L ∣F → ∣ L ∣F) → Type (ℓ-max ℓ₀ ℓ₁) 25 | isNuclear L j = N₀ × N₁ × N₂ 26 | where 27 | N₀ = (x y : ∣ L ∣F) → j (x ⊓[ L ] y) ≡ (j x) ⊓[ L ] (j y) 28 | N₁ = (x : ∣ L ∣F) → [ x ⊑[ pos L ] (j x) ] 29 | N₂ = (x : ∣ L ∣F) → [ j (j x) ⊑[ pos L ] j x ] 30 | ``` 31 | 32 | A nucleus is a nuclear endofunction on a frame. Notice that it is *not* 33 | required to be a homomorphism. 34 | 35 | ```agda 36 | Nucleus : Frame ℓ₀ ℓ₁ ℓ₂ → Type (ℓ-max ℓ₀ ℓ₁) 37 | Nucleus L = Σ (∣ L ∣F → ∣ L ∣F) (isNuclear L) 38 | ``` 39 | 40 | Some projections for convenience. 41 | 42 | ```agda 43 | 𝓃₀ : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ((j , _) : Nucleus F) 44 | → (x y : ∣ F ∣F) → j (x ⊓[ F ] y) ≡ (j x) ⊓[ F ] (j y) 45 | 𝓃₀ F (_ , n₀ , _) = n₀ 46 | 47 | 𝓃₁ : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ((j , _) : Nucleus F) 48 | → (x : ∣ F ∣F) → [ x ⊑[ pos F ] j x ] 49 | 𝓃₁ F (_ , _ , n₁ , _) = n₁ 50 | 51 | 𝓃₂ : (F : Frame ℓ₀ ℓ₁ ℓ₂) → ((j , _) : Nucleus F) 52 | → (x : ∣ F ∣F) → [ j (j x) ⊑[ pos F ] j x ] 53 | 𝓃₂ F (_ , _ , _ , n₂) = n₂ 54 | ``` 55 | 56 | The nuclearity property is an h-proposition. 57 | 58 | ```agda 59 | isNuclear-prop : (L : Frame ℓ₀ ℓ₁ ℓ₂) (j : ∣ L ∣F → ∣ L ∣F) 60 | → isProp (isNuclear L j) 61 | isNuclear-prop L j = isProp×2 N₀-prop N₁-prop N₂-prop 62 | where 63 | N₀-prop : isProp _ 64 | N₀-prop = isPropΠ2 (λ _ _ → carrier-is-set (pos L) _ _) 65 | 66 | N₁-prop : isProp _ 67 | N₁-prop = isPropΠ λ x → isProp[] (x ⊑[ pos L ] j x) 68 | 69 | N₂-prop : isProp _ 70 | N₂-prop = isPropΠ λ x → isProp[] (j (j x) ⊑[ pos L ] j x) 71 | 72 | Nucleus-set : (F : Frame ℓ₀ ℓ₁ ℓ₂) → isSet (Nucleus F) 73 | Nucleus-set F = isSetΣ 74 | (isSetΠ (λ _ → carrier-is-set (pos F))) λ j → 75 | isProp→isSet (isNuclear-prop F j) 76 | ``` 77 | 78 | ## Record-based definition of nuclei 79 | 80 | ```agda 81 | record Nucleusᵣ (F : Frame 𝓤 𝓥 𝓦) : (𝓤 ∨ 𝓥) ̇ where 82 | constructor nucl 83 | field 84 | j : ∣ F ∣F → ∣ F ∣F 85 | 86 | meet-preservation : (x y : ∣ F ∣F) → j (x ⊓[ F ] y) ≡ j x ⊓[ F ] j y 87 | inflationarity : (x : ∣ F ∣F) → [ x ⊑[ pos F ] j x ] 88 | idempotency : (x : ∣ F ∣F) → [ j (j x) ⊑[ pos F ] j x ] 89 | 90 | Nucleus≃Nucleusᵣ : (F : Frame 𝓤 𝓥 𝓦) → Nucleus F ≃ Nucleusᵣ F 91 | Nucleus≃Nucleusᵣ F = isoToEquiv (iso to from sec ret) 92 | where 93 | to : Nucleus F → Nucleusᵣ F 94 | to (j , mp , inf , i) = nucl j mp inf i 95 | 96 | from : Nucleusᵣ F → Nucleus F 97 | from (nucl j mp inf i) = j , mp , inf , i 98 | 99 | sec : section to from 100 | sec (nucl _ _ _ _) = refl 101 | 102 | ret : retract to from 103 | ret (_ , _ , _ , _) = refl 104 | ``` 105 | 106 | ## Properties of nuclei 107 | 108 | Every nucleus respects the top element. 109 | 110 | ```agda 111 | nuclei-resp-⊤ : (L : Frame ℓ₀ ℓ₁ ℓ₂) ((j , _) : Nucleus L) → j ⊤[ L ] ≡ ⊤[ L ] 112 | nuclei-resp-⊤ L (j , N₀ , N₁ , N₂) = ⊑[ pos L ]-antisym _ _ (⊤[ L ]-top _) (N₁ _) 113 | ``` 114 | 115 | Every nucleus is idempotent. 116 | 117 | ```agda 118 | idem : (L : Frame ℓ₀ ℓ₁ ℓ₂) → ((j , _) : Nucleus L) → (x : ∣ L ∣F) → j (j x) ≡ j x 119 | idem L (j , N₀ , N₁ , N₂) x = ⊑[ pos L ]-antisym _ _ (N₂ x) (N₁ (j x)) 120 | ``` 121 | 122 | Every nucleus is a monotonic map. 123 | 124 | ```agda 125 | mono : (L : Frame ℓ₀ ℓ₁ ℓ₂) ((j , _) : Nucleus L) 126 | → (x y : ∣ L ∣F) → [ x ⊑[ pos L ] y ] → [ (j x) ⊑[ pos L ] (j y) ] 127 | mono L (j , N₀ , N₁ , N₂) x y x⊑y = 128 | j x ⊑⟨ ≡⇒⊑ (pos L) (cong j (x⊑y⇒x=x∧y L x⊑y)) ⟩ 129 | j (x ⊓[ L ] y) ⊑⟨ ≡⇒⊑ (pos L) (N₀ x y) ⟩ 130 | j x ⊓[ L ] j y ⊑⟨ ⊓[ L ]-lower₁ (j x) (j y) ⟩ 131 | j y ■ 132 | where 133 | open PosetReasoning (pos L) 134 | ``` 135 | 136 | ## Some more lemmas about nuclei 137 | 138 | ```agda 139 | module Nucleus-∨-Lemmata (L : Frame ℓ₀ ℓ₁ ℓ₂) (𝒿 : Nucleus L) where 140 | 141 | open PosetReasoning (pos L) 142 | 143 | j = π₀ 𝒿 144 | n₁ = π₀ (π₁ (π₁ 𝒿)) 145 | n₂ = π₁ (π₁ (π₁ 𝒿)) 146 | 147 | nucleus-∨-lemma₀ : (x y : ∣ L ∣F) 148 | → [ j (x ∨[ L ] y) ⊑[ pos L ] j (x ∨[ L ] j y) ] 149 | nucleus-∨-lemma₀ x y = 150 | mono L 𝒿 _ _ (⊔[ L ]-least x y (x ∨[ L ] (j y)) (⊔[ L ]-upper₀ _ _) α) 151 | where 152 | α : [ y ⊑[ pos L ] x ∨[ L ] (j y) ] 153 | α = y ⊑⟨ n₁ y ⟩ j y ⊑⟨ ⊔[ L ]-upper₁ _ _ ⟩ x ∨[ L ] j y ■ 154 | 155 | nucleus-∨-lemma₁ : (x y : ∣ L ∣F) 156 | → [ j (x ∨[ L ] j y) ⊑[ pos L ] j (j x ∨[ L ] j y) ] 157 | nucleus-∨-lemma₁ x y = 158 | mono L 𝒿 _ _ (⊔[ L ]-least _ _ _ (x ⊑⟨ n₁ x ⟩ j x ⊑⟨ ⊔[ L ]-upper₀ _ _ ⟩ _ ■) (⊔[ L ]-upper₁ _ _)) 159 | 160 | nucleus-∨-lemma₂ : (x y : ∣ L ∣F) 161 | → [ j (j x ∨[ L ] j y) ⊑[ pos L ] j (j (x ∨[ L ] y)) ] 162 | nucleus-∨-lemma₂ x y = 163 | mono L 𝒿 _ _ (⊔[ L ]-least _ _ _ (mono L 𝒿 _ _ (⊔[ L ]-upper₀ x y)) (mono L 𝒿 _ _ (⊔[ L ]-upper₁ x y))) 164 | 165 | nucleus-∨-thm₀ : (x y : ∣ L ∣F) 166 | → j (x ∨[ L ] y) ≡ j (j x ∨[ L ] j y) 167 | nucleus-∨-thm₀ x y = 168 | ⊑[ pos L ]-antisym _ _ nts₀ nts₁ 169 | where 170 | nts₀ : _ 171 | nts₀ = j (x ∨[ L ] y) ⊑⟨ nucleus-∨-lemma₀ x y ⟩ 172 | j (x ∨[ L ] j y) ⊑⟨ nucleus-∨-lemma₁ x y ⟩ 173 | j (j x ∨[ L ] j y) ■ 174 | 175 | nts₁ : _ 176 | nts₁ = j (j x ∨[ L ] j y) ⊑⟨ nucleus-∨-lemma₂ x y ⟩ 177 | j (j (x ∨[ L ] y)) ⊑⟨ n₂ (x ∨[ L ] y) ⟩ 178 | j (x ∨[ L ] y) ■ 179 | ``` 180 | 181 | The set of fixed points for nucleus `j` is equivalent hence equal to its image. 182 | This is essentially due to the fact that j (j ()) 183 | 184 | ```agda 185 | nuclear-image : (L : Frame ℓ₀ ℓ₁ ℓ₂) 186 | → let ∣L∣ = ∣ L ∣F in (j : ∣L∣ → ∣L∣) 187 | → isNuclear L j 188 | → (Σ[ b ∈ ∣L∣ ] ∥ Σ[ a ∈ ∣L∣ ] (b ≡ j a) ∥) ≡ (Σ[ a ∈ ∣L∣ ] (j a ≡ a)) 189 | nuclear-image L j N@(n₀ , n₁ , n₂) = isoToPath (iso f g sec-f-g ret-f-g) 190 | where 191 | A-set = carrier-is-set (pos L) 192 | 193 | f : (Σ[ b ∈ ∣ L ∣F ] ∥ Σ[ a ∈ ∣ L ∣F ] (b ≡ j a) ∥) → Σ[ a ∈ ∣ L ∣F ] (j a ≡ a) 194 | f (b , img) = b , ∥∥-rec (A-set (j b) b) ind img 195 | where 196 | ind : Σ[ a ∈ ∣ L ∣F ](b ≡ j a) → j b ≡ b 197 | ind (a , img) = 198 | j b ≡⟨ cong j img ⟩ 199 | j (j a) ≡⟨ idem L (j , N) a ⟩ 200 | j a ≡⟨ sym img ⟩ 201 | b ∎ 202 | 203 | g : (Σ[ a ∈ ∣ L ∣F ] (j a ≡ a)) → (Σ[ b ∈ ∣ L ∣F ] ∥ Σ[ a ∈ ∣ L ∣F ] (b ≡ j a) ∥) 204 | g (a , a-fix) = a , ∣ a , (sym a-fix) ∣ 205 | 206 | sec-f-g : section f g 207 | sec-f-g (x , jx=x) = Σ≡Prop (λ y → A-set (j y) y) refl 208 | 209 | ret-f-g : retract f g 210 | ret-f-g (x , p) = Σ≡Prop (λ y → ∥∥-prop (Σ[ a ∈ ∣ L ∣F ] y ≡ j a)) refl 211 | ``` 212 | 213 | ## The frame of nuclei fixed points 214 | 215 | The set of fixed points for a nucleus `j` forms a poset. 216 | 217 | ```agda 218 | 𝔣𝔦𝔵-pos : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (N : Nucleus L) → Poset ℓ₀ ℓ₁ 219 | 𝔣𝔦𝔵-pos {ℓ₀ = ℓ₀} {ℓ₁} L (j , N₀ , N₁ , N₂) = 220 | 𝔽 , _≤_ , 𝔽-set , ≤-refl , ≤-trans , ≤-antisym 221 | where 222 | P = pos L 223 | A-set = carrier-is-set (pos L) 224 | 225 | 𝔽 : Type ℓ₀ 226 | 𝔽 = Σ[ a ∈ ∣ L ∣F ] j a ≡ a 227 | 228 | 𝔽-set : isSet 𝔽 229 | 𝔽-set = isSetΣ A-set (λ a → isProp→isSet (A-set (j a) a)) 230 | 231 | _≤_ : 𝔽 → 𝔽 → hProp ℓ₁ 232 | (a , _) ≤ (b , _) = [ a ⊑[ P ] b ] , is-true-prop (a ⊑[ P ] b) 233 | 234 | ≤-refl : [ isReflexive _≤_ ] 235 | ≤-refl (x , _) = ⊑[ P ]-refl x 236 | 237 | ≤-trans : [ isTransitive _≤_ ] 238 | ≤-trans (x , _) (y , _) (z , _) x≤y y≤x = ⊑[ P ]-trans x y z x≤y y≤x 239 | 240 | ≤-antisym : [ isAntisym 𝔽-set _≤_ ] 241 | ≤-antisym (x , _) (y , _) x≤y y≤x = 242 | Σ≡Prop (λ z → A-set (j z) z) (⊑[ P ]-antisym x y x≤y y≤x) 243 | ``` 244 | 245 | The set of fixed points of a nucleus `j` forms a frame. The join of this frame 246 | is define as ⊔ᵢ Uᵢ := j (⊔′ᵢ Uᵢ) where ⊔′ denotes the join of L. 247 | 248 | ```agda 249 | 𝔣𝔦𝔵 : (L : Frame ℓ₀ ℓ₁ ℓ₂) → (N : Nucleus L) → Frame ℓ₀ ℓ₁ ℓ₂ 250 | 𝔣𝔦𝔵 {ℓ₁ = ℓ₁} {ℓ₂ = ℓ₂} L N@(j , N₀ , N₁ , N₂) = 251 | ∣ 𝔣𝔦𝔵-pos L N ∣ₚ 252 | , (strₚ (𝔣𝔦𝔵-pos L N) , (⊤[ L ] , nuclei-resp-⊤ L N) , _∧_ , ⋁_) 253 | , ⊤-top 254 | , ( (λ x y → ⊓-lower₀ x y , ⊓-lower₁ x y) 255 | , λ { x y z (z⊑x , x⊑y) → ⊓-greatest x y z z⊑x x⊑y }) 256 | , ((⊔-upper , ⊔-least) , distr) 257 | where 258 | 𝒜 = π₀ (𝔣𝔦𝔵-pos L N) 259 | 260 | _⊑_ : ∣ pos L ∣ₚ → ∣ pos L ∣ₚ → hProp ℓ₁ 261 | _⊑_ = λ x y → x ⊑[ pos L ] y 262 | 263 | _⊑N_ : 𝒜 → 𝒜 → hProp ℓ₁ 264 | _⊑N_ = λ x y → x ⊑[ 𝔣𝔦𝔵-pos L N ] y 265 | 266 | ⋁L_ : Fam ℓ₂ ∣ L ∣F → ∣ L ∣F 267 | ⋁L x = ⋁[ L ] x 268 | 269 | ⊑N-antisym = ⊑[ 𝔣𝔦𝔵-pos L N ]-antisym 270 | A-set = carrier-is-set (𝔣𝔦𝔵-pos L N) 271 | 272 | open PosetReasoning (pos L) 273 | 274 | _∧_ : 𝒜 → 𝒜 → 𝒜 275 | _∧_ (x , x-f) (y , y-f) = 276 | x ⊓[ L ] y , NTS 277 | where 278 | NTS : j (x ⊓[ L ] y) ≡ x ⊓[ L ] y 279 | NTS = j (x ⊓[ L ] y) ≡⟨ N₀ x y ⟩ 280 | j x ⊓[ L ] j y ≡⟨ cong (λ - → - ⊓[ L ] _) x-f ⟩ 281 | x ⊓[ L ] j y ≡⟨ cong (λ - → _ ⊓[ L ] -) y-f ⟩ 282 | x ⊓[ L ] y ∎ 283 | 284 | ⋁_ : Fam ℓ₂ 𝒜 → 𝒜 285 | ⋁ (I , F) = j (⋁[ L ] 𝒢) , j⊔L-fixed 286 | where 287 | 𝒢 = I , π₀ ∘ F 288 | j⊔L-fixed : j (j (⋁[ L ] 𝒢)) ≡ j (⋁[ L ] 𝒢) 289 | j⊔L-fixed = ⊑[ pos L ]-antisym _ _ (N₂ (⋁[ L ] 𝒢)) (N₁ (j (⋁[ L ] 𝒢))) 290 | 291 | open JoinSyntax 𝒜 ⋁_ 292 | 293 | ⊤-top : (o : 𝒜) → [ o ⊑N (⊤[ L ] , nuclei-resp-⊤ L N) ] 294 | ⊤-top = ⊤[ L ]-top ∘ π₀ 295 | 296 | ⊓-lower₀ : (o p : 𝒜) → [ (o ∧ p) ⊑N o ] 297 | ⊓-lower₀ (o , _) (p , _) = ⊓[ L ]-lower₀ o p 298 | 299 | ⊓-lower₁ : (o p : 𝒜) → [ (o ∧ p) ⊑N p ] 300 | ⊓-lower₁ (o , _) (p , _) = ⊓[ L ]-lower₁ o p 301 | 302 | ⊓-greatest : (o p q : 𝒜) → [ q ⊑N o ] → [ q ⊑N p ] → [ q ⊑N (o ∧ p) ] 303 | ⊓-greatest (o , _) (p , _) (q , _) q⊑o q⊑p = ⊓[ L ]-greatest o p q q⊑o q⊑p 304 | 305 | ⊔-least : (U : Fam ℓ₂ 𝒜) (u : 𝒜) → [ ∀[ x ε U ] (x ⊑N u) ] → [ (⋁ U) ⊑N u ] 306 | ⊔-least U (u , fix) U⊑u = subst (λ - → [ j (⋁[ L ] 𝒢) ⊑ - ]) fix NTS₀ 307 | where 308 | 𝒢 : Fam ℓ₂ ∣ pos L ∣ₚ 309 | 𝒢 = π₀ ⟨$⟩ U 310 | 311 | 𝒢⊑u : [ ∀[ o ε 𝒢 ] (o ⊑ u) ] 312 | 𝒢⊑u o (i , eq′) = o ⊑⟨ ≡⇒⊑ (pos L) (sym eq′) ⟩ 313 | 𝒢 $ i ⊑⟨ U⊑u (𝒢 $ i , π₁ (U $ i)) (i , refl) ⟩ 314 | u ■ 315 | 316 | NTS₀ : [ j (⋁[ L ] 𝒢) ⊑ j u ] 317 | NTS₀ = mono L N (⋁[ L ] 𝒢) u (⋁[ L ]-least 𝒢 u 𝒢⊑u) 318 | 319 | ⊔-upper : (U : Fam ℓ₂ 𝒜) (x : 𝒜) → x ε U → [ x ⊑N (⋁ U) ] 320 | ⊔-upper U (x , _) o∈U@(i , eq) = 321 | x ⊑⟨ NTS ⟩ 322 | ⋁[ L ] (π₀ ⟨$⟩ U) ⊑⟨ N₁ (⋁[ L ] (π₀ ⟨$⟩ U)) ⟩ 323 | j (⋁[ L ] (π₀ ⟨$⟩ U)) ■ 324 | where 325 | NTS : [ x ⊑ (⋁[ L ] (π₀ ⟨$⟩ U)) ] 326 | NTS = ⋁[ L ]-upper (π₀ ⟨$⟩ U) x (i , λ j → π₀ (eq j)) 327 | 328 | distr : (x : Σ[ x ∈ ∣ L ∣F ] j x ≡ x) (U@(I , _) : Fam ℓ₂ 𝒜) 329 | → x ∧ (⋁ U) ≡ ⋁⟨ i ⟩ (x ∧ (U $ i)) 330 | distr 𝓍@(x , jx=x) U@(I , F) = Σ≡Prop (λ x → carrier-is-set (pos L) (j x) x) NTS 331 | where 332 | -- U is a family of inhabitants of ∣ L ∣F paired with proofs that they are fixed 333 | -- points for j. U₀ is the family obtained by discarding the proofs 334 | U₀ : Fam ℓ₂ ∣ L ∣F 335 | U₀ = ⁅ π₀ x ∣ x ε U ⁆ 336 | 337 | x=jx = sym jx=x 338 | 339 | NTS : π₀ (𝓍 ∧ (⋁ U)) ≡ π₀ (⋁⟨ i ⟩ (𝓍 ∧ (U $ i))) 340 | NTS = 341 | π₀ (𝓍 ∧ (⋁ U)) ≡⟨ refl ⟩ 342 | x ⊓[ L ] j (⋁L U₀) ≡⟨ cong (λ - → - ⊓[ L ] j (⋁L U₀)) x=jx ⟩ 343 | j x ⊓[ L ] j (⋁L U₀) ≡⟨ sym (N₀ x (⋁[ L ] U₀)) ⟩ 344 | j (x ⊓[ L ] (⋁L U₀)) ≡⟨ cong j (dist L x U₀) ⟩ 345 | j (⋁L ⁅ x ⊓[ L ] yᵢ ∣ yᵢ ε U₀ ⁆) ≡⟨ refl ⟩ 346 | π₀ (⋁⟨ i ⟩ (𝓍 ∧ (U $ i))) ∎ 347 | ``` 348 | 349 | ## Every nucleus is obviously a prenucleus 350 | 351 | ```agda 352 | nucleus⇒prenucleus : (F : Frame ℓ₀ ℓ₁ ℓ₂) → Nucleus F → Prenucleus F 353 | nucleus⇒prenucleus F (j , n₀ , n₁ , _)= j , n₀ , n₁ 354 | ``` 355 | 356 | ## The notion of sublocale 357 | 358 | ```agda 359 | isASublocaleOf : (S : Frame ℓ₀ ℓ₁ ℓ₂) (F : Frame ℓ₀ ℓ₁ ℓ₂) → Type _ 360 | isASublocaleOf S F = Σ[ j ∈ Nucleus F ] S ≡ 𝔣𝔦𝔵 F j 361 | ``` 362 | 363 | ## The identity nucleus 364 | 365 | ```agda 366 | idn : (F : Frame 𝓤 𝓥 𝓦) → Nucleus F 367 | π₀ (idn F) = id ∣ F ∣F 368 | π₀ (π₁ (idn F)) _ _ = refl 369 | π₀ (π₁ (π₁ (idn F))) = ⊑[ pos F ]-refl 370 | π₁ (π₁ (π₁ (idn F))) = ⊑[ pos F ]-refl 371 | ``` 372 | -------------------------------------------------------------------------------- /FormalTopology/UniversalProperty.lagda.md: -------------------------------------------------------------------------------- 1 | ``` 2 | {-# OPTIONS --cubical --safe #-} 3 | 4 | module UniversalProperty where 5 | 6 | open import Basis 7 | open import Frame 8 | open import Poset 9 | open import FormalTopology renaming (pos to pos′) 10 | open import CoverFormsNucleus 11 | 12 | compr : {X : Type ℓ₀} {Y : Type ℓ₁} → (g : X → Y) → 𝒫 X → Fam ℓ₀ Y 13 | compr g U = (index ⟪ U ⟫) , g ∘ (_$_ ⟪ U ⟫) 14 | 15 | syntax compr (λ x → e) ℱ = ⁅ e ∣ x ∈ ℱ ⁆ 16 | 17 | module _ (F : FormalTopology ℓ₀ ℓ₀) where 18 | 19 | P = pos′ F 20 | F↓ = DCFrame P 21 | P↓ = pos F↓ 22 | _⊑_ = λ (x y : stage F) → x ⊑[ P ] y 23 | 24 | open NucleusFrom F 25 | ``` 26 | 27 | ## Representation 28 | 29 | ``` 30 | represents : (R : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) → (P ─m→ pos R) → Type ℓ₀ 31 | represents R (f , _) = 32 | (a : ∣ P ∣ₚ) (b : exp F a) → 33 | [ f a ⊑[ pos R ] ⋁[ R ] ⁅ f (next F c) ∣ c ∶ outcome F b ⁆ ] 34 | ``` 35 | 36 | By the way, note that the converse is always true. 37 | 38 | ``` 39 | represents⁻¹ : (R : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) → (m : P ─m→ pos R) 40 | → Type ℓ₀ 41 | represents⁻¹ R (f , _) = 42 | (a : ∣ P ∣ₚ) (b : exp F a) → 43 | [ (⋁[ R ] ⁅ f (next F c) ∣ c ∶ outcome F b ⁆) ⊑[ pos R ] (f a) ] 44 | 45 | conv : (R : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) (f : P ─m→ pos R) → represents⁻¹ R f 46 | conv R (f , f-mono) a b = 47 | ⋁[ R ]-least (⁅ f (next F c) ∣ c ∶ outcome F b ⁆) (f a) NTS 48 | where 49 | NTS : [ ∀[ a′ ε ⁅ f (next F c) ∣ c ∶ outcome F b ⁆ ] (a′ ⊑[ pos R ] f a) ] 50 | NTS a′ (i , eq) = subst (λ - → [ rel (pos R) - (f a) ]) eq NTS′ 51 | where 52 | NTS′ : [ f (next F i) ⊑[ pos R ] f a ] 53 | NTS′ = f-mono (next F i) a (mono F a b i) 54 | ``` 55 | 56 | ## Flatness 57 | 58 | ``` 59 | _↓_↓ : ∣ P ∣ₚ → ∣ P ∣ₚ → 𝒫 ∣ P ∣ₚ 60 | _↓_↓ a b = λ - → - ⊑[ P ] a ⊓ - ⊑[ P ] b 61 | 62 | isFlat : (F : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) → (m : P ─m→ pos F) → Type (ℓ-suc ℓ₀) 63 | isFlat F (f , _) = (⊤[ F ] ≡ ⋁[ F ] ⁅ f a ∣ a ∶ ∣ P ∣ₚ ⁆) 64 | × ((a b : ∣ P ∣ₚ) → f a ⊓[ F ] f b ≡ ⋁[ F ] (f ⟨$⟩ ⟪ a ↓ b ↓ ⟫)) 65 | ``` 66 | 67 | ## The universal property 68 | 69 | Statement. 70 | 71 | ``` 72 | universal-prop : Type (ℓ-suc (ℓ-suc ℓ₀)) 73 | universal-prop = 74 | (R : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) (f : P ─m→ pos R) → isFlat R f → represents R f → 75 | isContr (Σ[ g ∈ (L ─f→ R) ] (_∘m_ {P = P} {Q = pos L} {R = pos R} (π₀ g) ηm) ≡ f) 76 | ``` 77 | 78 | Before the proof we will need some lemmas. 79 | 80 | ``` 81 | cover+ : {x y : ∣ P ∣ₚ} ((U , _) : ∣ F↓ ∣F) → [ x ∈ ⦅ η y ⦆ ] → [ y ∈ U ] → x ◁ U 82 | cover+ (_ , U-dc) x∈ηy y∈U = ◁-lem₄ _ _ (λ z z⊑y → dir (U-dc _ z y∈U z⊑y)) _ x∈ηy 83 | ``` 84 | 85 | ``` 86 | main-lemma : (𝔘 : ∣ L ∣F) → 𝔘 ≡ ⋁[ L ] ⁅ η u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ 87 | main-lemma 𝔘@((U , U-dc) , U-fix) = ⊑[ pos L ]-antisym _ _ down up 88 | where 89 | down : [ 𝔘 ⊑[ pos L ] (⋁[ L ] ⁅ η x ∣ x ∈ U ⁆) ] 90 | down x xεU = dir ∣ (x , xεU) , dir (⊑[ P ]-refl x) ∣ 91 | 92 | up : [ (⋁[ L ] ⁅ η x ∣ x ∈ U ⁆) ⊑[ pos L ] 𝔘 ] 93 | up x (dir xε⋁) = ∥∥-rec (is-true-prop (U x)) NTS xε⋁ 94 | where 95 | NTS : Σ[ y ∈ _ ] [ x ∈ ⦅ η (π₀ y) ⦆ ] → [ x ∈ U ] 96 | NTS ((y , yεU) , x◁y↓) = 97 | subst (λ V → [ π₀ V x ]) U-fix (cover+ (U , U-dc) x◁y↓ yεU) 98 | up x (branch b f) = subst (λ V → [ π₀ V x ]) U-fix (branch b (dir ∘ IH)) 99 | where 100 | IH : (c : outcome F b) → [ next F c ∈ U ] 101 | IH c = up (next F c) (f c) 102 | up x (squash x◁⋁₀ x◁⋁₁ i) = is-true-prop (U x) (up x x◁⋁₀) (up x x◁⋁₁) i 103 | ``` 104 | 105 | Proof. 106 | 107 | ``` 108 | module MainProof (R : Frame (ℓ-suc ℓ₀) ℓ₀ ℓ₀) 109 | (fm : P ─m→ pos R) 110 | (f-flat : isFlat R fm) 111 | (rep : represents R fm) where 112 | f = _$ₘ_ fm 113 | f-mono = π₁ fm 114 | 115 | _⊑R_ : ∣ R ∣F → ∣ R ∣F → hProp ℓ₀ 116 | x ⊑R y = x ⊑[ pos R ] y 117 | 118 | infix 9 _⊑R_ 119 | ``` 120 | 121 | ``` 122 | g : ∣ L ∣F → ∣ R ∣F 123 | g 𝔘 = ⋁[ R ] ⁅ f u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ 124 | ``` 125 | 126 | ``` 127 | g-mono : isMonotonic (pos L) (pos R) g 128 | g-mono ((U , _) , _) ((V , _) , _) U⊆V = 129 | ⋁[ R ]-least _ _ (λ o oεfU → ⋁[ R ]-upper _ _ (NTS o oεfU )) 130 | where 131 | NTS : (x : ∣ R ∣F) → x ε (∃ U , f ∘ π₀) → x ε (∃ V , f ∘ π₀) 132 | NTS x ((x′ , x′εfU) , fUεi=x) = (x′ , U⊆V x′ x′εfU) , fUεi=x 133 | 134 | gm : pos L ─m→ pos R 135 | gm = g , g-mono 136 | ``` 137 | 138 | ### `g` respects the top element 139 | 140 | ``` 141 | g-resp-𝟏 : g ⊤[ L ] ≡ ⊤[ R ] 142 | g-resp-𝟏 = g ⊤[ L ] ≡⟨ refl ⟩ 143 | ⋁[ R ] (f ⟨$⟩ (∃ ⦅ ⊤[ L ] ⦆ , π₀)) ≡⟨ family-iff R NTS ⟩ 144 | ⋁[ R ] (∣ P ∣ₚ , f) ≡⟨ sym (π₀ f-flat) ⟩ 145 | ⊤[ R ] ∎ 146 | where 147 | NTS : (x : ∣ R ∣F) 148 | → (x ε (f ⟨$⟩ (∃ ⦅ ⊤[ L ] ⦆ , π₀)) → x ε (∣ P ∣ₚ , f)) 149 | × (x ε (∣ P ∣ₚ , f) → x ε (f ⟨$⟩ (∃ ⦅ ⊤[ L ] ⦆ , π₀))) 150 | NTS x = (λ { ((y , _) , eq) → y , eq }) , (λ { (y , eq) → (y , tt) , eq }) 151 | ``` 152 | 153 | ### `g` respects the binary meets 154 | 155 | ``` 156 | g-resp-⊓ : (𝔘 𝔙 : ∣ L ∣F) → g (𝔘 ⊓[ L ] 𝔙) ≡ g 𝔘 ⊓[ R ] g 𝔙 157 | g-resp-⊓ 𝔘 𝔙 = 158 | g (𝔘 ⊓[ L ] 𝔙) 159 | ≡⟨ refl ⟩ 160 | ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ 𝔘 ⊓[ L ] 𝔙 ⦆ ⁆ 161 | ≡⟨ I ⟩ 162 | ⋁[ R ] ⁅ ⋁[ R ] (f ⟨$⟩ ⟪ u ↓ v ↓ ⟫) ∣ ((u , _) , (v , _)) ∶ (∃ ⦅ 𝔘 ⦆ × ∃ ⦅ 𝔙 ⦆) ⁆ 163 | ≡⟨ cong (λ - → (⋁[ R ] ((∃ ⦅ 𝔘 ⦆ × ∃ ⦅ 𝔙 ⦆) , -))) II ⟩ 164 | ⋁[ R ] (((∃ ⦅ 𝔘 ⦆) × (∃ ⦅ 𝔙 ⦆)) , λ { ((u , _) , (v , _)) → f u ⊓[ R ] f v }) 165 | ≡⟨ sym (sym-distr R (⁅ f u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆) (⁅ f v ∣ v ∈ ⦅ 𝔙 ⦆ ⁆)) ⟩ 166 | (⋁[ R ] ⁅ f u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆) ⊓[ R ] (⋁[ R ] ⁅ f v ∣ v ∈ ⦅ 𝔙 ⦆ ⁆) 167 | ≡⟨ refl ⟩ 168 | g 𝔘 ⊓[ R ] g 𝔙 169 | ∎ 170 | where 171 | II : (λ { ((u , _) , (v , _)) → ⋁[ R ] (f ⟨$⟩ ⟪ u ↓ v ↓ ⟫) }) 172 | ≡ (λ { ((u , _) , (v , _)) → (f u) ⊓[ R ] (f v) }) 173 | II = sym (funExt λ { ((u , _) , (v , _)) → π₁ f-flat u v }) 174 | I : _ 175 | I = ⊑[ pos R ]-antisym _ _ down up 176 | where 177 | LHS = ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ 𝔘 ⊓[ L ] 𝔙 ⦆ ⁆ 178 | RHS = ⋁[ R ] (∃ ⦅ 𝔘 ⦆ × ∃ ⦅ 𝔙 ⦆ , λ { ((u , _) , v , _) → ⋁[ R ] (f ⟨$⟩ ⟪ u ↓ v ↓ ⟫) }) 179 | 180 | down : [ LHS ⊑[ pos R ] RHS ] 181 | down = ⋁[ R ]-least _ _ isUB 182 | where 183 | isUB : _ 184 | isUB o ((a , (a∈U , a∈V)) , eq) = 185 | ⋁[ R ]-upper _ _ NTS 186 | where 187 | u : ∃ ⦅ 𝔘 ⦆ 188 | u = a , a∈U 189 | 190 | v : ∃ ⦅ 𝔙 ⦆ 191 | v = a , a∈V 192 | 193 | NTS : o ε (∃ ⦅ 𝔘 ⦆ × ∃ ⦅ 𝔙 ⦆ , (λ { ((u , _) , v , _) → ⋁[ R ] (f ⟨$⟩ ⟪ u ↓ v ↓ ⟫) })) 194 | NTS = (u , v) , subst (λ o′ → _ ≡ o′) eq NTS′ 195 | where 196 | down′ : [ ⋁[ R ] (f ⟨$⟩ ⟪ a ↓ a ↓ ⟫) ⊑[ pos R ] f a ] 197 | down′ = 198 | ⋁[ R ]-least _ _ λ { z ((_ , (k , _)) , eq′) → 199 | subst (λ - → [ - ⊑[ pos R ] _ ]) eq′ (f-mono _ _ k) } 200 | up′ : [ f a ⊑[ pos R ] ⋁[ R ] (f ⟨$⟩ ⟪ a ↓ a ↓ ⟫) ] 201 | up′ = ⋁[ R ]-upper _ _ ((a , (⊑[ P ]-refl a , ⊑[ P ]-refl a)) , refl) 202 | 203 | NTS′ : ⋁[ R ] (f ⟨$⟩ ⟪ a ↓ a ↓ ⟫) ≡ f a 204 | NTS′ = ⊑[ pos R ]-antisym _ _ down′ up′ 205 | 206 | up : [ LHS ⊒[ pos R ] RHS ] 207 | up = ⋁[ R ]-least _ _ isUB 208 | where 209 | isUB : _ 210 | isUB o (i@((x , xε𝔙) , (y , yε𝔘)) , eq) = 211 | subst (λ o′ → [ o′ ⊑[ pos R ] _ ]) eq (⋁[ R ]-least _ _ NTS) 212 | where 213 | NTS : _ 214 | NTS w (j@(z , (z⊑x , z⊑y)) , eq′) = ⋁[ R ]-upper _ _ ((z , φ) , eq′) 215 | where 216 | φ : [ z ∈ (⦅ 𝔘 ⦆ ∩ ⦅ 𝔙 ⦆) ] 217 | φ = (π₁ (π₀ 𝔘) x z xε𝔙 z⊑x) , (π₁ (π₀ 𝔙) y z yε𝔘 z⊑y) 218 | ``` 219 | 220 | ### `g` respects the joins 221 | 222 | ``` 223 | open PosetReasoning (pos R) 224 | 225 | resp-⋁-lem : (U@(I , _) : Fam ℓ₀ ∣ L ∣F) 226 | → ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ ⋁[ L ] U ⦆ ⁆ 227 | ≡ (⋁[ R ] ⁅ f a ∣ (_ , a , _) ∶ (Σ[ i ∈ I ] ∃ ⦅ U $ i ⦆) ⁆) 228 | resp-⋁-lem U@(I , _) = ⊑[ pos R ]-antisym _ _ down up 229 | where 230 | LHS = ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ ⋁[ L ] U ⦆ ⁆ 231 | RHS = ⋁[ R ] ⁅ f a ∣ (_ , a , _) ∶ (Σ[ i ∈ I ] ∃ ⦅ U $ i ⦆) ⁆ 232 | 233 | ϑ : (x : ∣ P ∣ₚ) → [ x ∈ ⦅ ⋁[ L ] U ⦆ ] → [ f x ⊑R RHS ] 234 | ϑ x (squash p q i) = is-true-prop (f x ⊑R _) (ϑ x p) (ϑ x q) i 235 | ϑ x (dir xε⋁U) = ∥∥-rec (is-true-prop (f x ⊑R _)) NTS xε⋁U 236 | where 237 | NTS : _ 238 | NTS (j , cov) = ⋁[ R ]-upper _ _ ((j , x , cov) , refl) 239 | ϑ x (branch b h) = 240 | f x ⊑⟨ rep x b ⟩ 241 | ⋁[ R ] (_ , f ∘ next F) ⊑⟨ ⋁[ R ]-least _ _ NTS ⟩ 242 | RHS ■ 243 | where 244 | NTS : (r : ∣ R ∣F) → r ε (_ , f ∘ next F) → [ r ⊑R _ ] 245 | NTS r (c , p) = subst (λ - → [ - ⊑R _ ]) p (ϑ (next F c) (h c)) 246 | 247 | down : [ LHS ⊑R RHS ] 248 | down = 249 | ⋁[ R ]-least _ _ λ r ((x , cov) , p) → subst (λ - → [ - ⊑R _ ]) p (ϑ x cov) 250 | 251 | up : [ LHS ⊒[ pos R ] RHS ] 252 | up = ⋁[ R ]-least _ _ NTS 253 | where 254 | NTS : _ 255 | NTS r ((i , (x , xεU)) , p) = ⋁[ R ]-upper _ _ ((x , dir ∣ i , xεU ∣) , p) 256 | ``` 257 | 258 | ``` 259 | g-resp-⊔ : (U : Fam ℓ₀ ∣ L ∣F) → g (⋁[ L ] U) ≡ ⋁[ R ] (g ⟨$⟩ U) 260 | g-resp-⊔ U@(I , h) = 261 | ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ ⋁[ L ] U ⦆ ⁆ 262 | ≡⟨ resp-⋁-lem U ⟩ 263 | ⋁[ R ] ⁅ f a ∣ (_ , (a , _)) ∶ Σ[ i ∈ I ] Σ[ x ∈ ∣ P ∣ₚ ] [ x ∈ ⦅ U $ i ⦆ ] ⁆ 264 | ≡⟨ flatten R I (λ i → Σ[ x ∈ ∣ P ∣ₚ ] [ x ∈ ⦅ U $ i ⦆ ]) (λ { _ (a , _) → f a }) ⟩ 265 | ⋁[ R ] ⁅ ⋁[ R ] ⁅ f a ∣ a ∈ ⦅ U $ i ⦆ ⁆ ∣ i ∶ I ⁆ 266 | ≡⟨ refl ⟩ 267 | ⋁[ R ] ⁅ g (U $ i) ∣ i ∶ I ⁆ 268 | ∎ 269 | ``` 270 | 271 | ### `g` is a frame homomorphism 272 | 273 | ``` 274 | g-frame-homo : isFrameHomomorphism L R gm 275 | g-frame-homo = g-resp-𝟏 , (g-resp-⊓ , g-resp-⊔) 276 | ``` 277 | 278 | ### `g` makes the diagram commute 279 | 280 | ``` 281 | lem : (a a′ : ∣ P ∣ₚ) → a′ ◁ π₀ (↓-clos a) → [ f a′ ⊑[ pos R ] f a ] 282 | lem a a′ (squash p q i) = is-true-prop (f a′ ⊑[ pos R ] f a) (lem _ _ p) (lem _ _ q) i 283 | lem a a′ (dir a′⊑a) = f-mono a′ a a′⊑a 284 | lem a a′ (branch b h) = 285 | f a′ ⊑⟨ rep a′ b ⟩ 286 | ⋁[ R ] (outcome F b , f ∘ next F) ⊑⟨ ⋁[ R ]-least _ _ isUB ⟩ 287 | f a ■ 288 | where 289 | isUB : ∀ a₀ → a₀ ε (outcome F b , f ∘ next F) → [ a₀ ⊑[ pos R ] f a ] 290 | isUB a₀ (c , p) = a₀ ⊑⟨ ≡⇒⊑ (pos R) (sym p) ⟩ 291 | f (next F c) ⊑⟨ lem a (next F c) (h c) ⟩ 292 | f a ■ 293 | 294 | gm∘ηm = _∘m_ {P = P} {Q = pos L} {R = pos R} gm ηm 295 | 296 | gm∘ηm~f : (x : ∣ P ∣ₚ) → gm $ₘ (ηm $ₘ x) ≡ fm $ₘ x 297 | gm∘ηm~f x = ⊑[ pos R ]-antisym _ _ down (⋁[ R ]-upper _ _ ((x , x◁x↓ x) , refl)) 298 | where 299 | down : [ (⋁[ R ] (∃ π₀ (e x) , f ∘ π₀)) ⊑[ pos R ] f x ] 300 | down = ⋁[ R ]-least _ _ λ { o ((y , φ) , eq) → subst (λ _ → _) eq (lem x y φ) } 301 | 302 | g∘η=f : gm∘ηm ≡ fm 303 | g∘η=f = forget-mono P (pos R) gm∘ηm fm (funExt gm∘ηm~f) 304 | 305 | g∘η=f′ : g ∘ η ≡ f 306 | g∘η=f′ = subst (λ { (h , _) → h ≡ f }) (sym g∘η=f) refl 307 | ``` 308 | 309 | ### `g` is uniquely determined 310 | 311 | ``` 312 | g-unique : (y : Σ[ g′ ∈ (L ─f→ R) ] 313 | (_∘m_ {P = P} {Q = pos L} {R = pos R} (π₀ g′) ηm ≡ fm)) 314 | → ((gm , g-frame-homo) , g∘η=f) ≡ y 315 | g-unique ((g′m , g′-frame-homo) , φ) = Σ≡Prop I II 316 | where 317 | g′ = _$ₘ_ g′m 318 | 319 | f=g′∘η : f ≡ g′ ∘ η 320 | f=g′∘η = subst (λ { (f′ , _) → f′ ≡ g′ ∘ η }) φ refl 321 | 322 | NTS₀ : (y : Σ (∣ pos L ∣ₚ → ∣ pos R ∣ₚ) (isMonotonic (pos L) (pos R))) 323 | → isProp ((_∘m_ {P = P} {Q = pos L} {R = pos R} y ηm) ≡ fm) 324 | NTS₀ y = isSetΣ 325 | (isSetΠ λ _ → carrier-is-set (pos R)) 326 | (λ h → isProp→isSet (isMonotonic-prop P (pos R) h)) 327 | (_∘m_ {P = P} {Q = pos L} {R = pos R} y ηm) fm 328 | 329 | I : (h : L ─f→ R) → isProp (_∘m_ {P = P} {Q = pos L} {R = pos R} (π₀ h) ηm ≡ fm) 330 | I h = isSetΣ 331 | (isSetΠ λ _ → carrier-is-set (pos R)) 332 | (λ h → isProp→isSet (isMonotonic-prop P (pos R) h)) 333 | (_∘m_ {P = P} {Q = pos L} {R = pos R} (π₀ h) ηm) fm 334 | 335 | g~g′ : (𝔘 : ∣ L ∣F) → g 𝔘 ≡ g′ 𝔘 336 | g~g′ 𝔘 = 337 | g 𝔘 ≡⟨ refl ⟩ 338 | ⋁[ R ] ⁅ f u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ ≡⟨ eq₀ ⟩ 339 | ⋁[ R ] ⁅ g′(η u) ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ ≡⟨ eq₁ ⟩ 340 | g′ (⋁[ L ] ⁅ η u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆) ≡⟨ cong g′ (sym (main-lemma 𝔘)) ⟩ 341 | g′ 𝔘 ∎ 342 | where 343 | eq₀ : ⋁[ R ] (f ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫) ≡ ⋁[ R ] ((g′ ∘ η) ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫) 344 | eq₀ = cong (λ - → ⋁[ R ] (- ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫)) f=g′∘η 345 | eq₁ : ⋁[ R ] ((g′ ∘ η) ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫) ≡ g′ (⋁[ L ] (η ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫)) 346 | eq₁ = sym (π₁ (π₁ g′-frame-homo) (η ⟨$⟩ ⟪ ⦅ 𝔘 ⦆ ⟫)) 347 | 348 | II : (gm , g-frame-homo) ≡ (g′m , g′-frame-homo) 349 | II = Σ≡Prop 350 | (isFrameHomomorphism-prop L R) 351 | (Σ≡Prop (isMonotonic-prop (pos L) (pos R)) (funExt g~g′)) 352 | ``` 353 | 354 | ### The final proof 355 | 356 | ``` 357 | main : universal-prop 358 | main R fm@(f , f-mono) f-flat rep = 359 | (((g , g-mono) , g-resp-𝟏 , g-resp-⊓ , g-resp-⊔) , g∘η=f) , g-unique 360 | where 361 | open MainProof R fm f-flat rep 362 | ``` 363 | -------------------------------------------------------------------------------- /FormalTopology/Spectral.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: Spectral Locales 3 | author: Ayberk Tosun 4 | --- 5 | 6 | 24 | 25 | ## Definition of spectrality 26 | 27 | **Definition.** A spectral locale is a locale for which the compact opens form 28 | a base closed under finite meets. 29 | 30 | **Definition (better).** Every open is the join of the set of compact opens 31 | below it and the meet of two compacts opens is compact. Also, the top element is 32 | compact. 33 | 34 | ```agda 35 | closedUnderFinMeets : Fam 𝓦 ∣ F ∣F → (𝓤 ∨ 𝓥 ∨ 𝓦) ̇ 36 | closedUnderFinMeets U = 37 | ∥ Σ[ i ∈ index U ] [ isTop (pos F) (U $ i) ] ∥ 38 | × (((i j : index U) → ∥ Σ[ k ∈ index U ] [ isInf (pos F) (U $ k) (U $ i) (U $ j) ] ∥)) 39 | 40 | closedUnderFinMeets′ : Fam 𝓦 ∣ F ∣F → hProp (𝓤 ∨ 𝓥 ∨ 𝓦) 41 | closedUnderFinMeets′ S = 42 | ∥ Σ[ i ∈ index S ] [ isTop (pos F) (S $ i) ] ∥Ω 43 | ∧ (∀[ i ∶ index S ] ∀[ j ∶ index S ] ∥ Σ[ k ∈ index S ] [ isInf (pos F) (S $ k) (S $ i) (S $ j) ] ∥Ω) 44 | 45 | isSpectralₛ : (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 46 | isSpectralₛ = 47 | Σ[ ℬ ∈ Fam 𝓦 ∣ F ∣F ] 48 | ((i : index ℬ) → [ isCompactOpen F (ℬ $ i) ]) 49 | × isDirBasisFor F ℬ 50 | × closedUnderFinMeets ℬ 51 | 52 | isSpectral′ : (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 53 | isSpectral′ = ∥ isSpectralₛ ∥ 54 | 55 | 56 | isSpectralMap : (G : Frame 𝓤′ 𝓥′ 𝓦) (f : F ─f→ G) 57 | → hProp (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺ ∨ 𝓤′ ∨ 𝓥′) 58 | isSpectralMap G ((f , _) , _) = 59 | ∀[ x ∶ ∣ F ∣F ] isCompactOpen F x ⇒ isCompactOpen G (f x) 60 | 61 | ∥∥-functorial : {A : Type 𝓤} {B : Type 𝓥} → ∥ (A → B) ∥ → ∥ A ∥ → ∥ B ∥ 62 | ∥∥-functorial {B = B} f x = ∥∥-rec (∥∥-prop B) (λ g → ∥∥-rec (∥∥-prop B) (λ y → ∣ g y ∣) x) f 63 | 64 | spec′→basis : isSpectral′ → ∥ hasBasis F ∥ 65 | spec′→basis sp = ∥∥-rec (∥∥-prop (hasBasis F)) nts sp 66 | where 67 | nts : Σ-syntax (Fam 𝓦 ∣ F ∣F) 68 | (λ ℬ → 69 | ((i : index ℬ) → [ isCompactOpen F (ℬ $ i) ]) × 70 | isDirBasisFor F ℬ × closedUnderFinMeets ℬ) → 71 | ∥ hasBasis F ∥ 72 | nts (ℬ , _ , fb , _)= ∣ ℬ , fb ∣ 73 | 74 | spec→compact : isSpectral′ → [ isCompact F ] 75 | spec→compact spec = ∥∥-rec (isProp[] (isCompact F)) nts spec 76 | where 77 | nts : Σ[ ℬ ∈ (Fam 𝓦 ∣ F ∣F) ] (((i : index ℬ) → [ isCompactOpen F (ℬ $ i) ]) × isDirBasisFor F ℬ × closedUnderFinMeets ℬ) 78 | → [ isCompact F ] 79 | nts (ℬ , κ , dir , fm) = ∥∥-rec (isProp[] (isCompact F)) nts′ (π₀ fm) 80 | where 81 | nts′ : _ 82 | nts′ (i , top) = subst (λ - → [ _≪_ F - - ]) (top-unique F (ℬ $ i) top) (κ i) 83 | 84 | spec→compacts-closed-under-∧ : isSpectral′ 85 | → (x y : ∣ F ∣F) 86 | → [ isCompactOpen F x ] 87 | → [ isCompactOpen F y ] 88 | → [ isCompactOpen F (x ⊓[ F ] y) ] 89 | spec→compacts-closed-under-∧ spec x y x-comp y-comp = 90 | ∥∥-rec (isProp[] (isCompactOpen F (x ⊓[ F ] y))) γ spec 91 | where 92 | γ : _ → [ isCompactOpen F (x ⊓[ F ] y) ] 93 | γ (ℬ , κ , p , q) = 94 | ∥∥-rec (isProp[] (isCompactOpen F (x ⊓[ F ] y))) δ (∥∥-× ϕ ψ) 95 | where 96 | ℐ : Fam 𝓦 (index ℬ) 97 | ℐ = π₀ (p x) 98 | 99 | 𝒥 : Fam 𝓦 (index ℬ) 100 | 𝒥 = π₀ (p y) 101 | 102 | x-eq : x ≡ ⋁[ F ] ⁅ ℬ $ i ∣ i ε ℐ ⁆ 103 | x-eq = uncurry (⋁-unique F ⁅ ℬ $ i ∣ i ε ℐ ⁆ x) (π₁ (π₁ (p x))) 104 | 105 | y-eq : y ≡ ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ 106 | y-eq = uncurry (⋁-unique F ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ y) (π₁ (π₁ (p y))) 107 | 108 | cover-x : [ x ⊑[ pos F ] ⋁[ F ] ⁅ ℬ $ i ∣ i ε ℐ ⁆ ] 109 | cover-x = ≡⇒⊑ (pos F) x-eq 110 | 111 | cover-y : [ y ⊑[ pos F ] ⋁[ F ] ⁅ ℬ $ i ∣ i ε 𝒥 ⁆ ] 112 | cover-y = ≡⇒⊑ (pos F) y-eq 113 | 114 | ϕ : ∥ Σ[ i ∈ index ℐ ] [ x ⊑[ pos F ] (ℬ $ ℐ $ i) ] ∥ 115 | ϕ = x-comp ⁅ ℬ $ i ∣ i ε ℐ ⁆ (π₀ (π₁ (p x))) cover-x 116 | 117 | ψ : ∥ Σ[ i ∈ index 𝒥 ] [ y ⊑[ pos F ] (ℬ $ 𝒥 $ i) ] ∥ 118 | ψ = y-comp ⁅ ℬ $ i ∣ i ε 𝒥 ⁆ (π₀ (π₁ (p y))) cover-y 119 | 120 | δ : _ 121 | δ ((i , x≤bᵢ) , (j , y≤bⱼ)) = 122 | ∥∥-rec (isProp[] (isCompactOpen F (x ⊓[ F ] y))) ν (π₁ q (ℐ $ i) (𝒥 $ j)) 123 | where 124 | x=bᵢ : x ≡ ℬ $ (ℐ $ i) 125 | x=bᵢ = ⊑[ pos F ]-antisym _ _ x≤bᵢ (subst (λ - → [ (ℬ $ ℐ $ i) ⊑[ pos F ] - ]) (sym x-eq) (⋁[ F ]-upper _ _ (i , refl))) 126 | 127 | y=bⱼ : y ≡ ℬ $ (𝒥 $ j) 128 | y=bⱼ = ⊑[ pos F ]-antisym _ _ y≤bⱼ (subst (λ - → [ (ℬ $ 𝒥 $ j) ⊑[ pos F ] - ]) (sym y-eq) (⋁[ F ]-upper _ _ (j , refl))) 129 | 130 | ∧-eq : x ⊓[ F ] y ≡ ((ℬ $ ℐ $ i) ⊓[ F ] (ℬ $ 𝒥 $ j)) 131 | ∧-eq = x ⊓[ F ] y ≡⟨ cong (λ - → - ⊓[ F ] y) x=bᵢ ⟩ 132 | (ℬ $ ℐ $ i) ⊓[ F ] y ≡⟨ cong (λ - → (ℬ $ ℐ $ i) ⊓[ F ] -) y=bⱼ ⟩ 133 | (ℬ $ ℐ $ i) ⊓[ F ] (ℬ $ 𝒥 $ j) ∎ 134 | 135 | ν : _ 136 | ν (k , inf) = 137 | subst (λ - → [ isCompactOpen F - ]) inf′ (κ k) 138 | where 139 | inf′ : ℬ $ k ≡ x ⊓[ F ] y 140 | inf′ = ℬ $ k ≡⟨ ⦅𝟏⦆ ⟩ 141 | (ℬ $ ℐ $ i) ⊓[ F ] (ℬ $ 𝒥 $ j) ≡⟨ sym ∧-eq ⟩ 142 | x ⊓[ F ] y ∎ 143 | where 144 | ⦅𝟏⦆ = ⊓-unique F _ _ _ (π₀ (π₀ inf)) (π₁ (π₀ inf)) λ w p q → π₁ inf w (p , q) 145 | 146 | -- spec′→spec : isSpectral′ → isSpectral 147 | -- spec′→spec spec′ = G𝟏 , G𝟐 , G𝟑 148 | -- where 149 | -- G𝟏 : (x : ∣ F ∣F) 150 | -- → ∥ Σ[ U ∈ Fam 𝓦 ∣ F ∣F ] [ isSup (pos F) U x ] × [ ∀[ x ε U ] (isCompactOpen F x) ] ∥ 151 | -- G𝟏 x = ∥∥-rec (∥∥-prop _) G𝟏a ? 152 | -- where 153 | -- G𝟏a : Σ-syntax (Fam 𝓦 ∣ F ∣F) 154 | -- (λ ℬ → 155 | -- ((i : index ℬ) → [ isCompactOpen F (ℬ $ i) ]) × 156 | -- ((x₁ : ∣ F ∣F) → 157 | -- Σ-syntax (Fam 𝓦 (index ℬ)) 158 | -- (λ I → 159 | -- [ isDirected (pos F) (fmap (_$_ ℬ) I) ] × 160 | -- [ isSup (pos F) (fmap (_$_ ℬ) I) x₁ ])) 161 | -- × 162 | -- Σ-syntax (index ℬ) (λ i → [ isTop (pos F) (ℬ $ i) ]) × 163 | -- ((i j : index ℬ) → 164 | -- Σ-syntax (index ℬ) 165 | -- (λ k → [ isInf (pos F) (ℬ $ k) (ℬ $ i) (ℬ $ j) ]))) → 166 | -- ∥ 167 | -- Σ-syntax (Fam 𝓦 ∣ F ∣F) 168 | -- (λ U → [ isSup (pos F) U x ] × [ fam-forall U (isCompactOpen F) ]) 169 | -- ∥ 170 | -- G𝟏a (ℬ , ϕ , J , _) = 171 | -- ∣ ⁅ ℬ $ j ∣ j ε π₀ (J x) ⁆ 172 | -- , (π₁ (π₁ (J x)) , λ b (i , p) → subst ([_] ∘ isCompactOpen F) p (ϕ (π₀ (J x) $ i))) ∣ 173 | 174 | -- G𝟐 : [ isCompactOpen F ⊤[ F ] ] 175 | -- G𝟐 = ∥∥-rec (isProp[] (isCompactOpen F ⊤[ F ])) G𝟐a spec′ 176 | -- where 177 | -- G𝟐a : _ → [ isCompactOpen F ⊤[ F ] ] 178 | -- G𝟐a (ℬ , (ϕ , _ , (i , p) , r)) = subst ([_] ∘ isCompactOpen F) G𝟐b (ϕ i) 179 | -- where 180 | -- G𝟐b : ℬ $ i ≡ ⊤[ F ] 181 | -- G𝟐b = top-unique F (ℬ $ i) p 182 | 183 | -- G𝟑 : (x y : ∣ F ∣F) 184 | -- → [ isCompactOpen F x ] → [ isCompactOpen F y ] → [ isCompactOpen F (x ⊓[ F ] y) ] 185 | -- G𝟑 x y x-comp y-comp = 186 | -- ∥∥-rec (isProp[] (isCompactOpen F (x ⊓[ F ] y))) G𝟑a spec′ 187 | -- where 188 | -- G𝟑a : _ → [ isCompactOpen F (x ⊓[ F ] y) ] 189 | -- G𝟑a (ℬ , κ , (ϕ , ψ)) = 190 | -- ∥∥-rec 191 | -- (isProp[] (isCompactOpen F (x ⊓[ F ] y))) 192 | -- G𝟑b 193 | -- (∥∥-× (x-comp _ dir₀ cover₀) (y-comp _ dir₁ cover₁)) 194 | -- where 195 | -- ℐ : Fam 𝓦 (index ℬ) 196 | -- ℐ = π₀ (ϕ x) 197 | 198 | -- 𝒥 : Fam 𝓦 (index ℬ) 199 | -- 𝒥 = π₀ (ϕ y) 200 | 201 | -- υ₀ : [ isSup (pos F) ⁅ ℬ $ i ∣ i ε ℐ ⁆ x ] 202 | -- υ₀ = π₁ (π₁ (ϕ x)) 203 | 204 | -- υ₁ : [ isSup (pos F) ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ y ] 205 | -- υ₁ = π₁ (π₁ (ϕ y)) 206 | 207 | -- dir₀ : [ isDirected (pos F) ⁅ ℬ $ i ∣ i ε ℐ ⁆ ] 208 | -- dir₀ = π₀ (π₁ (ϕ x)) 209 | 210 | -- dir₁ : [ isDirected (pos F) ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ] 211 | -- dir₁ = π₀ (π₁ (ϕ y)) 212 | 213 | -- cover₀ : [ x ⊑[ pos F ] ⋁[ F ] ⁅ ℬ $ i ∣ i ε ℐ ⁆ ] 214 | -- cover₀ = ≡⇒⊑ (pos F) (⋁-unique F _ _ (π₀ υ₀) (π₁ υ₀)) 215 | 216 | -- cover₁ : [ y ⊑[ pos F ] ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ] 217 | -- cover₁ = ≡⇒⊑ (pos F) (⋁-unique F _ _ (π₀ υ₁) (π₁ υ₁)) 218 | 219 | -- G𝟑b : _ → [ isCompactOpen F (x ⊓[ F ] y) ] 220 | -- G𝟑b ((𝒾 , p) , (𝒿 , q)) = subst ([_] ∘ isCompactOpen F) G𝟑c (κ k) 221 | -- where 222 | -- open PosetReasoning (pos F) 223 | 224 | -- i : index ℬ 225 | -- i = π₀ (ϕ x) $ 𝒾 226 | 227 | -- j : index ℬ 228 | -- j = π₀ (ϕ y) $ 𝒿 229 | 230 | -- k : index ℬ 231 | -- k = ? -- π₀ (π₁ ψ ((π₀ (ϕ x)) $ 𝒾) ((π₀ (ϕ y)) $ 𝒿)) 232 | 233 | -- foo : x ≡ ℬ $ i 234 | -- foo = ⊑[ pos F ]-antisym x (ℬ $ i) p nts 235 | -- where 236 | -- nts : [ (ℬ $ i) ⊑[ pos F ] x ] 237 | -- nts = ℬ $ i ⊑⟨ ⋁[ F ]-upper _ _ (𝒾 , refl) ⟩ 238 | -- ⋁[ F ] ⁅ ℬ $ j ∣ j ε π₀ (ϕ x) ⁆ ⊑⟨ ≡⇒⊑ (pos F) (sym (⋁-unique F _ _ (π₀ υ₀) (π₁ υ₀))) ⟩ 239 | -- x ■ 240 | 241 | -- bar : y ≡ ℬ $ j 242 | -- bar = ⊑[ pos F ]-antisym y (ℬ $ j) q nts 243 | -- where 244 | -- nts : [ (ℬ $ j) ⊑[ pos F ] y ] 245 | -- nts = 246 | -- ℬ $ j ⊑⟨ ⋁[ F ]-upper _ _ (𝒿 , refl) ⟩ 247 | -- ⋁[ F ] ⁅ ℬ $ j ∣ j ε π₀ (ϕ y) ⁆ ⊑⟨ ≡⇒⊑ (pos F) (sym (⋁-unique F _ _ (π₀ υ₁) (π₁ υ₁))) ⟩ 248 | -- y ■ 249 | 250 | -- G𝟑c : ℬ $ k ≡ x ⊓[ F ] y 251 | -- G𝟑c = ℬ $ k ≡⟨ ⊓-unique F _ _ _ (π₀ (π₀ (π₁ (π₁ ψ i j)))) ((π₁ (π₀ (π₁ (π₁ ψ i j))))) (λ w p q → π₁ (π₁ (π₁ ψ i j)) w (p , q)) ⟩ 252 | -- (ℬ $ i) ⊓[ F ] (ℬ $ j) ≡⟨ cong (λ - → - ⊓[ F ] (ℬ $ j)) (sym foo) ⟩ 253 | -- x ⊓[ F ] (ℬ $ j) ≡⟨ cong (λ - → x ⊓[ F ] -) (sym bar) ⟩ 254 | -- x ⊓[ F ] y ∎ 255 | 256 | 257 | 258 | -- TODO. 259 | -- The definition of spectral should be the same as Stone but the requirement of clopen 260 | -- basis replaced with the requirement of a compact basis. 261 | ``` 262 | 263 | ```agda 264 | compact-yoneda : isSpectral′ 265 | → (x y : ∣ F ∣F) 266 | → ((b : ∣ F ∣F) → [ isCompactOpen F b ] → 267 | [ b ⊑[ pos F ] x ] → [ b ⊑[ pos F ] y ]) 268 | → [ x ⊑[ pos F ] y ] 269 | compact-yoneda spec x y ϕ = 270 | ∥∥-rec (isProp[] (x ⊑[ pos F ] y)) G𝟏 spec 271 | where 272 | open PosetReasoning (pos F) 273 | 274 | G𝟏 : _ 275 | → [ x ⊑[ pos F ] y ] 276 | G𝟏 (ℬ , κ , basis , _) = 277 | x ⊑⟨ ⦅𝟏⦆ ⟩ 278 | ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ⊑⟨ ⦅𝟐⦆ ⟩ 279 | y ■ 280 | where 281 | 𝒥 : Fam 𝓦 (index ℬ) 282 | 𝒥 = π₀ (basis x) 283 | 284 | 𝒦 : Fam 𝓦 (index ℬ) 285 | 𝒦 = π₀ (basis y) 286 | 287 | cover-x : x ≡ ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ 288 | cover-x = uncurry (⋁-unique F ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ x) (π₁ (π₁ (basis x))) 289 | 290 | ⦅𝟏⦆ : [ x ⊑[ pos F ] ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ] 291 | ⦅𝟏⦆ = ≡⇒⊑ (pos F) cover-x 292 | 293 | ⦅𝟐⦆ : [ ⋁[ F ] ⁅ ℬ $ j ∣ j ε 𝒥 ⁆ ⊑[ pos F ] y ] 294 | ⦅𝟐⦆ = ⋁[ F ]-least _ _ ⦅𝟐a⦆ 295 | where 296 | ⦅𝟐a⦆ : _ 297 | ⦅𝟐a⦆ b (i , eq) = subst (λ - → [ - ⊑[ pos F ] y ]) eq (ϕ (ℬ $ (𝒥 $ i)) (κ (𝒥 $ i)) ⦅𝟐b⦆) 298 | where 299 | ⦅𝟐b⦆ : [ (ℬ $ 𝒥 $ i) ⊑[ pos F ] x ] 300 | ⦅𝟐b⦆ = subst 301 | (λ - → [ (ℬ $ 𝒥 $ i) ⊑[ pos F ] - ]) 302 | (sym cover-x) 303 | (⋁[ F ]-upper _ _ (i , refl)) 304 | 305 | compact-yoneda₁ : isSpectral′ 306 | → (x y : ∣ F ∣F) 307 | → [ x ⊑[ pos F ] y ] 308 | → ((b : ∣ F ∣F) → [ isCompactOpen F b ] → 309 | [ b ⊑[ pos F ] x ] → [ b ⊑[ pos F ] y ]) 310 | compact-yoneda₁ spec x y p b b-comp q = 311 | b ⊑⟨ q ⟩ x ⊑⟨ p ⟩ y ■ 312 | where 313 | open PosetReasoning (pos F) 314 | ``` 315 | 316 | ```agda 317 | -- spectral→stone : (F : Frame 𝓤 𝓥 𝓦) → isSpectral F → [ isStone F ] 318 | -- spectral→stone F F-spec = ? 319 | ``` 320 | 321 | ```agda 322 | resp-compactness : (f : ∣ F ∣F → ∣ F ∣F) → (𝓤 ∨ 𝓥 ∨ 𝓦 ⁺) ̇ 323 | resp-compactness f = 324 | (b x : ∣ F ∣F) → 325 | [ isCompactOpen F b ] → 326 | [ b ⊑[ pos F ] f x ] → 327 | Σ[ a ∈ ∣ F ∣F ] 328 | [ isCompactOpen F a ] × [ a ⊑[ pos F ] x ] × [ b ⊑[ pos F ] f a ] 329 | 330 | continuity-lemma : isSpectral′ 331 | → (f : ∣ F ∣F → ∣ F ∣F) 332 | → isMonotonic (pos F) (pos F) f 333 | → resp-compactness f 334 | → isScottCont′ F f 335 | continuity-lemma spec f mono comp U U-dir = 336 | ⊑[ pos F ]-antisym _ _ β γ 337 | where 338 | η : (b : ∣ F ∣F) 339 | → [ isCompactOpen F b ] 340 | → [ b ⊑[ pos F ] f (⋁[ F ] U) ] 341 | → [ b ⊑[ pos F ] ⋁[ F ] ⁅ f x ∣ x ε U ⁆ ] 342 | η b b-compact α = 343 | b ⊑⟨ π₁ (π₁ (π₁ (comp b _ b-compact α))) ⟩ 344 | f a ⊑⟨ nts ⟩ 345 | ⋁[ F ] ⁅ f x ∣ x ε U ⁆ ■ 346 | where 347 | open PosetReasoning (pos F) 348 | 349 | a = π₀ (comp b _ b-compact α) 350 | 351 | a-comp : [ isCompactOpen F a ] 352 | a-comp = π₀ (π₁ (comp b _ b-compact α)) 353 | 354 | q : [ a ⊑[ pos F ] ⋁[ F ] U ] 355 | q = π₀ (π₁ (π₁ (comp b _ b-compact α))) 356 | 357 | rem : Σ[ i ∈ index U ] [ a ⊑[ pos F ] (U $ i) ] 358 | → [ f a ⊑[ pos F ] ⋁[ F ] ⁅ f x ∣ x ε U ⁆ ] 359 | rem (i , p) = 360 | f a ⊑⟨ mono a (U $ i) p ⟩ 361 | f (U $ i) ⊑⟨ ⋁[ F ]-upper _ _ (i , refl) ⟩ 362 | ⋁[ F ] ⁅ f x ∣ x ε U ⁆ ■ 363 | 364 | nts : [ f a ⊑[ pos F ] ⋁[ F ] ⁅ f x ∣ x ε U ⁆ ] 365 | nts = ∥∥-rec (isProp[] (_ ⊑[ pos F ] _)) rem (a-comp U U-dir q) 366 | 367 | β : [ f (⋁[ F ] U) ⊑[ pos F ] (⋁[ F ] ⁅ f x ∣ x ε U ⁆) ] 368 | β = compact-yoneda spec (f (⋁[ F ] U)) (⋁[ F ] ⁅ f x ∣ x ε U ⁆) η 369 | 370 | δ : (z : ∣ F ∣F) → z ε ⁅ f x ∣ x ε U ⁆ → [ z ⊑[ pos F ] f (⋁[ F ] U) ] 371 | δ z (i , eq) = subst (λ - → [ - ⊑[ pos F ] f (⋁[ F ] U) ]) eq nts 372 | where 373 | nts : [ f (U $ i) ⊑[ pos F ] (f (⋁[ F ] U)) ] 374 | nts = mono (U $ i) (⋁[ F ] U) (⋁[ F ]-upper _ _ (i , refl)) 375 | 376 | γ : [ (⋁[ F ] ⁅ f x ∣ x ε U ⁆) ⊑[ pos F ] f (⋁[ F ] U) ] 377 | γ = ⋁[ F ]-least ⁅ f x ∣ x ε U ⁆ (f (⋁[ F ] U)) δ 378 | 379 | -- function-lemma : (f g : ∣ F ∣F → ∣ F ∣F) 380 | -- → isMonotonic (pos F) (pos F) f 381 | -- → isMonotonic (pos F) (pos F) g 382 | -- → ((b : ∣ F ∣F) → [ isCompactOpen F b ] → f b ≡ g b) 383 | -- → (x : ∣ F ∣F) 384 | -- → f x ≡ g x 385 | -- function-lemma f g f-sc g-sc ϕ x = 386 | -- f x ≡⟨ {!!} ⟩ f (⋁⟨ i ⟩ (b $ i)) ≡⟨ {!!} ⟩ {!f (⋁⟨ i ⟩ (b $ i))!} ≡⟨ {!!} ⟩ g x ∎ 387 | -- where 388 | -- open JoinSyntax ∣ F ∣F (λ - → ⋁[ F ] -) 389 | 390 | -- b = π₀ (π₀ spec x) 391 | ``` 392 | 393 | ```agda 394 | -- patch-is-stone : [ isStone Patch ] 395 | -- patch-is-stone = patch-is-compact , ∣ {!!} ∣ 396 | ``` 397 | 398 | TODO: 399 | 400 | 1. Prove 3.3.(i) 401 | 2. Patch(A) is a Stone locale for every spectral A. 402 | n 403 | -------------------------------------------------------------------------------- /FormalTopology/Sierpinski.lagda.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: The Sierpiński Frame 3 | author: Ayberk Tosun (j.w.w. Martín Escardó) 4 | --- 5 | 6 | 27 | 28 | We present a construction of the Sierpiński frame from a formal topology in 29 | cubical Agda. Essentially, we prove the following: 30 | 31 | > there exists a formal topology ℱ such that the frame generated from ℱ 32 | > classifies the opens of any locale. 33 | 34 | In Agda, we express this as follows: 35 | 36 | ```agda 37 | sierpiński-exists : Σ[ S ∈ FormalTopology 𝓤₀ 𝓤₀ ] 38 | ((A : Frame 𝓤₁ 𝓤₀ 𝓤₀) → ((to-frame S) ─f→ A) ≃ ∣ A ∣F) 39 | ``` 40 | 41 | You can click [here](#the-proof) to jump to the inhabitant of this type that we 42 | construct and follow the construction in a top-down manner by examining its 43 | constituents. Otherwise, you can continue reading to follow the construction in 44 | a bottom-up manner. 45 | 46 | ## Sierpiński formal topology 47 | 48 | We start by writing down our poset of basic opens, which is the following 49 | two element poset: 50 | 51 | ```text 52 | false 53 | | 54 | | 55 | | 56 | true 57 | ``` 58 | 59 | It is a bit counterintuitive that `true` is less than `false` but we are working 60 | with opposites of the usual “information ordering” posets from domain theory. 61 | 62 | ```agda 63 | S-pos : (𝓤 𝓥 : Universe) → Poset 𝓤 𝓥 64 | S-pos 𝓤 𝓥 = Bool 𝓤 , (_≤_ , Bool-set , ≤-refl , ≤-trans , ≤-antisym) 65 | where 66 | _≤_ : Bool 𝓤 → Bool 𝓤 → hProp 𝓥 67 | _ ≤ false = Unit 𝓥 , Unit-prop 68 | true ≤ true = Unit 𝓥 , Unit-prop 69 | _ ≤ _ = bot 𝓥 70 | 71 | ≤-refl : (x : Bool 𝓤) → [ x ≤ x ] 72 | ≤-refl false = tt 73 | ≤-refl true = tt 74 | 75 | ≤-trans : (x y z : Bool 𝓤) → [ x ≤ y ] → [ y ≤ z ] → [ x ≤ z ] 76 | ≤-trans _ true true p _ = p 77 | ≤-trans _ _ false _ _ = tt 78 | 79 | ≤-antisym : (x y : Bool 𝓤) → [ x ≤ y ] → [ y ≤ x ] → x ≡ y 80 | ≤-antisym false false p q = refl 81 | ≤-antisym true true p q = refl 82 | ``` 83 | 84 | The Sierpiński formal topology is obtained by equipping this poset with the 85 | empty interaction system, which ensures that the inductively generated covering 86 | relation is trivial. 87 | 88 | ```agda 89 | S : (𝓤 𝓥 : Universe) → FormalTopology 𝓤 𝓥 90 | S 𝓤 𝓥 = S-pos 𝓤 𝓥 , S-IS , S-has-mono , S-has-sim 91 | where 92 | S-exp : Bool 𝓤 → 𝓤 ̇ 93 | S-exp _ = 𝟘 𝓤 94 | 95 | S-out : {x : Bool 𝓤} → S-exp x → 𝓤 ̇ 96 | S-out () 97 | 98 | S-rev : {x : Bool 𝓤} {y : S-exp x} → S-out {x = x} y → Bool 𝓤 99 | S-rev {y = ()} 100 | 101 | S-IS : InteractionStr (Bool 𝓤) 102 | S-IS = S-exp , (λ {x} → S-out {x = x}) , S-rev 103 | 104 | S-has-mono : hasMono (S-pos 𝓤 𝓥) S-IS 105 | S-has-mono _ () _ 106 | 107 | S-has-sim : hasSimulation (S-pos 𝓤 𝓥) S-IS 108 | S-has-sim _ _ _ () 109 | ``` 110 | 111 | ## The Sierpiński frame 112 | 113 | The Sierpínski frame 𝕊 is defined simply as `to-frame S`: 114 | 115 | ```agda 116 | 𝕊 : Frame 𝓤₁ 𝓤₀ 𝓤₀ 117 | 𝕊 = to-frame (S 𝓤₀ 𝓤₀) 118 | ``` 119 | 120 | First of all, notice that the covering is trivial: 121 | 122 | 127 | 128 | ```agda 129 | ◁-triv : (U : 𝒫 (Bool 𝓤₀)) (b : Bool 𝓤₀) → b ◁ U → [ b ∈ U ] 130 | ◁-triv U b (dir p) = p 131 | ◁-triv U b (squash p q i) = isProp[] (b ∈ U) (◁-triv U b p) (◁-triv U b q) i 132 | ``` 133 | 134 | Let us write down the fact that equality in the Sierpiński frame reduces to 135 | equality of the underlying sets: 136 | 137 | ```agda 138 | 𝕊-equality : (𝔘 𝔙 : ∣ 𝕊 ∣F) → ⦅ 𝔘 ⦆ ≡ ⦅ 𝔙 ⦆ → 𝔘 ≡ 𝔙 139 | 𝕊-equality 𝔘 𝔙 p = Σ≡Prop nts₀ (Σ≡Prop nts₁ p) 140 | where 141 | nts₀ : (U : ∣ DCPoset (S-pos _ _) ∣ₚ) → isProp (𝕛 U ≡ U) 142 | nts₀ U = carrier-is-set (DCPoset (S-pos 𝓤₀ 𝓤₀)) (𝕛 U) U 143 | 144 | nts₁ : (U : 𝒫 ∣ S-pos 𝓤₀ 𝓤₀ ∣ₚ) → isProp [ isDownwardsClosed (S-pos 𝓤₀ 𝓤₀) U ] 145 | nts₁ U = isProp[] (isDownwardsClosed (S-pos 𝓤₀ 𝓤₀) U) 146 | ``` 147 | 148 | There are three inhabitants of the Sierpinski frame so let us write this down 149 | to make things a bit more concrete. 150 | 151 | The singleton set containing true: 152 | 153 | ```agda 154 | ⁅true⁆ : ∣ 𝕊 ∣F 155 | ⁅true⁆ = (Q , Q-dc) , fix 156 | where 157 | Q : 𝒫 ∣ S-pos 𝓤₀ 𝓤₀ ∣ₚ 158 | Q x = (x ≡ true) , Bool-set x true 159 | 160 | Q-dc : [ isDownwardsClosed (S-pos 𝓤₀ 𝓤₀) Q ] 161 | Q-dc true true x∈Q _ = x∈Q 162 | Q-dc false true x∈Q _ = ⊥-rec (true≠false (sym x∈Q)) 163 | Q-dc false false x∈Q _ = x∈Q 164 | 165 | fix : 𝕛 (Q , Q-dc) ≡ (Q , Q-dc) 166 | fix = Σ≡Prop 167 | (isProp[] ∘ isDownwardsClosed (S-pos _ _)) 168 | (⊆-antisym (◁-triv Q) (λ _ → dir)) 169 | ``` 170 | 171 | Note that this is the same thing as `η true` i.e. the set `_ ◁ ⁅ true ⁆`: 172 | 173 | ```agda 174 | ⊤-lemma : (𝔘 : ∣ 𝕊 ∣F) → [ true ∈ ⦅ 𝔘 ⦆ ] → [ ⁅true⁆ ⊑[ pos 𝕊 ] 𝔘 ] 175 | ⊤-lemma 𝔘 p true q = p 176 | ⊤-lemma 𝔘 p false q = ⊥-rec (true≠false (sym q)) 177 | 178 | ⁅true⁆=η-true : ⁅true⁆ ≡ η true 179 | ⁅true⁆=η-true = 𝕊-equality _ _ (⊆-antisym (⊤-lemma (η true) (dir tt)) goal) 180 | where 181 | goal : [ ⦅ η true ⦆ ⊆ ⦅ ⁅true⁆ ⦆ ] 182 | goal true (dir p) = refl 183 | goal b (squash p q i) = isProp[] (b ∈ ⦅ ⁅true⁆ ⦆) (goal b p) (goal b q) i 184 | ``` 185 | 186 | The top element `⊤[ 𝕊 ]` which is the set containing both `true` and `false`. It 187 | is the same thing as the downwards-closure of `η false`. 188 | 189 | ```agda 190 | 𝟏=η-false : ⊤[ 𝕊 ] ≡ η false 191 | 𝟏=η-false = 𝕊-equality ⊤[ 𝕊 ] (η false) (⊆-antisym goal λ _ _ → tt) 192 | where 193 | goal : [ ⦅ ⊤[ 𝕊 ] ⦆ ⊆ ⦅ η false ⦆ ] 194 | goal true _ = π₁ (π₀ (η false)) true _ (dir tt) tt 195 | goal false _ = dir (⊑[ S-pos 𝓤₀ 𝓤₀ ]-refl false) 196 | ``` 197 | 198 | We will sometimes how to talk about set being non-empty i.e. containing either 199 | `true` or `false`. To do that, we define the following function: 200 | 201 | ```agda 202 | _≠∅ : (U : ∣ 𝕊 ∣F) → 𝓤₀ ̇ 203 | 𝔘 ≠∅ = [ true ∈ ⦅ 𝔘 ⦆ ] ⊎ [ false ∈ ⦅ 𝔘 ⦆ ] 204 | ``` 205 | 206 | 213 | 214 | ## Is this the correct Sierpiński frame? 215 | 216 | Fix a frame `A` whose index types are small. 217 | 218 | ```agda 219 | module _ (A : Frame 𝓤 𝓥 𝓤₀) where 220 | ``` 221 | 222 | We need to show that `𝕊` classifies the opens of `A`. 223 | 224 | We will use the following shorthand for `A`'s operations: 225 | 226 | ``` 227 | private 228 | ⋁_ : Fam 𝓤₀ ∣ A ∣F → ∣ A ∣F 229 | ⋁ U = ⋁[ A ] U 230 | 231 | _∨_ : ∣ A ∣F → ∣ A ∣F → ∣ A ∣F 232 | x ∨ y = x ∨[ A ] y 233 | 234 | _∧_ : ∣ A ∣F → ∣ A ∣F → ∣ A ∣F 235 | x ∧ y = x ⊓[ A ] y 236 | 237 | _≤_ : ∣ A ∣F → ∣ A ∣F → hProp 𝓥 238 | x ≤ y = x ⊑[ pos A ] y 239 | 240 | 𝟏 : ∣ A ∣F 241 | 𝟏 = ⊤[ A ] 242 | ``` 243 | 244 | We now construct an isomorphism 245 | 246 | ```text 247 | to : (𝕊 ─f→ A) ≃ A : from 248 | ``` 249 | 250 | ### The forwards direction (easy) 251 | 252 | ```agda 253 | to : (𝕊 ─f→ A) → ∣ A ∣F 254 | to ((f , _) , _) = f ⁅true⁆ 255 | ``` 256 | 257 | ### The backwards direction 258 | 259 | Let us first define an auxiliary function that we will need in the definition 260 | of `from`, called `𝔨`, defined as: 261 | 262 | ```agda 263 | 𝔨 : ∣ A ∣F → ∣ 𝕊 ∣F → Fam 𝓤₀ ∣ A ∣F 264 | 𝔨 x 𝔘 = ⁅ x ∣ _ ∶ [ true ∈ ⦅ 𝔘 ⦆ ] ⁆ ∪f ⁅ 𝟏 ∣ _ ∶ [ false ∈ ⦅ 𝔘 ⦆ ] ⁆ 265 | ``` 266 | 267 | We then define the underlying function of `from` that we call `from₀`: 268 | 269 | ```agda 270 | from₀ : ∣ A ∣F → ∣ 𝕊 ∣F → ∣ A ∣F 271 | from₀ x 𝔘 = ⋁ 𝔨 x 𝔘 272 | ``` 273 | 274 | To be able to define `from`, which is supposed to be a frame homomorphism for 275 | every `x : ∣ A ∣`, we need to show that `from₀`: 276 | 277 | 1. is monotonic, 278 | 1. respects finite meets, and 279 | 1. respects joins, 280 | 281 | Let us start by proving some lemmas that we will need to prove those. 282 | 283 | First, we note that applying `from₀` to `⁅true⁆` gives back `x`: 284 | 285 | ```agda 286 | from-lemma₀ : (x : ∣ A ∣F) → from₀ x ⁅true⁆ ≡ x 287 | from-lemma₀ x = ⊑[ pos A ]-antisym _ _ nts₀ nts₁ 288 | where 289 | nts₀ : [ from₀ x ⁅true⁆ ⊑[ pos A ] x ] 290 | nts₀ = ⋁[ A ]-least _ _ λ { y (inl i , eq) → ≡⇒⊑ (pos A) (sym eq) 291 | ; y (inr i , eq) → ⊥-rec (true≠false (sym i)) 292 | } 293 | 294 | nts₁ : [ x ≤ from₀ x ⁅true⁆ ] 295 | nts₁ = ⋁[ A ]-upper _ _ (inl refl , refl) 296 | ``` 297 | 298 | `from₀` respects the top element: 299 | 300 | ```agda 301 | from₀-resp-⊤ : (x : ∣ A ∣F) → from₀ x ⊤[ 𝕊 ] ≡ ⊤[ A ] 302 | from₀-resp-⊤ x = 303 | ⊑[ pos A ]-antisym _ _ (⊤[ A ]-top _) (⋁[ A ]-upper _ _ (inr tt , refl)) 304 | ``` 305 | 306 | It also respects the bottom element: 307 | 308 | ```agda 309 | from-resp-⊥ : (x : ∣ A ∣F) → from₀ x ⊥[ 𝕊 ] ≡ ⊥[ A ] 310 | from-resp-⊥ x = 311 | ⊑[ pos A ]-antisym _ _ (⋁[ A ]-least _ _ nts) (⊥[ A ]-bottom _) 312 | where 313 | nts : [ ∀[ z ε _ ] (z ≤ ⊥[ A ]) ] 314 | nts z (inl (dir p) , eq) = ∥∥-rec (isProp[] (_ ≤ _)) (λ { (() , _) }) p 315 | nts z (inl (squash p q i) , eq) = isProp[] (_ ≤ _) (nts z (inl p , eq)) (nts z (inl q , eq)) i 316 | nts z (inr (dir p) , eq) = ∥∥-rec (isProp[] (_ ≤ _)) (λ { (() , _) }) p 317 | nts z (inr (squash p q i) , eq) = isProp[] (_ ≤ _) (nts z (inr p , eq)) (nts z (inr q , eq)) i 318 | ``` 319 | 320 | If some `𝔘 : 𝕊` contains `false`, then it has to be the entire subset: 321 | 322 | ```agda 323 | false∈𝔘→𝔘-top : (𝔘 : ∣ 𝕊 ∣F) → [ false ∈ ⦅ 𝔘 ⦆ ] → ⊤[ 𝕊 ] ≡ 𝔘 324 | false∈𝔘→𝔘-top 𝔘 false∈𝔘 = 𝕊-equality ⊤[ 𝕊 ] 𝔘 (sym (goal 𝔘 false∈𝔘)) 325 | where 326 | goal : (𝔘 : ∣ 𝕊 ∣F) → [ false ∈ ⦅ 𝔘 ⦆ ] → ⦅ 𝔘 ⦆ ≡ entire 327 | goal ((U , U-dc) , _) false∈𝔘 = ⊆-antisym nts₀ nts₁ 328 | where 329 | nts₀ : [ U ⊆ entire ] 330 | nts₀ _ _ = tt 331 | 332 | nts₁ : [ entire ⊆ U ] 333 | nts₁ true tt = U-dc false true false∈𝔘 tt 334 | nts₁ false tt = false∈𝔘 335 | ``` 336 | 337 | #### Monotonicity 338 | 339 | Monotonicity of `from₀` is easy to show. 340 | 341 | ```agda 342 | from₀-mono : (x : ∣ A ∣F) → isMonotonic (pos 𝕊) (pos A) (from₀ x) 343 | from₀-mono x 𝔘 𝔙 𝔘⊆𝔙 = ⋁[ A ]-least _ _ nts 344 | where 345 | nts : _ 346 | nts _ (inl i , eq) = ⋁[ A ]-upper _ _ (inl (𝔘⊆𝔙 true i) , eq) 347 | nts _ (inr j , eq) = ⋁[ A ]-upper _ _ (inr (𝔘⊆𝔙 false j) , eq) 348 | ``` 349 | 350 | ```agda 351 | from₀m : ∣ A ∣F → pos 𝕊 ─m→ pos A 352 | from₀m x = from₀ x , from₀-mono x 353 | ``` 354 | 355 | #### Continuity 356 | 357 | We now prove that `from₀` is a frame homomorphism. 358 | 359 | We have already shown that it respects the top element. 360 | 361 | ```agda 362 | resp-⊤ : (x : ∣ A ∣F) → from₀ x ⊤[ 𝕊 ] ≡ ⊤[ A ] 363 | resp-⊤ = from₀-resp-⊤ 364 | ``` 365 | 366 | To show meet preservation, we make use of the fact that bi-cofinal families have 367 | the same join. 368 | 369 | ```agda 370 | from₀-resp-∧ : (x : ∣ A ∣F) (𝔘 𝔙 : ∣ 𝕊 ∣F) 371 | → from₀ x (𝔘 ⊓[ 𝕊 ] 𝔙) ≡ from₀ x 𝔘 ∧ from₀ x 𝔙 372 | from₀-resp-∧ x 𝔘@((_ , 𝔘-dc) , _) 𝔙@((_ , 𝔙-dc) , _) = 373 | from₀ x (𝔘 ⊓[ 𝕊 ] 𝔙) ≡⟨ refl ⟩ 374 | ⋁ 𝔨 x (𝔘 ⊓[ 𝕊 ] 𝔙) ≡⟨ nts ⟩ 375 | (⋁ ⁅ _ ∧ _ ∣ (_ , _) ∶ 𝔘 ≠∅ × 𝔙 ≠∅ ⁆ ) ≡⟨ sym (sym-distr A _ _) ⟩ 376 | (⋁ 𝔨 x 𝔘) ∧ (⋁ 𝔨 x 𝔙) ≡⟨ refl ⟩ 377 | (from₀ x 𝔘) ∧ (from₀ x 𝔙) ∎ 378 | where 379 | nts₀ : _ 380 | nts₀ (inl (p , q)) = (inl p , inl q) , ≡⇒⊑ (pos A) (sym (x∧x=x A x)) 381 | nts₀ (inr (p , q)) = (inr p , inr q) , ≡⇒⊑ (pos A) (sym (x∧x=x A ⊤[ A ])) 382 | 383 | nts₁ : _ 384 | nts₁ (inl p , inl q) = inl (p , q) , ≡⇒⊑ (pos A) (x∧x=x A x) 385 | nts₁ (inl p , inr q) = inl (p , 𝔙-dc false true q tt) , ⊓[ A ]-lower₀ _ _ 386 | nts₁ (inr p , inl q) = inl (𝔘-dc false true p tt , q) , ⊓[ A ]-lower₁ _ _ 387 | nts₁ (inr p , inr q) = (inr (p , q)) , ≡⇒⊑ (pos A) (x∧x=x A ⊤[ A ]) 388 | 389 | nts : (⋁ _) ≡ (⋁ _) 390 | nts = bicofinal→same-join A _ _ (nts₀ , nts₁) 391 | ``` 392 | 393 | Preservation of joins is a bit more complicated. Let `W` be a family of 394 | Sierpiński opens and let `x : A`. We use the uniqueness of 395 | `⋁ ⁅ from₀ x 𝔘 ∣ 𝔘 ε W ⁆` by showing that `from₀ x (⋁[ 𝕊 ] W)` is the least 396 | upper bound of `⁅ from₀ x 𝔘 ∣ 𝔘 ε W ⁆`. The fact that it is an upper bound 397 | is given by `ub` and the fact that it is the least such is given in `least`. 398 | 399 | ```agda 400 | from₀-comm-⋁ : (x : ∣ A ∣F) (W : Fam 𝓤₀ ∣ 𝕊 ∣F) 401 | → from₀ x (⋁[ 𝕊 ] W) ≡ ⋁ ⁅ from₀ x 𝔘 ∣ 𝔘 ε W ⁆ 402 | from₀-comm-⋁ x W = ⋁-unique A _ _ ub least 403 | where 404 | ub : [ ∀[ z ε ⁅ from₀ x 𝔘 ∣ 𝔘 ε W ⁆ ] (z ≤ from₀ x (⋁[ 𝕊 ] W)) ] 405 | ub z (i , eq) = 406 | subst (λ - → [ - ≤ _ ]) eq (from₀-mono x (W $ i) (⋁[ 𝕊 ] W) rem) 407 | where 408 | rem : [ (W $ i) ⊑[ pos 𝕊 ] (⋁[ 𝕊 ] W) ] 409 | rem _ x∈Wᵢ = dir ∣ i , x∈Wᵢ ∣ 410 | 411 | least : (u : ∣ A ∣F) 412 | → (((z : ∣ A ∣F) → z ε ⁅ from₀ x 𝔘 ∣ 𝔘 ε W ⁆ → [ z ≤ u ])) 413 | → [ from₀ x (⋁[ 𝕊 ] W) ≤ u ] 414 | least u u-upper = ⋁[ A ]-least _ _ rem 415 | where 416 | open PosetReasoning (pos A) 417 | 418 | rem : (z : ∣ A ∣F) → z ε 𝔨 x (⋁[ 𝕊 ] W) → [ z ≤ u ] 419 | rem z (inl (dir p) , eq) = 420 | subst (λ - → [ - ≤ u ]) eq (∥∥-rec (isProp[] (_ ≤ _ )) goal p) 421 | where 422 | goal : _ 423 | goal (j , true∈Wⱼ) = 424 | x ⊑⟨ ≡⇒⊑ (pos A) (sym (from-lemma₀ x)) ⟩ 425 | from₀ x ⁅true⁆ ⊑⟨ from₀-mono x ⁅true⁆ (W $ j) nts ⟩ 426 | from₀ x (W $ j) ⊑⟨ u-upper (from₀ x (W $ j)) (j , refl) ⟩ 427 | u ■ 428 | where 429 | nts : [ ⁅true⁆ ⊑[ pos 𝕊 ] (W $ j) ] 430 | nts true _ = true∈Wⱼ 431 | nts false p = ⊥-rec (true≠false (sym p)) 432 | rem z (inr (dir p) , eq) = 433 | subst (λ - → [ - ≤ u ]) eq (∥∥-rec (isProp[] (_ ≤ _)) goal p) 434 | where 435 | goal : _ 436 | goal (j , false∈Wⱼ) = u-upper 𝟏 (j , nts) 437 | where 438 | Wⱼ=⊤ : W $ j ≡ ⊤[ 𝕊 ] 439 | Wⱼ=⊤ = sym (false∈𝔘→𝔘-top (W $ j) false∈Wⱼ) 440 | 441 | nts : from₀ x (W $ j) ≡ 𝟏 442 | nts = from₀ x (W $ j) ≡⟨ cong (from₀ x) Wⱼ=⊤ ⟩ 443 | from₀ x ⊤[ 𝕊 ] ≡⟨ resp-⊤ x ⟩ 444 | 𝟏 ∎ 445 | rem z (inl (squash p q i) , eq) = 446 | isProp[] (_ ≤ _) (rem z (inl p , eq)) (rem z (inl q , eq)) i 447 | rem z (inr (squash p q i) , eq) = 448 | isProp[] (_ ≤ _) (rem z (inr p , eq)) (rem z (inr q , eq)) i 449 | ``` 450 | 451 | The combination of these two proofs give us the fact that `from₀` is a frame 452 | homomorphism (i.e. a continuous function): 453 | 454 | ```agda 455 | from₀-cont : (x : ∣ A ∣F) → isFrameHomomorphism 𝕊 A (from₀m x) 456 | from₀-cont x = resp-⊤ x , from₀-resp-∧ x , from₀-comm-⋁ x 457 | ``` 458 | 459 | We are now ready to write down `from`: 460 | 461 | ```agda 462 | from : ∣ A ∣F → 𝕊 ─f→ A 463 | π₀ (π₀ (from x)) = from₀ x 464 | π₁ (π₀ (from x)) = from₀-mono x 465 | π₁ (from x) = from₀-cont x 466 | ``` 467 | 468 | #### `to` cancels `from` 469 | 470 | ```agda 471 | to∘from=id : (x : ∣ A ∣F) → to (from x) ≡ x 472 | to∘from=id x = to (from x) ≡⟨ refl ⟩ from₀ x ⁅true⁆ ≡⟨ from-lemma₀ x ⟩ x ∎ 473 | ``` 474 | 475 | #### `from` cancels `to` 476 | 477 | To prove this result, we will make use of the following, rather important 478 | lemma: 479 | 480 | ```agda 481 | useful-lemma : (((f , _) , _) : 𝕊 ─f→ A) (𝔘 : ∣ 𝕊 ∣F) 482 | → ⋁ ⁅ f (η u) ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ ≡ f 𝔘 483 | useful-lemma 𝒻@((f , _) , _ , _ , f-resp-⋁) 𝔘 = 484 | ⋁ ⁅ f (η u) ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ ≡⟨ sym (f-resp-⋁ (⁅ η u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆)) ⟩ 485 | f (⋁[ 𝕊 ] ⁅ η u ∣ u ∈ ⦅ 𝔘 ⦆ ⁆) ≡⟨ sym (cong f (main-lemma (S 𝓤₀ 𝓤₀) 𝔘)) ⟩ 486 | f 𝔘 ∎ 487 | ``` 488 | 489 | We now prove that `from` cancels `to`: 490 | 491 | ```agda 492 | from∘to=id : (𝒻 : 𝕊 ─f→ A) → from (to 𝒻) ≡ 𝒻 493 | from∘to=id 𝒻@((f , f-mono) , f-resp-⊤ , _) = 494 | forget-homo 𝕊 A (from (to 𝒻)) 𝒻 goal 495 | where 496 | goal : (𝔘 : ∣ 𝕊 ∣F) → from (to 𝒻) .π₀ .π₀ 𝔘 ≡ f 𝔘 497 | goal 𝔘 = sym (⋁-unique A _ _ ub least) 498 | where 499 | open PosetReasoning (pos A) 500 | 501 | ub : (x : ∣ A ∣F) → x ε 𝔨 (to 𝒻) 𝔘 → [ x ≤ (f 𝔘) ] 502 | ub x (inl i , eq) = subst (λ - → [ - ≤ f 𝔘 ]) eq nts 503 | where 504 | ⦅𝟏⦆ : [ f ⁅true⁆ ≤ f 𝔘 ] 505 | ⦅𝟏⦆ = f-mono _ _ (⊤-lemma 𝔘 i) 506 | 507 | nts : [ (𝔨 (f ⁅true⁆) 𝔘 $ inl i) ≤ f 𝔘 ] 508 | nts = ⁅ f ⁅true⁆ ∣ _ ∶ [ true ∈ ⦅ 𝔘 ⦆ ] ⁆ $ i ⊑⟨ ≡⇒⊑ (pos A) refl ⟩ 509 | f ⁅true⁆ ⊑⟨ ⦅𝟏⦆ ⟩ 510 | f 𝔘 ■ 511 | ub x (inr j , eq) = subst (λ - → [ - ≤ f 𝔘 ]) eq (≡⇒⊑ (pos A) nts) 512 | where 513 | nts : 𝔨 (to 𝒻) 𝔘 $ inr j ≡ f 𝔘 514 | nts = 𝔨 (to 𝒻) 𝔘 $ inr j ≡⟨ refl ⟩ 515 | ⊤[ A ] ≡⟨ sym f-resp-⊤ ⟩ 516 | f ⊤[ 𝕊 ] ≡⟨ cong f (false∈𝔘→𝔘-top 𝔘 j) ⟩ 517 | f 𝔘 ∎ 518 | 519 | least : (u : ∣ A ∣F) 520 | → ((x : ∣ A ∣F) → x ε 𝔨 (to 𝒻) 𝔘 → [ x ≤ u ]) → [ f 𝔘 ≤ u ] 521 | least u p = 522 | subst (λ - → [ - ≤ u ]) (useful-lemma 𝒻 𝔘) (⋁[ A ]-least _ _ rem) 523 | where 524 | π : [ false ∈ ⦅ 𝔘 ⦆ ] → [ ⊤[ A ] ≤ u ] 525 | π q = p ⊤[ A ] (inr q , refl) 526 | 527 | ρ : [ true ∈ ⦅ 𝔘 ⦆ ] → [ f ⁅true⁆ ≤ u ] 528 | ρ q = p (f ⁅true⁆) (inl q , refl) 529 | 530 | rem : [ ∀[ z ε ⁅ f (η u) ∣ u ∈ ⦅ 𝔘 ⦆ ⁆ ] (z ≤ u) ] 531 | rem z ((true , q) , eq) = subst (λ - → [ - ≤ u ]) eq nts 532 | where 533 | nts = f (η true) ⊑⟨ ≡⇒⊑ (pos A) (cong f (sym ⁅true⁆=η-true)) ⟩ 534 | f ⁅true⁆ ⊑⟨ ρ q ⟩ 535 | u ■ 536 | rem z ((false , q) , eq) = subst (λ - → [ - ≤ u ]) eq nts 537 | where 538 | nts = f (η false) ⊑⟨ ≡⇒⊑ (pos A) (cong f (sym 𝟏=η-false)) ⟩ 539 | f ⊤[ 𝕊 ] ⊑⟨ ≡⇒⊑ (pos A) f-resp-⊤ ⟩ 540 | ⊤[ A ] ⊑⟨ π q ⟩ 541 | u ■ 542 | ``` 543 | 544 | Finally, we write down the desired equivalence: 545 | 546 | ```agda 547 | 𝕊-correct : (𝕊 ─f→ A) ≃ ∣ A ∣F 548 | 𝕊-correct = isoToEquiv (iso to from to∘from=id from∘to=id) 549 | ``` 550 | 551 | #### The proof 552 | 553 | ```agda 554 | sierpiński-exists = S 𝓤₀ 𝓤₀ , 𝕊-correct 555 | ``` 556 | --------------------------------------------------------------------------------