├── cabal.project ├── .lycheeignore ├── .envrc ├── .gitignore ├── src-1lab ├── src-1lab.agda-lib ├── Lawvere.agda ├── Möbius.agda ├── Probability.agda ├── MonoidalFibres.agda ├── SplitMonoSet.agda ├── Madeleine.agda ├── Untruncate.agda ├── Mystery.agda ├── PresheafExponential.agda ├── YonedaColimit.agda ├── PointwiseMonoidal.agda ├── AdjunctionCommaIso.agda ├── Goat.agda ├── CoherentlyConstant.agda ├── DoubleCoverOfTheCircle.lagda.md ├── Applicative.agda ├── PostcomposeNotFull.agda ├── ObjectClassifiersOld.lagda.md ├── Skeletons.lagda.md ├── Hats.agda ├── EasyParametricity.lagda.md ├── SetsCover.lagda.md ├── TangentBundlesOfSpheres.lagda.md ├── SyntheticCategoricalDuality.lagda.md ├── ErasureOpen.lagda.md └── ObjectClassifiers.lagda.md ├── _typos.toml ├── src ├── src.agda-lib ├── DeMorKan.agda ├── NaiveFunext.agda ├── Torus.agda ├── NatChurchMonoid.agda ├── Shapes.agda └── Erasure.agda ├── .github └── workflows │ ├── check-links.yml │ └── web.yaml ├── main.js ├── index.html ├── module.html ├── shake ├── cubical-experiments-shake.cabal ├── LICENSE.agda ├── HTML │ ├── Backend.hs │ └── Base.hs └── Main.hs ├── diagram.template.tex ├── flake.nix ├── flake.lock └── style.css /cabal.project: -------------------------------------------------------------------------------- 1 | packages: shake 2 | -------------------------------------------------------------------------------- /.lycheeignore: -------------------------------------------------------------------------------- 1 | ^https?://ncatlab\.org 2 | -------------------------------------------------------------------------------- /.envrc: -------------------------------------------------------------------------------- 1 | use flake 2 | # use flake .#shakefile 3 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | result 2 | result-* 3 | *.agdai 4 | MAlonzo/** 5 | .shake 6 | dist-newstyle 7 | _build 8 | Everything*.agda 9 | -------------------------------------------------------------------------------- /src-1lab/src-1lab.agda-lib: -------------------------------------------------------------------------------- 1 | name: agda-stuff 2 | include: . 3 | depend: 4 | 1lab 5 | flags: 6 | --cubical 7 | --no-load-primitives 8 | --postfix-projections 9 | --allow-unsolved-metas 10 | --rewriting 11 | --guardedness 12 | --erasure 13 | -W noInteractionMetaBoundaries 14 | -W noUnsupportedIndexedMatch 15 | -------------------------------------------------------------------------------- /_typos.toml: -------------------------------------------------------------------------------- 1 | [files] 2 | extend-exclude = ["shake", "*.css", "*.bibtex"] 3 | 4 | [default.extend-words] 5 | identicals = "identicals" 6 | intensional = "intensional" 7 | operad = "operad" 8 | projectives = "projectives" 9 | raison = "raison" 10 | substituters = "substituters" 11 | 12 | # parts of identifiers 13 | hom = "hom" 14 | mor = "mor" 15 | tne = "tne" 16 | -------------------------------------------------------------------------------- /src/src.agda-lib: -------------------------------------------------------------------------------- 1 | name: agda-stuff 2 | include: . 3 | depend: 4 | standard-library 5 | cubical 6 | flags: 7 | --cubical 8 | --no-import-sorts 9 | --postfix-projections 10 | --hidden-argument-puns 11 | --allow-unsolved-metas 12 | --rewriting 13 | --guardedness 14 | --erasure 15 | -W noInteractionMetaBoundaries 16 | -W noUnsupportedIndexedMatch 17 | -------------------------------------------------------------------------------- /.github/workflows/check-links.yml: -------------------------------------------------------------------------------- 1 | name: Check links 2 | 3 | on: 4 | push: 5 | repository_dispatch: 6 | workflow_dispatch: 7 | schedule: 8 | - cron: "28 0 * * 0" 9 | 10 | jobs: 11 | check-links: 12 | runs-on: ubuntu-latest 13 | steps: 14 | - uses: actions/checkout@v4 15 | - name: Link Checker 16 | uses: lycheeverse/lychee-action@v2 17 | with: 18 | args: --insecure --verbose --no-progress 'src/**/*.md' 'src/**/*.html' 'src-1lab/**/*.md' 'src-1lab/**/*.html' 'index.html' 19 | -------------------------------------------------------------------------------- /main.js: -------------------------------------------------------------------------------- 1 | function scrollToHash() { 2 | if (window.location.hash === '') return; 3 | 4 | const id = decodeURI(window.location.hash.slice(1)); 5 | 6 | // #id doesn't work with numerical IDs 7 | const elem = document.querySelector(`[id="${id}"]`); 8 | if (!elem) return; 9 | 10 | // If the element is in a
tag, open it and scroll to it. 11 | const details = elem.closest('details'); 12 | if (details) { 13 | details.setAttribute("open", ""); 14 | elem.scrollIntoView(); 15 | } 16 | } 17 | 18 | window.addEventListener("DOMContentLoaded", scrollToHash); 19 | window.addEventListener("hashchange", scrollToHash); 20 | -------------------------------------------------------------------------------- /index.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | 6 | 7 | agda-stuff 8 | 9 | 10 | 11 | 12 | 13 | 14 |

15 | index ∙ 16 | source 17 |

18 | Naïm Camille Favier's Agda blog/lab/playground. 19 |
20 | @contents@
21 | 22 | 23 | -------------------------------------------------------------------------------- /src-1lab/Lawvere.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | open import Data.Power 3 | 4 | -- Lawvere's fixed point theorem and related results. 5 | module Lawvere where 6 | 7 | private variable 8 | ℓ : Level 9 | A B : Type ℓ 10 | 11 | module _ (f : A → A → B) (g : B → B) where 12 | δ : A → B 13 | δ a = g (f a a) 14 | 15 | lawvere-worker : fibre f δ → Σ[ b ∈ B ] g b ≡ b 16 | lawvere-worker (a , p) = f a a , sym (p $ₚ a) 17 | 18 | module _ (f : A → A → B) (f-surj : is-surjective f) where 19 | lawvere : (g : B → B) → ∃[ b ∈ B ] g b ≡ b 20 | lawvere g = lawvere-worker f g <$> f-surj _ 21 | 22 | Curry : (A ≃ (A → B)) → B 23 | Curry e = lawvere-worker (e .fst) id (equiv-centre e _) .fst 24 | 25 | module _ (f : A → ℙ A) (f-surj : is-surjective f) where 26 | cantor : ⊥ 27 | cantor = ∥-∥-out! do 28 | A , fixed ← lawvere f f-surj ¬Ω_ 29 | pure (Curry (path→equiv (ap ⌞_⌟ (sym fixed)))) 30 | -------------------------------------------------------------------------------- /src-1lab/Möbius.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | 3 | open import Data.Int renaming (Int to ℤ) 4 | 5 | open import Homotopy.Space.Circle 6 | 7 | -- https://math.stackexchange.com/questions/4940313/giving-calculating-explicit-homomorphism-between-fundamental-groups 8 | module Möbius where 9 | 10 | data Möbius : Type where 11 | up down : Möbius 12 | seam : up ≡ down 13 | top : up ≡ down 14 | bottom : down ≡ up 15 | surf : PathP (λ i → top i ≡ bottom i) seam (sym seam) 16 | 17 | Cover : Möbius → Type 18 | Cover up = ℤ 19 | Cover down = ℤ 20 | Cover (seam i) = ℤ 21 | Cover (top i) = ua suc-equiv i 22 | Cover (bottom i) = ua suc-equiv i 23 | Cover (surf i j) = ua suc-equiv i 24 | 25 | decode : ∀ {x} → up ≡ x → Cover x 26 | decode p = subst Cover p 0 27 | 28 | ι : S¹ → Möbius 29 | ι = S¹-rec up (top ∙ bottom) 30 | 31 | ι* : ℤ → ℤ 32 | ι* = decode ∘ ap ι ∘ loopⁿ 33 | 34 | _ : ι* 1 ≡ 2 35 | _ = refl 36 | -------------------------------------------------------------------------------- /src-1lab/Probability.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | 3 | open import Data.Dec 4 | open import Data.Fin 5 | open import Data.Fin.Closure 6 | 7 | module Probability where 8 | 9 | -- “I have two children, (at least) one of whom is a boy born on a Tuesday - 10 | -- what is the probability that both children are boys?” 11 | 12 | -- Simplifying assumptions: gender is binary; gender and day of birth are 13 | -- uniformly distributed. 14 | Gender = Fin 2 15 | Day = Fin 7 16 | Child = Fin 2 17 | 18 | Sample = Child → Day × Gender 19 | 20 | count : (A : Sample → Type) → ⦃ ∀ {s} → Finite (A s) ⦄ → Nat 21 | count A = cardinality {A = Σ Sample A} 22 | 23 | condition : Sample → Type 24 | condition s = ∃[ i ∈ Child ] s i ≡ (1 , 1) 25 | 26 | event : Sample → Type 27 | event s = (i : Child) → s i .snd ≡ 1 28 | 29 | answer : Nat × Nat -- formal fraction 30 | answer = count (λ s → event s × condition s) , count condition 31 | 32 | _ : answer ≡ (13 , 27) 33 | _ = refl 34 | -------------------------------------------------------------------------------- /src-1lab/MonoidalFibres.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | open import Homotopy.Connectedness 3 | open import Cat.Prelude 4 | open import Cat.Functor.Base 5 | open import Cat.Functor.Properties 6 | 7 | -- eso + full-on-isos functors have monoidal fibres. 8 | module MonoidalFibres where 9 | 10 | private variable 11 | o ℓ : Level 12 | C D : Precategory o ℓ 13 | 14 | monoidal-fibres 15 | : ∀ {F : Functor C D} 16 | → is-category C → is-category D 17 | → is-eso F → is-full-on-isos F 18 | → ∀ y → is-connected (Essential-fibre F y) 19 | monoidal-fibres {D = D} {F = F} ccat dcat eso full≅ y = 20 | case eso y of λ y′ Fy′≅y → is-connected∙→is-connected {X = _ , y′ , Fy′≅y} λ (x , Fx≅y) → do 21 | (x≅y′ , eq) ← full≅ (Fy′≅y Iso⁻¹ ∘Iso Fx≅y) 22 | pure (Σ-pathp (ccat .to-path x≅y′) 23 | (≅-pathp _ _ (transport (λ i → PathP (λ j → Hom (F-map-path F ccat dcat x≅y′ (~ i) j) y) (Fx≅y .to) (Fy′≅y .to)) 24 | (Hom-pathp-refll-iso (sym (ap from (Iso-swapl (sym eq)))))))) 25 | where open Univalent dcat 26 | -------------------------------------------------------------------------------- /module.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | 6 | 7 | @moduleName@ 8 | 9 | 10 | 11 | 12 | 13 | 14 | 15 | 16 | 17 |

18 | index ∙ 19 | source 20 |

21 | @contents@ 22 | 23 | 24 | -------------------------------------------------------------------------------- /src-1lab/SplitMonoSet.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Prelude 2 | import Cat.Reasoning 3 | open import 1Lab.Classical 4 | open import Data.Dec 5 | open import Data.Sum 6 | 7 | -- If every split monomorphism with inhabited domain splits in Sets then excluded middle holds. 8 | module SplitMonoSet where 9 | 10 | module Sets = Cat.Reasoning (Sets lzero) 11 | 12 | module _ (every-mono-splits : ∀ {A B} (a : ∥ ⌞ A ⌟ ∥) (f : Sets.Hom A B) (f-mono : Sets.is-monic {a = A} {B} f) → Sets.is-split-monic {a = A} {B} f) where 13 | lem : LEM 14 | lem P = ∥-∥-out! do 15 | Sets.make-retract f⁻¹ ret ← f-split 16 | pure (go (f⁻¹ (inr tt)) λ p → ret $ₚ inr p) 17 | where 18 | A : Set lzero 19 | A = el! (⊤ ⊎ ⌞ P ⌟) 20 | B : Set lzero 21 | B = el! (⊤ ⊎ ⊤) 22 | f : Sets.Hom A B 23 | f = ⊎-map _ _ 24 | f-mono : Sets.is-monic {a = A} {B} f 25 | f-mono = embedding→monic {f = f} $ injective→is-embedding! λ where 26 | {inl _} {inl _} p → refl 27 | {inl _} {inr _} p → absurd (inl≠inr p) 28 | {inr _} {inl _} p → absurd (inr≠inl p) 29 | {inr _} {inr _} p → ap inr prop! 30 | f-split : Sets.is-split-monic {a = A} {B} f 31 | f-split = every-mono-splits (inc (inl _)) f λ {c} → f-mono {c} 32 | go : (f⁻¹r : ∣ A ∣) → (∀ p → f⁻¹r ≡ inr p) → Dec ∣ P ∣ 33 | go (inl _) l = no λ p → inl≠inr (l p) 34 | go (inr p) l = yes p 35 | -------------------------------------------------------------------------------- /src-1lab/Madeleine.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | open import 1Lab.Classical 3 | open import Data.Sum 4 | open import Data.Dec 5 | open import Data.Bool 6 | open import Meta.Invariant 7 | 8 | module Madeleine where 9 | 10 | axiom = ∀ {ℓ} {P Q : Type ℓ} ⦃ _ : H-Level P 1 ⦄ ⦃ _ : H-Level Q 1 ⦄ → ∥ P ⊎ Q ∥ → P ⊎ Q 11 | 12 | lem→axiom : LEM → axiom 13 | lem→axiom lem {P = P} {Q} pq with lem (elΩ P) 14 | ... | yes a = inl (□-out! a) 15 | ... | no ¬a = inr (rec! (λ { (inl p) → absurd (¬a (inc p)); (inr q) → q }) pq) 16 | 17 | module _ (ε : axiom) where 18 | 19 | module _ {ℓ} {X : Type ℓ} ⦃ _ : H-Level X 2 ⦄ (a₀ a₁ : X) where 20 | E : X → Type _ 21 | E x = (a₀ ≡ x) ⊎ (a₁ ≡ x) 22 | 23 | E' : Type _ 24 | E' = Σ X λ x → ∥ E x ∥ 25 | 26 | r : Bool → E' 27 | r true = a₀ , inc (inl refl) 28 | r false = a₁ , inc (inr refl) 29 | 30 | s : E' → Bool 31 | s (x , e) with ε e 32 | ... | inl _ = true 33 | ... | inr _ = false 34 | 35 | r-s : ∀ e → r (s e) ≡ e 36 | r-s (x , e) with ε e 37 | ... | inl p = Σ-prop-path! p 38 | ... | inr p = Σ-prop-path! p 39 | 40 | discrete : Dec (a₀ ≡ a₁) 41 | discrete = invmap 42 | (λ p → ap fst (right-inverse→injective r r-s p)) 43 | (λ p → ap s (Σ-prop-path! p)) 44 | (s (r true) ≡? s (r false)) 45 | 46 | lem : LEM 47 | lem P = invmap (λ p → subst ∣_∣ (sym p) _) (λ p → Ω-ua _ (λ _ → p)) 48 | (discrete P ⊤Ω) 49 | -------------------------------------------------------------------------------- /src-1lab/Untruncate.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | 3 | -- The identity function on homogeneous types "factors" through the propositional truncation 4 | -- (https://homotopytypetheory.org/2013/10/28/the-truncation-map-_-%E2%84%95-%E2%80%96%E2%84%95%E2%80%96-is-nearly-invertible) 5 | module Untruncate where 6 | 7 | point : ∀ {ℓ} (X : Type ℓ) → X → Type∙ ℓ 8 | point X x = X , x 9 | 10 | is-homogeneous : ∀ {ℓ} → Type ℓ → Type (lsuc ℓ) 11 | is-homogeneous X = ∀ x y → point X x ≡ point X y 12 | 13 | ∥-∥-rec-const 14 | : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} 15 | → (f : A → B) 16 | → (b : B) 17 | → (∀ x → b ≡ f x) 18 | → ∥ A ∥ → B 19 | ∥-∥-rec-const {A = A} {B} f b f-const x = 20 | ∥-∥-elim {P = λ _ → Singleton b} (λ _ → is-contr→is-prop Singleton-is-contr) 21 | (λ x → f x , f-const x) x .fst 22 | 23 | module old {ℓ} (X : Type ℓ) (x : X) (hom : is-homogeneous X) where 24 | point' : ∥ X ∥ → Type∙ ℓ 25 | point' = ∥-∥-rec-const (point X) (point X x) (hom x) 26 | 27 | myst : (x : ∥ X ∥) → point' x .fst 28 | myst x = point' x .snd 29 | 30 | _ : myst ∘ inc ≡ id 31 | _ = refl 32 | 33 | -- Simplification by David Wärn https://gist.github.com/dwarn/31d7002a5ca8df0443b31501056e357f 34 | module new {ℓ : Level} {X : Type ℓ} where 35 | fam : ∥ X ∥ → n-Type ℓ 0 36 | fam = rec! λ x → el! (Singleton x) 37 | 38 | magic : X → X 39 | magic = fst ∘ centre ∘ is-tr ∘ fam ∘ inc 40 | 41 | _ : magic ≡ id 42 | _ = refl 43 | -------------------------------------------------------------------------------- /.github/workflows/web.yaml: -------------------------------------------------------------------------------- 1 | name: web 2 | on: [push, pull_request, workflow_dispatch] 3 | jobs: 4 | build: 5 | runs-on: ubuntu-latest 6 | steps: 7 | - uses: actions/checkout@v4 8 | - uses: cachix/install-nix-action@v27 9 | with: 10 | extra_nix_config: | 11 | access-tokens = github.com=${{ secrets.GITHUB_TOKEN }} 12 | extra-substituters = https://nix.monade.li https://1lab.cachix.org 13 | extra-trusted-public-keys = nix.monade.li:2Zgy59ai/edDBizXByHMqiGgaHlE04G6Nzuhx1RPFgo= 1lab.cachix.org-1:eYjd9F9RfibulS4OSFBYeaTMxWojPYLyMqgJHDvG1fs= 14 | - uses: cachix/cachix-action@v15 15 | with: 16 | name: ncfavier 17 | authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} 18 | extraPullNames: 1lab 19 | - name: Build 20 | run: | 21 | web=$(nix -L build --print-out-paths) 22 | cp -rL --no-preserve=mode,ownership,timestamps "$web" pages 23 | - uses: actions/upload-pages-artifact@v3 24 | with: 25 | path: pages 26 | retention-days: 1 27 | deploy: 28 | if: github.ref_name == 'main' 29 | needs: build 30 | permissions: 31 | pages: write 32 | id-token: write 33 | environment: 34 | name: github-pages 35 | url: ${{ steps.deployment.outputs.page_url }} 36 | runs-on: ubuntu-latest 37 | steps: 38 | - id: deployment 39 | uses: actions/deploy-pages@v4 40 | -------------------------------------------------------------------------------- /src-1lab/Mystery.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Classical 2 | open import 1Lab.Prelude 3 | 4 | open import Data.Dec 5 | 6 | -- A weird constructive principle equivalent to Markov's principle + WLEM + ¬¬-shift 7 | -- https://types.pl/@ncf/112854858302377592 8 | -- https://x.com/ncfavier/status/1817846740944314834 9 | module Mystery where 10 | 11 | TNE : ∀ {ℓ} {P : Type ℓ} → ¬ ¬ ¬ P → ¬ P 12 | TNE h p = h λ k → k p 13 | 14 | case01 : ∀ {ℓ} {A : Type ℓ} → A → A → Nat → A 15 | case01 z s zero = z 16 | case01 z s (suc n) = s 17 | 18 | mystery MP DNS : Type 19 | 20 | mystery = (P : Nat → Ω) → (¬ ∀ n → ∣ P n ∣) → ∃ _ λ n → ¬ ∣ P n ∣ 21 | 22 | MP = (P : Nat → Ω) → (∀ n → Dec ∣ P n ∣) → (¬ ∀ n → ∣ P n ∣) → ∃ _ λ n → ¬ ∣ P n ∣ 23 | 24 | DNS = (P : Nat → Ω) → (∀ n → ¬ ¬ ∣ P n ∣) → ¬ ¬ ∀ n → ∣ P n ∣ 25 | 26 | mystery→MP : mystery → MP 27 | mystery→MP m P _ = m P 28 | 29 | mystery→WLEM : mystery → WLEM 30 | mystery→WLEM m P = case m (case01 P (¬Ω P)) (λ h → h 1 (h 0)) of λ where 31 | zero p → yes p 32 | (suc _) p → no p 33 | 34 | mystery→DNS : mystery → DNS 35 | mystery→DNS m P h k = case m P k of h 36 | 37 | MP+WLEM+DNS→mystery : MP × WLEM × DNS → mystery 38 | MP+WLEM+DNS→mystery (mp , wlem , dns) P h = 39 | mp (λ n → ¬Ω ¬Ω P n) (λ n → wlem (¬Ω P n)) (λ k → dns P k h) 40 | <&> Σ-map id TNE 41 | 42 | mystery≃MP+WLEM+DNS : mystery ≃ (MP × WLEM × DNS) 43 | mystery≃MP+WLEM+DNS = prop-ext! 44 | ⟨ mystery→MP , ⟨ mystery→WLEM , mystery→DNS ⟩ ⟩ 45 | MP+WLEM+DNS→mystery 46 | -------------------------------------------------------------------------------- /src/DeMorKan.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Prelude 2 | 3 | -- A silly attempt at implementing composition for the interval, 4 | -- for https://proofassistants.stackexchange.com/questions/2043/is-the-de-morgan-interval-kan 5 | module DeMorKan where 6 | 7 | -- The built-in I lives in its own "non-fibrant" universe, so Agda won't let 8 | -- us express partial elements and subtypes. 9 | -- Hence we define a "wrapper" HIT, but do not make use of its Kan structure! 10 | data Interval : Type where 11 | i0' : Interval 12 | i1' : Interval 13 | inI : i0' ≡ i1' 14 | 15 | module 2D (i j : I) where 16 | ⊔ : I 17 | ⊔ = ~ j ∨ i ∨ ~ i 18 | B L R : I → I 19 | B i = i ∨ ~ i -- 20 | L j = i1 -- replace these with anything (as long as they agree on endpoints) 21 | R j = ~ j -- 22 | horn : Partial ⊔ Interval 23 | horn (j = i0) = inI (B i) 24 | horn (i = i0) = inI (L j) 25 | horn (i = i1) = inI (R j) 26 | filler : Interval [ ⊔ ↦ horn ] 27 | filler = inS (inI ((~ j ∧ B i) ∨ (~ i ∧ L j) ∨ (i ∧ R j))) 28 | 29 | module 3D (i j k : I) where 30 | ⊔ : I 31 | ⊔ = ~ k ∨ ~ i ∨ i ∨ ~ j ∨ j 32 | B L R D U : I → I → I 33 | B i j = i0 34 | L j k = i0 35 | R j k = i0 36 | D i k = i0 37 | U i k = i0 38 | horn : Partial ⊔ Interval 39 | horn (k = i0) = inI (B i j) 40 | horn (i = i0) = inI (L j k) 41 | horn (i = i1) = inI (R j k) 42 | horn (j = i0) = inI (D i k) 43 | horn (j = i1) = inI (U i k) 44 | filler : Interval [ ⊔ ↦ horn ] 45 | filler = inS (inI ((~ k ∧ B i j) ∨ (~ i ∧ L j k) ∨ (i ∧ R j k) ∨ (~ j ∧ D i k) ∨ (j ∧ U i k))) 46 | -------------------------------------------------------------------------------- /src-1lab/PresheafExponential.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Prelude 2 | open import Cat.Functor.Base 3 | open import Cat.Functor.Naturality 4 | open import Cat.Instances.Presheaf.Limits 5 | open import Cat.Instances.Presheaf.Exponentials 6 | open import Cat.Diagram.Exponential 7 | open import Cat.Diagram.Product 8 | import Cat.Reasoning 9 | 10 | module PresheafExponential {ℓ} {C : Precategory ℓ ℓ} where 11 | 12 | module C = Cat.Reasoning C 13 | module PSh = Cat.Reasoning (PSh ℓ C) 14 | open Binary-products (PSh ℓ C) (PSh-products _ C) 15 | open Cartesian-closed (PSh-closed C) 16 | 17 | open Functor 18 | open _=>_ 19 | 20 | module _ b a where 21 | open Exponential (has-exp a b) 22 | renaming (B^A to infixr 50 _^_) 23 | using () public 24 | 25 | module _ (K L M : PSh.Ob) where 26 | 27 | internal-currying : M ^ (K ⊗₀ L) PSh.≅ (M ^ K) ^ L 28 | internal-currying = PSh.make-iso 29 | (λ where 30 | .η n f .η q (v , y) .η p (u , x) → f .η p (v C.∘ u , x , L .F₁ u y) 31 | .η n f .η q (v , y) .is-natural → {! !} 32 | .η n f .is-natural → {! !} 33 | .is-natural → {! !}) 34 | (λ where 35 | .η n g .η q (v , x , y) → g .η q (v , y) .η q (C.id , x) 36 | .η n g .is-natural → {! !} 37 | .is-natural → {! !}) 38 | (ext λ n g q v y p u x → 39 | ⌜ g .η p (v C.∘ u , L .F₁ u y) ⌝ .η p (C.id , x) ≡⟨ g .is-natural _ _ u $ₚ (v , y) ηₚ p $ₚ (C.id , x) ⟩ 40 | (M ^ K) .F₁ u (g .η q (v , y)) .η p (C.id , x) ≡⟨⟩ 41 | g .η q (v , y) .η p (⌜ u C.∘ C.id ⌝ , x) ≡⟨ ap! (C.idr u) ⟩ 42 | g .η q (v , y) .η p (u , x) ∎) 43 | {! !} 44 | -------------------------------------------------------------------------------- /src/NaiveFunext.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | open import Agda.Primitive renaming (Set to Type) 3 | open import Data.Product 4 | open import Data.Product.Properties 5 | open import Relation.Binary.PropositionalEquality 6 | 7 | -- Naïve function extensionality implies function extensionality (HoTT book exercise 4.9). 8 | -- This is actually weaker as we assume ~ → ≡ for *dependent* functions. 9 | module NaiveFunext where 10 | 11 | private variable 12 | ℓ ℓ' : Level 13 | A : Type ℓ 14 | B : A → Type ℓ 15 | f g : (a : A) → B a 16 | 17 | _~_ : (f g : (a : A) → B a) → Type _ 18 | f ~ g = ∀ a → f a ≡ g a 19 | 20 | happly : f ≡ g → f ~ g 21 | happly {f = f} p = subst (λ x → f ~ x) p λ _ → refl 22 | 23 | cong-proj₂ : ∀ {a b} {c : B a} {d : B b} → (p : (a , c) ≡ (b , d)) → subst B (cong proj₁ p) c ≡ d 24 | cong-proj₂ {c = c} refl = refl 25 | 26 | singleton-is-contr : ∀ {a : A} {s : Σ A (a ≡_)} → (a , refl) ≡ s 27 | singleton-is-contr {s = _ , refl} = refl 28 | 29 | module _ 30 | (ext : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {f g : (a : A) → B a} → f ~ g → f ≡ g) 31 | where 32 | 33 | module _ (f : (a : A) → B a) where 34 | from : ((a : A) → Σ _ (f a ≡_)) → Σ _ (f ~_) 35 | from p = (λ a → p a .proj₁) , (λ a → p a .proj₂) 36 | 37 | to : Σ _ (f ~_) → ((a : A) → Σ _ (f a ≡_)) 38 | to g = λ a → g .proj₁ a , g .proj₂ a 39 | 40 | -- Homotopies form an identity system, which is equivalent to function extensionality. 41 | htpy-is-contr : (g : Σ _ (f ~_)) → (f , λ _ → refl) ≡ g 42 | htpy-is-contr g = cong from p 43 | where 44 | p : (λ a → f a , refl) ≡ to g 45 | p = ext λ _ → singleton-is-contr 46 | -------------------------------------------------------------------------------- /src-1lab/YonedaColimit.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Prelude 2 | open import Cat.Functor.Hom 3 | open import Cat.Functor.Base 4 | open import Cat.Functor.Constant 5 | open import Cat.Diagram.Colimit.Base 6 | open import Cat.Diagram.Limit.Base 7 | open import Cat.Diagram.Terminal 8 | open import Cat.Diagram.Initial 9 | open import Cat.Instances.Presheaf.Limits 10 | open import Cat.Instances.Presheaf.Exponentials 11 | 12 | import Cat.Reasoning 13 | 14 | open _=>_ 15 | open make-is-colimit 16 | 17 | module YonedaColimit {o ℓ} (C : Precategory o ℓ) where 18 | 19 | open Cat.Reasoning C 20 | 21 | Δ1 : Terminal (PSh ℓ C) 22 | Δ1 = PSh-terminal _ C 23 | 24 | open Terminal Δ1 25 | 26 | よ-colimit : Colimit (よ C) 27 | よ-colimit = to-colimit (to-is-colimit colim) where 28 | colim : make-is-colimit (よ C) top 29 | colim .ψ c = ! 30 | colim .commutes f = ext λ _ _ → refl 31 | colim .universal eta comm .η x _ = eta x .η x id 32 | colim .universal eta comm .is-natural x y f = ext λ _ → 33 | sym (comm f ηₚ y $ₚ id) ∙∙ ap (eta x .η y) id-comm ∙∙ eta x .is-natural _ _ f $ₚ id 34 | colim .factors eta comm = ext λ x f → 35 | sym (comm f ηₚ x $ₚ id) ∙ ap (eta _ .η x) (idr _) 36 | colim .unique eta comm univ' fac' = ext λ x _ → fac' x ηₚ x $ₚ id 37 | 38 | Δ0 : Initial (PSh ℓ C) 39 | Δ0 = {! Const ? !} 40 | 41 | よ-limit : Limit (よ C) 42 | よ-limit = to-limit (to-is-limit lim) where 43 | lim : make-is-limit (よ C) (Const (el! (Lift _ ⊥))) 44 | lim .make-is-limit.ψ c .η x () 45 | lim .make-is-limit.ψ c .is-natural _ _ _ = ext λ () 46 | lim .make-is-limit.commutes f = ext λ _ () 47 | lim .make-is-limit.universal eps comm .η = {! !} 48 | lim .make-is-limit.universal eps comm .is-natural = {! !} 49 | lim .make-is-limit.factors = {! !} 50 | lim .make-is-limit.unique = {! !} 51 | -------------------------------------------------------------------------------- /src-1lab/PointwiseMonoidal.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Prelude 2 | open import Cat.Functor.Base 3 | open import Cat.Functor.Compose 4 | open import Cat.Functor.Constant 5 | open import Cat.Functor.Equivalence.Path 6 | open import Cat.Monoidal.Base 7 | open import Cat.Monoidal.Diagram.Monoid 8 | open import Cat.Instances.Product 9 | open import Cat.Displayed.Base 10 | open import Cat.Displayed.Total 11 | 12 | open Monoidal-category 13 | open Precategory 14 | open Functor 15 | open _=>_ 16 | 17 | -- ⚠️ WIP ⚠️ 18 | module PointwiseMonoidal 19 | {o o′ ℓ ℓ′} (C : Precategory o ℓ) (D : Precategory o′ ℓ′) 20 | (M : Monoidal-category D) 21 | where 22 | 23 | Pointwise : Monoidal-category Cat[ C , D ] 24 | Pointwise = pw where 25 | prod : Functor (Cat[ C , D ] ×ᶜ Cat[ C , D ]) Cat[ C , D ] 26 | prod .F₀ (a , b) = M .-⊗- F∘ Cat⟨ a , b ⟩ 27 | prod .F₁ {x = x} {y = y} (na , nb) = M .-⊗- ▸ nat where 28 | nat : Cat⟨ x .fst , x .snd ⟩ => Cat⟨ y .fst , y .snd ⟩ 29 | nat .η x = (na .η x) , (nb .η x) 30 | nat .is-natural x y f i = (na .is-natural x y f i) , (nb .is-natural x y f i) 31 | prod .F-id = ext λ _ → M .-⊗- .F-id 32 | prod .F-∘ f g = ext λ _ → M .-⊗- .F-∘ _ _ 33 | pw : Monoidal-category Cat[ C , D ] 34 | pw .-⊗- = prod 35 | pw .Unit = Const (M .Unit) 36 | pw .unitor-l = {! M .unitor-l !} 37 | pw .unitor-r = {! !} 38 | pw .associator = {! !} 39 | pw .triangle = {! !} 40 | pw .pentagon = {! !} 41 | 42 | MonCD→CMonD : Functor (∫ Mon[ Pointwise ]) (Cat[ C , ∫ Mon[ M ] ]) 43 | MonCD→CMonD .F₀ (F , mon) .F₀ c = F .F₀ c , {! !} 44 | MonCD→CMonD .F₀ (F , mon) .F₁ = {! !} 45 | MonCD→CMonD .F₀ (F , mon) .F-id = {! !} 46 | MonCD→CMonD .F₀ (F , mon) .F-∘ = {! !} 47 | MonCD→CMonD .F₁ = {! !} 48 | MonCD→CMonD .F-id = {! !} 49 | MonCD→CMonD .F-∘ = {! !} 50 | 51 | MonCD≡CMonD : ∫ Mon[ Pointwise ] ≡ Cat[ C , ∫ Mon[ M ] ] 52 | MonCD≡CMonD = Precategory-path MonCD→CMonD {! !} 53 | -------------------------------------------------------------------------------- /src/Torus.agda: -------------------------------------------------------------------------------- 1 | module Torus where 2 | 3 | open import Cubical.Foundations.Prelude 4 | open import Cubical.Foundations.Isomorphism 5 | open import Cubical.Foundations.Equiv 6 | open import Cubical.Foundations.GroupoidLaws 7 | open import Cubical.HITs.Torus 8 | 9 | private 10 | variable 11 | ℓ : Level 12 | A : Type ℓ 13 | 14 | -- 🍩 15 | data T² : Type where 16 | base : T² 17 | p q : base ≡ base 18 | surf : p ∙ q ≡ q ∙ p 19 | 20 | hcomp-inv : {φ : I} (u : I → Partial φ A) (u0 : A [ φ ↦ u i1 ]) 21 | → hcomp u (hcomp (λ k → u (~ k)) (outS u0)) ≡ outS u0 22 | hcomp-inv u u0 i = hcomp-equivFiller (λ k → u (~ k)) u0 (~ i) 23 | 24 | T²≃Torus : T² ≃ Torus 25 | T²≃Torus = isoToEquiv (iso to from to-from from-to) 26 | where 27 | sides : {a : A} (p1 p2 : a ≡ a) (i j k : I) → Partial (i ∨ ~ i ∨ j ∨ ~ j) A 28 | sides p1 p2 i j k (i = i0) = compPath-filler p2 p1 (~ k) j 29 | sides p1 p2 i j k (i = i1) = compPath-filler' p1 p2 (~ k) j 30 | sides p1 p2 i j k (j = i0) = p1 (i ∧ k) 31 | sides p1 p2 i j k (j = i1) = p1 (i ∨ ~ k) 32 | 33 | to : T² → Torus 34 | to base = point 35 | to (p i) = line1 i 36 | to (q j) = line2 j 37 | to (surf i j) = hcomp (λ k → sides line1 line2 (~ i) j (~ k)) (square (~ i) j) 38 | 39 | from : Torus → T² 40 | from point = base 41 | from (line1 i) = p i 42 | from (line2 j) = q j 43 | from (square i j) = hcomp (sides p q i j) (surf (~ i) j) 44 | 45 | to-from : ∀ x → to (from x) ≡ x 46 | to-from point = refl 47 | to-from (line1 i) = refl 48 | to-from (line2 i) = refl 49 | to-from (square i j) = hcomp-inv (sides line1 line2 i j) (inS (square i j)) 50 | 51 | from-to : ∀ x → from (to x) ≡ x 52 | from-to base = refl 53 | from-to (p i) = refl 54 | from-to (q i) = refl 55 | from-to (surf i j) = {! hcomp-inv (λ k → sides p q (~ i) j (~ k)) (inS (surf i j)) !} 56 | -- see https://github.com/agda/cubical/pull/912 for the full proof 57 | -------------------------------------------------------------------------------- /shake/cubical-experiments-shake.cabal: -------------------------------------------------------------------------------- 1 | cabal-version: 3.4 2 | name: agda-stuff-shake 3 | version: 0.1.0.0 4 | license: AGPL-3.0-or-later 5 | author: Naïm Favier 6 | maintainer: n@monade.li 7 | category: Development 8 | build-type: Simple 9 | 10 | executable agda-stuff-shake 11 | ghc-options: -Wall -Wno-name-shadowing 12 | main-is: Main.hs 13 | other-modules: HTML.Base, 14 | HTML.Backend, 15 | build-depends: base, 16 | Agda, 17 | shake, 18 | blaze-html, 19 | bytestring, 20 | containers, 21 | deepseq, 22 | directory, 23 | filepath, 24 | mtl, 25 | pandoc, 26 | pandoc-types, 27 | regex-tdfa, 28 | SHA, 29 | silently, 30 | split, 31 | tagsoup, 32 | text, 33 | transformers, 34 | uri-encode, 35 | hs-source-dirs: . 36 | default-language: GHC2021 37 | default-extensions: 38 | BangPatterns 39 | BlockArguments 40 | ConstraintKinds 41 | DefaultSignatures 42 | DeriveFoldable 43 | DeriveFunctor 44 | DeriveGeneric 45 | DeriveTraversable 46 | DerivingStrategies 47 | ExistentialQuantification 48 | FlexibleContexts 49 | FlexibleInstances 50 | FunctionalDependencies 51 | GADTs 52 | GeneralizedNewtypeDeriving 53 | InstanceSigs 54 | LambdaCase 55 | MultiParamTypeClasses 56 | MultiWayIf 57 | NamedFieldPuns 58 | OverloadedStrings 59 | PatternSynonyms 60 | RankNTypes 61 | RecordWildCards 62 | ScopedTypeVariables 63 | StandaloneDeriving 64 | TupleSections 65 | TypeFamilies 66 | TypeOperators 67 | TypeSynonymInstances 68 | ViewPatterns 69 | -------------------------------------------------------------------------------- /diagram.template.tex: -------------------------------------------------------------------------------- 1 | \documentclass[varwidth=\maxdimen,multi=page,12pt]{standalone} 2 | \usepackage{pgfplots} 3 | \usepackage{xcolor} 4 | \usepackage{tikz-cd} 5 | \usepackage{amssymb} 6 | \usepackage{amsmath} 7 | \usepackage{amsopn} 8 | \usepackage{mathpazo} 9 | \usepackage[T1]{fontenc} 10 | 11 | \usetikzlibrary{calc} 12 | \usetikzlibrary{fit} 13 | \usetikzlibrary{decorations.pathmorphing} 14 | 15 | 16 | \ExplSyntaxOn 17 | \NewDocumentCommand{\definealphabet}{mmmm} 18 | {% #1 = prefix, #2 = command, #3 = start, #4 = end 19 | \int_step_inline:nnn { `#3 } { `#4 } 20 | { 21 | \cs_new_protected:cpx { #1 \char_generate:nn { ##1 }{ 11 } } 22 | { 23 | \exp_not:N #2 { \char_generate:nn { ##1 } { 11 } } 24 | } 25 | } 26 | } 27 | 28 | \ExplSyntaxOff 29 | 30 | \definecolor{accent}{HTML}{c15cff} 31 | \definecolor{accent2}{HTML}{c59efd} 32 | 33 | \color{white} 34 | 35 | \tikzset{curve/.style={settings={#1},to path={(\tikztostart) 36 | .. controls ($(\tikztostart)!\pv{pos}!(\tikztotarget)!\pv{height}!270:(\tikztotarget)$) 37 | and ($(\tikztostart)!1-\pv{pos}!(\tikztotarget)!\pv{height}!270:(\tikztotarget)$) 38 | .. (\tikztotarget)\tikztonodes}}, 39 | settings/.code={\tikzset{quiver/.cd,#1} 40 | \def\pv##1{\pgfkeysvalueof{/tikz/quiver/##1}}}, 41 | quiver/.cd,pos/.initial=0.35,height/.initial=0} 42 | 43 | \tikzset{tail reversed/.code={\pgfsetarrowsstart{tikzcd to}}} 44 | \tikzset{2tail/.code={\pgfsetarrowsstart{Implies[reversed]}}} 45 | \tikzset{2tail reversed/.code={\pgfsetarrowsstart{Implies}}} 46 | \tikzset{no body/.style={/tikz/dash pattern=on 0 off 1mm}} 47 | \tikzset{ 48 | lies over/.style = { 49 | |-{Triangle[open, length=1.875mm, width=1.875mm]} 50 | }, 51 | category over/.style = { 52 | -{Triangle[open, length=1.875mm, width=1.875mm]} 53 | }, 54 | 3cell/.style={-,preaction={draw,Rightarrow}}, 55 | 2cell/.style={double=transparent,-Implies}, 56 | Rightarrow/.style={double equal sign distance, >={Implies}, ->}, 57 | commutative diagrams/background color=transparent 58 | } 59 | 60 | \begin{document} 61 | \begin{page} 62 | 63 | __BODY__ 64 | 65 | \end{page} 66 | \end{document} 67 | -------------------------------------------------------------------------------- /src-1lab/AdjunctionCommaIso.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Functor.Adjoint 2 | open import Cat.Functor.Equivalence 3 | open import Cat.Instances.Comma 4 | open import Cat.Prelude 5 | 6 | import Cat.Reasoning 7 | 8 | open ↓Hom 9 | open ↓Obj 10 | open Functor 11 | open is-iso 12 | open is-precat-iso 13 | 14 | -- An adjunction F ⊣ G induces an isomorphism of comma categories F ↓ 1 ≅ 1 ↓ G 15 | module AdjunctionCommaIso where 16 | 17 | module _ 18 | {oc ℓc od ℓd} {C : Precategory oc ℓc} {D : Precategory od ℓd} 19 | {F : Functor C D} {G : Functor D C} (F⊣G : F ⊣ G) 20 | where 21 | 22 | module C = Cat.Reasoning C 23 | module D = Cat.Reasoning D 24 | 25 | to : Functor (F ↓ Id) (Id ↓ G) 26 | to .F₀ o .dom = o .dom 27 | to .F₀ o .cod = o .cod 28 | to .F₀ o .map = L-adjunct F⊣G (o .map) 29 | to .F₁ f .top = f .top 30 | to .F₁ f .bot = f .bot 31 | to .F₁ {a} {b} f .com = 32 | L-adjunct F⊣G (b .map) C.∘ f .top ≡˘⟨ L-adjunct-naturall F⊣G _ _ ⟩ 33 | L-adjunct F⊣G (b .map D.∘ F .F₁ (f .top)) ≡⟨ ap (L-adjunct F⊣G) (f .com) ⟩ 34 | L-adjunct F⊣G (f .bot D.∘ a .map) ≡⟨ L-adjunct-naturalr F⊣G _ _ ⟩ 35 | G .F₁ (f .bot) C.∘ L-adjunct F⊣G (a .map) ∎ 36 | to .F-id = trivial! 37 | to .F-∘ _ _ = trivial! 38 | 39 | to-is-precat-iso : is-precat-iso to 40 | to-is-precat-iso .has-is-ff = is-iso→is-equiv is where 41 | is : ∀ {a b} → is-iso (to .F₁ {a} {b}) 42 | is .from f .top = f .top 43 | is .from f .bot = f .bot 44 | is {a} {b} .from f .com = Equiv.injective (adjunct-hom-equiv F⊣G) $ 45 | L-adjunct F⊣G (b .map D.∘ F .F₁ (f .top)) ≡⟨ L-adjunct-naturall F⊣G _ _ ⟩ 46 | L-adjunct F⊣G (b .map) C.∘ f .top ≡⟨ f .com ⟩ 47 | G .F₁ (f .bot) C.∘ L-adjunct F⊣G (a .map) ≡˘⟨ L-adjunct-naturalr F⊣G _ _ ⟩ 48 | L-adjunct F⊣G (f .bot D.∘ a .map) ∎ 49 | is .rinv f = trivial! 50 | is .linv f = trivial! 51 | to-is-precat-iso .has-is-iso = is-iso→is-equiv is where 52 | is : is-iso (to .F₀) 53 | is .from o .dom = o .dom 54 | is .from o .cod = o .cod 55 | is .from o .map = R-adjunct F⊣G (o .map) 56 | is .rinv o = ↓Obj-path _ _ refl refl (L-R-adjunct F⊣G _) 57 | is .linv o = ↓Obj-path _ _ refl refl (R-L-adjunct F⊣G _) 58 | -------------------------------------------------------------------------------- /src-1lab/Goat.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude 2 | open import Data.Dec 3 | open import Data.Nat 4 | 5 | {- 6 | A formalisation of https://en.wikipedia.org/wiki/Wolf,_goat_and_cabbage_problem 7 | to demonstrate proof by reflection. 8 | -} 9 | module Goat where 10 | 11 | holds : ∀ {ℓ} (A : Type ℓ) ⦃ _ : Dec A ⦄ → Type 12 | holds _ ⦃ yes _ ⦄ = ⊤ 13 | holds _ ⦃ no _ ⦄ = ⊥ 14 | 15 | data Side : Type where 16 | left right : Side 17 | 18 | left≠right : ¬ left ≡ right 19 | left≠right p = subst (λ { left → ⊤; right → ⊥ }) p tt 20 | 21 | instance 22 | Discrete-Side : Discrete Side 23 | Discrete-Side .decide left left = yes refl 24 | Discrete-Side .decide left right = no left≠right 25 | Discrete-Side .decide right left = no (left≠right ∘ sym) 26 | Discrete-Side .decide right right = yes refl 27 | 28 | cross : Side → Side 29 | cross left = right 30 | cross right = left 31 | 32 | record State : Type where 33 | constructor state 34 | field 35 | farmer wolf goat cabbage : Side 36 | 37 | open State 38 | 39 | count-left : Side → Nat 40 | count-left left = 1 41 | count-left right = 0 42 | count-lefts : State → Nat 43 | count-lefts (state f w g c) = count-left f + count-left w + count-left g + count-left c 44 | 45 | is-valid : State → Type 46 | is-valid s@(state f w g c) = count-lefts s ≡ 2 → f ≡ g 47 | 48 | record Valid-state : Type where 49 | constructor valid 50 | field 51 | has-state : State 52 | ⦃ has-valid ⦄ : holds (is-valid has-state) 53 | 54 | open Valid-state 55 | 56 | data Move : State → State → Type where 57 | go-alone : ∀ {s} → Move s (record s { farmer = cross (s .farmer) }) 58 | take-wolf : ∀ {s} → Move s (record s { farmer = cross (s .farmer); wolf = cross (s .wolf) }) 59 | take-goat : ∀ {s} → Move s (record s { farmer = cross (s .farmer); goat = cross (s .goat) }) 60 | take-cabbage : ∀ {s} → Move s (record s { farmer = cross (s .farmer); cabbage = cross (s .cabbage) }) 61 | 62 | data Moves : Valid-state → Valid-state → Type where 63 | done : ∀ {s} → Moves s s 64 | _∷_ : ∀ {a b c} ⦃ _ : holds (is-valid b) ⦄ → Move (a .has-state) b → Moves (valid b) c → Moves a c 65 | 66 | infixr 6 _∷_ 67 | 68 | initial final : Valid-state 69 | initial = valid (state left left left left) 70 | final = valid (state right right right right) 71 | 72 | goal : Moves initial final 73 | goal = take-goat ∷ go-alone ∷ take-wolf ∷ take-goat ∷ take-cabbage ∷ go-alone ∷ take-goat ∷ done 74 | -- goal = {! take-wolf ∷ ? !} -- No instance of type holds (is-valid ...) was found in scope. 75 | -------------------------------------------------------------------------------- /shake/LICENSE.agda: -------------------------------------------------------------------------------- 1 | The files under HTML/ are modified versions of Agda's HTML backend. 2 | Agda is distributed with the following license: 3 | 4 | Copyright (c) 2005-2024 remains with the authors. 5 | Agda 2 was originally written by Ulf Norell, 6 | partially based on code from Agda 1 by Catarina Coquand and Makoto Takeyama, 7 | and from Agdalight by Ulf Norell and Andreas Abel. 8 | Cubical Agda was originally contributed by Andrea Vezzosi. 9 | 10 | Agda 2 is currently actively developed mainly by Andreas Abel, 11 | Guillaume Allais, Liang-Ting Chen, Jesper Cockx, Matthew Daggitt, 12 | Nils Anders Danielsson, Amélia Liao, Ulf Norell, and 13 | Andrés Sicard-Ramírez. 14 | 15 | Further, Agda 2 has received contributions by, amongst others, 16 | Arthur Adjedj, Stevan Andjelkovic, 17 | Marcin Benke, Jean-Philippe Bernardy, Guillaume Brunerie, 18 | James Chapman, Jonathan Coates, 19 | Dominique Devriese, Péter Diviánszky, Robert Estelle, 20 | Olle Fredriksson, Adam Gundry, Daniel Gustafsson, Philipp Hausmann, 21 | Alan Jeffrey, Phil de Joux, 22 | Wolfram Kahl, Wen Kokke, John Leo, Fredrik Lindblad, 23 | Víctor López Juan, Ting-Gan Lua, Francesco Mazzoli, Stefan Monnier, 24 | Guilhem Moulin, Konstantin Nisht, Fredrik Nordvall Forsberg, 25 | Josselin Poiret, Nicolas Pouillard, Jonathan Prieto, Christian Sattler, 26 | Makoto Takeyama, Andrea Vezzosi, Noam Zeilberger, and Tesla Ice Zhang. 27 | The full list of contributors is available at 28 | https://github.com/agda/agda/graphs/contributors or from the git 29 | repository via ``git shortlog -sne``. 30 | 31 | Permission is hereby granted, free of charge, to any person obtaining 32 | a copy of this software and associated documentation files (the 33 | "Software"), to deal in the Software without restriction, including 34 | without limitation the rights to use, copy, modify, merge, publish, 35 | distribute, sublicense, and/or sell copies of the Software, and to 36 | permit persons to whom the Software is furnished to do so, subject to 37 | the following conditions: 38 | 39 | The above copyright notice and this permission notice shall be 40 | included in all copies or substantial portions of the Software. 41 | 42 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, 43 | EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF 44 | MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. 45 | IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY 46 | CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, 47 | TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION WITH THE 48 | SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. 49 | -------------------------------------------------------------------------------- /src-1lab/CoherentlyConstant.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Prelude hiding (∥_∥³; ∥-∥³-elim-set; ∥-∥³-elim-prop; ∥-∥³-rec; ∥-∥³-is-prop; ∥-∥-rec-groupoid) 2 | open import 1Lab.Path.Reasoning 3 | 4 | -- Coherently constant maps into groupoids, now at https://1lab.dev/1Lab.HIT.Truncation.html#maps-into-groupoids 5 | module CoherentlyConstant where 6 | 7 | data ∥_∥³ {ℓ} (A : Type ℓ) : Type ℓ where 8 | inc : A → ∥ A ∥³ 9 | iconst : ∀ a b → inc a ≡ inc b 10 | icoh : ∀ a b c → PathP (λ i → inc a ≡ iconst b c i) (iconst a b) (iconst a c) 11 | squash : is-groupoid ∥ A ∥³ 12 | 13 | ∥-∥³-elim-set 14 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ∥ A ∥³ → Type ℓ'} 15 | → (∀ a → is-set (P a)) 16 | → (f : (a : A) → P (inc a)) 17 | → (∀ a b → PathP (λ i → P (iconst a b i)) (f a) (f b)) 18 | → ∀ a → P a 19 | ∥-∥³-elim-set {P = P} sets f fconst = go where 20 | go : ∀ a → P a 21 | go (inc x) = f x 22 | go (iconst a b i) = fconst a b i 23 | go (icoh a b c i j) = is-set→squarep (λ i j → sets (icoh a b c i j)) 24 | refl (λ i → go (iconst a b i)) (λ i → go (iconst a c i)) (λ i → go (iconst b c i)) 25 | i j 26 | go (squash a b p q r s i j k) = is-hlevel→is-hlevel-dep 2 (λ _ → is-hlevel-suc 2 (sets _)) 27 | (go a) (go b) 28 | (λ k → go (p k)) (λ k → go (q k)) 29 | (λ j k → go (r j k)) (λ j k → go (s j k)) 30 | (squash a b p q r s) i j k 31 | 32 | ∥-∥³-elim-prop 33 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ∥ A ∥³ → Type ℓ'} 34 | → (∀ a → is-prop (P a)) 35 | → (f : (a : A) → P (inc a)) 36 | → ∀ a → P a 37 | ∥-∥³-elim-prop props f = ∥-∥³-elim-set (λ _ → is-hlevel-suc 1 (props _)) f 38 | (λ _ _ → is-prop→pathp (λ _ → props _) _ _) 39 | 40 | ∥-∥³-rec 41 | : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} 42 | → is-groupoid B 43 | → (f : A → B) 44 | → (fconst : ∀ x y → f x ≡ f y) 45 | → (∀ x y z → fconst x y ∙ fconst y z ≡ fconst x z) 46 | → ∥ A ∥³ → B 47 | ∥-∥³-rec {A = A} {B} bgrpd f fconst fcoh = go where 48 | go : ∥ A ∥³ → B 49 | go (inc x) = f x 50 | go (iconst a b i) = fconst a b i 51 | go (icoh a b c i j) = ∙→square (sym (fcoh a b c)) i j 52 | go (squash x y a b p q i j k) = bgrpd 53 | (go x) (go y) 54 | (λ i → go (a i)) (λ i → go (b i)) 55 | (λ i j → go (p i j)) (λ i j → go (q i j)) 56 | i j k 57 | 58 | ∥-∥³-is-prop : ∀ {ℓ} {A : Type ℓ} → is-prop ∥ A ∥³ 59 | ∥-∥³-is-prop = is-contr-if-inhabited→is-prop $ 60 | ∥-∥³-elim-prop (λ _ → hlevel 1) 61 | (λ a → contr (inc a) (∥-∥³-elim-set (λ _ → squash _ _) (iconst a) (icoh a))) 62 | 63 | ∥-∥-rec-groupoid 64 | : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} 65 | → is-groupoid B 66 | → (f : A → B) 67 | → (fconst : ∀ x y → f x ≡ f y) 68 | → (∀ x y z → fconst x y ∙ fconst y z ≡ fconst x z) 69 | → ∥ A ∥ → B 70 | ∥-∥-rec-groupoid bgrpd f fconst fcoh = 71 | ∥-∥³-rec bgrpd f fconst fcoh ∘ ∥-∥-rec ∥-∥³-is-prop inc 72 | -------------------------------------------------------------------------------- /src/NatChurchMonoid.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Algebra.Monoid 2 | open import Cubical.Algebra.Monoid.Instances.Nat 3 | open import Cubical.Algebra.Semigroup 4 | open import Cubical.Data.Nat 5 | open import Cubical.Data.Sigma 6 | open import Cubical.Foundations.Prelude 7 | open import Cubical.Foundations.Function 8 | open import Cubical.Foundations.Isomorphism 9 | open import Cubical.Foundations.Structure 10 | 11 | -- ℕ ≃ (m : Monoid) → ⟨ m ⟩ → ⟨ m ⟩ 12 | module NatChurchMonoid where 13 | 14 | MEndo : Type₁ 15 | MEndo = (m : Monoid ℓ-zero) → ⟨ m ⟩ → ⟨ m ⟩ 16 | 17 | isNatural : MEndo → Type₁ 18 | isNatural me = {m1 m2 : Monoid ℓ-zero} (f : MonoidHom m1 m2) → me m2 ∘ f .fst ≡ f .fst ∘ me m1 19 | 20 | isPropIsNatural : (me : MEndo) → isProp (isNatural me) 21 | isPropIsNatural me a b i {m1} {m2} f j x = m2 .snd .MonoidStr.isMonoid .IsMonoid.isSemigroup .IsSemigroup.is-set (me m2 (f .fst x)) (f .fst (me m1 x)) (funExt⁻ (a f) x) (funExt⁻ (b f) x) i j 22 | 23 | MEndoNatural : Type₁ 24 | MEndoNatural = Σ MEndo isNatural 25 | 26 | -- A generalised Church encoding for ℕ. This boils down to the fact that the forgetful functor 27 | -- U : Mon → Set is represented by ℕ ≃ F 1, followed by the Yoneda lemma. 28 | ℕ≃MEndoNatural : Iso ℕ MEndoNatural 29 | ℕ≃MEndoNatural = iso mtimes on1 mtimes-on1 on1-mtimes where 30 | 31 | mtimes : ℕ → MEndoNatural 32 | mtimes zero .fst (_ , monoidstr ε _·_ m) = (λ _ → ε) 33 | mtimes zero .snd f = funExt λ _ → sym (f .snd .IsMonoidHom.presε) 34 | mtimes (suc n) .fst ms@(t , monoidstr ε _·_ m) = (λ x → x · mtimes n .fst ms x) 35 | mtimes (suc n) .snd {m1} {m2@(_ , monoidstr _ _·_ m)} f = funExt λ x → cong (f .fst x ·_) (λ i → mtimes n .snd f i x) ∙ sym (f .snd .IsMonoidHom.pres· x (mtimes n .fst m1 x)) 36 | 37 | mtimes-hom : (m : Monoid ℓ-zero) (x : ⟨ m ⟩) → MonoidHom NatMonoid m 38 | mtimes-hom m x = (λ n → mtimes n .fst m x) , monoidequiv refl (λ n n' → mtimes-+ n n') where 39 | mtimes-+ : (n n' : ℕ) {m : Monoid ℓ-zero} {x : ⟨ m ⟩} → mtimes (n + n') .fst m x ≡ m .snd .MonoidStr._·_ (mtimes n .fst m x) (mtimes n' .fst m x) 40 | mtimes-+ zero n' {m} = sym (m .snd .MonoidStr.isMonoid .IsMonoid.·IdL _) 41 | mtimes-+ (suc n) n' {m} {x} = cong (m .snd .MonoidStr._·_ x) (mtimes-+ n n') ∙ m .snd .MonoidStr.isMonoid .IsMonoid.isSemigroup .IsSemigroup.·Assoc _ _ _ 42 | 43 | on1 : MEndoNatural → ℕ 44 | on1 me = me .fst NatMonoid 1 45 | 46 | on1-mtimes : (n : ℕ) → on1 (mtimes n) ≡ n 47 | on1-mtimes zero = refl 48 | on1-mtimes (suc n) = cong suc (on1-mtimes n) 49 | 50 | mtimes-on1 : (me : MEndoNatural) → mtimes (on1 me) ≡ me 51 | mtimes-on1 me = Σ≡Prop isPropIsNatural (λ i m x → p m x i) where 52 | p : (m : Monoid ℓ-zero) (x : ⟨ m ⟩) → mtimes (on1 me) .fst m x ≡ me .fst m x 53 | p m x = sym (funExt⁻ (me .snd (mtimes-hom m x)) 1) 54 | ∙ cong (me .fst m) (m .snd .MonoidStr.isMonoid .IsMonoid.·IdR _) 55 | -------------------------------------------------------------------------------- /src-1lab/DoubleCoverOfTheCircle.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | open import 1Lab.Prelude 3 | open import 1Lab.Reflection.Induction 4 | 5 | open import Data.Bool 6 | 7 | open import Homotopy.Space.Circle 8 | 9 | module DoubleCoverOfTheCircle where 10 | ``` 11 | 12 | # double cover of the circle 13 | 14 | We define the connected double cover of the circle in two ways, as a 15 | higher inductive family and by recursion, and show that the two 16 | definitions are equivalent. 17 | 18 | ~~~{.quiver} 19 | % stolen from the Symmetry book https://github.com/UniMath/SymmetryBook 20 | \begin{tikzpicture} 21 | \node at (2,-1) {$S^1$}; 22 | \node at (2,0.5) {$\mathsf{Cover}$}; 23 | \draw[line width=1pt] (0,-1) ellipse (1 and .3); 24 | \draw[line width=1pt] (-1,0) 25 | .. controls ++( 90:-.3) and ++(210: .4) .. (0,0.1) 26 | .. controls ++(210:-.4) and ++(270: .3) .. (1,1) 27 | .. controls ++(270:-.3) and ++( 0: .1) .. (0,1.3) 28 | .. controls ++( 0:-.1) and ++( 90: .3) .. (-1,1) 29 | .. controls ++( 90:-.3) and ++(150: .4) .. (0,0.1) 30 | .. controls ++(150:-.4) and ++(270: .3) .. (1,0) 31 | .. controls ++(270:-.3) and ++( 0: .1) .. (0,0.3) 32 | .. controls ++( 0:-.1) and ++( 90: .3) .. (-1,0); 33 | \node[fill,circle,inner sep=1.5pt] at (-1,-1) {}; 34 | \node[fill,circle,inner sep=1.5pt] at (-1,0) {}; 35 | \node[fill,circle,inner sep=1.5pt] at (-1,1) {}; 36 | \end{tikzpicture} 37 | ~~~ 38 | 39 | The higher inductive definition has two $n$-cells for every $n$-cell 40 | in the circle: two points over $\mathsf{base}$ and two loops over 41 | $\mathsf{loop}$. 42 | 43 | ```agda 44 | data Cover : S¹ → Type where 45 | base0 base1 : Cover base 46 | loop0 : PathP (λ i → Cover (loop i)) base0 base1 47 | loop1 : PathP (λ i → Cover (loop i)) base1 base0 48 | 49 | unquoteDecl Cover-elim = 50 | make-elim-with default-elim-visible Cover-elim (quote Cover) 51 | ``` 52 | 53 | The recursive definition defines a $2$-element set bundle directly by sending 54 | $\mathsf{base}$ to the booleans and $\mathsf{loop}$ to the identification 55 | that swaps the booleans, using univalence. 56 | 57 | ```agda 58 | cover : S¹ → Type 59 | cover = S¹-elim Bool (ua not≃) 60 | ``` 61 | 62 | We can then write functions in both directions and prove that they are 63 | inverses using the introduction and elimination helpers for `ua`. 64 | 65 | ```agda 66 | recover : ∀ s → Cover s → cover s 67 | recover = Cover-elim _ true false 68 | (path→ua-pathp _ refl) (path→ua-pathp _ refl) 69 | 70 | discover-base : cover base → Cover base 71 | discover-base true = base0 72 | discover-base false = base1 73 | 74 | discover-loop 75 | : ∀ c → PathP (λ i → Cover (loop i)) (discover-base c) (discover-base (not c)) 76 | discover-loop true = loop0 77 | discover-loop false = loop1 78 | 79 | discover : ∀ s → cover s → Cover s 80 | discover = S¹-elim discover-base (ua→ discover-loop) 81 | 82 | rediscover : ∀ s c → recover s (discover s c) ≡ c 83 | rediscover = S¹-elim (Bool-elim _ refl refl) hlevel! 84 | 85 | disrecover : ∀ s c → discover s (recover s c) ≡ c 86 | disrecover = Cover-elim _ refl refl 87 | (transposeP (ua→-β not≃ {f₁ = discover-base} discover-loop refl ∙ ▷-idr loop0)) 88 | (transposeP (ua→-β not≃ {f₁ = discover-base} discover-loop refl ∙ ▷-idr loop1)) 89 | 90 | Cover≃cover : ∀ s → Cover s ≃ cover s 91 | Cover≃cover s = recover s , is-iso→is-equiv λ where 92 | .is-iso.from → discover s 93 | .is-iso.rinv → rediscover s 94 | .is-iso.linv → disrecover s 95 | ``` 96 | -------------------------------------------------------------------------------- /flake.nix: -------------------------------------------------------------------------------- 1 | { 2 | inputs = { 3 | nixpkgs.url = "nixpkgs/nixpkgs-unstable"; 4 | agda.url = "github:agda/agda/5cb475a67135ca4ce42428c6f0294cea58a3ca2b"; 5 | agda-stdlib = { 6 | url = "github:agda/agda-stdlib"; 7 | flake = false; 8 | }; 9 | the1lab = { 10 | url = "github:ncfavier/1lab/stuff"; 11 | flake = false; 12 | }; 13 | }; 14 | 15 | outputs = inputs@{ self, nixpkgs, ... }: let 16 | system = "x86_64-linux"; 17 | pkgs = import nixpkgs { 18 | inherit system; 19 | overlays = [ 20 | inputs.agda.overlays.default 21 | (final: prev: { 22 | haskell = prev.haskell // { 23 | packageOverrides = final.lib.composeExtensions prev.haskell.packageOverrides 24 | (hfinal: hprev: { 25 | Agda = final.haskell.lib.enableCabalFlag hprev.Agda "debug"; 26 | }); 27 | }; 28 | }) 29 | (self: super: { 30 | agdaPackages = super.agdaPackages.overrideScope (aself: asuper: { 31 | standard-library = asuper.standard-library.overrideAttrs { 32 | version = "unstable-${inputs.agda-stdlib.shortRev}"; 33 | src = inputs.agda-stdlib; 34 | }; 35 | _1lab = asuper._1lab.overrideAttrs { 36 | version = "unstable-${inputs.the1lab.shortRev}"; 37 | src = inputs.the1lab; 38 | }; 39 | }); 40 | }) 41 | ]; 42 | }; 43 | 44 | agdaLibs = libs: [ 45 | libs.standard-library 46 | libs.cubical 47 | libs._1lab 48 | ]; 49 | agda = pkgs.agda.withPackages agdaLibs; 50 | AGDA_LIBRARIES_FILE = pkgs.agdaPackages.mkLibraryFile agdaLibs; 51 | 52 | texlive = pkgs.texlive.combine { 53 | inherit (pkgs.texlive) 54 | collection-basic 55 | collection-latex 56 | xcolor 57 | preview 58 | pgf tikz-cd pgfplots 59 | mathpazo 60 | varwidth xkeyval standalone; 61 | }; 62 | 63 | deps = [ 64 | pkgs.pandoc-katex 65 | texlive 66 | pkgs.poppler-utils 67 | ]; 68 | 69 | PANDOC_KATEX_CONFIG_FILE = pkgs.writeText "katex-config.toml" '' 70 | trust = true 71 | throw_on_error = true 72 | 73 | [macros] 74 | "\\bN" = "\\mathbb{N}" 75 | "\\bZ" = "\\mathbb{Z}" 76 | ''; 77 | 78 | shakefile = pkgs.haskellPackages.callCabal2nix "agda-stuff-shake" ./shake {}; 79 | in { 80 | devShells.${system} = { 81 | default = self.packages.${system}.default.overrideAttrs (old: { 82 | nativeBuildInputs = old.nativeBuildInputs or [] ++ [ agda ]; 83 | Agda = pkgs.haskellPackages.Agda; # prevent garbage collection 84 | }); 85 | 86 | shakefile = pkgs.haskellPackages.shellFor { 87 | packages = _: [ shakefile ]; 88 | nativeBuildInputs = deps ++ [ 89 | pkgs.haskell-language-server 90 | ]; 91 | inherit AGDA_LIBRARIES_FILE PANDOC_KATEX_CONFIG_FILE; 92 | }; 93 | }; 94 | 95 | packages.${system} = { 96 | default = pkgs.stdenv.mkDerivation { 97 | name = "agda-stuff"; 98 | src = self; 99 | nativeBuildInputs = deps ++ [ shakefile ]; 100 | inherit AGDA_LIBRARIES_FILE PANDOC_KATEX_CONFIG_FILE; 101 | LC_ALL = "C.UTF-8"; 102 | buildPhase = '' 103 | agda-stuff-shake 104 | mv _build/site "$out" 105 | ''; 106 | }; 107 | 108 | inherit shakefile; 109 | }; 110 | }; 111 | } 112 | -------------------------------------------------------------------------------- /src-1lab/Applicative.agda: -------------------------------------------------------------------------------- 1 | open import 1Lab.Equiv 2 | open import 1Lab.Extensionality 3 | open import 1Lab.HLevel 4 | open import 1Lab.HLevel.Closure 5 | open import 1Lab.Path 6 | open import 1Lab.Reflection.HLevel 7 | open import 1Lab.Reflection.Record 8 | open import 1Lab.Type 9 | open import 1Lab.Type.Sigma 10 | 11 | -- Applicative fully determines the underlying Functor. 12 | module Applicative {ℓ} where 13 | 14 | private variable 15 | A B C : Type ℓ 16 | 17 | _∘'_ : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Type ℓ₁} {B : Type ℓ₂} {C : Type ℓ₃} 18 | → (B → C) → (A → B) → A → C 19 | f ∘' g = λ z → f (g z) 20 | 21 | record applicative (F : Type ℓ → Type ℓ) : Type (lsuc ℓ) where 22 | infixl 8 _<*>_ 23 | field 24 | sets : is-set (F A) 25 | pure : A → F A 26 | _<*>_ : F (A → B) → F A → F B 27 | <*>-identity : ∀ {u : F A} 28 | → pure id <*> u ≡ u 29 | <*>-composition : ∀ {u : F (B → C)} {v : F (A → B)} {w : F A} 30 | → pure _∘'_ <*> u <*> v <*> w ≡ u <*> (v <*> w) 31 | <*>-homomorphism : ∀ {f : A → B} {x : A} 32 | → pure f <*> pure x ≡ pure (f x) 33 | <*>-interchange : ∀ {u : F (A → B)} {x : A} 34 | → u <*> pure x ≡ pure (λ f → f x) <*> u 35 | 36 | record applicative-functor (F : Type ℓ → Type ℓ) (app : applicative F) : Type (lsuc ℓ) where 37 | open applicative app 38 | infixl 8 _<$>_ 39 | field 40 | _<$>_ : (A → B) → F A → F B 41 | <$>-identity : ∀ {x : F A} 42 | → id <$> x ≡ x 43 | <$>-composition : ∀ {f : B → C} {g : A → B} {x : F A} 44 | → f <$> (g <$> x) ≡ f ∘' g <$> x 45 | pure-natural : ∀ {f : A → B} {x : A} 46 | → f <$> pure x ≡ pure (f x) 47 | <*>-extranatural-A : ∀ {f : F (B → C)} {g : A → B} {x : F A} 48 | → f <*> (g <$> x) ≡ (_∘' g) <$> f <*> x 49 | <*>-natural-B : ∀ {g : B → C} {f : F (A → B)} {x : F A} 50 | → g <$> (f <*> x) ≡ (g ∘'_) <$> f <*> x 51 | 52 | open applicative-functor 53 | unquoteDecl eqv = declare-record-iso eqv (quote applicative-functor) 54 | 55 | applicative-functor-path 56 | : ∀ {F : Type ℓ → Type ℓ} {app} {a b : applicative-functor F app} 57 | → (∀ {A B} (f : A → B) x → a ._<$>_ f x ≡ b ._<$>_ f x) 58 | → a ≡ b 59 | applicative-functor-path {F = F} {app = app} p = Iso.injective eqv (Σ-prop-path! (ext λ f → p f)) 60 | where instance 61 | F-sets : ∀ {x} → H-Level (F x) 2 62 | F-sets = hlevel-instance (app .applicative.sets) 63 | 64 | applicative-determines-functor : ∀ {F} (app : applicative F) 65 | → is-contr (applicative-functor F app) 66 | applicative-determines-functor {F} app = p where 67 | open applicative app 68 | p : is-contr (applicative-functor F app) 69 | p .centre ._<$>_ f x = pure f <*> x 70 | p .centre .<$>-identity = <*>-identity 71 | p .centre .<$>-composition {f = f} {g = g} {x = x} = 72 | pure f <*> (pure g <*> x) ≡⟨ sym <*>-composition ⟩ 73 | pure _∘'_ <*> pure f <*> pure g <*> x ≡⟨ ap (λ y → y <*> pure g <*> x) <*>-homomorphism ⟩ 74 | pure (f ∘'_) <*> pure g <*> x ≡⟨ ap (_<*> x) <*>-homomorphism ⟩ 75 | pure (f ∘' g) <*> x ∎ 76 | p .centre .pure-natural = <*>-homomorphism 77 | p .centre .<*>-extranatural-A {f = f} {g = g} {x = x} = 78 | f <*> (pure g <*> x) ≡⟨ sym <*>-composition ⟩ 79 | pure _∘'_ <*> f <*> pure g <*> x ≡⟨ ap (_<*> x) <*>-interchange ⟩ 80 | pure (_$ g) <*> (pure _∘'_ <*> f) <*> x ≡⟨ ap (_<*> x) (p .centre .<$>-composition) ⟩ 81 | pure (_∘' g) <*> f <*> x ∎ 82 | p .centre .<*>-natural-B {g = g} {f = f} {x = x} = 83 | pure g <*> (f <*> x) ≡⟨ sym <*>-composition ⟩ 84 | pure _∘'_ <*> pure g <*> f <*> x ≡⟨ ap (λ y → y <*> f <*> x) <*>-homomorphism ⟩ 85 | pure (g ∘'_) <*> f <*> x ∎ 86 | p .paths app' = applicative-functor-path λ f x → 87 | pure f <*> x ≡⟨ ap (_<*> x) (sym A.pure-natural) ⟩ 88 | (f ∘'_) A.<$> pure id <*> x ≡˘⟨ A.<*>-natural-B ⟩ 89 | f A.<$> (pure id <*> x) ≡⟨ ap (f A.<$>_) <*>-identity ⟩ 90 | f A.<$> x ∎ 91 | where module A = applicative-functor app' 92 | -------------------------------------------------------------------------------- /flake.lock: -------------------------------------------------------------------------------- 1 | { 2 | "nodes": { 3 | "agda": { 4 | "inputs": { 5 | "flake-parts": "flake-parts", 6 | "nixpkgs": "nixpkgs" 7 | }, 8 | "locked": { 9 | "lastModified": 1756135068, 10 | "narHash": "sha256-tVzLa/wuGqo+4LdHB3DI9S9xIblWIIEHg1+G9m2nTCA=", 11 | "owner": "agda", 12 | "repo": "agda", 13 | "rev": "5cb475a67135ca4ce42428c6f0294cea58a3ca2b", 14 | "type": "github" 15 | }, 16 | "original": { 17 | "owner": "agda", 18 | "repo": "agda", 19 | "rev": "5cb475a67135ca4ce42428c6f0294cea58a3ca2b", 20 | "type": "github" 21 | } 22 | }, 23 | "agda-stdlib": { 24 | "flake": false, 25 | "locked": { 26 | "lastModified": 1753005360, 27 | "narHash": "sha256-9m62MQHEOVxv61Nv9JlFnksn2lPq6QTuLZ0saPTOIJQ=", 28 | "owner": "agda", 29 | "repo": "agda-stdlib", 30 | "rev": "0e97e2ed1f999ea0d2cce2a3b2395d5a9a7bd36a", 31 | "type": "github" 32 | }, 33 | "original": { 34 | "owner": "agda", 35 | "repo": "agda-stdlib", 36 | "type": "github" 37 | } 38 | }, 39 | "flake-parts": { 40 | "inputs": { 41 | "nixpkgs-lib": "nixpkgs-lib" 42 | }, 43 | "locked": { 44 | "lastModified": 1701473968, 45 | "narHash": "sha256-YcVE5emp1qQ8ieHUnxt1wCZCC3ZfAS+SRRWZ2TMda7E=", 46 | "owner": "hercules-ci", 47 | "repo": "flake-parts", 48 | "rev": "34fed993f1674c8d06d58b37ce1e0fe5eebcb9f5", 49 | "type": "github" 50 | }, 51 | "original": { 52 | "owner": "hercules-ci", 53 | "repo": "flake-parts", 54 | "type": "github" 55 | } 56 | }, 57 | "nixpkgs": { 58 | "locked": { 59 | "lastModified": 1763948260, 60 | "narHash": "sha256-dY9qLD0H0zOUgU3vWacPY6Qc421BeQAfm8kBuBtPVE0=", 61 | "owner": "NixOS", 62 | "repo": "nixpkgs", 63 | "rev": "1c8ba8d3f7634acac4a2094eef7c32ad9106532c", 64 | "type": "github" 65 | }, 66 | "original": { 67 | "owner": "NixOS", 68 | "ref": "nixos-25.05", 69 | "repo": "nixpkgs", 70 | "type": "github" 71 | } 72 | }, 73 | "nixpkgs-lib": { 74 | "locked": { 75 | "dir": "lib", 76 | "lastModified": 1701253981, 77 | "narHash": "sha256-ztaDIyZ7HrTAfEEUt9AtTDNoCYxUdSd6NrRHaYOIxtk=", 78 | "owner": "NixOS", 79 | "repo": "nixpkgs", 80 | "rev": "e92039b55bcd58469325ded85d4f58dd5a4eaf58", 81 | "type": "github" 82 | }, 83 | "original": { 84 | "dir": "lib", 85 | "owner": "NixOS", 86 | "ref": "nixos-unstable", 87 | "repo": "nixpkgs", 88 | "type": "github" 89 | } 90 | }, 91 | "nixpkgs_2": { 92 | "locked": { 93 | "lastModified": 1764230294, 94 | "narHash": "sha256-Z63xl5Scj3Y/zRBPAWq1eT68n2wBWGCIEF4waZ0bQBE=", 95 | "owner": "NixOS", 96 | "repo": "nixpkgs", 97 | "rev": "0d59e0290eefe0f12512043842d7096c4070f30e", 98 | "type": "github" 99 | }, 100 | "original": { 101 | "id": "nixpkgs", 102 | "ref": "nixpkgs-unstable", 103 | "type": "indirect" 104 | } 105 | }, 106 | "root": { 107 | "inputs": { 108 | "agda": "agda", 109 | "agda-stdlib": "agda-stdlib", 110 | "nixpkgs": "nixpkgs_2", 111 | "the1lab": "the1lab" 112 | } 113 | }, 114 | "the1lab": { 115 | "flake": false, 116 | "locked": { 117 | "lastModified": 1766187858, 118 | "narHash": "sha256-/AtWB0iq4XOmp/q2BENN+4bQiLuuqufYGN8EIdbi+fY=", 119 | "owner": "ncfavier", 120 | "repo": "1lab", 121 | "rev": "67cb39fc3300f8a140440d0eeba03dedf735235a", 122 | "type": "github" 123 | }, 124 | "original": { 125 | "owner": "ncfavier", 126 | "ref": "stuff", 127 | "repo": "1lab", 128 | "type": "github" 129 | } 130 | } 131 | }, 132 | "root": "root", 133 | "version": 7 134 | } 135 | -------------------------------------------------------------------------------- /src-1lab/PostcomposeNotFull.agda: -------------------------------------------------------------------------------- 1 | open import Cat.Instances.Shape.Involution 2 | open import Cat.Instances.Shape.Interval 3 | open import Cat.Functor.Properties 4 | open import Cat.Functor.Compose 5 | open import Cat.Prelude 6 | 7 | open import Data.Bool 8 | 9 | open Precategory 10 | open Functor 11 | open _=>_ 12 | 13 | module PostcomposeNotFull where 14 | 15 | {- 16 | We prove that it is NOT the case that, for every full functor p, the 17 | postcomposition functor p ∘ — is full. 18 | -} 19 | 20 | claim = 21 | ∀ {o ℓ o' ℓ' od ℓd} {C : Precategory o ℓ} {C' : Precategory o' ℓ'} {D : Precategory od ℓd} 22 | → (p : Functor C C') → is-full p → is-full (postcompose p {D = D}) 23 | 24 | module _ (assume : claim) where 25 | {- 26 | The counterexample consists of the following category (identities omitted): 27 | https://q.uiver.app/#q=WzAsMixbMCwwLCJhIl0sWzAsMSwiYiJdLFswLDBdLFsxLDEsIiIsMix7InJhZGl1cyI6LTN9XSxbMCwxLCIiLDAseyJjdXJ2ZSI6LTF9XSxbMCwxLCIiLDEseyJjdXJ2ZSI6MX1dXQ== 28 | where the loops on a and b are involutions, the involution on a swaps 29 | the two morphisms a ⇉ b, and the involution on b leaves them alone. 30 | There is a full functor p from C to the walking arrow that collapses 31 | all the morphisms, and there are two inclusion functors F and G from 32 | the walking involution into C. A natural transformation p ∘ F ⇒ p ∘ G 33 | is trivial, but a natural transformation F ⇒ G is a "ℤ/2ℤ-equivariant" 34 | morphism a → b, that is one that commutes with the involutions on a and b. 35 | There is no such thing in C, hence the action of p ∘ — on natural 36 | transformations (whiskering) cannot be surjective. 37 | -} 38 | 39 | C : Precategory lzero lzero 40 | C .Ob = Bool 41 | C .Hom true true = Bool 42 | C .Hom true false = Bool 43 | C .Hom false true = ⊥ 44 | C .Hom false false = Bool 45 | C .Hom-set true true = hlevel 2 46 | C .Hom-set true false = hlevel 2 47 | C .Hom-set false true = hlevel 2 48 | C .Hom-set false false = hlevel 2 49 | C .id {true} = false 50 | C .id {false} = false 51 | C ._∘_ {true} {true} {true} = xor 52 | C ._∘_ {false} {false} {false} = xor 53 | C ._∘_ {true} {true} {false} f g = xor g f 54 | C ._∘_ {true} {false} {false} f g = g 55 | C .idr {true} {true} f = xor-falser f 56 | C .idr {true} {false} f = refl 57 | C .idr {false} {false} f = xor-falser f 58 | C .idl {true} {true} f = refl 59 | C .idl {true} {false} f = refl 60 | C .idl {false} {false} f = refl 61 | C .assoc {true} {true} {true} {true} f g h = xor-associative f g h 62 | C .assoc {false} {false} {false} {false} f g h = xor-associative f g h 63 | C .assoc {true} {true} {true} {false} f true true = sym (not-involutive f) 64 | C .assoc {true} {true} {true} {false} f true false = refl 65 | C .assoc {true} {true} {true} {false} f false h = refl 66 | C .assoc {true} {true} {false} {false} f g h = refl 67 | C .assoc {true} {false} {false} {false} f g h = refl 68 | 69 | p : Functor C (0≤1 ^op) 70 | p .F₀ o = o 71 | p .F₁ {true} {true} = _ 72 | p .F₁ {true} {false} = _ 73 | p .F₁ {false} {true} () 74 | p .F₁ {false} {false} = _ 75 | p .F-id {true} = refl 76 | p .F-id {false} = refl 77 | p .F-∘ {true} {true} {true} f g = refl 78 | p .F-∘ {true} {true} {false} f g = refl 79 | p .F-∘ {true} {false} {false} f g = refl 80 | p .F-∘ {false} {false} {false} f g = refl 81 | 82 | p-is-full : is-full p 83 | p-is-full {true} {true} _ = inc (false , refl) 84 | p-is-full {true} {false} _ = inc (false , refl) 85 | p-is-full {false} {false} _ = inc (false , refl) 86 | 87 | p*-is-full : is-full (postcompose p {D = ∙⤮∙}) 88 | p*-is-full = assume p p-is-full 89 | 90 | F G : Functor ∙⤮∙ C 91 | F .F₀ _ = true 92 | F .F₁ f = f 93 | F .F-id = refl 94 | F .F-∘ _ _ = refl 95 | G .F₀ _ = false 96 | G .F₁ f = f 97 | G .F-id = refl 98 | G .F-∘ _ _ = refl 99 | 100 | impossible : F => G → ⊥ 101 | impossible θ = not-no-fixed (sym (θ .is-natural _ _ true)) 102 | 103 | pθ : p F∘ F => p F∘ G 104 | pθ .η = _ 105 | pθ .is-natural _ _ _ = refl 106 | 107 | θ : ∥ F => G ∥ 108 | θ = fst <$> p*-is-full pθ 109 | 110 | contradiction : ⊥ 111 | contradiction = rec! impossible θ 112 | -------------------------------------------------------------------------------- /src-1lab/ObjectClassifiersOld.lagda.md: -------------------------------------------------------------------------------- 1 |
2 | 3 | An improved version of this note is at [`ObjectClassifiers`](ObjectClassifiers). 4 | 5 | 6 | ```agda 7 | open import 1Lab.Type 8 | open import 1Lab.Type.Sigma 9 | open import 1Lab.Type.Pointed 10 | open import 1Lab.Path 11 | open import 1Lab.HLevel 12 | open import 1Lab.HLevel.Closure 13 | open import 1Lab.Equiv 14 | 15 | -- Univalence from object classifiers in the sense of higher topos theory. 16 | module ObjectClassifiersOld where 17 | 18 | -- The type of arrows/bundles/fibrations. 19 | Bundle : ∀ ℓ → Type (lsuc ℓ) 20 | Bundle ℓ 21 | = Σ (Type ℓ) λ A 22 | → Σ (Type ℓ) λ B 23 | → A → B 24 | 25 | -- The standard pullback construction (HoTT book 2.15.11). 26 | record Pullback {ℓ ℓ'} {B : Type ℓ} {C D : Type ℓ'} (s : B → D) (q : C → D) : Type (ℓ ⊔ ℓ') where 27 | constructor pb 28 | field 29 | pb₁ : B 30 | pb₂ : C 31 | pbeq : s pb₁ ≡ q pb₂ 32 | 33 | open Pullback 34 | 35 | pb-path : ∀ {ℓ ℓ'} {B : Type ℓ} {C D : Type ℓ'} {s : B → D} {q : C → D} → {a b : Pullback s q} 36 | → (p1 : a .pb₁ ≡ b .pb₁) → (p2 : a .pb₂ ≡ b .pb₂) → PathP (λ i → s (p1 i) ≡ q (p2 i)) (a .pbeq) (b .pbeq) 37 | → a ≡ b 38 | pb-path p1 p2 pe i = pb (p1 i) (p2 i) (pe i) 39 | 40 | -- The morphisms of interest between bundles p : A → B and q : C → D are pairs 41 | -- r : A → C, s : B → D that make a *pullback square*: 42 | -- r 43 | -- A ------> C 44 | -- | ⯾ | 45 | -- p | | q 46 | -- v v 47 | -- B ------> D 48 | -- s 49 | -- Rather than laboriously state the universal property of a pullback, we take it 50 | -- on faith that the construction above has the universal property (this is 51 | -- exercise 2.11 in the HoTT book) and simply define "being a pullback" as having 52 | -- r and p factor through an equivalence A ≃ Pullback s q. 53 | -- This completely determines r by the factorisation r = pb₂ ∘ e .fst, so we can 54 | -- omit it by contractibility of singletons. 55 | _⇒_ : ∀ {ℓ} {ℓ'} → Bundle ℓ → Bundle ℓ' → Type (ℓ ⊔ ℓ') 56 | (A , B , p) ⇒ (C , D , q) 57 | = Σ (B → D) λ s 58 | → Σ (A ≃ Pullback s q) λ e 59 | → p ≡ pb₁ ∘ e .fst 60 | 61 | -- An object classifier is a *universal* bundle U∙ → U such that any other 62 | -- bundle has a unique map (i.e. pullback square) into it. 63 | -- Categorically, it is a terminal object in the category of arrows and pullback squares. 64 | is-classifier : ∀ {ℓ} → Bundle (lsuc ℓ) → Type (lsuc ℓ) 65 | is-classifier {ℓ} u = ∀ (p : Bundle ℓ) → is-contr (p ⇒ u) 66 | 67 | -- The projection from the type of pointed types to the type of types is our 68 | -- universal bundle: the fibre above A : Type ℓ is equivalent to A itself. 69 | Type↓ : ∀ ℓ → Bundle (lsuc ℓ) 70 | Type↓ ℓ = Type∙ ℓ , Type ℓ , fst 71 | 72 | Type↓-fibre : ∀ {ℓ} (A : Type ℓ) → A ≃ Pullback {B = Lift ℓ ⊤} {C = Type∙ ℓ} (λ _ → A) fst 73 | Type↓-fibre A = Iso→Equiv λ where 74 | .fst a → pb _ (A , a) refl 75 | .snd .is-iso.from (pb _ (A' , a) eq) → transport (sym eq) a 76 | .snd .is-iso.linv → transport-refl 77 | .snd .is-iso.rinv (pb _ (A' , a) eq) → pb-path refl (sym (Σ-path (sym eq) refl)) λ i j → eq (i ∧ j) 78 | 79 | postulate 80 | -- We assume that Type↓ is an object classifier in the sense above, and show that 81 | -- this makes it a univalent universe. 82 | Type↓-is-classifier : ∀ {ℓ} → is-classifier (Type↓ ℓ) 83 | 84 | -- Every type is trivially a bundle over the unit type. 85 | ! : ∀ {ℓ} → Type ℓ → Bundle ℓ 86 | ! A = A , Lift _ ⊤ , _ 87 | 88 | -- The key observation is that the type of pullback squares from ! A to Type↓ is 89 | -- equivalent to the type of types equipped with an equivalence to A. 90 | -- Since the former was assumed to be contractible, so is the latter. 91 | lemma : ∀ {ℓ} (A : Type ℓ) → (! A ⇒ Type↓ ℓ) ≃ Σ (Type ℓ) (λ B → A ≃ B) 92 | lemma {ℓ} A = Iso→Equiv λ where 93 | .fst (ty , e , _) → ty _ , e ∙e Type↓-fibre (ty _) e⁻¹ 94 | .snd .is-iso.from (B , e) → (λ _ → B) , e ∙e Type↓-fibre B , refl 95 | .snd .is-iso.linv (ty , e , _) → Σ-pathp refl (Σ-pathp (Σ-prop-path is-equiv-is-prop 96 | (funext λ _ → Equiv.ε (Type↓-fibre (ty _)) _)) refl) 97 | .snd .is-iso.rinv (B , e) → Σ-pathp refl (Σ-prop-path is-equiv-is-prop 98 | (funext λ _ → Equiv.η (Type↓-fibre B) _)) 99 | 100 | -- Equivalences form an identity system, which is another way to state univalence. 101 | univalence : ∀ {ℓ} (A : Type ℓ) → is-contr (Σ (Type ℓ) λ B → A ≃ B) 102 | univalence {ℓ} A = Equiv→is-hlevel 0 (lemma A e⁻¹) (Type↓-is-classifier (! A)) 103 | ``` 104 |
105 | -------------------------------------------------------------------------------- /src/Shapes.agda: -------------------------------------------------------------------------------- 1 | open import Cubical.Foundations.Prelude 2 | open import Cubical.Foundations.Path 3 | open import Cubical.Foundations.Isomorphism renaming (Iso to _≃_) 4 | open import Cubical.Foundations.Univalence 5 | open import Cubical.Data.Unit renaming (Unit to ⊤) 6 | open import Cubical.Data.Sigma 7 | open import Cubical.Data.Int 8 | open import Cubical.Relation.Nullary 9 | 10 | -- — 11 | data Interval : Type where 12 | l : Interval 13 | r : Interval 14 | seg : l ≡ r 15 | 16 | Interval-isContr : isContr Interval 17 | Interval-isContr = l , paths where 18 | paths : (x : Interval) → l ≡ x 19 | paths l = refl 20 | paths r = seg 21 | paths (seg i) j = seg (i ∧ j) 22 | 23 | Interval-loops : (x : Interval) → x ≡ x 24 | Interval-loops l = refl 25 | Interval-loops r = refl 26 | Interval-loops (seg i) j = seg i 27 | 28 | -- ○ 29 | data S¹ : Type where 30 | base : S¹ 31 | loop : base ≡ base 32 | 33 | S¹→⊤ : S¹ → ⊤ 34 | S¹→⊤ base = tt 35 | S¹→⊤ (loop i) = tt 36 | ⊤→S¹ : ⊤ → S¹ 37 | ⊤→S¹ tt = base 38 | ⊤→S¹→⊤ : (t : ⊤) → S¹→⊤ (⊤→S¹ t) ≡ t 39 | ⊤→S¹→⊤ tt = refl 40 | -- S¹→⊤→S¹ : (x : S¹) → ⊤→S¹ (S¹→⊤ x) ≡ x 41 | -- S¹→⊤→S¹ base = refl 42 | -- S¹→⊤→S¹ (loop i) j = {! IMPOSSIBLE the point doesn't retract onto the circle! !} 43 | 44 | always-loop : (x : S¹) → x ≡ x 45 | always-loop base = loop 46 | always-loop (loop i) j = 47 | hcomp (λ where k (i = i0) → loop (j ∨ ~ k) 48 | k (i = i1) → loop (j ∧ k) 49 | k (j = i0) → loop (i ∨ ~ k) 50 | k (j = i1) → loop (i ∧ k)) 51 | base 52 | 53 | loop-induction : {ℓ : Level} {P : base ≡ base → Type ℓ} 54 | → (pprop : ∀ p → isProp (P p)) 55 | → (prefl : P refl) 56 | → (ploop : ∀ p → P p → P (p ∙ loop)) 57 | → (ppool : ∀ p → P p → P (p ∙ sym loop)) 58 | → (p : base ≡ base) → P p 59 | loop-induction {ℓ} {P} pprop prefl ploop ppool = J Q prefl 60 | where 61 | bridge : PathP (λ i → base ≡ loop i → Type ℓ) P P 62 | bridge = toPathP (funExt λ p → isoToPath 63 | (iso (λ x → subst P (compPathr-cancel _ _) (ploop _ x)) 64 | (ppool p) 65 | (λ _ → pprop _ _ _) 66 | (λ _ → pprop _ _ _))) 67 | Q : (x : S¹) → base ≡ x → Type ℓ 68 | Q base p = P p 69 | Q (loop i) p = bridge i p 70 | 71 | data Bool₁ : Type₁ where 72 | false true : Bool₁ 73 | 74 | S¹⋆ : Σ Type (λ A → A) 75 | S¹⋆ = S¹ , base 76 | 77 | flip : S¹ → S¹ 78 | flip base = base 79 | flip (loop i) = loop (~ i) 80 | flip≡ : S¹ ≡ S¹ 81 | flip≡ = isoToPath (iso flip flip inv inv) where 82 | inv : section flip flip 83 | inv base = refl 84 | inv (loop i) = refl 85 | flip⋆ : S¹⋆ ≡ S¹⋆ 86 | flip⋆ i = flip≡ i , base≡base i where 87 | base≡base : PathP (λ i → flip≡ i) base base 88 | base≡base = ua-gluePath _ refl 89 | 90 | Cover : S¹ → Type 91 | Cover base = ℤ 92 | Cover (loop i) = sucPathℤ i 93 | 94 | S¹⋆-auto : (S¹⋆ ≡ S¹⋆) ≡ Bool₁ 95 | S¹⋆-auto = isoToPath (iso to from sec ret) where 96 | isPos : ℤ → Bool₁ 97 | isPos (pos _) = true 98 | isPos _ = false 99 | to : S¹⋆ ≡ S¹⋆ → Bool₁ 100 | to p = isPos (transport (λ i → Cover (loop' i)) 0) where 101 | loop' : base ≡ base 102 | loop' i = comp (λ j → p j .fst) 103 | (λ where j (i = i0) → p j .snd 104 | j (i = i1) → p j .snd) 105 | (loop i) 106 | from : Bool₁ → S¹⋆ ≡ S¹⋆ 107 | from false = flip⋆ 108 | from true = refl 109 | sec : section to from 110 | sec false = refl 111 | sec true = refl 112 | ret : retract to from 113 | ret p = {! !} 114 | 115 | -- ● 116 | data D² : Type where 117 | base² : D² 118 | loop² : base² ≡ base² 119 | disk : refl ≡ loop² 120 | 121 | D²-isContr : isContr D² 122 | D²-isContr = base² , paths where 123 | paths : (x : D²) → base² ≡ x 124 | paths base² = refl 125 | paths (loop² i) j = disk j i 126 | paths (disk i j) k = disk (i ∧ k) j 127 | 128 | D²-isProp : isProp D² 129 | D²-isProp x y = sym (D²-isContr .snd x) ∙ D²-isContr .snd y 130 | 131 | data coeq (X : Type) : Type where 132 | inc : X → coeq X 133 | eq : ∀ x → inc x ≡ inc x 134 | 135 | lemma : ∀ X → coeq X ≃ (X × S¹) 136 | lemma X = iso to from to-from from-to where 137 | to : coeq X → X × S¹ 138 | to (inc x) = x , base 139 | to (eq x i) = x , loop i 140 | from : X × S¹ → coeq X 141 | from (x , base) = inc x 142 | from (x , loop i) = eq x i 143 | to-from : ∀ x → to (from x) ≡ x 144 | to-from (_ , base) = refl 145 | to-from (_ , loop i) = refl 146 | from-to : ∀ x → from (to x) ≡ x 147 | from-to (inc x) = refl 148 | from-to (eq x i) = refl 149 | -------------------------------------------------------------------------------- /src-1lab/Skeletons.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | open import Cat.Functor.Base 3 | open import Cat.Functor.Equivalence 4 | open import Cat.Functor.FullSubcategory 5 | open import Cat.Functor.Properties 6 | open import Cat.Instances.FinSets 7 | open import Cat.Instances.Sets 8 | open import Cat.Prelude 9 | open import Cat.Skeletal 10 | open import Data.Bool 11 | open import Data.Fin 12 | open import Data.Nat 13 | 14 | import Cat.Reasoning 15 | 16 | open Functor 17 | open is-precat-iso 18 | 19 | module Skeletons where 20 | 21 | module Sets {ℓ} = Cat.Reasoning (Sets ℓ) 22 | ``` 23 | 24 | Formalising parts of [this math.SE answer](https://math.stackexchange.com/a/4943344), with 25 | finite-dimensional real vector spaces replaced with finite sets; 26 | the situation is entirely analogous. 27 | 28 | Instead of the skeletal category whose objects are natural numbers 29 | representing $ℝⁿ$ and whose morphisms are linear maps, we use the skeletal 30 | category whose objects are natural numbers representing the standard 31 | finite sets $[n]$ and whose morphisms are functions. 32 | 33 | ```agda 34 | S : Precategory lzero lzero 35 | S = FinSets 36 | 37 | S-is-skeletal : is-skeletal S 38 | S-is-skeletal = FinSets-is-skeletal 39 | ``` 40 | 41 | Instead of the univalent category of finite-dimensional real vector 42 | spaces, we use the univalent category of finite sets, here realised as 43 | the *essential image* $C$ of the inclusion of $S$ into sets. 44 | Explicitly, an object of $C$ is a set $X$ such that there merely exists a 45 | natural number $n$ such that $X ≃ [n]$. 46 | Equivalently, an object of $C$ is a set $X$ equipped with a natural number 47 | $n$ such that $∥ X ≃ [n] ∥$ (we can extract $n$ from the truncation because 48 | the statements $X ≃ [n]$ are mutually exclusive for distinct $n$). 49 | 50 | $C$ is a Rezk completion of $S$. 51 | 52 | ```agda 53 | C : Precategory (lsuc lzero) lzero 54 | C = Essential-image Fin→Sets 55 | 56 | C-is-category : is-category C 57 | C-is-category = Essential-image-is-category Fin→Sets Sets-is-category 58 | ``` 59 | 60 | If we remove the truncation (but do not change the morphisms), 61 | we get a skeletal category *isomorphic* to $S$, because we can contract $X$ 62 | away. This is entirely analogous to the way that the naïve definition 63 | of the image of a function using $Σ$ instead of $∃$ 64 | [yields](https://1lab.dev/1Lab.Counterexamples.Sigma.html) the domain of the function. 65 | 66 | ```agda 67 | C' : Precategory (lsuc lzero) lzero 68 | C' = Restrict {C = Sets _} λ X → Σ[ n ∈ Nat ] Fin→Sets .F₀ n Sets.≅ X 69 | 70 | S→C' : Functor S C' 71 | S→C' .F₀ n = el! (Fin n) , n , Sets.id-iso 72 | S→C' .F₁ f = f 73 | S→C' .F-id = refl 74 | S→C' .F-∘ _ _ = refl 75 | 76 | S≡C' : is-precat-iso S→C' 77 | S≡C' .has-is-ff = id-equiv 78 | S≡C' .has-is-iso = inverse-is-equiv (e .snd) where 79 | e : (Σ[ X ∈ Set lzero ] Σ[ n ∈ Nat ] Fin→Sets .F₀ n Sets.≅ X) ≃ Nat 80 | e = Σ-swap₂ ∙e Σ-contr-snd λ n → is-contr-ΣR Sets-is-category 81 | ``` 82 | 83 | Since $C$ is a Rezk completion of $S$, we should expect to have a fully 84 | faithful and essentially surjective functor $S → C$. 85 | 86 | ```agda 87 | S→C : Functor S C 88 | S→C = Essential-inc Fin→Sets 89 | 90 | S→C-is-ff : is-fully-faithful S→C 91 | S→C-is-ff = ff→Essential-inc-ff Fin→Sets Fin→Sets-is-ff 92 | 93 | S→C-is-eso : is-eso S→C 94 | S→C-is-eso = Essential-inc-eso Fin→Sets 95 | ``` 96 | 97 | However, this functor is *not* an equivalence of categories: in order 98 | to obtain a functor going the other way, we would have to choose an 99 | enumeration of every finite set in a coherent way. This is a form of 100 | [global choice](https://1lab.dev/1Lab.Counterexamples.GlobalChoice.html), 101 | which is just false in homotopy type theory. 102 | 103 | ```agda 104 | module _ (S≃C : is-equivalence S→C) where private 105 | open is-equivalence S≃C renaming (F⁻¹ to C→S) 106 | module C = Cat.Reasoning C 107 | 108 | module _ (X : Set lzero) (e : ∥ ⌞ X ⌟ ≃ Fin 2 ∥) where 109 | c : C.Ob 110 | c = X , ((λ e → 2 , equiv→iso (e e⁻¹)) <$> e) 111 | 112 | chosen : ⌞ X ⌟ 113 | chosen with C→S .F₀ c | counit.ε c | counit-iso c 114 | ... | suc n | ε | _ = ε 0 115 | ... | zero | ε | ε-inv = absurd (case e of λ e → 116 | zero≠suc (Fin-injective (iso→equiv (sub-iso→super-iso _ (C.invertible→iso ε ε-inv)) ∙e e))) 117 | 118 | b : Bool 119 | b = chosen (el! Bool) enumeration 120 | 121 | swap : Bool ≡ Bool 122 | swap = ua (not , not-is-equiv) 123 | 124 | p : PathP (λ i → swap i) b b 125 | p = ap₂ chosen (n-ua _) prop! 126 | 127 | ¬S≃C : ⊥ 128 | ¬S≃C = not-no-fixed (from-pathp⁻ p) 129 | ``` 130 | 131 | Alternatively, an equivalence $S \simeq C$ would imply that $C$ is 132 | strict (see [SetsCover](/SetsCover)), but the category of finite sets 133 | is of course not gaunt. 134 | -------------------------------------------------------------------------------- /style.css: -------------------------------------------------------------------------------- 1 | @font-face { 2 | font-family: JuliaMono; 3 | src: url("https://cdn.jsdelivr.net/gh/cormullion/juliamono/webfonts/JuliaMono-Medium.woff2"); 4 | font-display: swap; 5 | } 6 | 7 | @font-face { 8 | font-family: JuliaMono; 9 | font-weight: bold; 10 | src: url("https://cdn.jsdelivr.net/gh/cormullion/juliamono/webfonts/JuliaMono-ExtraBold.woff2"); 11 | font-display: swap; 12 | } 13 | 14 | :root { 15 | --background: #080709; 16 | --foreground: white; 17 | --accent: #ae4bff; 18 | --comment: hsl(0, 0%, 60%); 19 | --defined: var(--foreground); 20 | --literal: var(--accent); 21 | --keyword: var(--foreground); 22 | --symbol: var(--foreground); 23 | --bound: #c59efd; 24 | --module: var(--literal); 25 | --constructor: var(--literal); 26 | 27 | font-family: Inter, "Source Han Sans", "Source Sans 3", "Gill Sans", "Trebuchet MS", "Open Sans", Helvetica, sans-serif; 28 | font-feature-settings: 'liga' 1, 'calt' 1; /* fix for Chrome */ 29 | font-optical-sizing: auto; 30 | } 31 | 32 | @supports (font-variation-settings: normal) { 33 | :root { 34 | font-family: InterVariable, "Source Han Sans", "Source Sans 3", "Gill Sans", "Trebuchet MS", "Open Sans", Helvetica, sans-serif; 35 | } 36 | } 37 | 38 | ::selection { 39 | background-color: var(--accent); 40 | color: white; 41 | } 42 | 43 | html { 44 | padding: 0 1em; 45 | } 46 | 47 | body { 48 | margin: 0 auto; 49 | max-width: 50em; 50 | text-align: justify; 51 | background-color: var(--background); 52 | color: var(--foreground); 53 | } 54 | 55 | pre { 56 | /* Otherwise Firefox takes ages trying to justify
 blocks... */
 57 |   text-align: initial;
 58 | 
 59 |   overflow-x: auto;
 60 | }
 61 | 
 62 | .katex-html {
 63 |   /* Prevents browsers from inserting a line break just before punctuation following
 64 |      inline maths https://github.com/KaTeX/KaTeX/issues/1233 */
 65 |   white-space: nowrap;
 66 | }
 67 | 
 68 | pre, code {
 69 |   font-family: JuliaMono, monospace;
 70 | }
 71 | 
 72 | .diagram-container {
 73 |   text-align: center;
 74 |   overflow-x: auto;
 75 | }
 76 | 
 77 | details {
 78 |   border-inline-start: 4px solid var(--accent);
 79 |   padding-inline-start: 8px;
 80 | }
 81 | 
 82 | details > summary::after {
 83 |   content: '(click to unfold)';
 84 |   color: var(--comment);
 85 | }
 86 | 
 87 | details[open] > summary::after {
 88 |   content: '(click to fold)';
 89 | }
 90 | 
 91 | .anchor::before {
 92 |   content: '🔗';
 93 |   display: inline-block;
 94 |   font-size: 80%;
 95 |   margin-left: 10px;
 96 |   opacity: 0;
 97 | }
 98 | 
 99 | .anchor:hover {
100 |   text-decoration: none;
101 | }
102 | 
103 | h1:hover > .anchor::before, h2:hover > .anchor::before, h3:hover > .anchor::before, h4:hover > .anchor::before, h5:hover > .anchor::before, h6:hover > .anchor::before {
104 |   opacity: 1;
105 | }
106 | 
107 | :any-link {
108 |   text-decoration: none;
109 |   color: var(--module);
110 | }
111 | 
112 | /* Aspects. */
113 | .Agda .Comment       { color: var(--comment) }
114 | .Agda .Background    { }
115 | .Agda .Markup        { }
116 | .Agda .Keyword       { color: var(--keyword); font-weight: bold; }
117 | .Agda .String        { color: var(--literal) }
118 | .Agda .Number        { color: var(--literal) }
119 | .Agda .Symbol        { color: var(--symbol) }
120 | .Agda .PrimitiveType { color: var(--defined) }
121 | .Agda .Pragma        { color: var(--keyword)   }
122 | .Agda .Operator      {}
123 | .Agda .Hole          { background: #490764 }
124 | .Agda .Macro         { color: var(--defined) }
125 | 
126 | /* NameKinds. */
127 | .Agda .Bound                  { color: var(--bound)   }
128 | .Agda .Generalizable          { color: var(--bound)   }
129 | .Agda .InductiveConstructor   { color: var(--constructor) }
130 | .Agda .CoinductiveConstructor { color: var(--constructor) }
131 | .Agda .Datatype               { color: var(--defined) }
132 | .Agda .Field                  { color: var(--constructor) }
133 | .Agda .Function               { color: var(--defined) }
134 | .Agda .Module                 { color: var(--module) }
135 | .Agda .Postulate              { color: var(--defined) }
136 | .Agda .Primitive              { color: var(--defined) }
137 | .Agda .Record                 { color: var(--defined) }
138 | 
139 | /* OtherAspects. */
140 | .Agda .DottedPattern        {}
141 | .Agda .UnsolvedMeta         { color: var(--foreground); background: yellow         }
142 | .Agda .UnsolvedConstraint   { color: var(--foreground); background: yellow         }
143 | .Agda .TerminationProblem   { color: var(--foreground); background: #FFA07A        }
144 | .Agda .IncompletePattern    { color: var(--foreground); background: #F5DEB3        }
145 | .Agda .Error                { color: red;   text-decoration: underline }
146 | .Agda .TypeChecks           { color: var(--foreground); background: #ADD8E6        }
147 | .Agda .Deadcode             { color: var(--foreground); background: #808080        }
148 | .Agda .ShadowingInTelescope { color: var(--foreground); background: #808080        }
149 | 
150 | /* Standard attributes. */
151 | .Agda a { text-decoration: none }
152 | .Agda a[href]:hover { background-color: #444; }
153 | .Agda [href].hover-highlight { background-color: #444; }
154 | 


--------------------------------------------------------------------------------
/src-1lab/Hats.agda:
--------------------------------------------------------------------------------
  1 | open import 1Lab.Path.Reasoning
  2 | open import 1Lab.Prelude
  3 | 
  4 | open import Algebra.Monoid.Category
  5 | open import Algebra.Group
  6 | 
  7 | open import Data.List hiding (lookup)
  8 | open import Data.Fin
  9 | open import Data.Fin.Closure
 10 | 
 11 | open is-iso
 12 | 
 13 | {-
 14 | This is a formalisation of a hat puzzle whose statement you can read at any
 15 | of these places:
 16 | 
 17 |   https://www.jsoftware.com/pipermail/general/2007-June/030272.html
 18 |   https://www.cut-the-knot.org/blue/PuzzleWithHats.shtml
 19 |   https://twitter.com/gro_tsen/status/1618989823100096512
 20 | 
 21 | 🎩 SPOILERS AHEAD! 🎩
 22 | -}
 23 | module Hats where
 24 | 
 25 | -- We assume there's at least one person, and that everyone has a unique
 26 | -- "name" between 0 and n - 1, known to everyone else.
 27 | module _ (n-1 : Nat) where
 28 |   private
 29 |     n = suc n-1
 30 |     Person = Fin n
 31 |     Hat = Fin n
 32 |     Hats = Person → Hat
 33 | 
 34 |   record Strategy : Type where
 35 |     -- `guess i xs` is the guess made by the ith person upon seeing the n - 1
 36 |     -- other hats given by `xs` (a vector of numbers from 0 to n - 1).
 37 |     field
 38 |       guess : Person → (Fin n-1 → Hat) → Hat
 39 | 
 40 |     -- The relation `hats ✓ i` means that the ith person guesses correctly,
 41 |     -- given the assignment `hats`.
 42 |     _✓_ : Hats → Person → Type
 43 |     hats ✓ i = guess i (delete hats i) ≡ hats i
 44 | 
 45 |     -- We are interested in strategies where at least one person guesses
 46 |     -- correctly for every assignment of hats.
 47 |     field
 48 |       one-right : ∀ hats → Σ Person λ i → hats ✓ i
 49 | 
 50 |     -- First, note that, given this requirement, there can be at *most* one
 51 |     -- correct guess for every assignment of hats.
 52 |     -- This follows from a probabilistic argument: every person guesses correctly
 53 |     -- with probability 1/n, so the total number of correct guesses across all
 54 |     -- hat assignments is nⁿ.
 55 |     -- In order to conclude, we use the fact that any surjection between finite
 56 |     -- sets of equal cardinality is an equivalence.
 57 |     exactly-one-right : ∀ hats → is-contr (Σ Person λ i → hats ✓ i)
 58 |     exactly-one-right hats = Equiv→is-hlevel 0 (Fibre-equiv _ hats e⁻¹) (p-is-equiv .is-eqv hats)
 59 |       where
 60 |         probability : ∀ i → Iso (Σ Hats (_✓ i)) (Fin n-1 → Hat)
 61 |         probability i .fst (hats , _) = delete hats i
 62 |         probability i .snd .from other .fst = other [ i ≔ guess i other ]
 63 |         probability i .snd .from other .snd =
 64 |           guess i ⌜ delete (other [ i ≔ guess i other ]) i ⌝ ≡⟨ ap! (funext (delete-insert _ i _)) ⟩
 65 |           guess i other                                      ≡˘⟨ insert-lookup _ i _ ⟩
 66 |           (other [ i ≔ guess i other ]) i                    ∎
 67 |         probability i .snd .rinv _ = funext (delete-insert _ i _)
 68 |         probability i .snd .linv (hats , r) = Σ-prop-path! (funext (insert-delete _ i _ (sym r)))
 69 | 
 70 |         only-one : Σ Hats (λ hats → Σ Person (hats ✓_)) ≃ Hats
 71 |         only-one =
 72 |           Σ _ (λ hats → Σ _ λ i → hats ✓ i) ≃⟨ Σ-swap₂ ⟩
 73 |           Σ _ (λ i → Σ _ λ hats → hats ✓ i) ≃⟨ Σ-ap-snd (Iso→Equiv ∘ probability) ⟩
 74 |           Fin n × (Fin n-1 → Fin n)         ≃˘⟨ Fin-suc-Π ⟩
 75 |           (Fin n → Fin n)                   ≃∎
 76 | 
 77 |         p : Σ Hats (λ hats → Σ Person (hats ✓_)) → Hats
 78 |         p = fst
 79 |         p-is-equiv : is-equiv p
 80 |         p-is-equiv = Finite-surjection→equiv (inc only-one) p
 81 |           λ other → inc ((other , one-right other) , refl)
 82 | 
 83 |   open Strategy public
 84 | 
 85 |   -- n-hypercubes of order m. We won't use the extra degree of generality,
 86 |   -- but it doesn't hurt.
 87 |   Hypercube : Nat → Type
 88 |   Hypercube m = (Fin n → Fin m) → Fin m
 89 | 
 90 |   -- Latin hypercubes, or n-ary quasigroups.
 91 |   -- Every number appears exactly once on each "line"; equivalently,
 92 |   -- every partial application to n - 1 coordinates is an automorphism.
 93 |   is-latin : ∀ {m} → Hypercube m → Type
 94 |   is-latin {m} h = ∀ (i : Fin n) (xs : Fin n-1 → Fin m) → is-equiv λ x → h (xs [ i ≔ x ])
 95 | 
 96 |   Latin-hypercube : Nat → Type
 97 |   Latin-hypercube m = Σ (Hypercube m) is-latin
 98 | 
 99 |   -- Every latin n-hypercube h of order n induces a strategy where everyone
100 |   -- guesses that the multiplication of all the hats is equal to their index.
101 |   latin→strategy : Latin-hypercube n → Strategy
102 |   latin→strategy (h , lat) .guess i other = equiv→inverse (lat i other) i
103 |   latin→strategy (h , lat) .one-right hats =
104 |     h hats , Equiv.from (eqv (h hats)) refl
105 |     where
106 |       module L i = Equiv (_ , lat i (delete hats i))
107 |       eqv : ∀ (i : Fin n) → (L.from i i ≡ hats i) ≃ (h hats ≡ i)
108 |       eqv i =
109 |         L.from i i ≡ hats i ≃⟨ Equiv.adjunct (L.inverse i) ⟩
110 |         i ≡ L.to i (hats i) ≃⟨ sym-equiv ⟩
111 |         L.to i (hats i) ≡ i ≃⟨ ∙-pre-equiv (sym (ap h (funext (insert-delete hats i _ refl)))) ⟩
112 |         h hats ≡ i          ≃∎
113 | 
114 |   -- In particular, every group structure on Fin n induces a strategy since
115 |   -- group multiplication is an n-ary equivalence.
116 |   group→latin : Group-on (Fin n) → Latin-hypercube n
117 |   group→latin G = mul , mul-equiv
118 |     where
119 |       open Group-on G hiding (magma-hlevel)
120 | 
121 |       mul : ∀ {m} → (Fin m → Fin n) → Fin n
122 |       mul xs = fold underlying-monoid (tabulate xs)
123 | 
124 |       mul-equiv : ∀ {m} (i : Fin (suc m)) (xs : Fin m → Fin n)
125 |                 → is-equiv (λ x → mul (xs [ i ≔ x ]))
126 |       mul-equiv i xs with fin-view i
127 |       mul-equiv _ xs | zero = ⋆-equivr _
128 |       mul-equiv {suc m} _ xs | suc i = ∘-is-equiv (⋆-equivl _) (mul-equiv i (xs ∘ fsuc))
129 | 
130 |   group→strategy : Group-on (Fin n) → Strategy
131 |   group→strategy = latin→strategy ∘ group→latin
132 | 


--------------------------------------------------------------------------------
/src/Erasure.agda:
--------------------------------------------------------------------------------
  1 | open import Agda.Primitive renaming (Set to Type; Setω to Typeω)
  2 | open import Relation.Binary.PropositionalEquality hiding ([_])
  3 | open import Axiom.Extensionality.Propositional
  4 | {-# BUILTIN REWRITE _≡_ #-}
  5 | 
  6 | -- Investigating the erasure modality. See also ErasureOpen
  7 | module Erasure where
  8 | 
  9 | private variable
 10 |   a b : Level
 11 |   A : Type a
 12 | 
 13 | -- The erased path induction principle J₀
 14 | 
 15 | J₀-type =
 16 |   ∀ {a b} {@0 A : Type a} {@0 x : A} (B : (@0 y : A) → @0 x ≡ y → Type b)
 17 |   → {@0 y : A} (@0 p : x ≡ y) → B x refl → B y p
 18 | 
 19 | -- The Erased monadic modality
 20 | 
 21 | record Erased (@0 A : Type a) : Type a where
 22 |   constructor [_]
 23 |   field
 24 |     @0 erased : A
 25 | 
 26 | open Erased
 27 | 
 28 | η : {@0 A : Type a} → A → Erased A
 29 | η x = [ x ]
 30 | 
 31 | μ : {@0 A : Type a} → Erased (Erased A) → Erased A
 32 | μ [ [ x ] ] = [ x ]
 33 | 
 34 | -- Paths (Erased A) → Erased (Paths A)
 35 | erased-cong : ∀ {a} {@0 A : Type a} {@0 x y : A} → [ x ] ≡ [ y ] → Erased (x ≡ y)
 36 | erased-cong p = [ cong erased p ]
 37 | 
 38 | -- Erased (Paths A) → Paths (Erased A) ("erasure extensionality")
 39 | []-cong-type = ∀ {a} {@0 A : Type a} {@0 x y : A} → Erased (x ≡ y) → [ x ] ≡ [ y ]
 40 | 
 41 | -- J₀ and []-cong (with their respective computation rules) are interderivable
 42 | 
 43 | module J₀→[]-cong where
 44 |   postulate
 45 |     J₀ : J₀-type
 46 |     J₀-refl
 47 |       : ∀ {a b} {@0 A : Type a} {@0 x : A} (B : (@0 y : A) → @0 x ≡ y → Type b) (r : B x refl)
 48 |       → J₀ B refl r ≡ r
 49 |     {-# REWRITE J₀-refl #-}
 50 | 
 51 |   []-cong : []-cong-type
 52 |   []-cong {x} [ p ] = J₀ (λ y _ → [ x ] ≡ [ y ]) p refl
 53 | 
 54 |   []-cong-refl
 55 |     : ∀ {a} {@0 A : Type a} {@0 x : A}
 56 |     → []-cong {x = x} [ refl ] ≡ refl
 57 |   []-cong-refl = refl
 58 | 
 59 | module []-cong→J₀ where
 60 |   postulate
 61 |     []-cong : []-cong-type
 62 |     []-cong-refl
 63 |       : ∀ {a} {@0 A : Type a} {@0 x : A}
 64 |       → []-cong {x = x} [ refl ] ≡ refl
 65 |     {-# REWRITE []-cong-refl #-}
 66 | 
 67 |   --                        []-cong                        μ
 68 |   -- Erased (Paths (Erased A)) → Paths (Erased (Erased A)) → Paths (Erased A)
 69 |   stable-≡ : ∀ {@0 A : Type a} {x y : Erased A} → Erased (x ≡ y) → x ≡ y
 70 |   stable-≡ p = cong μ ([]-cong p)
 71 | 
 72 |   --         η               []-cong          erased-cong
 73 |   -- Paths A → Erased (Paths A) → Paths (Erased A) → Erased (Paths A)
 74 |   --           Erased (Paths A)           →          Erased (Paths A)
 75 |   --                                      id
 76 |   []-cong-section'
 77 |     : ∀ {@0 A : Type a} {@0 x y : A} (p : x ≡ y)
 78 |     → erased-cong ([]-cong (η p)) ≡ η p
 79 |   []-cong-section' refl = refl
 80 | 
 81 |   -- We can cancel out η by unique elimination and stability of paths in Erased
 82 |   []-cong-section
 83 |     : ∀ {@0 A : Type a} {@0 x y : A} (@0 p : x ≡ y)
 84 |     → erased-cong ([]-cong [ p ]) ≡ [ p ]
 85 |   []-cong-section p = stable-≡ [ []-cong-section' p ]
 86 | 
 87 |   J₀ : J₀-type
 88 |   J₀ B {y} p r = subst (λ ([ p ]) → B y p) ([]-cong-section p) b'
 89 |     where
 90 |       b' : B y (cong erased ([]-cong [ p ]))
 91 |       b' = J (λ ([ y ]) p → B y (cong erased p)) ([]-cong [ p ]) r
 92 | 
 93 |   J₀-refl
 94 |     : ∀ {a b} {@0 A : Type a} {@0 x : A} (B : (@0 y : A) → @0 x ≡ y → Type b) (r : B x refl)
 95 |     → J₀ B refl r ≡ r
 96 |   J₀-refl B r = refl
 97 | 
 98 | -- Function extensionality implies erasure extensionality
 99 | module funext→[]-cong where
100 |   postulate
101 |     funext : ∀ {a b} → Extensionality a b
102 | 
103 |   -- Direct proof, extracted from "Logical properties of a modality for erasure" (Danielsson 2019)
104 | 
105 |   -- id : Paths (Erased A) → Paths (Erased A)
106 |   --    → {funext}
107 |   --      Paths (Paths (Erased A) → Erased A)
108 |   --    → {uniquely eliminating}
109 |   --      Paths (Erased (Paths (Erased A)) → Erased A)
110 |   --    → {apply p}
111 |   --      Paths (Erased A)
112 |   stable-≡ : ∀ {@0 A : Type a} {x y : Erased A} → Erased (x ≡ y) → x ≡ y
113 |   stable-≡ {A} {x} {y} [ p ] =
114 |     cong (λ (f : x ≡ y → Erased A) → [ f p .erased ])
115 |          (funext (λ (p : x ≡ y) → p))
116 | 
117 |   --                  η                        stable-≡
118 |   -- Erased (Paths A) → Erased (Paths (Erased A)) → Paths (Erased A)
119 |   []-cong : []-cong-type
120 |   []-cong [ p ] = stable-≡ [ cong η p ]
121 | 
122 |   -- Alternative proof: ignoring some details, the types of funext and []-cong look very similar:
123 |   --   funext  : Functions (Paths A) → Paths (Functions A)
124 |   --   []-cong : Erased    (Paths A) → Paths (Erased    A)
125 |   --
126 |   -- If we have inductive types with erased constructors, then we can
127 |   -- present erasure as an *open modality* generated by the subterminal
128 |   -- object with a single erased point (see ErasureOpen):
129 | 
130 |   data Compiling : Type where
131 |     @0 compiling : Compiling
132 | 
133 |   ○_ : Type a → Type a
134 |   ○ A = Compiling → A
135 | 
136 |   ○'_ : ○ Type a → Type a
137 |   ○' A = (n : Compiling) → A n
138 | 
139 |   E→○ : {A : ○ Type a} → Erased (A compiling) → ○' A
140 |   E→○ a compiling = a .erased
141 | 
142 |   ○→E : {A : ○ Type a} → ○' A → Erased (A compiling)
143 |   ○→E f .erased = f compiling
144 | 
145 |   E→○→E : {A : ○ Type a} → (a : Erased (A compiling)) → ○→E (E→○ {A = A} a) ≡ a
146 |   E→○→E _ = refl
147 | 
148 |   -- We don't actually need this
149 |   ○→E→○ : {A : ○ Type a} → (f : ○' A) → E→○ (○→E f) ≡ f
150 |   ○→E→○ f = funext (E→○ [ refl {x = f compiling} ])
151 | 
152 |   -- Since Erased is (equivalent to) a function type, erasure extensionality/[]-cong
153 |   -- is a special case of function extensionality:
154 |   --
155 |   --                                       funext
156 |   -- Erased (Paths A) ≃ (Compiling → Paths A) → Paths (Compiling → A) ≃ Paths (Erased A)
157 |   []-cong' : []-cong-type
158 |   []-cong' {A} {x} {y} p = cong ○→E x'≡y'
159 |     where
160 |       x' y' : ○' E→○ [ A ]
161 |       x' = E→○ [ x ]
162 |       y' = E→○ [ y ]
163 | 
164 |       x'≡y' : x' ≡ y'
165 |       x'≡y' = funext (E→○ p)
166 | 


--------------------------------------------------------------------------------
/shake/HTML/Backend.hs:
--------------------------------------------------------------------------------
  1 | {-# OPTIONS_GHC -Wunused-imports #-}
  2 | 
  3 | -- | Backend for generating highlighted, hyperlinked HTML from Agda sources.
  4 | 
  5 | module HTML.Backend
  6 |   ( htmlBackend
  7 |   , htmlBackend'
  8 |   , HtmlFlags(..)
  9 |   , initialHtmlFlags
 10 |   ) where
 11 | 
 12 | import HTML.Base
 13 | 
 14 | import Prelude hiding ((!!), concatMap)
 15 | 
 16 | import Control.DeepSeq
 17 | import Control.Monad.Trans ( MonadIO )
 18 | import Control.Monad.Except ( MonadError(throwError) )
 19 | 
 20 | import Data.Maybe
 21 | import qualified Data.Map as Map
 22 | import Data.Map (Map)
 23 | 
 24 | import GHC.Generics (Generic)
 25 | 
 26 | import Agda.Interaction.Options
 27 |     ( ArgDescr(ReqArg, NoArg)
 28 |     , OptDescr(..)
 29 |     )
 30 | import Agda.Compiler.Backend
 31 | import Agda.Compiler.Common (curIF)
 32 | import Agda.Interaction.Library.Base (_libName, LibName)
 33 | import Agda.Utils.FileName
 34 | 
 35 | -- | Options for HTML generation
 36 | 
 37 | data HtmlFlags = HtmlFlags
 38 |   { htmlFlagEnabled              :: Bool
 39 |   , htmlFlagDir                  :: FilePath
 40 |   , htmlFlagHighlight            :: HtmlHighlight
 41 |   , htmlFlagHighlightOccurrences :: Bool
 42 |   , htmlFlagCssFile              :: Maybe FilePath
 43 |   , htmlFlagLibToURL             :: LibName -> Maybe String
 44 |   } deriving (Generic)
 45 | 
 46 | instance NFData HtmlFlags
 47 | 
 48 | data HtmlCompileEnv = HtmlCompileEnv
 49 |   { htmlCompileEnvOpts :: HtmlOptions
 50 |   , htmlModuleToURL :: ModuleToURL
 51 |   }
 52 | 
 53 | data HtmlModuleEnv = HtmlModuleEnv
 54 |   { htmlModEnvCompileEnv :: HtmlCompileEnv
 55 |   , htmlModEnvName       :: TopLevelModuleName
 56 |   }
 57 | 
 58 | data HtmlModule = HtmlModule
 59 | data HtmlDef = HtmlDef
 60 | 
 61 | htmlBackend :: Backend
 62 | htmlBackend = Backend htmlBackend'
 63 | 
 64 | htmlBackend' :: Backend' HtmlFlags HtmlCompileEnv HtmlModuleEnv HtmlModule HtmlDef
 65 | htmlBackend' = Backend'
 66 |   { backendName           = "HTML"
 67 |   , backendVersion        = Nothing
 68 |   , options               = initialHtmlFlags
 69 |   , commandLineFlags      = htmlFlags
 70 |   , isEnabled             = htmlFlagEnabled
 71 |   , preCompile            = preCompileHtml
 72 |   , preModule             = preModuleHtml
 73 |   , compileDef            = compileDefHtml
 74 |   , postModule            = postModuleHtml
 75 |   , postCompile           = postCompileHtml
 76 |   -- --only-scope-checking works, but with the caveat that cross-module links
 77 |   -- will not have their definition site populated.
 78 |   , scopeCheckingSuffices = True
 79 |   , mayEraseType          = const $ return False
 80 |   , backendInteractTop    = Nothing
 81 |   , backendInteractHole   = Nothing
 82 |   }
 83 | 
 84 | initialHtmlFlags :: HtmlFlags
 85 | initialHtmlFlags = HtmlFlags
 86 |   { htmlFlagEnabled   = False
 87 |   , htmlFlagDir       = defaultHTMLDir
 88 |   , htmlFlagHighlight = HighlightAll
 89 |   -- Don't enable by default because it causes potential
 90 |   -- performance problems
 91 |   , htmlFlagHighlightOccurrences = False
 92 |   , htmlFlagCssFile              = Nothing
 93 |   , htmlFlagLibToURL             = const Nothing
 94 |   }
 95 | 
 96 | htmlOptsOfFlags :: HtmlFlags -> HtmlOptions
 97 | htmlOptsOfFlags flags = HtmlOptions
 98 |   { htmlOptDir = htmlFlagDir flags
 99 |   , htmlOptHighlight = htmlFlagHighlight flags
100 |   , htmlOptHighlightOccurrences = htmlFlagHighlightOccurrences flags
101 |   , htmlOptCssFile = htmlFlagCssFile flags
102 |   }
103 | 
104 | -- | The default output directory for HTML.
105 | 
106 | defaultHTMLDir :: FilePath
107 | defaultHTMLDir = "html"
108 | 
109 | htmlFlags :: [OptDescr (Flag HtmlFlags)]
110 | htmlFlags =
111 |     [ Option []     ["html"] (NoArg htmlFlag)
112 |                     "generate HTML files with highlighted source code"
113 |     , Option []     ["html-dir"] (ReqArg htmlDirFlag "DIR")
114 |                     ("directory in which HTML files are placed (default: " ++
115 |                      defaultHTMLDir ++ ")")
116 |     , Option []     ["highlight-occurrences"] (NoArg highlightOccurrencesFlag)
117 |                     ("highlight all occurrences of hovered symbol in generated " ++
118 |                      "HTML files")
119 |     , Option []     ["css"] (ReqArg cssFlag "URL")
120 |                     "the CSS file used by the HTML files (can be relative)"
121 |     , Option []     ["html-highlight"] (ReqArg htmlHighlightFlag "[code,all,auto]")
122 |                     ("whether to highlight only the code parts (code) or " ++
123 |                      "the file as a whole (all) or " ++
124 |                      "decide by source file type (auto)")
125 |     ]
126 | 
127 | htmlFlag :: Flag HtmlFlags
128 | htmlFlag o = return $ o { htmlFlagEnabled = True }
129 | 
130 | htmlDirFlag :: FilePath -> Flag HtmlFlags
131 | htmlDirFlag d o = return $ o { htmlFlagDir = d }
132 | 
133 | cssFlag :: FilePath -> Flag HtmlFlags
134 | cssFlag f o = return $ o { htmlFlagCssFile = Just f }
135 | 
136 | highlightOccurrencesFlag :: Flag HtmlFlags
137 | highlightOccurrencesFlag o = return $ o { htmlFlagHighlightOccurrences = True }
138 | 
139 | parseHtmlHighlightFlag :: MonadError String m => String -> m HtmlHighlight
140 | parseHtmlHighlightFlag "code" = return HighlightCode
141 | parseHtmlHighlightFlag "all"  = return HighlightAll
142 | parseHtmlHighlightFlag "auto" = return HighlightAuto
143 | parseHtmlHighlightFlag opt    = throwError $ concat ["Invalid option <", opt, ">, expected ,  or "]
144 | 
145 | htmlHighlightFlag :: String -> Flag HtmlFlags
146 | htmlHighlightFlag opt    o = do
147 |   flag <- parseHtmlHighlightFlag opt
148 |   return $ o { htmlFlagHighlight = flag }
149 | 
150 | runLogHtmlWithMonadDebug :: MonadDebug m => LogHtmlT m a -> m a
151 | runLogHtmlWithMonadDebug = runLogHtmlWith $ reportS "html" 2 -- be less noisy by default
152 | 
153 | preCompileHtml
154 |   :: HtmlFlags
155 |   -> TCM HtmlCompileEnv
156 | preCompileHtml flags = do
157 |   moduleToSourceId <- useTC stModuleToSourceId
158 |   modulesToURL <- Map.traverseWithKey moduleToURL moduleToSourceId
159 |   runLogHtmlWithMonadDebug $ do
160 |     logHtml $ unlines
161 |       [ "Warning: HTML is currently generated for ALL files which can be"
162 |       , "reached from the given module, including library files."
163 |       ]
164 |     let opts = htmlOptsOfFlags flags
165 |     prepareCommonDestinationAssets opts
166 |     return $ HtmlCompileEnv opts modulesToURL
167 |   where
168 |     moduleToURL :: TopLevelModuleName -> SourceFile -> TCM String
169 |     moduleToURL m sf = do
170 |       p <- srcFilePath sf
171 |       agdaLibs <- getAgdaLibFiles p m
172 |       case mapMaybe (htmlFlagLibToURL flags . _libName) agdaLibs of
173 |         u:_ -> pure u
174 |         _ -> fail ("Failed to find link target for file " <> filePath p)
175 | 
176 | preModuleHtml
177 |   :: Applicative m
178 |   => HtmlCompileEnv
179 |   -> IsMain
180 |   -> TopLevelModuleName
181 |   -> Maybe FilePath
182 |   -> m (Recompile HtmlModuleEnv HtmlModule)
183 | preModuleHtml cenv _isMain modName _ifacePath = pure $ Recompile (HtmlModuleEnv cenv modName)
184 | 
185 | compileDefHtml
186 |   :: Applicative m
187 |   => HtmlCompileEnv
188 |   -> HtmlModuleEnv
189 |   -> IsMain
190 |   -> Definition
191 |   -> m HtmlDef
192 | compileDefHtml _env _menv _isMain _def = pure HtmlDef
193 | 
194 | postModuleHtml
195 |   :: (MonadIO m, MonadDebug m, ReadTCState m)
196 |   => HtmlCompileEnv
197 |   -> HtmlModuleEnv
198 |   -> IsMain
199 |   -> TopLevelModuleName
200 |   -> [HtmlDef]
201 |   -> m HtmlModule
202 | postModuleHtml env menv _isMain _modName _defs = do
203 |   let generatePage = defaultPageGen (htmlModuleToURL env) . htmlCompileEnvOpts . htmlModEnvCompileEnv $ menv
204 |   htmlSrc <- srcFileOfInterface (htmlModEnvName menv) <$> curIF
205 |   runLogHtmlWithMonadDebug $ generatePage htmlSrc
206 |   return HtmlModule
207 | 
208 | postCompileHtml
209 |   :: Applicative m
210 |   => HtmlCompileEnv
211 |   -> IsMain
212 |   -> Map TopLevelModuleName HtmlModule
213 |   -> m ()
214 | postCompileHtml _cenv _isMain _modulesByName = pure ()
215 | 


--------------------------------------------------------------------------------
/src-1lab/EasyParametricity.lagda.md:
--------------------------------------------------------------------------------
  1 | 
2 | Imports 3 | 4 | ```agda 5 | open import Cat.Prelude hiding (J) 6 | open import Cat.Diagram.Limit.Base 7 | open import Cat.Instances.Discrete 8 | open import Cat.Instances.Shape.Join 9 | open import Cat.Instances.Product 10 | 11 | open import Data.Sum 12 | 13 | import Cat.Reasoning 14 | import Cat.Functor.Reasoning 15 | import Cat.Functor.Bifunctor 16 | 17 | open Precategory 18 | open Functor 19 | open make-is-limit 20 | ``` 21 |
22 | 23 | This module formalises a few very interesting results from Jem Lord's recent work on 24 | [*Easy Parametricity*](https://hott-uf.github.io/2025/abstracts/HoTTUF_2025_paper_21.pdf), 25 | presented at [HoTT/UF 2025](https://hott-uf.github.io/2025/). 26 | 27 | ```agda 28 | module EasyParametricity {u} where 29 | 30 | U = Type u 31 | 𝟘 = Lift u ⊥ 32 | 𝟙 = Lift u ⊤ 33 | 34 | -- We think of functions f : U → A as "bridges" from f 𝟘 to f 𝟙. 35 | record Bridge {ℓ} (A : Type ℓ) (x y : A) : Type (ℓ ⊔ lsuc u) where 36 | no-eta-equality 37 | constructor bridge 38 | pattern 39 | field 40 | app : U → A 41 | app𝟘 : app 𝟘 ≡ x 42 | app𝟙 : app 𝟙 ≡ y 43 | 44 | -- Every function preserves bridges. 45 | ap-bridge 46 | : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) {x y : A} 47 | → Bridge A x y → Bridge B (f x) (f y) 48 | ap-bridge f (bridge app app𝟘 app𝟙) = bridge (f ⊙ app) (ap f app𝟘) (ap f app𝟙) 49 | 50 | postulate 51 | -- An immediate consequence of Jem Lord's parametricity axiom: a function 52 | -- out of U into a U-small type cannot tell 0 and 1 apart; this is all we need here. 53 | -- In other words, U-small types are bridge-discrete. 54 | parametricity : ∀ {A : U} {x y : A} → Bridge A x y → x ≡ y 55 | 56 | -- The type of formal composites r ∘ l : A → B in C. We want to think of this 57 | -- as the type of factorisations of some morphism f : A → B, but it turns out 58 | -- to be unnecessary to track f in the type. 59 | record Factorisation {o ℓ} (C : Precategory o ℓ) (A B : C .Ob) : Type (o ⊔ ℓ) where 60 | constructor factor 61 | module C = Precategory C 62 | field 63 | X : C.Ob 64 | l : C.Hom A X 65 | r : C.Hom X B 66 | 67 | module _ 68 | {o ℓ} {C : Precategory o ℓ} 69 | (let module C = Precategory C) 70 | where 71 | 72 | module _ {A B : C.Ob} (f : C.Hom A B) where 73 | 74 | -- The two factorisations id ∘ f and f ∘ id. 75 | _∘id id∘_ : Factorisation C A B 76 | _∘id = factor A C.id f 77 | id∘_ = factor B f C.id 78 | 79 | module _ 80 | (C-complete : is-complete u lzero C) 81 | (C-category : is-category C) 82 | {A B : C.Ob} (f : C.Hom A B) 83 | where 84 | 85 | -- In a U-complete univalent category, every type of factorisations is bridge-codiscrete. 86 | -- We define a bridge from id ∘ f to f ∘ id. 87 | 88 | factorisation-bridge : Bridge (Factorisation C A B) (id∘ f) (f ∘id) 89 | factorisation-bridge = bridge b b0 b1 where 90 | 91 | b : U → Factorisation C A B 92 | b P = fac (C-complete diagram) module b where 93 | 94 | -- This is the interesting part: given a type P : U, we construct the 95 | -- wide pullback of P-many copies of f. 96 | -- Since we only care about the cases where P is a proposition, we 97 | -- can just take the discrete or codiscrete category on P and adjoin a 98 | -- terminal object to get our diagram shape. 99 | J : Precategory u lzero 100 | J = Codisc' P ▹ 101 | 102 | diagram : Functor J C 103 | diagram .F₀ (inl _) = A 104 | diagram .F₀ (inr _) = B 105 | diagram .F₁ {inl _} {inl _} _ = C.id 106 | diagram .F₁ {inl _} {inr _} _ = f 107 | diagram .F₁ {inr _} {inr _} _ = C.id 108 | diagram .F-id {inl _} = refl 109 | diagram .F-id {inr _} = refl 110 | diagram .F-∘ {inl _} {inl _} {inl _} _ _ = sym (C.idl _) 111 | diagram .F-∘ {inl _} {inl _} {inr _} _ _ = sym (C.idr _) 112 | diagram .F-∘ {inl _} {inr _} {inr _} _ _ = sym (C.idl _) 113 | diagram .F-∘ {inr _} {inr _} {inr _} _ _ = sym (C.idl _) 114 | 115 | -- Given a limit of this diagram (which exists by the assumption of U-completeness), 116 | -- we get a factorisation of f as the universal map followed by the projection to B. 117 | fac : Limit diagram → Factorisation C A B 118 | fac lim = factor X l r where 119 | module lim = Limit lim 120 | X : C.Ob 121 | X = lim.apex 122 | l : C.Hom A X 123 | l = lim.universal (λ { (inl _) → C.id; (inr _) → f }) λ where 124 | {inl _} {inl _} _ → C.idl _ 125 | {inl _} {inr _} _ → C.idr _ 126 | {inr _} {inr _} _ → C.idl _ 127 | r : C.Hom X B 128 | r = lim.cone ._=>_.η (inr tt) 129 | 130 | -- We check that the endpoints of the bridge are what we expect: when P 131 | -- is empty, we are taking the limit of the single-object diagram B, so 132 | -- our factorisation is A → B → B. 133 | b0 : b 𝟘 ≡ id∘ f 134 | b0 = ap (b.fac 𝟘) (Limit-is-prop C-category (C-complete _) (to-limit lim)) where 135 | lim : is-limit (b.diagram 𝟘) B _ 136 | lim = to-is-limit λ where 137 | .ψ (inr _) → C.id 138 | .commutes {inr _} {inr _} _ → C.idl _ 139 | .universal eps comm → eps (inr _) 140 | .factors {inr _} eps comm → C.idl _ 141 | .unique eps comm other fac → sym (C.idl _) ∙ fac (inr _) 142 | 143 | -- When P is contractible, we are taking the limit of the arrow diagram 144 | -- A → B, so our factorisation is A → A → B. 145 | b1 : b 𝟙 ≡ f ∘id 146 | b1 = ap (b.fac 𝟙) (Limit-is-prop C-category (C-complete _) (to-limit lim)) where 147 | lim : is-limit (b.diagram 𝟙) A _ 148 | lim = to-is-limit λ where 149 | .ψ (inl _) → C.id 150 | .ψ (inr _) → f 151 | .commutes {inl _} {inl _} _ → C.idl _ 152 | .commutes {inl _} {inr _} _ → C.idr _ 153 | .commutes {inr _} {inr _} _ → C.idl _ 154 | .universal eps comm → eps (inl (lift tt)) 155 | .factors {inl _} eps comm → C.idl _ 156 | .factors {inr _} eps comm → comm {inl _} {inr _} _ 157 | .unique eps comm other fac → sym (C.idl _) ∙ fac (inl _) 158 | 159 | -- Theorem 1: let C be a U-complete univalent category and D a locally 160 | -- U-small category. 161 | module _ 162 | {o o' ℓ} {C : Precategory o ℓ} {D : Precategory o' u} 163 | (let module C = Cat.Reasoning C) (let module D = Cat.Reasoning D) 164 | (C-complete : is-complete u lzero C) 165 | (C-category : is-category C) 166 | where 167 | 168 | -- 1.a: naturality of transformations between functors C → D is free. 169 | -- (This is a special case of 1.b.) 170 | module _ 171 | (F G : Functor C D) 172 | (let module F = Cat.Functor.Reasoning F) (let module G = Cat.Functor.Reasoning G) 173 | (η : ∀ x → D.Hom (F.₀ x) (G.₀ x)) 174 | where 175 | 176 | natural : is-natural-transformation F G η 177 | natural A B f = G.introl refl ∙ z0≡z1 ∙ (D.refl⟩∘⟨ F.elimr refl) where 178 | 179 | -- Given a factorisation A → X → B, we define the map 180 | -- F A 181 | -- ↓ 182 | -- η X : F X → G X 183 | -- ↓ 184 | -- G B 185 | -- which recovers the naturality square for f as the factorisation varies 186 | -- from id ∘ f to f ∘ id. 187 | z : Factorisation C A B → D.Hom (F.₀ A) (G.₀ B) 188 | z (factor X l r) = G.₁ r D.∘ η X D.∘ F.₁ l 189 | 190 | -- As a result, we get a bridge from one side of the naturality square 191 | -- to the other; since D is locally U-small, the Hom-sets of D are bridge-discrete, 192 | -- so we get the desired equality. 193 | z0≡z1 : z (id∘ f) ≡ z (f ∘id) 194 | z0≡z1 = parametricity (ap-bridge z (factorisation-bridge C-complete C-category f)) 195 | 196 | -- 1.b: dinaturality of transformations between bifunctors C^op × C → D is free. 197 | module _ 198 | (F G : Functor (C ^op ×ᶜ C) D) 199 | (let module F = Cat.Functor.Bifunctor F) (let module G = Cat.Functor.Bifunctor G) 200 | (η : ∀ x → D.Hom (F.₀ (x , x)) (G.₀ (x , x))) 201 | where 202 | 203 | dinatural 204 | : ∀ A B (f : C.Hom A B) 205 | → G.first f D.∘ η B D.∘ F.second f ≡ G.second f D.∘ η A D.∘ F.first f 206 | dinatural A B f = z0≡z1 where 207 | 208 | -- Given a factorisation A → X → B, we define the map 209 | -- F B A → F X X → G X X → G A B 210 | -- which interpolates between the two sides of the dinaturality hexagon. 211 | z : Factorisation C A B → D.Hom (F.₀ (B , A)) (G.₀ (A , B)) 212 | z (factor X l r) = G.₁ (l , r) D.∘ η X D.∘ F.₁ (r , l) 213 | 214 | z0≡z1 : z (id∘ f) ≡ z (f ∘id) 215 | z0≡z1 = parametricity (ap-bridge z (factorisation-bridge C-complete C-category f)) 216 | ``` 217 | -------------------------------------------------------------------------------- /src-1lab/SetsCover.lagda.md: -------------------------------------------------------------------------------- 1 |
2 | Imports 3 | 4 | ```agda 5 | open import Cat.Prelude 6 | open import Cat.Strict 7 | open import Cat.Instances.StrictCat 8 | open import Cat.Skeletal 9 | open import Cat.Gaunt 10 | open import Cat.Functor.Properties 11 | open import Cat.Functor.Equivalence 12 | open import Cat.Functor.Equivalence.Path 13 | open import Cat.Functor.Naturality 14 | open import Cat.Groupoid 15 | open import Cat.Univalent.Rezk 16 | open import Cat.Univalent.Rezk.HIT hiding (C^-elim-set) 17 | open import Cat.Instances.Discrete 18 | 19 | open import 1Lab.Classical 20 | open import 1Lab.Path.Reasoning 21 | open import 1Lab.Function.Surjection 22 | open import 1Lab.Reflection.Induction 23 | 24 | import Cat.Morphism 25 | import Cat.Reasoning 26 | 27 | open is-precat-iso 28 | open is-gaunt 29 | open Functor 30 | 31 | module SetsCover where 32 | 33 | private 34 | unquoteDecl C^-rec = make-rec-n 3 C^-rec (quote C^) 35 | unquoteDecl C^-elim-set = make-elim-n 2 C^-elim-set (quote C^) 36 | ``` 37 |
38 | 39 | # sets cover 40 | 41 | We say that sets cover (types) if every type merely admits a surjection 42 | from a set. 43 | 44 | ```agda 45 | sets-cover : Typeω 46 | sets-cover = ∀ {ℓ} (A : Type ℓ) → ∃[ X ∈ Set ℓ ] ∣ X ∣ ↠ A 47 | ``` 48 | 49 | ## categories 50 | 51 | If sets cover types, then strict categories cover categories in the sense 52 | that every (pre)category is weakly equivalent to (admits a fully faithful, 53 | essentially surjective on objects functor from) a strict one. 54 | 55 | Given a category $C$ and a map $f : X \to C$ where $X$ is a set, we 56 | think of $f$ as a functor from a discrete category and take its 57 | [(bijective on objects, fully faithful) factorisation](https://ncatlab.org/nlab/show/%28bo%2C+ff%29+factorization+system): 58 | when $f$ is surjective, the fully faithful part is also essentially surjective 59 | on objects, so it is a weak equivalence from a strict category. 60 | 61 | ```agda 62 | module _ 63 | {o ℓ} (C : Precategory o ℓ) 64 | ((X , f) : Σ[ X ∈ Set o ] (∣ X ∣ → ⌞ C ⌟)) 65 | where 66 | open Precategory C 67 | 68 | full-image : Precategory o ℓ 69 | full-image = record 70 | { Ob = ⌞ X ⌟ 71 | ; Hom = λ x y → Hom (f x) (f y) 72 | ; Hom-set = λ x y → Hom-set (f x) (f y) 73 | ; id = id ; _∘_ = _∘_ ; idr = idr ; idl = idl ; assoc = assoc } 74 | 75 | forget : Functor full-image C 76 | forget .F₀ = f 77 | forget .F₁ x = x 78 | forget .F-id = refl 79 | forget .F-∘ _ _ = refl 80 | 81 | module _ 82 | (sc : sets-cover) 83 | {o ℓ} (C : Precategory o ℓ) 84 | where 85 | 86 | sets-cover→strict-cats-cover 87 | : ∃[ C' ∈ Precategory o ℓ ] 88 | is-strict C' 89 | × Σ[ F ∈ Functor C' C ] 90 | is-fully-faithful F × is-eso F 91 | sets-cover→strict-cats-cover = do 92 | X , f , f-surj ← sc ⌞ C ⌟ 93 | let 94 | C' = full-image C (X , f) 95 | F = forget C (X , f) 96 | F-eso : is-eso F 97 | F-eso c = f-surj c <&> λ (x , p) → x , path→iso p 98 | pure (C' , hlevel 2 , F , (λ {x y} → id-equiv) , F-eso) 99 | ``` 100 | 101 | Note that we cannot strengthen this to an equivalence in general: 102 | if $F : C \cong C'$ is an equivalence of precategories where $C$ is 103 | univalent and $C'$ is strict, then $C$ is also strict. 104 | If $C'$ is furthermore skeletal, then $F$ is an isomorphism of 105 | precategories (thus $C$ and $C'$ are both gaunt). 106 | 107 | ```agda 108 | module _ 109 | {oc ℓc oc' ℓd'} 110 | (C : Precategory oc ℓc) (C' : Precategory oc' ℓd') 111 | (F : Functor C C') (F-eqv : is-equivalence F) 112 | (C-cat : is-category C) 113 | where 114 | 115 | open is-equivalence F-eqv 116 | private 117 | module C-cat = Univalent C-cat 118 | 119 | univalent≅strict→strict : is-strict C' → is-strict C 120 | univalent≅strict→strict = retract→is-hlevel 2 (F⁻¹ .F₀) (F .F₀) 121 | λ c → sym $ C-cat.iso→path (isoⁿ→iso Id≅F⁻¹∘F c) 122 | 123 | univalent≅skeletal→iso : is-skeletal C' → is-precat-iso F 124 | univalent≅skeletal→iso C'-skeletal .has-is-ff = is-equivalence→is-ff F F-eqv 125 | univalent≅skeletal→iso C'-skeletal .has-is-iso = is-iso→is-equiv λ where 126 | .is-iso.from → F⁻¹ .F₀ 127 | .is-iso.rinv c' → C'-skeletal.to-path (inc (isoⁿ→iso F∘F⁻¹≅Id c')) 128 | .is-iso.linv c → sym $ C-cat.iso→path (isoⁿ→iso Id≅F⁻¹∘F c) 129 | where module C'-skeletal = is-identity-system C'-skeletal 130 | ``` 131 | 132 | ## choice 133 | 134 | If every set is projective ($\mathrm{AC}_\infty$), then every type is 135 | covered by its set truncation. 136 | 137 | ```agda 138 | is-projective : ∀ {ℓ} (A : Type ℓ) → (κ : Level) → Type _ 139 | is-projective A κ = 140 | ∀ (P : A → Type κ) 141 | → (∀ a → ∥ P a ∥) 142 | → ∥ (∀ a → P a) ∥ 143 | 144 | AC∞ = ∀ {ℓ} (A : Set ℓ) κ → is-projective ∣ A ∣ κ 145 | 146 | AC∞→sets-cover 147 | : AC∞ 148 | → ∀ {ℓ} (A : Type ℓ) 149 | → Σ[ X ∈ Set ℓ ] ∥ ∣ X ∣ ↠ A ∥ 150 | AC∞→sets-cover ac∞ A = el! ∥ A ∥₀ , do 151 | split ← ac∞ (el! ∥ A ∥₀) _ (fibre inc) (elim! λ a → inc (a , refl)) 152 | pure $ fst ⊙ split , λ a → do 153 | p ← ∥-∥₀-path.to (split (inc a) .snd) 154 | inc (inc a , p) 155 | ``` 156 | 157 | In the other direction, if choice holds and sets cover, then 158 | $\mathrm{AC}_\infty$ holds. We show that all surjections into a set 159 | split: given a surjection $B \twoheadrightarrow A$ and assuming $X$ 160 | covers $B$, we get a composite surjection $X \twoheadrightarrow A$ where 161 | $X$ and $A$ are both sets, hence a split surjection; this implies that 162 | our original surjection splits. 163 | 164 | ```agda 165 | surjections-split→AC∞ 166 | : (∀ {ℓ κ} (A : Set ℓ) (B : Type κ) (f : B → ∣ A ∣) 167 | → is-surjective f 168 | → is-split-surjective f) 169 | → AC∞ 170 | surjections-split→AC∞ split A κ P h = do 171 | s ← split A (Σ _ P) fst λ a → h a <&> Equiv.from (Fibre-equiv P a) 172 | pure λ a → Equiv.to (Fibre-equiv P a) (s a) 173 | 174 | AC+sets-cover→AC∞ : Axiom-of-choice → sets-cover → AC∞ 175 | AC+sets-cover→AC∞ ac sc = surjections-split→AC∞ λ A B f f-surj → do 176 | X , g , g-surj ← sc B 177 | s ← ac (hlevel 2) (λ _ → hlevel 2) (∘-is-surjective f-surj g-surj) 178 | pure λ a → let (x , p) = s a in g x , p 179 | ``` 180 | 181 | If sets cover, then every projective type is a set, by closure under 182 | retracts. 183 | 184 | ```agda 185 | module _ 186 | (sc : sets-cover) 187 | where 188 | 189 | projective→is-set : ∀ {ℓ} (A : Type ℓ) → is-projective A ℓ → is-set A 190 | projective→is-set A A-proj = ∥-∥-out! do 191 | X , f , f-surj ← sc A 192 | s ← A-proj (fibre f) f-surj 193 | pure (retract→is-hlevel 2 f (fst ⊙ s) (snd ⊙ s) (hlevel 2)) 194 | ``` 195 | 196 | ## truncation 197 | 198 | On the other hand, the fully untruncated version is inconsistent: it 199 | implies that the type of groupoids is a retract of the type of strict 200 | pregroupoids, hence a groupoid itself, which is 201 | [known to be false](https://doi.org/10.48550/arXiv.1311.4002). 202 | 203 | ```agda 204 | StrictGrpd : ∀ ℓ → Type (lsuc ℓ) 205 | StrictGrpd ℓ = Σ[ C ∈ Precategory ℓ ℓ ] (is-strict C × is-pregroupoid C) 206 | 207 | StrictGrpd-is-groupoid : ∀ {ℓ} → is-groupoid (StrictGrpd ℓ) 208 | StrictGrpd-is-groupoid = Equiv→is-hlevel 3 Σ-assoc 209 | (Σ-is-hlevel 3 Strict-cat-is-groupoid λ _ → hlevel 3) 210 | 211 | module _ 212 | (sc 213 | : ∀ {ℓ} (A : Type ℓ) 214 | → Σ[ X ∈ Set ℓ ] ∣ X ∣ ↠ A) 215 | where 216 | 217 | _↓ : ∀ {ℓ} → n-Type ℓ 3 → Precategory ℓ ℓ 218 | G ↓ using X , f , f-surj ← sc ∣ G ∣ 219 | = full-image (Disc ∣ G ∣ (hlevel 3)) (X , f) 220 | 221 | strictify : ∀ {ℓ} → n-Type ℓ 3 → StrictGrpd ℓ 222 | strictify G .fst = G ↓ 223 | strictify G .snd .fst = hlevel 2 224 | strictify G .snd .snd p = make-invertible (sym p) (∙-invl _) (∙-invr _) 225 | where open Cat.Reasoning (G ↓) 226 | 227 | rezk : ∀ {ℓ} → StrictGrpd ℓ → n-Type ℓ 3 228 | rezk (G , G-grpd , G-strict) = el (C^ G) squash 229 | 230 | rezk-strictify : ∀ {ℓ} (G : n-Type ℓ 3) → rezk (strictify G) ≡ G 231 | rezk-strictify G using X , f , f-surj ← sc ∣ G ∣ 232 | = n-ua (compare , compare-is-equiv) 233 | where 234 | open Cat.Reasoning (G ↓) 235 | open is-pregroupoid (strictify G .snd .snd) 236 | 237 | commutes→glueᵀ 238 | : ∀ {x y z} {p : x ≅ y} {q : x ≅ z} {r : y ≅ z} 239 | → to q ≡ to p ∙ to r 240 | → Triangle (glue p) (glue q) (glue r) 241 | commutes→glueᵀ {p = p} {q} {r} t = glueᵀ p r ▷ ap C^.glue (ext (sym t)) 242 | 243 | compare : C^ (G ↓) → ∣ G ∣ 244 | compare = C^-rec (hlevel 3) f to λ _ _ → ∙-filler _ _ 245 | 246 | compare-is-equiv : is-equiv compare 247 | compare-is-equiv .is-eqv g = ∥-∥-out! do 248 | (x , p) ← f-surj g 249 | pure $ contr (inc x , p) $ uncurry $ C^-elim-set 250 | (λ _ → Π-is-hlevel 2 λ _ → fibre-is-hlevel 3 squash (hlevel 3) compare g _ _) 251 | (λ y q → glue (hom→iso (p ∙ sym q)) ,ₚ flip₂ (∙-filler'' p (sym q))) 252 | λ q → funext-dep-i0 λ r → 253 | Σ-set-square (λ _ → hlevel 2) $ commutes→glueᵀ $ 254 | p ∙ sym ⌜ transport (λ z → to q z ≡ g) r ⌝ ≡⟨ ap! (subst-path-left r (to q)) ⟩ 255 | p ∙ sym (sym (to q) ∙ r) ≡⟨ ap (p ∙_) (sym-∙ _ _) ⟩ 256 | p ∙ sym r ∙ to q ≡⟨ ∙-assoc _ _ _ ⟩ 257 | (p ∙ sym r) ∙ to q ∎ 258 | 259 | groupoid-of-groupoids : ∀ {ℓ} → is-groupoid (n-Type ℓ 3) 260 | groupoid-of-groupoids = retract→is-hlevel 3 rezk strictify rezk-strictify 261 | StrictGrpd-is-groupoid 262 | ``` 263 | -------------------------------------------------------------------------------- /src-1lab/TangentBundlesOfSpheres.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | open import 1Lab.Path.Cartesian 3 | open import 1Lab.Path.Reasoning 4 | open import 1Lab.Prelude 5 | 6 | open import Data.Dec 7 | open import Data.Int 8 | open import Data.Nat 9 | open import Data.Sum 10 | 11 | open import Homotopy.Space.Suspension.Properties 12 | open import Homotopy.Connectedness.Automation 13 | open import Homotopy.Space.Sphere.Degree 14 | open import Homotopy.Space.Suspension 15 | open import Homotopy.Connectedness 16 | open import Homotopy.Space.Sphere 17 | open import Homotopy.Conjugation 18 | open import Homotopy.Loopspace 19 | open import Homotopy.Base 20 | ``` 21 | 22 | # tangent bundles of spheres 23 | 24 | A formalisation of [The tangent bundles of spheres](https://www.youtube.com/watch?v=9T9B9XBjVpk) 25 | by David Jaz Myers, Ulrik Buchholtz, Dan Christensen and Egbert Rijke, up until 26 | the proof of the hairy ball theorem. 27 | 28 | ```agda 29 | module TangentBundlesOfSpheres where 30 | 31 | -- Some generalities about wild functors and natural transformations. 32 | 33 | record Functorial (M : Effect) : Typeω where 34 | private module M = Effect M 35 | field 36 | ⦃ Map-Functorial ⦄ : Map M 37 | map-id : ∀ {ℓ} {A : Type ℓ} → map {M} {A = A} id ≡ id 38 | map-∘ 39 | : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} 40 | → {f : B → C} {g : A → B} 41 | → map {M} (f ∘ g) ≡ map f ∘ map g 42 | 43 | map-iso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} 44 | → (e : A ≃ B) → is-iso (map (Equiv.to e)) 45 | map-iso e .is-iso.from = map (Equiv.from e) 46 | map-iso e .is-iso.rinv mb = 47 | map (Equiv.to e) (map (Equiv.from e) mb) ≡˘⟨ map-∘ $ₚ mb ⟩ 48 | map ⌜ Equiv.to e ∘ Equiv.from e ⌝ mb ≡⟨ ap! (funext (Equiv.ε e)) ⟩ 49 | map id mb ≡⟨ map-id $ₚ mb ⟩ 50 | mb ∎ 51 | map-iso e .is-iso.linv ma = 52 | map (Equiv.from e) (map (Equiv.to e) ma) ≡˘⟨ map-∘ $ₚ ma ⟩ 53 | map ⌜ Equiv.from e ∘ Equiv.to e ⌝ ma ≡⟨ ap! (funext (Equiv.η e)) ⟩ 54 | map id ma ≡⟨ map-id $ₚ ma ⟩ 55 | ma ∎ 56 | 57 | map≃ 58 | : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} 59 | → (e : A ≃ B) → M.₀ A ≃ M.₀ B 60 | map≃ e = _ , is-iso→is-equiv (map-iso e) 61 | 62 | map-transport 63 | : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} 64 | → (p : A ≡ B) → map (transport p) ≡ transport (ap M.₀ p) 65 | map-transport {A = A} p i = comp (λ i → M.₀ A → M.₀ (p i)) (∂ i) λ where 66 | j (j = i0) → map-id i 67 | j (i = i0) → map (funextP (transport-filler p) j) 68 | j (i = i1) → funextP (transport-filler (ap M.₀ p)) j 69 | 70 | open Functorial ⦃ ... ⦄ 71 | 72 | is-natural 73 | : ∀ {M N : Effect} (let module M = Effect M; module N = Effect N) ⦃ _ : Map M ⦄ ⦃ _ : Map N ⦄ 74 | → (f : ∀ {ℓ} {A : Type ℓ} → M.₀ A → N.₀ A) → Typeω 75 | is-natural f = 76 | ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {g : A → B} 77 | → ∀ a → map g (f a) ≡ f (map g a) 78 | 79 | -- Operations on suspensions: functorial action, flipping 80 | 81 | instance 82 | Map-Susp : Map (eff Susp) 83 | Map-Susp .Map.map = Susp-map 84 | 85 | Functorial-Susp : Functorial (eff Susp) 86 | Functorial-Susp .Functorial.Map-Functorial = Map-Susp 87 | Functorial-Susp .Functorial.map-id = funext $ Susp-elim _ refl refl λ _ _ → refl 88 | Functorial-Susp .Functorial.map-∘ = funext $ Susp-elim _ refl refl λ _ _ → refl 89 | 90 | flipΣ : ∀ {ℓ} {A : Type ℓ} → Susp A → Susp A 91 | flipΣ north = south 92 | flipΣ south = north 93 | flipΣ (merid a i) = merid a (~ i) 94 | 95 | flipΣ∙ : ∀ {n} → Sⁿ (suc n) →∙ Sⁿ (suc n) 96 | flipΣ∙ = flipΣ , sym (merid north) 97 | 98 | flipΣ-involutive : ∀ {ℓ} {A : Type ℓ} → (p : Susp A) → flipΣ (flipΣ p) ≡ p 99 | flipΣ-involutive = Susp-elim _ refl refl λ _ _ → refl 100 | 101 | flipΣ≃ : ∀ {ℓ} {A : Type ℓ} → Susp A ≃ Susp A 102 | flipΣ≃ = flipΣ , is-involutive→is-equiv flipΣ-involutive 103 | 104 | flipΣ-natural : is-natural flipΣ 105 | flipΣ-natural = Susp-elim _ refl refl λ _ _ → refl 106 | 107 | twist : ∀ {ℓ} {A : Type ℓ} {a b : A} {p q : a ≡ b} (α : p ≡ q) 108 | → PathP (λ i → PathP (λ j → α i j ≡ α j (~ i)) 109 | (λ k → p (~ i ∧ k)) 110 | (λ k → q (~ i ∨ ~ k))) 111 | (λ j k → p (j ∨ k)) 112 | (λ j k → q (j ∧ ~ k)) 113 | twist α i j k = hcomp (∂ i ∨ ∂ j ∨ ∂ k) λ where 114 | l (l = i0) → α (I-interp k i j) (I-interp k j (~ i)) 115 | l (i = i0) → α (~ l ∧ k ∧ j) (k ∨ j) 116 | l (i = i1) → α (l ∨ ~ k ∨ j) (~ k ∧ j) 117 | l (j = i0) → α (~ l ∧ ~ k ∧ i) (k ∧ ~ i) 118 | l (j = i1) → α (l ∨ k ∨ i) (~ k ∨ ~ i) 119 | l (k = i0) → α i j 120 | l (k = i1) → α j (~ i) 121 | 122 | -- Flipping ΣΣA along the first axis is homotopic to flipping along the second axis, 123 | -- by rotating 180°. 124 | rotateΣ : ∀ {ℓ} {A : Type ℓ} → map flipΣ ≡ flipΣ {A = Susp A} 125 | rotateΣ = funext $ Susp-elim _ (merid north) (sym (merid south)) ( 126 | Susp-elim _ (flip₁ (double-connection _ _)) (double-connection _ _) 127 | λ a i j k → hcomp (∂ j ∨ ∂ k) λ where 128 | l (l = i0) → merid (merid a j) i 129 | l (j = i0) → merid north (I-interp l i k) 130 | l (j = i1) → merid south (I-interp l i (~ k)) 131 | l (k = i0) → twist (λ i j → merid (merid a i) j) (~ i) j (~ l) 132 | l (k = i1) → twist (λ i j → merid (merid a i) j) j i l) 133 | 134 | Susp-map∙-flipΣ∙ : ∀ n → flipΣ∙ {n = suc n} ≡ Susp-map∙ flipΣ∙ 135 | Susp-map∙-flipΣ∙ n = sym rotateΣ ,ₚ λ i j → merid north (~ i ∧ ~ j) 136 | 137 | opaque 138 | unfolding Ωⁿ≃∙Sⁿ-map Ω¹-map conj 139 | 140 | degree∙-flipΣ : ∀ n → degree∙ n flipΣ∙ ≡ -1 141 | degree∙-flipΣ zero = refl 142 | degree∙-flipΣ (suc n) = 143 | ap (degree∙ (suc n)) (Susp-map∙-flipΣ∙ n) 144 | ∙∙ degree∙-Susp-map∙ n flipΣ∙ 145 | ∙∙ degree∙-flipΣ n 146 | 147 | flip≠id : ∀ n → ¬ flipΣ ≡ id {A = Sⁿ⁻¹ (suc n)} 148 | flip≠id zero h = subst (Susp-elim _ ⊤ ⊥ (λ ())) (h $ₚ south) _ 149 | flip≠id (suc n) h = negsuc≠pos $ 150 | -1 ≡˘⟨ degree∙-flipΣ n ⟩ 151 | degree∙ n flipΣ∙ ≡˘⟨ degree≡degree∙ n _ ⟩ 152 | degree n flipΣ ≡⟨ ap (degree n) h ⟩ 153 | degree n id ≡⟨ degree≡degree∙ n _ ⟩ 154 | degree∙ n id∙ ≡⟨ degree∙-id∙ n ⟩ 155 | 1 ∎ 156 | 157 | Susp-ua→ 158 | : ∀ {ℓ ℓ'} {A B : Type ℓ} {C : Type ℓ'} 159 | → {e : A ≃ B} {f : Susp A → C} {g : Susp B → C} 160 | → (∀ sa → f sa ≡ g (map (e .fst) sa)) 161 | → PathP (λ i → Susp (ua e i) → C) f g 162 | Susp-ua→ h i north = h north i 163 | Susp-ua→ h i south = h south i 164 | Susp-ua→ {g = g} h i (merid a j) = hcomp (∂ i ∨ ∂ j) λ where 165 | k (k = i0) → g (merid (unglue a) j) 166 | k (i = i0) → h (merid a j) (~ k) 167 | k (i = i1) → g (merid a j) 168 | k (j = i0) → h north (i ∨ ~ k) 169 | k (j = i1) → h south (i ∨ ~ k) 170 | 171 | -- The tangent bundles of spheres 172 | 173 | Tⁿ⁻¹ : ∀ n → Sⁿ⁻¹ n → Type 174 | θⁿ⁻¹ : ∀ n → (p : Sⁿ⁻¹ n) → Susp (Tⁿ⁻¹ n p) ≃ Sⁿ⁻¹ n 175 | 176 | Tⁿ⁻¹ zero () 177 | Tⁿ⁻¹ (suc n) = Susp-elim _ 178 | (Sⁿ⁻¹ n) 179 | (Sⁿ⁻¹ n) 180 | λ p → ua (θⁿ⁻¹ n p e⁻¹ ∙e flipΣ≃ ∙e θⁿ⁻¹ n p) 181 | 182 | θⁿ⁻¹ zero () 183 | θⁿ⁻¹ (suc n) = Susp-elim _ 184 | id≃ 185 | flipΣ≃ 186 | λ p → Σ-prop-pathp! $ Susp-ua→ $ happly $ sym $ 187 | let module θ = Equiv (θⁿ⁻¹ n p) in 188 | flipΣ ∘ map (θ.to ∘ flipΣ ∘ θ.from) ≡⟨ flipΣ ∘⟨ map-∘ ⟩ 189 | flipΣ ∘ map θ.to ∘ map (flipΣ ∘ θ.from) ≡⟨ flipΣ ∘ map _ ∘⟨ map-∘ ⟩ 190 | flipΣ ∘ map θ.to ∘ map flipΣ ∘ map θ.from ≡⟨ flipΣ ∘ map _ ∘⟨ rotateΣ ⟩∘ map _ ⟩ 191 | flipΣ ∘ map θ.to ∘ flipΣ ∘ map θ.from ≡⟨ flipΣ ∘⟨ funext flipΣ-natural ⟩∘ map _ ⟩ 192 | flipΣ ∘ flipΣ ∘ map θ.to ∘ map θ.from ≡⟨ funext flipΣ-involutive ⟩∘⟨refl ⟩ 193 | map θ.to ∘ map θ.from ≡⟨ funext (is-iso.rinv (map-iso (θⁿ⁻¹ n p))) ⟩ 194 | id ∎ 195 | 196 | antipodeⁿ⁻¹ : ∀ n → Sⁿ⁻¹ n ≃ Sⁿ⁻¹ n 197 | antipodeⁿ⁻¹ zero = id≃ 198 | antipodeⁿ⁻¹ (suc n) = map≃ (antipodeⁿ⁻¹ n) ∙e flipΣ≃ 199 | 200 | θN : ∀ n → (p : Sⁿ⁻¹ n) → θⁿ⁻¹ n p .fst north ≡ p 201 | θN (suc n) = Susp-elim _ refl refl λ p → transpose $ 202 | ap sym (∙-idl _ ∙ ∙-idl _ ∙ ∙-elimr (∙-idl _ ∙ ∙-idl _ ∙ ∙-idr _ ∙ ∙-idl _ ∙ ∙-idl _ ∙ ∙-idl _)) 203 | ∙ ap merid (θN n p) 204 | 205 | θS : ∀ n → (p : Sⁿ⁻¹ n) → θⁿ⁻¹ n p .fst south ≡ Equiv.to (antipodeⁿ⁻¹ n) p 206 | θS (suc n) = Susp-elim _ refl refl λ p → transpose $ 207 | ap sym (∙-idl _ ∙ ∙-idl _ ∙ ∙-elimr (∙-idl _ ∙ ∙-idl _ ∙ ∙-idr _ ∙ ∙-idl _ ∙ ∙-idl _ ∙ ∙-idl _)) 208 | ∙ ap (sym ∘ merid) (θS n p) 209 | 210 | cⁿ⁻¹ : ∀ n → (p : Sⁿ⁻¹ n) → Tⁿ⁻¹ n p → p ≡ Equiv.to (antipodeⁿ⁻¹ n) p 211 | cⁿ⁻¹ n p t = sym (θN n p) ∙ ap (θⁿ⁻¹ n p .fst) (merid t) ∙ θS n p 212 | 213 | flipΣⁿ : ∀ n → Sⁿ⁻¹ n → Sⁿ⁻¹ n 214 | flipΣⁿ zero = id 215 | flipΣⁿ (suc n) = if⁺ even-or-odd n then flipΣ else id 216 | 217 | flipΣⁿ⁺² : ∀ n → map (map (flipΣⁿ n)) ≡ flipΣⁿ (suc (suc n)) 218 | flipΣⁿ⁺² zero = ap map map-id ∙ map-id 219 | flipΣⁿ⁺² (suc n) with even-or-odd n 220 | ... | inl e = ap map rotateΣ ∙ rotateΣ 221 | ... | inr o = ap map map-id ∙ map-id 222 | 223 | antipode≡flip : ∀ n → Equiv.to (antipodeⁿ⁻¹ n) ≡ flipΣⁿ n 224 | antipode≡flip zero = refl 225 | antipode≡flip (suc zero) = ap (flipΣ ∘_) map-id 226 | antipode≡flip (suc (suc n)) = 227 | flipΣ ∘ map (flipΣ ∘ map (antipodeⁿ⁻¹ n .fst)) ≡⟨ flipΣ ∘⟨ map-∘ ⟩ 228 | flipΣ ∘ map flipΣ ∘ map (map (antipodeⁿ⁻¹ n .fst)) ≡⟨ flipΣ ∘⟨ rotateΣ ⟩∘ map _ ⟩ 229 | flipΣ ∘ flipΣ ∘ map (map (antipodeⁿ⁻¹ n .fst)) ≡⟨ funext flipΣ-involutive ⟩∘⟨refl ⟩ 230 | map (map (antipodeⁿ⁻¹ n .fst)) ≡⟨ ap (map ∘ map) (antipode≡flip n) ⟩ 231 | map (map (flipΣⁿ n)) ≡⟨ flipΣⁿ⁺² n ⟩ 232 | flipΣⁿ (suc (suc n)) ∎ 233 | 234 | -- If the tangent bundle of the n-sphere admits a section, then we get 235 | -- a homotopy between flipΣⁿ and the identity. 236 | section→homotopy : ∀ n → ((p : Sⁿ⁻¹ n) → Tⁿ⁻¹ n p) → flipΣⁿ n ≡ id 237 | section→homotopy n sec = sym $ funext (λ p → cⁿ⁻¹ n p (sec p)) ∙ antipode≡flip n 238 | 239 | -- Therefore, n must be even. 240 | hairy-ball : ∀ n → ((p : Sⁿ⁻¹ n) → Tⁿ⁻¹ n p) → is-even n 241 | hairy-ball zero sec = ∣-zero 242 | hairy-ball (suc n) sec with even-or-odd n | section→homotopy (suc n) sec 243 | ... | inl e | h = absurd (flip≠id n h) 244 | ... | inr o | _ = o 245 | 246 | T-connected : ∀ n (p : Sⁿ⁻¹ (3 + n)) → is-connected (Tⁿ⁻¹ (3 + n) p) 247 | T-connected n = Susp-elim-prop (λ _ → hlevel 1) (n-connected 2) (n-connected 2) 248 | 249 | -- As a corollary, not all connected types are pointed. 250 | ¬connected→pointed 251 | : ¬ ∀ {A : Type} → is-connected A → A 252 | ¬connected→pointed h = decide! (hairy-ball 3 λ p → h (T-connected 0 p)) 253 | ``` 254 | -------------------------------------------------------------------------------- /shake/Main.hs: -------------------------------------------------------------------------------- 1 | import Agda.Compiler.Backend hiding (getEnv) 2 | import Agda.Interaction.Imports 3 | import Agda.Interaction.Library.Base 4 | import Agda.Interaction.Options (CommandLineOptions(..), defaultOptions) 5 | import Agda.TypeChecking.Errors 6 | import Agda.Utils.FileName 7 | import Agda.Utils.Monad 8 | 9 | import Control.Monad.Error.Class 10 | 11 | import Data.ByteString.Lazy qualified as LBS 12 | import Data.Digest.Pure.SHA 13 | import Data.Foldable 14 | import Data.List 15 | import Data.Map qualified as Map 16 | import Data.Maybe 17 | import Data.Text qualified as T 18 | import Data.Text.IO qualified as T 19 | import Data.Text.Encoding qualified as T 20 | 21 | import Development.Shake 22 | import Development.Shake.Classes 23 | import Development.Shake.FilePath 24 | 25 | import HTML.Backend 26 | import HTML.Base 27 | 28 | import System.Console.GetOpt 29 | import System.Directory 30 | 31 | import Text.HTML.TagSoup 32 | import Text.Pandoc 33 | import Text.Pandoc.Filter 34 | import Text.Pandoc.Walk 35 | 36 | newtype CompileDirectory = CompileDirectory (FilePath, FilePath) 37 | deriving (Show, Typeable, Eq, Hashable, Binary, NFData) 38 | type instance RuleResult CompileDirectory = () 39 | 40 | newtype WriteDiagram = WriteDiagram T.Text 41 | deriving (Show, Typeable, Eq, Hashable, Binary, NFData) 42 | type instance RuleResult WriteDiagram = FilePath 43 | 44 | sourceDir, source1labDir, buildDir, htmlDir, siteDir, diagramsDir, everything, everything1lab :: FilePath 45 | sourceDir = "src" 46 | source1labDir = "src-1lab" 47 | buildDir = "_build" 48 | htmlDir = buildDir "html" 49 | siteDir = buildDir "site" 50 | diagramsDir = buildDir "diagrams" 51 | everything = sourceDir "Everything.agda" 52 | everything1lab = source1labDir "Everything-1lab.agda" 53 | 54 | myHtmlBackend :: Backend 55 | myHtmlBackend = Backend htmlBackend' 56 | { options = initialHtmlFlags 57 | { htmlFlagDir = htmlDir 58 | , htmlFlagHighlightOccurrences = True 59 | , htmlFlagCssFile = Just "style.css" 60 | , htmlFlagHighlight = HighlightCode 61 | , htmlFlagLibToURL = \ (LibName lib version) -> 62 | let v = intercalate "." (map show version) in 63 | Map.lookup lib $ Map.fromList 64 | [ ("agda-builtins", "https://agda.github.io/agda-stdlib/master") 65 | , ("standard-library", "https://agda.github.io/agda-stdlib/v" <> v) 66 | , ("cubical", "https://agda.github.io/cubical") 67 | , ("1lab", "https://1lab.dev") 68 | , ("agda-stuff", "") 69 | ] 70 | } 71 | } 72 | 73 | filenameToModule :: FilePath -> String 74 | filenameToModule f = dropExtensions f 75 | 76 | makeEverythingFile :: [FilePath] -> String 77 | makeEverythingFile = unlines . map (\ m -> "import " <> filenameToModule m) 78 | 79 | readFileText :: FilePath -> Action T.Text 80 | readFileText x = need [x] >> liftIO (T.readFile x) 81 | 82 | importToModule :: String -> String 83 | importToModule s = innerText tags 84 | where tags = parseTags s 85 | 86 | shakeOpts :: ShakeOptions 87 | shakeOpts = shakeOptions { 88 | shakeColor = True 89 | } 90 | 91 | data Flags = SkipAgda deriving Eq 92 | 93 | optDescrs :: [OptDescr (Either String Flags)] 94 | optDescrs = 95 | [ Option [] ["skip-agda"] (NoArg (Right SkipAgda)) "Skip typechecking Agda." 96 | ] 97 | 98 | data SourceType = HTML FilePath | Markdown FilePath 99 | 100 | -- | Takes a base source directory and a source filename and returns the 101 | -- corresponding Markdown source and the path to the output file generated 102 | -- by Agda. 103 | getSourceFile :: FilePath -> FilePath -> Action (T.Text, SourceType) 104 | getSourceFile sourceDir sourceFile = do 105 | source <- readFileText (sourceDir sourceFile) 106 | case takeExtensions sourceFile of 107 | ".agda" -> pure (T.unlines ["```agda", source, "```"], HTML $ htmlDir sourceFile -<.> "html") 108 | ".lagda.md" -> pure (source, Markdown $ htmlDir dropExtensions sourceFile <.> "md") 109 | _ -> fail ("unsupported extension for file " <> sourceFile) 110 | 111 | -- | Compute the scaled height of a diagram (given in SVG), to use as a @style@ tag. 112 | diagramHeight :: FilePath -> Action Double 113 | diagramHeight fp = do 114 | contents <- readFile' fp 115 | let 116 | height (TagOpen "svg" attrs:_) | Just h <- lookup "height" attrs = 117 | fromMaybe h (T.unpack <$> T.stripSuffix "pt" (T.pack h)) 118 | height (_:t) = height t 119 | height [] = error $ "Diagram SVG has no height: " <> fp 120 | it :: Double 121 | it = read (height (parseTags contents)) * 1.9 -- 🔮 122 | pure $! it 123 | 124 | main :: IO () 125 | main = shakeArgsWith shakeOpts optDescrs \ flags targets -> pure $ Just do 126 | let 127 | skipAgda = SkipAgda `elem` flags 128 | 129 | diagramTemplate <- newCache \ () -> readFileText "diagram.template.tex" 130 | 131 | -- Write a diagram's source code to a .tex file with a filename computed 132 | -- from a hash of the code, and return the filename of the corresponding 133 | -- .svg file which can then be `need`ed. 134 | writeDiagram <- (. WriteDiagram) <$> addOracle \ (WriteDiagram contents) -> do 135 | let 136 | diagramDigest = take 12 . showDigest . sha1 . LBS.fromStrict $ T.encodeUtf8 contents 137 | tex = diagramsDir diagramDigest <.> "tex" 138 | template <- diagramTemplate () 139 | writeFileChanged tex 140 | $ T.unpack 141 | $ T.replace "__BODY__" contents 142 | $ template 143 | pure (diagramDigest <.> "svg") 144 | 145 | diagramsDir "*.pdf" %> \ pdf -> do 146 | let texPath = pdf -<.> "tex" 147 | need [texPath] 148 | command_ [EchoStdout False] "pdflatex" 149 | ["-output-directory", diagramsDir, "-interaction=nonstopmode", texPath] 150 | 151 | siteDir "*.svg" %> \ svg -> do 152 | let pdf = diagramsDir takeFileName svg -<.> "pdf" 153 | need [pdf] 154 | liftIO $ createDirectoryIfMissing True siteDir 155 | command_ [] "pdftocairo" ["-svg", pdf, svg] 156 | 157 | let 158 | patchBlock :: Block -> Action Block 159 | -- Add anchor links next to headers 160 | patchBlock (Header i a@(ident, _, _) inl) | ident /= "" = pure $ Header i a $ 161 | inl ++ [Link ("", ["anchor"], [("aria-hidden", "true")]) [] ("#" <> ident, "")] 162 | -- Render commutative diagrams 163 | patchBlock (CodeBlock (id, classes, attrs) contents) | "quiver" `elem` classes = do 164 | let 165 | title = fromMaybe "commutative diagram" (lookup "title" attrs) 166 | diagramName <- writeDiagram contents 167 | need [siteDir diagramName] 168 | 169 | height <- diagramHeight (siteDir diagramName) 170 | let attrs' = ("style", "height: " <> T.pack (show height) <> "px;"):attrs 171 | 172 | pure $ Div ("", ["diagram-container"], []) 173 | [ Plain [ Image (id, "diagram":classes, attrs') [] (T.pack diagramName, title) ] 174 | ] 175 | patchBlock b = pure b 176 | 177 | -- Render a single type-checked module to HTML. 178 | let 179 | renderModule (sourceDir, sourceFile) = do 180 | let 181 | renderMarkdown markdown = do 182 | pandoc <- traced ("pandoc read: " <> sourceFile) $ runIOorExplode do 183 | readMarkdown def { 184 | readerExtensions = foldr enableExtension pandocExtensions [Ext_autolink_bare_uris] 185 | } markdown 186 | pandoc <- walkM patchBlock pandoc 187 | pandoc <- applyJSONFilter def [] "pandoc-katex" pandoc 188 | traced ("pandoc write: " <> sourceFile) $ runIOorExplode do 189 | writeHtml5String def pandoc 190 | (markdown, agdaOutputFile) <- getSourceFile sourceDir sourceFile 191 | if skipAgda then 192 | renderMarkdown markdown 193 | else do 194 | case agdaOutputFile of 195 | HTML file -> do 196 | html <- liftIO $ T.readFile file 197 | pure $ "
" <> html <> "
" 198 | Markdown file -> do 199 | markdown <- liftIO $ T.readFile file 200 | renderMarkdown markdown 201 | 202 | -- Invoke Agda on a source directory and render all the modules in it. 203 | compileDirectory <- (. CompileDirectory) <$> addOracle \ (CompileDirectory (sourceDir, everything)) -> do 204 | librariesFile <- getEnv "AGDA_LIBRARIES_FILE" 205 | sourceFiles <- filter (not . ("Everything*" ?==)) <$> 206 | getDirectoryFiles sourceDir ["//*.agda", "//*.lagda.md"] 207 | -- Write Everything.agda 208 | writeFile' everything (makeEverythingFile sourceFiles) 209 | -- Run Agda on Everything.agda 210 | unless skipAgda $ traced "agda" do 211 | root <- absolute sourceDir 212 | runTCMTopPrettyErrors do 213 | setCommandLineOptions' root defaultOptions 214 | { optOverrideLibrariesFile = librariesFile 215 | , optDefaultLibs = False 216 | } 217 | stBackends `setTCLens` [myHtmlBackend] 218 | sourceFile <- srcFromPath =<< liftIO (absolute everything) 219 | source <- parseSource sourceFile 220 | checkResult <- typeCheckMain TypeCheck source 221 | callBackend "HTML" IsMain checkResult 222 | -- Render modules as needed and insert the results into the module template. 223 | moduleTemplate <- readFileText "module.html" 224 | for_ sourceFiles \ sourceFile -> do 225 | let htmlFile = dropExtensions sourceFile <.> "html" 226 | html <- renderModule (sourceDir, sourceFile) 227 | writeFile' (siteDir htmlFile) 228 | $ T.unpack 229 | $ T.replace "@contents@" html 230 | $ T.replace "@moduleName@" (T.pack $ filenameToModule sourceFile) 231 | $ T.replace "@path@" (T.pack $ sourceDir sourceFile) 232 | $ moduleTemplate 233 | 234 | -- Compile all modules. 235 | "modules" ~> do 236 | compileDirectory (sourceDir, everything) 237 | compileDirectory (source1labDir, everything1lab) 238 | 239 | -- Generate the index. 240 | siteDir "index.html" %> \ index -> do 241 | need ["modules"] 242 | indexTemplate <- readFileText "index.html" 243 | everythingAgda <- 244 | if skipAgda then pure [] 245 | else (<>) 246 | <$> readFileLines (htmlDir "Everything.html") 247 | <*> readFileLines (htmlDir "Everything-1lab.html") 248 | writeFile' index 249 | $ T.unpack 250 | $ T.replace "@contents@" (T.pack $ unlines $ sortOn importToModule $ everythingAgda) 251 | $ indexTemplate 252 | unless skipAgda do 253 | copyFile' (htmlDir "highlight-hover.js") (siteDir "highlight-hover.js") 254 | 255 | siteDir "style.css" %> \ out -> do 256 | copyFileChanged "style.css" out 257 | 258 | siteDir "main.js" %> \ out -> do 259 | copyFileChanged "main.js" out 260 | 261 | "all" ~> do 262 | need 263 | [ siteDir "index.html" 264 | , siteDir "style.css" 265 | , siteDir "main.js" 266 | ] 267 | 268 | want (if null targets then ["all"] else targets) 269 | 270 | runTCMTopPrettyErrors :: TCM a -> IO a 271 | runTCMTopPrettyErrors tcm = do 272 | r <- runTCMTop' $ (Just <$> tcm) `catchError` \err -> do 273 | warnings <- fmap (map show) . prettyTCWarnings' =<< getAllWarningsOfTCErr err 274 | errors <- show <$> prettyError err 275 | let everything = filter (not . null) $ warnings ++ [errors] 276 | unless (null errors) . liftIO . putStr $ unlines everything 277 | pure Nothing 278 | 279 | maybe (fail "Agda compilation failed") pure r 280 | -------------------------------------------------------------------------------- /src-1lab/SyntheticCategoricalDuality.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | open import 1Lab.Reflection.Regularity 3 | open import 1Lab.Path.Cartesian 4 | open import 1Lab.Reflection hiding (absurd) 5 | 6 | open import Cat.Functor.Equivalence.Path 7 | open import Cat.Functor.Equivalence 8 | open import Cat.Prelude hiding (_[_↦_]) 9 | 10 | open import Data.Fin 11 | ``` 12 | 13 | # synthetic categorical duality 14 | 15 | A synthetic account of categorical duality, based on an idea by [**David Wärn**](https://dwarn.se/). 16 | 17 | The theory of categories has a fundamental S₂-symmetry that swaps "source" 18 | and "target", which can be expressed synthetically by defining categories 19 | in the context of the delooping BS₂. 20 | By choosing as our delooping the type of 2-element types, this amounts 21 | to defining categories relative to an arbitrary 2-element type X, which 22 | we can think of as the set {source, target} except we've forgotten 23 | which is which. 24 | Then, instantiating this with a chosen 2-element type recovers usual 25 | categories, and the non-trivial symmetry of BS₂ automatically gives 26 | a symmetry of the type of categories which coincides with the usual 27 | categorical opposite. 28 | 29 | ```agda 30 | module SyntheticCategoricalDuality where 31 | ``` 32 | 33 |
Some auxiliary definitions 34 | 35 | ```agda 36 | private variable 37 | ℓ o h : Level 38 | X O : Type ℓ 39 | H : O → Type ℓ 40 | a b c : O 41 | i j k : X 42 | 43 | excluded-middle : ∀ {x y z : Bool} → x ≠ y → y ≠ z → x ≡ z 44 | excluded-middle {true} {y} {true} x≠y y≠z = refl 45 | excluded-middle {true} {y} {false} x≠y y≠z = absurd (x≠y (sym (x≠false→x≡true y y≠z))) 46 | excluded-middle {false} {y} {true} x≠y y≠z = absurd (x≠y (sym (x≠true→x≡false y y≠z))) 47 | excluded-middle {false} {y} {false} x≠y y≠z = refl 48 | 49 | instance 50 | Extensional-Bool-map 51 | : ∀ {ℓ ℓr} {C : Bool → Type ℓ} → ⦃ e : ∀ {b} → Extensional (C b) ℓr ⦄ 52 | → Extensional ((b : Bool) → C b) ℓr 53 | Extensional-Bool-map ⦃ e ⦄ .Pathᵉ f g = 54 | e .Pathᵉ (f false) (g false) × e .Pathᵉ (f true) (g true) 55 | Extensional-Bool-map ⦃ e ⦄ .reflᵉ f = 56 | e .reflᵉ (f false) , e .reflᵉ (f true) 57 | Extensional-Bool-map ⦃ e ⦄ .idsᵉ .to-path (false≡ , true≡) = funext λ where 58 | true → e .idsᵉ .to-path true≡ 59 | false → e .idsᵉ .to-path false≡ 60 | Extensional-Bool-map ⦃ e ⦄ .idsᵉ .to-path-over (false≡ , true≡) = 61 | Σ-pathp (e .idsᵉ .to-path-over false≡) (e .idsᵉ .to-path-over true≡) 62 | 63 | Extensional-Bool-homotopy 64 | : ∀ {ℓ ℓr} {C : Bool → Type ℓ} → ⦃ e : ∀ {b} {x y : C b} → Extensional (x ≡ y) ℓr ⦄ 65 | → {f g : (b : Bool) → C b} 66 | → Extensional (f ≡ g) ℓr 67 | Extensional-Bool-homotopy ⦃ e ⦄ {f} {g} .Pathᵉ p q = 68 | e .Pathᵉ (p $ₚ false) (q $ₚ false) × e .Pathᵉ (p $ₚ true) (q $ₚ true) 69 | Extensional-Bool-homotopy ⦃ e ⦄ .reflᵉ p = 70 | e .reflᵉ (p $ₚ false) , e .reflᵉ (p $ₚ true) 71 | Extensional-Bool-homotopy ⦃ e ⦄ .idsᵉ .to-path (false≡ , true≡) = funext-square λ where 72 | true → e .idsᵉ .to-path true≡ 73 | false → e .idsᵉ .to-path false≡ 74 | Extensional-Bool-homotopy ⦃ e ⦄ .idsᵉ .to-path-over (false≡ , true≡) = 75 | Σ-pathp (e .idsᵉ .to-path-over false≡) (e .idsᵉ .to-path-over true≡) 76 | 77 | Bool-η : (b : Bool → O) → if (b true) (b false) ≡ b 78 | Bool-η b = ext (refl , refl) 79 | ``` 80 |
81 | 82 | ```agda 83 | -- We define X-(pre)categories relative to a 2-element type X. 84 | module X (o h : Level) (X : Type) (e : ∥ X ≃ Bool ∥) where 85 | ``` 86 | 87 |
Some more auxiliary definitions 88 | 89 | ```agda 90 | private instance 91 | Finite-X : Finite X 92 | Finite-X = ⦇ Equiv→listing (e <&> _e⁻¹) auto ⦈ 93 | 94 | Discrete-X : Discrete X 95 | Discrete-X = Finite→Discrete 96 | 97 | abstract 98 | H-Level-X : H-Level X 2 99 | H-Level-X = Finite→H-Level 100 | 101 | _[_↦_] : (X → O) → X → O → X → O 102 | _[_↦_] b x m i = ifᵈ i ≡? x then m else b i 103 | 104 | assign-id : (b : X → O) → (x : X) → b [ x ↦ b x ] ≡ b 105 | assign-id b x = ext go where 106 | go : ∀ i → (b [ x ↦ b x ]) i ≡ b i 107 | go i with i ≡? x 108 | ... | yes p = ap b (sym p) 109 | ... | no _ = refl 110 | 111 | assign-const : (b : X → O) (i j : X) → j ≠ i → b [ j ↦ b i ] ≡ λ _ → b i 112 | assign-const b i j j≠i = ext go where 113 | go : ∀ k → (b [ j ↦ b i ]) k ≡ b i 114 | go k with k ≡? j 115 | ... | yes _ = refl 116 | ... | no k≠j = ap b $ ∥-∥-out! do 117 | e ← e 118 | pure (subst (λ X → {x y z : X} → x ≠ y → y ≠ z → x ≡ z) 119 | (ua (e e⁻¹)) excluded-middle k≠j j≠i) 120 | 121 | degenerate 122 | : (H : (X → O) → Type h) (b : X → O) (x : X) (f : H b) (id : H (λ _ → b x)) (i : X) 123 | → H (b [ i ↦ b x ]) 124 | -- degenerate H b x f id i with i ≡ᵢ? x 125 | -- ... | yes reflᵢ = subst H (sym (assign-id b x)) f 126 | -- ... | no i≠x = subst H (sym (assign-const b x i (i≠x ⊙ Id≃path.from))) id 127 | -- NOTE performing the with-translation manually somehow results in fewer transports when X = Bool and x = i. 128 | -- I'm not sure what's happening here... 129 | degenerate H b x f id i = go (i ≡ᵢ? x) where 130 | go : Dec (i ≡ᵢ x) → H (b [ i ↦ b x ]) 131 | go (yes reflᵢ) = subst H (sym (assign-id b x)) f 132 | go (no i≠x) = subst H (sym (assign-const b x i (i≠x ⊙ Id≃path.from))) id 133 | ``` 134 |
135 | 136 | ```agda 137 | record XPrecategory : Type (lsuc (o ⊔ h)) where 138 | no-eta-equality 139 | 140 | field 141 | Ob : Type o 142 | 143 | -- Hom is a family indexed over "X-pairs" of objects, or boundaries. 144 | Hom : (X → Ob) → Type h 145 | Hom-set : (b : X → Ob) → is-set (Hom b) 146 | 147 | -- The identity lives over the constant pair. 148 | id : ∀ {x} → Hom λ _ → x 149 | 150 | -- Composition takes an outer boundary b, a middle object and an 151 | -- X-pair of morphisms with the appropriate boundaries and returns 152 | -- a morphism with boundary b. 153 | compose : (b : X → Ob) (m : Ob) → ((x : X) → Hom (b [ x ↦ m ])) → Hom b 154 | 155 | -- We can (and must) state both unit laws at once: given a "direction" x : X 156 | -- and a morphism f with boundary b, we can form the X-pair {f, id} 157 | -- where id lies in the direction x from f, and ask that the 158 | -- composite equal f. 159 | compose-id 160 | : (b : X → Ob) (f : Hom b) (x : X) 161 | → compose b (b x) (degenerate Hom b x f id) ≡ f 162 | 163 | -- TODO: associativity 164 | -- assoc 165 | -- : (b : X → Ob) (m n : Ob) (x : X) 166 | -- → compose b m (λ i → {! !}) ≡ compose b n {! !} 167 | ``` 168 | 169 |
Some lemmas about paths between X-precategories 170 | 171 | ```agda 172 | private 173 | hom-set : ∀ (C : XPrecategory) {b} → is-set (C .XPrecategory.Hom b) 174 | hom-set C = C .XPrecategory.Hom-set _ 175 | 176 | instance 177 | hlevel-proj-xhom : hlevel-projection (quote XPrecategory.Hom) 178 | hlevel-proj-xhom .hlevel-projection.has-level = quote hom-set 179 | hlevel-proj-xhom .hlevel-projection.get-level _ = pure (lit (nat 2)) 180 | hlevel-proj-xhom .hlevel-projection.get-argument (c v∷ _) = pure c 181 | hlevel-proj-xhom .hlevel-projection.get-argument _ = typeError [] 182 | 183 | private unquoteDecl record-iso = declare-record-iso record-iso (quote XPrecategory) 184 | 185 | XPrecategory-path 186 | : ∀ {C D : XPrecategory} (let module C = XPrecategory C; module D = XPrecategory D) 187 | → (ob≡ : C.Ob ≡ D.Ob) 188 | → (hom≡ : PathP (λ i → (X → ob≡ i) → Type h) C.Hom D.Hom) 189 | → (id≡ : PathP (λ i → ∀ {x} → hom≡ i (λ _ → x)) C.id D.id) 190 | → (compose≡ : PathP (λ i → ∀ (b : X → ob≡ i) (m : ob≡ i) (f : ∀ x → hom≡ i (b [ x ↦ m ])) → hom≡ i b) C.compose D.compose) 191 | → C ≡ D 192 | XPrecategory-path ob≡ hom≡ id≡ compose≡ = Iso.injective record-iso 193 | $ Σ-pathp ob≡ $ Σ-pathp hom≡ $ Σ-pathp prop! 194 | $ Σ-pathp id≡ $ Σ-pathp compose≡ $ hlevel 0 .centre 195 | ``` 196 |
197 | 198 | ```agda 199 | open X using (XPrecategory; XPrecategory-path) 200 | 201 | -- We recover categories by choosing a 2-element type X with designated 202 | -- source and target elements. Here we pick the booleans with 203 | -- the convention that true = source and false = target. 204 | 2Precategory : (o h : Level) → Type (lsuc (o ⊔ h)) 205 | 2Precategory o h = XPrecategory o h Bool (inc id≃) 206 | 207 | module _ {o h : Level} where 208 | module B = X o h Bool (inc id≃) 209 | 210 | Precategory→2Precategory : Precategory o h → 2Precategory o h 211 | Precategory→2Precategory C = C' where 212 | module C = Precategory C 213 | open XPrecategory 214 | C' : 2Precategory o h 215 | C' .Ob = C.Ob 216 | C' .Hom b = C.Hom (b true) (b false) 217 | C' .Hom-set b = C.Hom-set _ _ 218 | C' .id = C.id 219 | C' .compose b m f = f true C.∘ f false 220 | C' .compose-id b f true = ap₂ C._∘_ (transport-refl f) (transport-refl C.id) ∙ C.idr f 221 | C' .compose-id b f false = ap₂ C._∘_ (transport-refl C.id) (transport-refl f) ∙ C.idl f 222 | -- C' .assoc = ? 223 | 224 | 2Precategory→Precategory : 2Precategory o h → Precategory o h 225 | 2Precategory→Precategory C' = C where 226 | module C' = XPrecategory C' 227 | open Precategory 228 | C : Precategory o h 229 | C .Ob = C'.Ob 230 | C .Hom a b = C'.Hom (if a b) 231 | C .Hom-set a b = C'.Hom-set _ 232 | C .id = subst C'.Hom (ext (refl , refl)) C'.id 233 | C ._∘_ {a} {b} {c} f g = C'.compose (if a c) b λ where 234 | true → subst C'.Hom (ext (refl , refl)) f 235 | false → subst C'.Hom (ext (refl , refl)) g 236 | C .idr {x} {y} f = 237 | ap (C'.compose (if x y) x) (ext 238 | ( sym (subst-∙ C'.Hom _ _ C'.id) 239 | ∙ ap (λ p → subst C'.Hom p C'.id) (ext (∙-idr refl , ∙-idr refl)) 240 | , ap (λ p → subst C'.Hom p f) (ext (refl , refl)))) 241 | ∙ C'.compose-id (if x y) f true 242 | C .idl {x} {y} f = 243 | ap (C'.compose (if x y) y) (ext 244 | ( ap (λ p → subst C'.Hom p f) (ext (refl , refl)) 245 | , sym (subst-∙ C'.Hom _ _ C'.id) 246 | ∙ ap (λ p → subst C'.Hom p C'.id) (ext (∙-idr refl , ∙-idr refl)))) 247 | ∙ C'.compose-id (if x y) f false 248 | C .assoc = {! !} 249 | 250 | Precategory→2Precategory-is-iso : is-iso Precategory→2Precategory 251 | Precategory→2Precategory-is-iso .is-iso.from = 2Precategory→Precategory 252 | Precategory→2Precategory-is-iso .is-iso.rinv C' = XPrecategory-path _ _ _ _ 253 | refl 254 | (ext λ b → ap C'.Hom (Bool-η b)) 255 | (funextP' λ {a} → to-pathp⁻ (ap (λ p → subst C'.Hom p C'.id) (ext (refl , refl)))) 256 | (funextP λ b → funextP λ m → funext-dep-i1 λ f → 257 | let 258 | path : PathP (λ i → C'.Hom (Bool-η b i)) 259 | (C'.compose (if (b true) (b false)) m λ x → coe1→0 (λ i → C'.Hom (Bool-η b i B.[ x ↦ m ])) (f x)) 260 | (C'.compose b m f) 261 | path i = C'.compose (Bool-η b i) m 262 | λ x → coe1→i (λ i → C'.Hom (Bool-η b i B.[ x ↦ m ])) i (f x) 263 | in 264 | ap (C'.compose (if (b true) (b false)) m) (ext 265 | ( sym (subst-∙ C'.Hom _ _ (f false)) 266 | ∙ ap (λ p → subst C'.Hom p (f false)) (ext (∙-idr refl , ∙-idr refl)) 267 | , sym (subst-∙ C'.Hom _ _ (f true)) 268 | ∙ ap (λ p → subst C'.Hom p (f true)) (ext (∙-idr refl , ∙-idr refl)))) 269 | ◁ path) 270 | where module C' = XPrecategory C' 271 | Precategory→2Precategory-is-iso .is-iso.linv C = Precategory-path F (iso id-equiv id-equiv) 272 | where 273 | module C = Precategory C 274 | open Functor 275 | F : Functor (2Precategory→Precategory (Precategory→2Precategory C)) C 276 | F .F₀ o = o 277 | F .F₁ f = f 278 | F .F-id = transport-refl C.id 279 | F .F-∘ f g = ap₂ C._∘_ (transport-refl f) (transport-refl g) 280 | 281 | Precategory≃2Precategory : Precategory o h ≃ 2Precategory o h 282 | Precategory≃2Precategory = Iso→Equiv (Precategory→2Precategory , Precategory→2Precategory-is-iso) 283 | 284 | -- We get categorical duality from the action of the X-category construction 285 | -- on the non-trivial path Bool ≡ Bool, and we check that this agrees 286 | -- with the usual categorical duality. 287 | duality : 2Precategory o h ≡ 2Precategory o h 288 | duality = ap₂ (XPrecategory _ _) (ua not≃) prop! 289 | 290 | _^Xop : 2Precategory o h → 2Precategory o h 291 | _^Xop = transport duality 292 | 293 | dualities-agree 294 | : (C : Precategory o h) 295 | → Precategory→2Precategory C ^Xop ≡ Precategory→2Precategory (C ^op) 296 | dualities-agree C = XPrecategory-path _ _ _ _ 297 | refl 298 | (ext λ b → ap₂ C.Hom (transport-refl _) (transport-refl _)) 299 | Regularity.reduce! 300 | (to-pathp (ext λ b m f → {! Regularity.reduce! !})) -- this takes too long 301 | where module C = Precategory C 302 | ``` 303 | -------------------------------------------------------------------------------- /src-1lab/ErasureOpen.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | open import 1Lab.Prelude hiding (map) 3 | open import 1Lab.Reflection.Induction 4 | ``` 5 | 6 | # erasure as an open modality 7 | 8 | This post investigates the fact that Agda's erasure modality is an open modality. 9 | Terminology is borrowed and some proofs are extracted from the paper 10 | [Modalities in homotopy type theory](https://arxiv.org/abs/1706.07526) 11 | by Rijke, Shulman and Spitters. 12 | The erasure modality was previously investigated in 13 | [Logical properties of a modality for erasure](https://www.cse.chalmers.se/~nad/publications/danielsson-erased.pdf) 14 | by Danielsson. 15 | 16 | ```agda 17 | module ErasureOpen where 18 | 19 | private variable 20 | ℓ ℓ' : Level 21 | A B : Type ℓ 22 | ``` 23 | 24 | The `Erased` monadic modality, internalising `@0`: 25 | 26 | ```agda 27 | record Erased (@0 A : Type ℓ) : Type ℓ where 28 | constructor [_] 29 | field 30 | @0 erased : A 31 | 32 | open Erased 33 | 34 | η : {@0 A : Type ℓ} → A → Erased A 35 | η x = [ x ] 36 | 37 | μ : {@0 A : Type ℓ} → Erased (Erased A) → Erased A 38 | μ [ [ x ] ] = [ x ] 39 | ``` 40 | 41 | ...is equivalent to the **open** modality `○` induced by the following subsingleton: 42 | 43 | ```agda 44 | data Compiling : Type where 45 | @0 compiling : Compiling 46 | 47 | Compiling-is-prop : is-prop Compiling 48 | Compiling-is-prop compiling compiling = refl 49 | 50 | ○_ : Type ℓ → Type ℓ 51 | ○ A = Compiling → A 52 | 53 | ○'_ : ○ Type ℓ → Type ℓ 54 | ○' A = (c : Compiling) → A c 55 | 56 | infix 30 ○_ ○'_ 57 | 58 | ○→Erased : ○ A → Erased A 59 | ○→Erased a .erased = a compiling 60 | 61 | -- Agda considers clauses that match on erased constructors as erased. 62 | Erased→○ : Erased A → ○ A 63 | Erased→○ a compiling = a .erased 64 | 65 | ○≃Erased : ○ A ≃ Erased A 66 | ○≃Erased = Iso→Equiv (○→Erased , 67 | iso Erased→○ (λ _ → refl) (λ _ → funext λ where compiling → refl)) 68 | 69 | η○ : A → ○ A 70 | η○ a _ = a 71 | ``` 72 | 73 | Since Agda allows erased matches for the empty type, the empty type is 74 | modal; in other words, we are not not `Compiling`. 75 | 76 | ```agda 77 | ¬¬compiling : ¬ ¬ Compiling 78 | ¬¬compiling ¬c with ○→Erased ¬c 79 | ... | () 80 | ``` 81 | 82 | ## open and closed modalities 83 | 84 | The corresponding **closed** modality `●` is given by the join with `Compiling`, 85 | which is equivalent to the following higher inductive type. 86 | 87 | ```agda 88 | data ●_ (A : Type ℓ) : Type ℓ where 89 | -- At runtime, we only have A. 90 | η● : A → ● A 91 | -- At compile time, we also have an erased "cone" that glues all of A together, 92 | -- so that ● A is contractible. 93 | @0 tip : ● A 94 | @0 cone : (a : A) → η● a ≡ tip 95 | 96 | infix 30 ●_ 97 | 98 | unquoteDecl ●-elim = make-elim ●-elim (quote ●_) 99 | 100 | @0 ●-contr : is-contr (● A) 101 | ●-contr {A = A} = contr tip λ a → sym (ps a) where 102 | ps : (a : ● A) → a ≡ tip 103 | ps = ●-elim cone refl λ a i j → cone a (i ∨ j) 104 | ``` 105 | 106 | The rest of this file investigates some properties of open and closed 107 | modalities that are not specific to the `Compiling` proposition we use here. 108 | 109 |
110 | Some common definitions about higher modalities 111 | 112 | ```agda 113 | module Modality 114 | {○_ : ∀ {ℓ} → Type ℓ → Type ℓ} 115 | (η○ : ∀ {ℓ} {A : Type ℓ} → A → ○ A) 116 | (○-elim : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ○ A → Type ℓ'} 117 | → ((a : A) → ○ P (η○ a)) → (a : ○ A) → ○ P a) 118 | (○-elim-β : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ○ A → Type ℓ'} {pη : (a : A) → ○ P (η○ a)} 119 | → (a : A) → ○-elim {P = P} pη (η○ a) ≡ pη a) 120 | (○-≡-modal : ∀ {ℓ} {A : Type ℓ} {x y : ○ A} → is-equiv (η○ {A = x ≡ y})) 121 | where 122 | 123 | modal : Type ℓ → Type ℓ 124 | modal A = is-equiv (η○ {A = A}) 125 | 126 | modal-map : (A → B) → Type _ 127 | modal-map {B = B} f = (b : B) → modal (fibre f b) 128 | 129 | connected : Type ℓ → Type ℓ 130 | connected A = is-contr (○ A) 131 | 132 | connected-map : (A → B) → Type _ 133 | connected-map {B = B} f = (b : B) → connected (fibre f b) 134 | 135 | modal+connected→contr : modal A → connected A → is-contr A 136 | modal+connected→contr A-mod A-conn = Equiv→is-hlevel 0 (η○ , A-mod) A-conn 137 | 138 | modal+connected→equiv : {f : A → B} → modal-map f → connected-map f → is-equiv f 139 | modal+connected→equiv f-mod f-conn .is-eqv b = modal+connected→contr (f-mod b) (f-conn b) 140 | 141 | elim-modal 142 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ○ A → Type ℓ'} 143 | → (∀ a → modal (P a)) 144 | → ((a : A) → P (η○ a)) → (a : ○ A) → P a 145 | elim-modal P-modal pη a = equiv→inverse (P-modal a) (○-elim (λ a → η○ (pη a)) a) 146 | 147 | elim-modal-β 148 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ○ A → Type ℓ'} P-modal {pη : (a : A) → P (η○ a)} 149 | → (a : A) → elim-modal {P = P} P-modal pη (η○ a) ≡ pη a 150 | elim-modal-β P-modal {pη} a = 151 | ap (equiv→inverse (P-modal (η○ a))) (○-elim-β a) 152 | ∙ equiv→unit (P-modal (η○ a)) (pη a) 153 | 154 | map : (A → B) → ○ A → ○ B 155 | map f = ○-elim (η○ ∘ f) 156 | 157 | map-≃ : A ≃ B → (○ A) ≃ (○ B) 158 | map-≃ e = map (e .fst) , is-iso→is-equiv λ where 159 | .is-iso.from → map (Equiv.from e) 160 | .is-iso.rinv → elim-modal (λ _ → ○-≡-modal) λ b → 161 | ap (map (e .fst)) (○-elim-β b) ∙ ○-elim-β (Equiv.from e b) ∙ ap η○ (Equiv.ε e b) 162 | .is-iso.linv → elim-modal (λ _ → ○-≡-modal) λ a → 163 | ap (map (Equiv.from e)) (○-elim-β a) ∙ ○-elim-β (e .fst a) ∙ ap η○ (Equiv.η e a) 164 | 165 | retract-○→modal : (η⁻¹ : ○ A → A) → is-left-inverse η⁻¹ η○ → modal A 166 | retract-○→modal η⁻¹ ret = is-iso→is-equiv $ 167 | iso η⁻¹ (elim-modal (λ _ → ○-≡-modal) λ a → ap η○ (ret a)) ret 168 | 169 | retract→modal 170 | : (f : A → B) (g : B → A) 171 | → is-left-inverse f g → modal A → modal B 172 | retract→modal {B = B} f g ret A-modal = retract-○→modal η⁻¹ linv where 173 | η⁻¹ : ○ B → B 174 | η⁻¹ = f ∘ elim-modal (λ _ → A-modal) g 175 | linv : is-left-inverse η⁻¹ η○ 176 | linv b = ap f (elim-modal-β (λ _ → A-modal) b) ∙ ret b 177 | 178 | modal-≃ : B ≃ A → modal A → modal B 179 | modal-≃ e = retract→modal (Equiv.from e) (Equiv.to e) (Equiv.η e) 180 | 181 | connected-≃ : B ≃ A → connected A → connected B 182 | connected-≃ e A-conn = Equiv→is-hlevel 0 (map-≃ e) A-conn 183 | 184 | ≡-modal : modal A → ∀ {x y : A} → modal (x ≡ y) 185 | ≡-modal A-modal = modal-≃ (ap-equiv (η○ , A-modal)) ○-≡-modal 186 | 187 | PathP-modal : {A : I → Type ℓ} → modal (A i0) → ∀ {x y} → modal (PathP A x y) 188 | PathP-modal {A = A} A-modal {x} {y} = subst modal (sym (PathP≡Path⁻ A x y)) (≡-modal A-modal) 189 | 190 | reflection-modal : modal (○ A) 191 | reflection-modal = is-iso→is-equiv λ where 192 | .is-iso.from → ○-elim id 193 | .is-iso.rinv → elim-modal (λ _ → ○-≡-modal) λ a → ap η○ (○-elim-β a) 194 | .is-iso.linv → ○-elim-β 195 | 196 | Π-modal : {B : A → Type ℓ} → (∀ a → modal (B a)) → modal ((a : A) → B a) 197 | Π-modal B-modal = retract-○→modal 198 | (λ f a → elim-modal (λ _ → B-modal _) (_$ a) f) 199 | (λ f → funext λ a → elim-modal-β (λ _ → B-modal _) f) 200 | 201 | Σ-modal : {B : A → Type ℓ} → modal A → (∀ a → modal (B a)) → modal (Σ A B) 202 | Σ-modal {B = B} A-modal B-modal = retract-○→modal 203 | (Equiv.from Σ-Π-distrib 204 | ( elim-modal (λ _ → A-modal) fst 205 | , elim-modal (λ _ → B-modal _) λ (a , b) → 206 | subst B (sym (elim-modal-β (λ _ → A-modal) (a , b))) b)) 207 | λ (a , b) → 208 | elim-modal-β (λ _ → A-modal) (a , b) 209 | ,ₚ elim-modal-β (λ _ → B-modal _) (a , b) ◁ to-pathp⁻ refl 210 | 211 | η-connected : connected-map (η○ {A = A}) 212 | η-connected a = contr 213 | (○-elim {P = fibre η○} (λ a → η○ (a , refl)) a) 214 | (elim-modal (λ _ → ○-≡-modal) λ (a' , p) → 215 | J (λ a p → ○-elim (λ x → η○ (x , refl)) a ≡ η○ (a' , p)) (○-elim-β a') p) 216 | 217 | ○Σ○≃○Σ : {B : A → Type ℓ} → (○ (Σ A λ a → ○ B a)) ≃ (○ (Σ A B)) 218 | ○Σ○≃○Σ .fst = ○-elim λ (a , b) → map (a ,_) b 219 | ○Σ○≃○Σ .snd = is-iso→is-equiv λ where 220 | .is-iso.from → map (Σ-map₂ η○) 221 | .is-iso.rinv → elim-modal (λ _ → ○-≡-modal) λ (a , b) → 222 | ap (○-elim _) (○-elim-β (a , b)) ∙ ○-elim-β (a , η○ b) ∙ ○-elim-β b 223 | .is-iso.linv → elim-modal (λ _ → ○-≡-modal) λ (a , b) → 224 | ap (map _) (○-elim-β (a , b)) ∙ elim-modal 225 | {P = λ b → ○-elim _ (○-elim _ b) ≡ η○ (a , b)} (λ _ → ○-≡-modal) 226 | (λ b → ap (○-elim _) (○-elim-β b) ∙ ○-elim-β (a , b)) b 227 | 228 | Σ-connected : {B : A → Type ℓ} → connected A → (∀ a → connected (B a)) → connected (Σ A B) 229 | Σ-connected A-conn B-conn = Equiv→is-hlevel 0 (○Σ○≃○Σ e⁻¹) 230 | (connected-≃ (Σ-contr-snd B-conn) A-conn) 231 | 232 | -- Additional properties of *lex* modalities 233 | 234 | module _ (○-lex : ∀ {ℓ} {A : Type ℓ} {a b : A} → (○ (a ≡ b)) ≃ (η○ a ≡ η○ b)) where 235 | ≡-connected : connected A → {x y : A} → connected (x ≡ y) 236 | ≡-connected A-conn = Equiv→is-hlevel 0 ○-lex (Path-is-hlevel 0 A-conn) 237 | 238 | PathP-connected : {A : I → Type ℓ} → connected (A i0) → ∀ {x y} → connected (PathP A x y) 239 | PathP-connected {A = A} A-conn {x} {y} = 240 | subst connected (sym (PathP≡Path⁻ A x y)) (≡-connected A-conn) 241 | ``` 242 |
243 | 244 | `○` and `●` are higher modalities, so we can instantiate this module 245 | for both of them. 246 | 247 | ```agda 248 | ○-elim-○ 249 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ○ A → Type ℓ'} 250 | → ((a : A) → ○ P (η○ a)) → (a : ○ A) → ○ P a 251 | ○-elim-○ {P = P} pη a c = 252 | subst P (funext λ _ → ap a (Compiling-is-prop _ _)) (pη (a c) c) 253 | 254 | ○-≡-modal : {x y : ○ A} → is-equiv (η○ {A = x ≡ y}) 255 | ○-≡-modal = is-iso→is-equiv λ where 256 | .is-iso.from p i compiling → p compiling i compiling 257 | .is-iso.rinv p i compiling j compiling → p compiling j compiling 258 | .is-iso.linv p i j compiling → p j compiling 259 | 260 | ●-elim-● 261 | : ∀ {ℓ ℓ'} {A : Type ℓ} {P : ● A → Type ℓ'} 262 | → ((a : A) → ● P (η● a)) → (a : ● A) → ● P a 263 | ●-elim-● pη = ●-elim pη tip λ _ → is-contr→pathp (λ _ → ●-contr) _ _ 264 | 265 | ●-≡-modal : {x y : ● A} → is-equiv (η● {A = x ≡ y}) 266 | ●-≡-modal = is-iso→is-equiv λ where 267 | .is-iso.from → ●-elim id (is-contr→is-prop ●-contr _ _) 268 | λ p → is-contr→is-set ●-contr _ _ _ _ 269 | .is-iso.rinv → ●-elim (λ _ → refl) (sym (●-contr .paths _)) 270 | λ p → is-set→squarep (λ _ _ → is-contr→is-set ●-contr) _ _ _ _ 271 | .is-iso.linv _ → refl 272 | 273 | module ○ = Modality η○ ○-elim-○ (λ _ → funext λ _ → transport-refl _) ○-≡-modal 274 | module ● = Modality η● ●-elim-● (λ _ → refl) ●-≡-modal 275 | ``` 276 | 277 | Open and closed modalities are lex. 278 | 279 | ```agda 280 | ○-lex : {a b : A} → ○ (a ≡ b) ≃ (η○ a ≡ η○ b) 281 | ○-lex = funext≃ 282 | 283 | module ●-ids {A : Type ℓ} {a : A} where 284 | code : ● A → Type ℓ 285 | code = ●-elim (λ b → ● (a ≡ b)) (Lift _ ⊤) (λ b → ua (is-contr→≃ ●-contr (hlevel 0))) 286 | 287 | code-refl : code (η● a) 288 | code-refl = η● refl 289 | 290 | decode : ∀ b → code b → η● a ≡ b 291 | decode = ●.elim-modal (λ _ → ●.Π-modal λ _ → ●-≡-modal) 292 | λ a → ●.elim-modal (λ _ → ●-≡-modal) (ap η●) 293 | 294 | decode-over : ∀ b (c : code b) → PathP (λ i → code (decode b c i)) code-refl c 295 | decode-over = ●.elim-modal (λ _ → ●.Π-modal λ _ → ●.PathP-modal ●.reflection-modal) 296 | λ a → ●.elim-modal (λ _ → ●.PathP-modal ●.reflection-modal) 297 | λ p i → η● λ j → p (i ∧ j) 298 | 299 | ids : is-based-identity-system (η● a) code code-refl 300 | ids .to-path {b} = decode b 301 | ids .to-path-over {b} = decode-over b 302 | 303 | ●-lex : {a b : A} → ● (a ≡ b) ≃ (η● a ≡ η● b) 304 | ●-lex = based-identity-system-gives-path ●-ids.ids 305 | ``` 306 | 307 | Some equivalences specific to open and closed modalities: 308 | 309 |
310 | `●-modal A ≃ ○ (is-contr A) ≃ is-contr (○ A) = ○-connected A` 311 |
312 | 313 | ```agda 314 | @0 ●-modal→contr : ●.modal A → is-contr A 315 | ●-modal→contr A-modal = Equiv→is-hlevel 0 (η● , A-modal) ●-contr 316 | 317 | contr→●-modal : @0 is-contr A → ●.modal A 318 | contr→●-modal A-contr = ●.retract-○→modal 319 | (●-elim id (A-contr .centre) λ a → sym (A-contr .paths a)) 320 | λ _ → refl 321 | 322 | contr→○-connected : @0 is-contr A → ○.connected A 323 | contr→○-connected A-contr = contr (Erased→○ [ A-contr .centre ]) λ a → 324 | funext λ where compiling → A-contr .paths _ 325 | 326 | @0 ○-connected→contr : ○.connected A → is-contr A 327 | ○-connected→contr A-conn = contr (A-conn .centre compiling) λ a → 328 | A-conn .paths (η○ a) $ₚ compiling 329 | 330 | ○-connected→●-modal : ○.connected A → ●.modal A 331 | ○-connected→●-modal A-conn = contr→●-modal (○-connected→contr A-conn) 332 | ``` 333 | 334 | ## Artin gluing 335 | 336 | We prove an **Artin gluing** theorem: every type `A` is equivalent to a 337 | certain pullback of `○ A` and `● A` over `● ○ A`, which we call `Fracture A`. 338 | Handwaving, this corresponds to decomposing a type into its "compile time" 339 | part and its "runtime" part. 340 | 341 | ```agda 342 | ○→●○ : ○ A → ● ○ A 343 | ○→●○ = η● 344 | 345 | ●→●○ : ● A → ● ○ A 346 | ●→●○ = ●.map η○ 347 | 348 | Fracture : Type ℓ → Type ℓ 349 | Fracture A = Σ (○ A × ● A) λ (o , c) → ○→●○ o ≡ ●→●○ c 350 | 351 | module _ {A : Type ℓ} where 352 | fracture : A → Fracture A 353 | fracture a = (η○ a , η● a) , refl 354 | ``` 355 | 356 | The idea is to prove that the fibres of the `fracture` map are both 357 | `●`-modal and `●`-connected, and hence contractible. 358 | 359 | For the modal part, we observe that an element of the fibre of `fracture` 360 | at a triple `(o : ○ A, c : ● A, p)` can be rearranged into an element 361 | of the fibre of `η○` at `o` (which is `○`-connected, hence `●`-modal) together with 362 | a dependent path whose type is `●`-modal by standard results about higher modalities. 363 | 364 | ```agda 365 | fracture-modal : ●.modal-map fracture 366 | fracture-modal ((o , c) , p) = ●.modal-≃ e $ 367 | ●.Σ-modal (○-connected→●-modal (○.η-connected _)) λ _ → 368 | ●.PathP-modal $ ●.Σ-modal ●.reflection-modal λ _ → ●-≡-modal 369 | where 370 | e : fibre fracture ((o , c) , p) 371 | ≃ Σ (fibre η○ o) λ (a , q) → 372 | PathP (λ i → Σ (● A) λ c → ○→●○ (q i) ≡ ●→●○ c) (η● a , refl) (c , p) 373 | e = Σ-ap-snd (λ _ → ap-equiv (Σ-assoc e⁻¹) ∙e Σ-pathp≃ e⁻¹) ∙e Σ-assoc 374 | ``` 375 | 376 | Almost symmetrically, for the connected part, we rearrange the fibre 377 | into an element of the fibre of `η●` at `c` (which is `●`-connected) together 378 | with a dependent path in the fibres of `○→●○`. Since the latter is 379 | defined as `η●` its fibres are `●`-connected as well, hence the path type 380 | is `●`-connected because `●` is lex. 381 | 382 | ```agda 383 | fracture-connected : ●.connected-map fracture 384 | fracture-connected ((o , c) , p) = ●.connected-≃ e $ 385 | ●.Σ-connected (●.η-connected _) λ _ → 386 | ●.PathP-connected ●-lex (●.η-connected _) 387 | where 388 | e : fibre fracture ((o , c) , p) 389 | ≃ Σ (fibre η● c) λ (a , q) → 390 | PathP (λ i → Σ (○ A) λ o → ○→●○ o ≡ ●→●○ (q i)) (η○ a , refl) (o , p) 391 | e = Σ-ap-snd (λ _ → ap-equiv (Σ-ap-fst ×-swap ∙e Σ-assoc e⁻¹) ∙e Σ-pathp≃ e⁻¹) ∙e Σ-assoc 392 | 393 | fracture-is-equiv : is-equiv fracture 394 | fracture-is-equiv = ●.modal+connected→equiv fracture-modal fracture-connected 395 | 396 | Artin-gluing : A ≃ Fracture A 397 | Artin-gluing = fracture , fracture-is-equiv 398 | ``` 399 | -------------------------------------------------------------------------------- /src-1lab/ObjectClassifiers.lagda.md: -------------------------------------------------------------------------------- 1 | ```agda 2 | module ObjectClassifiers where 3 | ``` 4 | 5 | # object classifiers and univalence 6 | 7 |
8 | 9 | Imports and basic definitions. 10 | 11 | 12 | ```agda 13 | open import 1Lab.Type hiding (⊤) 14 | open import 1Lab.Type.Pointed 15 | open import 1Lab.Type.Sigma 16 | open import 1Lab.Path 17 | open import 1Lab.Equiv 18 | open import 1Lab.HLevel 19 | open import 1Lab.HLevel.Closure 20 | open import 1Lab.Extensionality 21 | open import Data.Id.Base 22 | 23 | open import 1Lab.Univalence 24 | using (Fibre-equiv) -- this does not use univalence 25 | 26 | -- The (homotopy) pullback, defined using the inductive identity type to 27 | -- make some things easier 28 | record 29 | Pullback 30 | {ℓ ℓ' ℓ''} {B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} 31 | (s : B → D) (q : C → D) 32 | : Type (ℓ ⊔ ℓ' ⊔ ℓ'') where 33 | constructor pb 34 | field 35 | pb₁ : B 36 | pb₂ : C 37 | pbeq : s pb₁ ≡ᵢ q pb₂ 38 | 39 | open Pullback 40 | 41 | -- A universe-polymorphic unit type to avoid Lift noise 42 | record ⊤ {u} : Type u where 43 | constructor tt 44 | ``` 45 |
46 | 47 | It is well-known that univalent universes in type theory correspond to 48 | object classifiers in higher topos theory. However, there are a few different ways 49 | to make sense of this internally to type theory itself. 50 | 51 | In what follows, we fix a (Russell) universe $\mathcal{U}$ and call types 52 | in $\mathcal{U}$ *small*. We write 53 | $\mathcal{U}^\mathsf{p} : \mathcal{U}^\bullet \to \mathcal{U}$ for the 54 | universal projection from pointed small types to small 55 | types. We assume function extensionality (implicitly by working in Cubical Agda), 56 | but *not* univalence. 57 | 58 | ```agda 59 | module _ {u} where 60 | 61 | U = Type u 62 | U∙ = Type∙ u 63 | U⁺ = Type (lsuc u) 64 | 65 | Uᵖ : U∙ → U 66 | Uᵖ (A , _) = A 67 | ``` 68 | 69 | We *define* univalence for $\mathcal{U}$ as the statement that the type 70 | $(Y : \mathcal{U}) \times X \simeq Y$ is contractible for all $X : \mathcal{U}$ 71 | (thus that equivalences form an identity system on $\mathcal{U}$). 72 | We will also refer to $n$-univalence, the restriction of that statement 73 | to the sub-universe of $n$-types in $\mathcal{U}$. 74 | 75 | ```agda 76 | univalence = {X : U} → is-contr (Σ U λ Y → X ≃ Y) 77 | ``` 78 | 79 | Given a "base" type $B : \mathcal{U}$, there are two main ways we can express 80 | "type families over $B$" in type theory: as maps *into* $B$ from a small 81 | domain (this is often called the "fibred" perspective, and such maps may be 82 | called "fibrations" or "bundles" over $B$), or as maps *out* of $B$ into 83 | the universe (the "indexed" perspective). We thus define: 84 | 85 | $$ 86 | \begin{align*} 87 | \mathsf{Bun}(B) &= (A : \mathcal{U}) \times \begin{matrix}A \\ \downarrow \\ B\end{matrix}\\ 88 | \mathsf{Fam}(B) &= B \to \mathcal{U} 89 | \end{align*} 90 | $$ 91 | 92 | ```agda 93 | module _ (B : U) where 94 | 95 | Bun : U⁺ 96 | Bun = Σ[ A ∈ U ] (A → B) 97 | 98 | Fam : U⁺ 99 | Fam = B → U 100 | ``` 101 | 102 | We have maps between those in both directions: the map $\chi_B : \mathsf{Bun}(B) \to 103 | \mathsf{Fam}(B)$ assigns to a bundle $A \to B$ its family of *fibres*, while 104 | the map $\phi_B : \mathsf{Fam}(B) \to \mathsf{Bun}(B)$ assigns to a $B$-indexed 105 | family the first projection out of its *total space*. 106 | 107 | ```agda 108 | χ : Bun → Fam 109 | χ (A , p) = fibre p 110 | 111 | φ : Fam → Bun 112 | φ a = Σ B a , fst 113 | ``` 114 | 115 | The [HoTT book](https://homotopytypetheory.org/book/) proves (theorem 4.8.3) that 116 | $\chi_B$ and $\phi_B$ are inverses, assuming that $\mathcal{U}$ is univalent. This is also 117 | formalised in the 1Lab: 118 | 119 | ```agda 120 | _ : is-equiv χ 121 | _ = 1Lab.Univalence.Fibration-equiv {ℓ' = u} .snd 122 | 123 | _ : is-equiv φ 124 | _ = (1Lab.Univalence.Fibration-equiv {ℓ' = u} e⁻¹) .snd 125 | ``` 126 | 127 | A partial converse to this result was formalised independently by 128 | Martín Escardó (in [TypeTopology](https://martinescardo.github.io/TypeTopology/UF.Classifiers.html)) and by Amélia Liao (in [this math.SE answer](https://math.stackexchange.com/a/4364416)): both of them show that univalence follows from $\chi_B$ being a fibrewise equivalence 129 | **if one also assumes $(-2)$-univalence**, i.e. that any two contractible 130 | types are equal^[This is implied by propositional extensionality, i.e. $(-1)$-univalence.] (as well as function extensionality). 131 | In [this Mastodon post](https://types.pl/@ncf/114737741485982172) I gave 132 | an example of a Tarski universe containing two distinct unit types and which 133 | can be closed under $\Sigma$- and identity types in such a way that $\chi_B$ and 134 | $\phi_B$ are both fibrewise equivalences, but which is not $n$-univalent for any $n$. 135 | This suggests that one cannot get rid of the $(-2)$-univalence assumption. 136 | 137 | In this post I would like to suggest an explanation for *why* the converse of theorem 4.8.3 fails. 138 | The key is to notice that $\mathsf{Bun}(B)$ and $\mathsf{Fam}(B)$ can be thought of as 139 | the endpoints of a *span* of types whose apex is the type of *pullback squares* with 140 | bottom-left corner $B$ and right-hand side map $\mathcal{U}^\mathsf{p}$: 141 | 142 | ~~~{.quiver} 143 | % https://q.uiver.app/#q=WzAsNCxbMCwxLCJCIl0sWzEsMSwiXFxtYXRoY2Fse1V9Il0sWzEsMCwiXFxtYXRoY2Fse1V9XlxcYnVsbGV0Il0sWzAsMCwiQSIsWzI3MCw2MCw2MCwxXV0sWzAsMSwiYSIsMix7ImNvbG91ciI6WzI3MCw2MCw2MF19LFsyNzAsNjAsNjAsMV1dLFsyLDEsIlxcbWF0aGNhbHtVfV5cXG1hdGhzZntwfSJdLFszLDAsInAiLDIseyJjb2xvdXIiOlsyNzAsNjAsNjBdfSxbMjcwLDYwLDYwLDFdXSxbMywyLCJhXisiLDAseyJjb2xvdXIiOlsyNzAsNjAsNjBdfSxbMjcwLDYwLDYwLDFdXSxbMywxLCIiLDEseyJjb2xvdXIiOlsyNzAsNjAsNjBdLCJzdHlsZSI6eyJuYW1lIjoiY29ybmVyIn19XV0= 144 | \[\begin{tikzcd} 145 | \textcolor{accent2}{A} & {\mathcal{U}^\bullet} \\ 146 | B & {\mathcal{U}} 147 | \arrow["{a^+}", color={accent2}, from=1-1, to=1-2] 148 | \arrow["p"', color={accent2}, from=1-1, to=2-1] 149 | \arrow["\lrcorner"{anchor=center, pos=0.125}, color={accent2}, draw=none, from=1-1, to=2-2] 150 | \arrow["{\mathcal{U}^\mathsf{p}}", from=1-2, to=2-2] 151 | \arrow["a"', color={accent2}, from=2-1, to=2-2] 152 | \end{tikzcd}\] 153 | ~~~ 154 | 155 | Let us call such a pullback square a *classification* over $B$: it consists of 156 | a fibration over $B$ that is classified by a given map $a : B \to \mathcal{U}$. 157 | 158 | To define the type $\mathsf{Cls}(B)$ of classifications, we take a shortcut: instead of recording three maps and a 159 | universal property, we only record the top-left corner $A$ and the bottom map $a$, 160 | and ask for an equivalence between $A$ and a given pullback of $\mathcal{U}^\mathsf{p}$ 161 | along $a$. This fully determines the projection maps out of $A$, 162 | so we are justified in leaving them out by function extensionality and contractibility of singletons. 163 | Note that we *cannot* similarly contract $A$ away as that would implicitly assume univalence. 164 | 165 | ```agda 166 | record Cls : U⁺ where 167 | constructor cls 168 | field 169 | A : U 170 | a : B → U 171 | classifies : A ≃ Pullback a Uᵖ 172 | 173 | module c = Equiv classifies 174 | 175 | p : A → B 176 | p = pb₁ ∘ c.to 177 | 178 | a⁺ : A → U∙ 179 | a⁺ = pb₂ ∘ c.to 180 | ``` 181 | 182 | We obtain the span promised above by projecting the left and bottom maps of the 183 | pullback square, respectively. 184 | 185 | ~~~{.quiver} 186 | % https://q.uiver.app/#q=WzAsMyxbMCwwLCJcXG1hdGhzZntCdW59KEIpIl0sWzEsMCwiXFxtYXRoc2Z7Q2xzfShCKSJdLFsyLDAsIlxcbWF0aHNme0ZhbX0oQikiXSxbMSwwLCJcXHBpX1xcZG93bmFycm93IiwyXSxbMSwyLCJcXHBpX1xccmlnaHRhcnJvdyJdXQ== 187 | \[\begin{tikzcd} 188 | {\mathsf{Bun}(B)} & {\mathsf{Cls}(B)} & {\mathsf{Fam}(B)} 189 | \arrow["{\pi_\downarrow}"', from=1-2, to=1-1] 190 | \arrow["{\pi_\rightarrow}", from=1-2, to=1-3] 191 | \end{tikzcd}\] 192 | ~~~ 193 | 194 | ```agda 195 | π↓ : Cls → Bun 196 | π↓ p = _ , Cls.p p 197 | 198 | π→ : Cls → Fam 199 | π→ = Cls.a 200 | ``` 201 | 202 | Observe that both $\pi_\downarrow$ and $\pi_\to$ have sections, given by 203 | taking fibres and total spaces respectively, and that composing one 204 | section with the other projection recovers the maps $\chi_B$ and $\phi_B$. 205 | 206 | ~~~{.quiver} 207 | % https://q.uiver.app/#q=WzAsMyxbMCwwLCJcXG1hdGhzZntCdW59KEIpIl0sWzEsMCwiXFxtYXRoc2Z7Q2xzfShCKSJdLFsyLDAsIlxcbWF0aHNme0ZhbX0oQikiXSxbMSwwLCJcXHBpX1xcZG93bmFycm93IiwwLHsib2Zmc2V0IjotMn1dLFsxLDIsIlxccGlfXFxyaWdodGFycm93IiwwLHsib2Zmc2V0IjotMn1dLFswLDEsIlxcc2lnbWFfXFxkb3duYXJyb3ciLDAseyJvZmZzZXQiOi0yfV0sWzIsMSwiXFxzaWdtYV9cXHRvIiwwLHsib2Zmc2V0IjotMn1dLFswLDIsIlxcY2hpX0IiLDAseyJjdXJ2ZSI6LTN9XSxbMiwwLCJcXHBoaV9CIiwwLHsiY3VydmUiOi0zfV1d 208 | \[\begin{tikzcd} 209 | {\mathsf{Bun}(B)} & {\mathsf{Cls}(B)} & {\mathsf{Fam}(B)} 210 | \arrow["{\sigma_\downarrow}", shift left=2, from=1-1, to=1-2] 211 | \arrow["{\chi_B}", curve={height=-35pt}, from=1-1, to=1-3] 212 | \arrow["{\pi_\downarrow}", shift left=2, from=1-2, to=1-1] 213 | \arrow["{\pi_\rightarrow}", shift left=2, from=1-2, to=1-3] 214 | \arrow["{\phi_B}", curve={height=-35pt}, from=1-3, to=1-1] 215 | \arrow["{\sigma_\to}", shift left=2, from=1-3, to=1-2] 216 | \end{tikzcd}\] 217 | ~~~ 218 | 219 | ```agda 220 | σ↓ : Bun → Cls 221 | σ↓ (A , p) = λ where 222 | .Cls.A → A 223 | .Cls.a → fibre p 224 | .Cls.classifies .fst a → pb (p a) (fibre p (p a) , a , refl) reflᵢ 225 | .Cls.classifies .snd → is-iso→is-equiv λ where 226 | .is-iso.from (pb b (_ , (a , _)) reflᵢ) → a 227 | .is-iso.rinv (pb b (_ , (a , eq)) reflᵢ) i → 228 | pb (eq i) (fibre p (eq i) , a , λ j → eq (i ∧ j)) reflᵢ 229 | .is-iso.linv _ → refl 230 | 231 | section↓ : is-right-inverse σ↓ π↓ 232 | section↓ (A , p) = refl 233 | 234 | _ : π→ ∘ σ↓ ≡ χ 235 | _ = refl 236 | 237 | σ→ : Fam → Cls 238 | σ→ a = λ where 239 | .Cls.A → Σ B a 240 | .Cls.a → a 241 | .Cls.classifies .fst (b , x) → pb b (a b , x) reflᵢ 242 | .Cls.classifies .snd → is-iso→is-equiv λ where 243 | .is-iso.from (pb b (_ , x) reflᵢ) → b , x 244 | .is-iso.rinv (pb b (_ , x) reflᵢ) → refl 245 | .is-iso.linv _ → refl 246 | 247 | section→ : is-right-inverse σ→ π→ 248 | section→ a = refl 249 | 250 | _ : π↓ ∘ σ→ ≡ φ 251 | _ = refl 252 | ``` 253 | 254 | Now for the main point of this post: **each of $\pi_\downarrow$ and $\pi_\to$ 255 | is an equivalence if and only if univalence holds**. 256 | This is reflected in the 1Lab's proof of [`Fibration-equiv`{.Agda}](https://1lab.dev/1Lab.Univalence#Fibration-equiv), which uses 257 | univalence *twice*: once in the left inverse proof and once in the right inverse proof. 258 | This suggests that the two uses of univalence *cancel each other out* in a certain 259 | sense, so that we lose information if we only ask for a composite equivalence 260 | between $\mathsf{Bun}(B)$ and $\mathsf{Fam}(B)$. 261 | 262 | In more detail, the statement that $\pi_\downarrow$ is an equivalence says that 263 | every bundle $A \to B$ is classified by a unique^[That is, a contractible space of.] pullback square into 264 | $\mathcal{U}^\mathsf{p}$; this looks a lot like the definition of an 265 | [object classifier](https://ncatlab.org/nlab/show/%28sub%29object+classifier+in+an+%28infinity%2C1%29-topos) in higher topos theory! 266 | On the other hand, the statement that $\pi_\to$ is an equivalence says that 267 | every map $B \to \mathcal{U}$ has a unique pullback along $\mathcal{U}^\mathsf{p}$; 268 | in other words, that this pullback is determined uniquely up to *identity* 269 | rather than just equivalence. 270 | Notice that both statements are roughly of the form "there is a unique pullback square", 271 | but that they quantify over different *parts* of the pullback square. 272 | 273 | To see the forward implications, we specialise $B$ to the unit type; 274 | in this case, the span simplifies to the two projections out of the type of 275 | equivalences in $\mathcal{U}$: 276 | 277 | ~~~{.quiver} 278 | % https://q.uiver.app/#q=WzAsNixbMCwwLCJcXG1hdGhzZntCdW59KFxcdG9wKSJdLFsxLDAsIlxcbWF0aHNme0Nsc30oXFx0b3ApIl0sWzIsMCwiXFxtYXRoc2Z7RmFtfShcXHRvcCkiXSxbMCwxLCJcXG1hdGhjYWx7VX0iXSxbMSwxLCIoWCwgWSA6IFxcbWF0aGNhbHtVfSkgXFx0aW1lcyBYIFxcc2ltZXEgWSJdLFsyLDEsIlxcbWF0aGNhbHtVfSJdLFsxLDAsIlxccGlfXFxkb3duYXJyb3ciLDJdLFsxLDIsIlxccGlfXFx0byJdLFswLDMsIlxcc2ltIiwyXSxbMSw0LCJcXHNpbSIsMl0sWzIsNSwiXFxzaW0iLDJdLFs0LDUsIlxccGleXFxzaW1lcV8yIiwyXSxbNCwzLCJcXHBpXlxcc2ltZXFfMSJdXQ== 279 | \[\begin{tikzcd} 280 | {\mathsf{Bun}(\top)} & {\mathsf{Cls}(\top)} & {\mathsf{Fam}(\top)} \\ 281 | {\mathcal{U}} & {(X, Y : \mathcal{U}) \times X \simeq Y} & {\mathcal{U}} 282 | \arrow["\sim"', from=1-1, to=2-1] 283 | \arrow["{\pi_\downarrow}"', from=1-2, to=1-1] 284 | \arrow["{\pi_\to}", from=1-2, to=1-3] 285 | \arrow["\sim"', from=1-2, to=2-2] 286 | \arrow["\sim"', from=1-3, to=2-3] 287 | \arrow["{\pi^\simeq_1}", from=2-2, to=2-1] 288 | \arrow["{\pi^\simeq_2}"', from=2-2, to=2-3] 289 | \end{tikzcd}\] 290 | ~~~ 291 | 292 | ```agda 293 | Equiv : U⁺ 294 | Equiv = Σ[ X ∈ U ] Σ[ Y ∈ U ] X ≃ Y 295 | 296 | π₁≃ : Equiv → U 297 | π₁≃ (X , Y , e) = X 298 | 299 | π₂≃ : Equiv → U 300 | π₂≃ (X , Y , e) = Y 301 | 302 | Fam⊤≃U : Fam ⊤ ≃ U 303 | Fam⊤≃U .fst a = a _ 304 | Fam⊤≃U .snd = is-iso→is-equiv λ where 305 | .is-iso.from X _ → X 306 | .is-iso.rinv X → refl 307 | .is-iso.linv a → refl 308 | 309 | Bun⊤≃U : Bun ⊤ ≃ U 310 | Bun⊤≃U .fst (A , _) = A 311 | Bun⊤≃U .snd = is-iso→is-equiv λ where 312 | .is-iso.from X → X , _ 313 | .is-iso.rinv _ → refl 314 | .is-iso.linv _ → refl 315 | 316 | Pullback≃a : {a : ⊤ {u} → U} → Pullback a Uᵖ ≃ a _ 317 | Pullback≃a .fst (pb _ (_ , b) reflᵢ) = b 318 | Pullback≃a .snd = is-iso→is-equiv λ where 319 | .is-iso.from b → pb _ (_ , b) reflᵢ 320 | .is-iso.rinv _ → refl 321 | .is-iso.linv (pb _ (_ , b) reflᵢ) → refl 322 | 323 | Cls⊤≃Equiv : Cls ⊤ ≃ Equiv 324 | Cls⊤≃Equiv .fst (cls A a e) = A , a _ , e ∙e Pullback≃a 325 | Cls⊤≃Equiv .snd = is-iso→is-equiv λ where 326 | .is-iso.from (X , Y , e) → cls X (λ _ → Y) (e ∙e Pullback≃a e⁻¹) 327 | .is-iso.rinv (X , Y , e) → refl ,ₚ refl ,ₚ ext λ _ → refl 328 | .is-iso.linv (cls A a e) → 329 | let 330 | h : (e ∙e Pullback≃a) ∙e Pullback≃a e⁻¹ ≡ e 331 | h = ext λ _ → Equiv.η Pullback≃a _ 332 | in λ i → cls A a (h i) 333 | 334 | _ : π₁≃ ≡ Equiv.to Bun⊤≃U ∘ π↓ ⊤ ∘ Equiv.from Cls⊤≃Equiv 335 | _ = refl 336 | 337 | _ : π₂≃ ≡ Equiv.to Fam⊤≃U ∘ π→ ⊤ ∘ Equiv.from Cls⊤≃Equiv 338 | _ = refl 339 | ``` 340 | 341 | This makes the proof immediate after noting that the fibre of $\pi^\simeq_1$ at 342 | $X$ is equivalent to $(Y : \mathcal{U}) \times X \simeq Y$ (a general fact 343 | that does not require univalence or even function extensionality), and 344 | symmetrically for $\pi^\simeq_2$. 345 | 346 | ```agda 347 | π₁≃-equiv→univalence : is-equiv π₁≃ → univalence 348 | π₁≃-equiv→univalence h {X} = Equiv→is-hlevel 0 349 | (Fibre-equiv (λ X → Σ U λ Y → X ≃ Y) X e⁻¹) 350 | (h .is-eqv X) 351 | 352 | π↓-equiv→univalence : is-left-inverse (σ↓ ⊤) (π↓ ⊤) → univalence 353 | π↓-equiv→univalence h = π₁≃-equiv→univalence $ 354 | ∘-is-equiv (Bun⊤≃U .snd) 355 | (∘-is-equiv π↓-is-equiv 356 | ((Cls⊤≃Equiv e⁻¹) .snd)) 357 | where 358 | π↓-is-equiv : is-equiv (π↓ ⊤) 359 | π↓-is-equiv = is-iso→is-equiv (iso (σ↓ _) (section↓ _) h) 360 | 361 | sym≃ : Equiv ≃ Equiv 362 | sym≃ .fst (X , Y , e) = Y , X , e e⁻¹ 363 | sym≃ .snd = is-involutive→is-equiv λ (X , Y , e) → 364 | refl ,ₚ refl ,ₚ ext λ _ → refl 365 | 366 | π→-equiv→univalence : is-left-inverse (σ→ ⊤) (π→ ⊤) → univalence 367 | π→-equiv→univalence h = π₁≃-equiv→univalence $ 368 | ∘-is-equiv (Fam⊤≃U .snd) 369 | (∘-is-equiv π→-is-equiv 370 | (∘-is-equiv ((Cls⊤≃Equiv e⁻¹) .snd) 371 | (sym≃ .snd))) 372 | where 373 | π→-is-equiv : is-equiv (π→ ⊤) 374 | π→-is-equiv = is-iso→is-equiv (iso (σ→ _) (section→ _) h) 375 | ``` 376 | 377 | It also makes it easy to see the complete loss of information in assuming 378 | an equivalence $\mathsf{Bun}(\top) \simeq \mathsf{Fam}(\top)$, since this 379 | amounts to the trivial $\mathcal{U} \simeq \mathcal{U}$! Assuming that 380 | $\chi_\top$ or $\phi_\top$ is an equivalence doesn't get us very far either, 381 | as in both cases we just get that forming products with a contractible type is an 382 | equivalence. 383 | 384 | ```agda 385 | _ : Equiv.to Bun⊤≃U ∘ π↓ ⊤ ∘ σ→ ⊤ ∘ Equiv.from Fam⊤≃U 386 | ≡ λ X → ⊤ × X 387 | _ = refl 388 | 389 | _ : Equiv.to Fam⊤≃U ∘ π→ ⊤ ∘ σ↓ ⊤ ∘ Equiv.from Bun⊤≃U 390 | ≡ λ X → X × tt ≡ tt 391 | _ = refl 392 | ``` 393 | -------------------------------------------------------------------------------- /shake/HTML/Base.hs: -------------------------------------------------------------------------------- 1 | {-# OPTIONS_GHC -Wunused-imports #-} 2 | 3 | -- | Function for generating highlighted, hyperlinked HTML from Agda 4 | -- sources. 5 | 6 | module HTML.Base 7 | ( HtmlOptions(..) 8 | , HtmlHighlight(..) 9 | , prepareCommonDestinationAssets 10 | , srcFileOfInterface 11 | , defaultPageGen 12 | , MonadLogHtml(logHtml) 13 | , LogHtmlT 14 | , runLogHtmlWith 15 | , ModuleToURL 16 | ) where 17 | 18 | import Prelude hiding ((!!), concatMap) 19 | 20 | import Control.DeepSeq 21 | import Control.Monad 22 | import Control.Monad.Trans ( MonadIO(..), lift ) 23 | import Control.Monad.Trans.Reader ( ReaderT(runReaderT), ask ) 24 | 25 | import Data.Foldable (toList, concatMap) 26 | import Data.Maybe 27 | import qualified Data.IntMap as IntMap 28 | import qualified Data.Map as Map 29 | import Data.Map (Map) 30 | import Data.List.Split (splitWhen) 31 | import Data.Text.Lazy (Text) 32 | import qualified Data.Text.Lazy as T 33 | 34 | import GHC.Generics (Generic) 35 | 36 | import qualified Network.URI.Encode 37 | 38 | import System.FilePath 39 | import System.Directory 40 | 41 | import Text.Blaze.Html5 42 | ( preEscapedToHtml 43 | , toHtml 44 | , stringValue 45 | , Html 46 | , (!) 47 | , Attribute 48 | ) 49 | import qualified Text.Blaze.Html5 as Html5 50 | import qualified Text.Blaze.Html5.Attributes as Attr 51 | import Text.Blaze.Html.Renderer.Text ( renderHtml ) 52 | -- The imported operator (!) attaches an Attribute to an Html value 53 | -- The defined operator (!!) attaches a list of such Attributes 54 | 55 | import Agda.Interaction.Highlighting.Precise hiding (toList) 56 | 57 | import Agda.Syntax.Common 58 | import Agda.Syntax.TopLevelModuleName 59 | 60 | import qualified Agda.TypeChecking.Monad as TCM 61 | ( Interface(..) 62 | ) 63 | 64 | import Agda.Utils.Function 65 | import Agda.Utils.List1 (String1, pattern (:|)) 66 | import qualified Agda.Utils.List1 as List1 67 | import qualified Agda.Utils.IO.UTF8 as UTF8 68 | import Agda.Syntax.Common.Pretty 69 | 70 | import Agda.Utils.Impossible 71 | import Agda.Utils.String (rtrim) 72 | import Agda.Main (printAgdaDataDir) 73 | 74 | import System.IO.Silently 75 | 76 | type ModuleToURL = Map TopLevelModuleName String 77 | 78 | getDataFileName :: FilePath -> IO FilePath 79 | getDataFileName name = do 80 | dir <- capture_ printAgdaDataDir -- ... 81 | return (rtrim dir name) 82 | 83 | -- | The Agda data directory containing the files for the HTML backend. 84 | 85 | htmlDataDir :: FilePath 86 | htmlDataDir = "html" 87 | 88 | -- | The name of the default CSS file. 89 | 90 | defaultCSSFile :: FilePath 91 | defaultCSSFile = "Agda.css" 92 | 93 | -- | The name of the occurrence-highlighting JS file. 94 | 95 | occurrenceHighlightJsFile :: FilePath 96 | occurrenceHighlightJsFile = "highlight-hover.js" 97 | 98 | -- | The directive inserted before the rendered code blocks 99 | 100 | rstDelimiter :: String 101 | rstDelimiter = ".. raw:: html\n" 102 | 103 | -- | The directive inserted before rendered code blocks in org 104 | 105 | orgDelimiterStart :: String 106 | orgDelimiterStart = "#+BEGIN_EXPORT html\n
\n"
107 | 
108 | -- | The directive inserted after rendered code blocks in org
109 | 
110 | orgDelimiterEnd :: String
111 | orgDelimiterEnd = "
\n#+END_EXPORT\n" 112 | 113 | -- | Determine how to highlight the file 114 | 115 | data HtmlHighlight = HighlightAll | HighlightCode | HighlightAuto 116 | deriving (Show, Eq, Generic) 117 | 118 | instance NFData HtmlHighlight 119 | 120 | highlightOnlyCode :: HtmlHighlight -> FileType -> Bool 121 | highlightOnlyCode HighlightAll _ = False 122 | highlightOnlyCode HighlightCode _ = True 123 | highlightOnlyCode HighlightAuto AgdaFileType = False 124 | highlightOnlyCode HighlightAuto MdFileType = True 125 | highlightOnlyCode HighlightAuto RstFileType = True 126 | highlightOnlyCode HighlightAuto OrgFileType = True 127 | highlightOnlyCode HighlightAuto TypstFileType = True 128 | highlightOnlyCode HighlightAuto TreeFileType = True 129 | highlightOnlyCode HighlightAuto TexFileType = False 130 | 131 | -- | Determine the generated file extension 132 | 133 | highlightedFileExt :: HtmlHighlight -> FileType -> String 134 | highlightedFileExt hh ft 135 | | not $ highlightOnlyCode hh ft = "html" 136 | | otherwise = case ft of 137 | AgdaFileType -> "html" 138 | MdFileType -> "md" 139 | RstFileType -> "rst" 140 | TexFileType -> "tex" 141 | OrgFileType -> "org" 142 | TypstFileType -> "typ" 143 | TreeFileType -> "tree" 144 | 145 | -- | Options for HTML generation 146 | 147 | data HtmlOptions = HtmlOptions 148 | { htmlOptDir :: FilePath 149 | , htmlOptHighlight :: HtmlHighlight 150 | , htmlOptHighlightOccurrences :: Bool 151 | , htmlOptCssFile :: Maybe FilePath 152 | } deriving Eq 153 | 154 | -- | Internal type bundling the information related to a module source file 155 | 156 | data HtmlInputSourceFile = HtmlInputSourceFile 157 | { _srcFileModuleName :: TopLevelModuleName 158 | , _srcFileType :: FileType 159 | -- ^ Source file type 160 | , _srcFileText :: Text 161 | -- ^ Source text 162 | , _srcFileHighlightInfo :: HighlightingInfo 163 | -- ^ Highlighting info 164 | } 165 | 166 | -- | Bundle up the highlighting info for a source file 167 | 168 | srcFileOfInterface :: 169 | TopLevelModuleName -> TCM.Interface -> HtmlInputSourceFile 170 | srcFileOfInterface m i = HtmlInputSourceFile m (TCM.iFileType i) (TCM.iSource i) (TCM.iHighlighting i) 171 | 172 | -- | Logging during HTML generation 173 | 174 | type HtmlLogMessage = String 175 | type HtmlLogAction m = HtmlLogMessage -> m () 176 | 177 | class MonadLogHtml m where 178 | logHtml :: HtmlLogAction m 179 | 180 | type LogHtmlT m = ReaderT (HtmlLogAction m) m 181 | 182 | instance Monad m => MonadLogHtml (LogHtmlT m) where 183 | logHtml message = do 184 | doLog <- ask 185 | lift $ doLog message 186 | 187 | runLogHtmlWith :: Monad m => HtmlLogAction m -> LogHtmlT m a -> m a 188 | runLogHtmlWith = flip runReaderT 189 | 190 | renderSourceFile :: ModuleToURL -> HtmlOptions -> HtmlInputSourceFile -> Text 191 | renderSourceFile moduleToURL opts = renderSourcePage 192 | where 193 | cssFile = fromMaybe defaultCSSFile (htmlOptCssFile opts) 194 | highlightOccur = htmlOptHighlightOccurrences opts 195 | htmlHighlight = htmlOptHighlight opts 196 | renderSourcePage (HtmlInputSourceFile moduleName fileType sourceCode hinfo) = 197 | page cssFile highlightOccur onlyCode moduleName pageContents 198 | where 199 | tokens = tokenStream sourceCode hinfo 200 | onlyCode = highlightOnlyCode htmlHighlight fileType 201 | pageContents = code moduleToURL onlyCode fileType tokens 202 | 203 | defaultPageGen :: (MonadIO m, MonadLogHtml m) => ModuleToURL -> HtmlOptions -> HtmlInputSourceFile -> m () 204 | defaultPageGen moduleToURL opts srcFile@(HtmlInputSourceFile moduleName ft _ _) = do 205 | logHtml $ render $ "Generating HTML for" <+> pretty moduleName <+> ((parens (pretty target)) <> ".") 206 | writeRenderedHtml html target 207 | where 208 | ext = highlightedFileExt (htmlOptHighlight opts) ft 209 | target = (htmlOptDir opts) modToFile moduleName ext 210 | html = renderSourceFile moduleToURL opts srcFile 211 | 212 | prepareCommonDestinationAssets :: MonadIO m => HtmlOptions -> m () 213 | prepareCommonDestinationAssets options = liftIO $ do 214 | -- There is a default directory given by 'defaultHTMLDir' 215 | let htmlDir = htmlOptDir options 216 | createDirectoryIfMissing True htmlDir 217 | 218 | -- If the default CSS file should be used, then it is copied to 219 | -- the output directory. 220 | let cssFile = htmlOptCssFile options 221 | when (isNothing $ cssFile) $ do 222 | defCssFile <- getDataFileName $ 223 | htmlDataDir defaultCSSFile 224 | copyFile defCssFile (htmlDir defaultCSSFile) 225 | 226 | let highlightOccurrences = htmlOptHighlightOccurrences options 227 | when highlightOccurrences $ do 228 | highlightJsFile <- getDataFileName $ 229 | htmlDataDir occurrenceHighlightJsFile 230 | copyFile highlightJsFile (htmlDir occurrenceHighlightJsFile) 231 | 232 | -- | Converts module names to the corresponding HTML file names. 233 | 234 | modToFile :: TopLevelModuleName -> String -> FilePath 235 | modToFile m ext = render (pretty m) <.> ext 236 | 237 | -- | Generates a highlighted, hyperlinked version of the given module. 238 | 239 | writeRenderedHtml 240 | :: MonadIO m 241 | => Text -- ^ Rendered page 242 | -> FilePath -- ^ Output path. 243 | -> m () 244 | writeRenderedHtml html target = liftIO $ UTF8.writeTextToFile target html 245 | 246 | 247 | -- | Attach multiple Attributes 248 | 249 | (!!) :: Html -> [Attribute] -> Html 250 | h !! as = h ! mconcat as 251 | 252 | -- | Constructs the web page, including headers. 253 | 254 | page :: FilePath -- ^ URL to the CSS file. 255 | -> Bool -- ^ Highlight occurrences 256 | -> Bool -- ^ Whether to reserve literate 257 | -> TopLevelModuleName -- ^ Module to be highlighted. 258 | -> Html 259 | -> Text 260 | page css 261 | highlightOccurrences 262 | htmlHighlight 263 | modName 264 | pageContent = 265 | renderHtml $ if htmlHighlight 266 | then pageContent 267 | else Html5.docTypeHtml $ hdr <> rest 268 | where 269 | 270 | hdr = Html5.head $ mconcat 271 | [ Html5.meta !! [ Attr.charset "utf-8" ] 272 | , Html5.title (toHtml . render $ pretty modName) 273 | , Html5.link !! [ Attr.rel "stylesheet" 274 | , Attr.href $ stringValue css 275 | ] 276 | , if highlightOccurrences 277 | then Html5.script mempty !! 278 | [ Attr.src $ stringValue occurrenceHighlightJsFile 279 | ] 280 | else mempty 281 | ] 282 | 283 | rest = Html5.body $ (Html5.pre ! Attr.class_ "Agda") pageContent 284 | 285 | -- | Position, Contents, Infomation 286 | 287 | type TokenInfo = 288 | ( Int 289 | , String1 290 | , Aspects 291 | ) 292 | 293 | -- | Constructs token stream ready to print. 294 | 295 | tokenStream 296 | :: Text -- ^ The contents of the module. 297 | -> HighlightingInfo -- ^ Highlighting information. 298 | -> [TokenInfo] 299 | tokenStream contents info = 300 | map (\ ((mi, (pos, c)) :| xs) -> 301 | (pos, c :| map (snd . snd) xs, fromMaybe mempty mi)) $ 302 | List1.groupBy ((==) `on` fst) $ 303 | zipWith (\pos c -> (IntMap.lookup pos infoMap, (pos, c))) [1..] (T.unpack contents) 304 | where 305 | infoMap = toMap info 306 | 307 | -- | Constructs the HTML displaying the code. 308 | 309 | code :: ModuleToURL 310 | -> Bool -- ^ Whether to generate non-code contents as-is 311 | -> FileType -- ^ Source file type 312 | -> [TokenInfo] 313 | -> Html 314 | code moduleToURL onlyCode fileType = mconcat . if onlyCode 315 | then case fileType of 316 | -- Explicitly written all cases, so people 317 | -- get compile error when adding new file types 318 | -- when they forget to modify the code here 319 | RstFileType -> map mkRst . splitByMarkup 320 | MdFileType -> map mkMd . splitByMarkup 321 | AgdaFileType -> map mkHtml 322 | OrgFileType -> map mkOrg . splitByMarkup 323 | TreeFileType -> map mkMd . splitByMarkup 324 | -- Two useless cases, probably will never used by anyone 325 | TexFileType -> map mkMd . splitByMarkup 326 | TypstFileType -> map mkMd . splitByMarkup 327 | else map mkHtml 328 | where 329 | trd (_, _, a) = a 330 | 331 | splitByMarkup :: [TokenInfo] -> [[TokenInfo]] 332 | splitByMarkup = splitWhen $ (== Just Markup) . aspect . trd 333 | 334 | mkHtml :: TokenInfo -> Html 335 | mkHtml (pos, s, mi) = 336 | -- Andreas, 2017-06-16, issue #2605: 337 | -- Do not create anchors for whitespace. 338 | applyUnless (mi == mempty) (annotate pos mi) $ toHtml $ List1.toList s 339 | 340 | backgroundOrAgdaToHtml :: TokenInfo -> Html 341 | backgroundOrAgdaToHtml token@(_, s, mi) = case aspect mi of 342 | Just Background -> preEscapedToHtml $ List1.toList s 343 | Just Markup -> __IMPOSSIBLE__ 344 | _ -> mkHtml token 345 | 346 | -- Proposed in #3373, implemented in #3384 347 | mkRst :: [TokenInfo] -> Html 348 | mkRst = mconcat . (toHtml rstDelimiter :) . map backgroundOrAgdaToHtml 349 | 350 | -- The assumption here and in mkOrg is that Background tokens and Agda tokens are always 351 | -- separated by Markup tokens, so these runs only contain one kind. 352 | mkMd :: [TokenInfo] -> Html 353 | mkMd tokens = if containsCode then formatCode else formatNonCode 354 | where 355 | containsCode = any ((/= Just Background) . aspect . trd) tokens 356 | 357 | formatCode = Html5.pre ! Attr.class_ "Agda" $ mconcat $ backgroundOrAgdaToHtml <$> tokens 358 | formatNonCode = mconcat $ backgroundOrAgdaToHtml <$> tokens 359 | 360 | mkOrg :: [TokenInfo] -> Html 361 | mkOrg tokens = mconcat $ if containsCode then formatCode else formatNonCode 362 | where 363 | containsCode = any ((/= Just Background) . aspect . trd) tokens 364 | 365 | startDelimiter = preEscapedToHtml orgDelimiterStart 366 | endDelimiter = preEscapedToHtml orgDelimiterEnd 367 | 368 | formatCode = startDelimiter : foldr (\x -> (backgroundOrAgdaToHtml x :)) [endDelimiter] tokens 369 | formatNonCode = map backgroundOrAgdaToHtml tokens 370 | 371 | -- Put anchors that enable referencing that token. 372 | -- We put a fail safe numeric anchor (file position) for internal references 373 | -- (issue #2756), as well as a heuristic name anchor for external references 374 | -- (issue #2604). 375 | annotate :: Int -> Aspects -> Html -> Html 376 | annotate pos mi = 377 | applyWhen hereAnchor (anchorage nameAttributes mempty <>) . anchorage posAttributes 378 | where 379 | -- Warp an anchor ( tag) with the given attributes around some HTML. 380 | anchorage :: [Attribute] -> Html -> Html 381 | anchorage attrs html = Html5.a html !! attrs 382 | 383 | -- File position anchor (unique, reliable). 384 | posAttributes :: [Attribute] 385 | posAttributes = concat 386 | [ [Attr.id $ stringValue $ show pos ] 387 | , toList $ link <$> definitionSite mi 388 | , Attr.class_ (stringValue $ unwords classes) <$ guard (not $ null classes) 389 | ] 390 | 391 | -- Named anchor (not reliable, but useful in the general case for outside refs). 392 | nameAttributes :: [Attribute] 393 | nameAttributes = [ Attr.id $ stringValue $ fromMaybe __IMPOSSIBLE__ $ mDefSiteAnchor ] 394 | 395 | classes = concat 396 | [ concatMap noteClasses (note mi) 397 | , otherAspectClasses (toList $ otherAspects mi) 398 | , concatMap aspectClasses (aspect mi) 399 | ] 400 | 401 | aspectClasses (Name mKind op) = kindClass ++ opClass 402 | where 403 | kindClass = toList $ fmap showKind mKind 404 | 405 | showKind (Constructor Inductive) = "InductiveConstructor" 406 | showKind (Constructor CoInductive) = "CoinductiveConstructor" 407 | showKind k = show k 408 | 409 | opClass = ["Operator" | op] 410 | aspectClasses a = [show a] 411 | 412 | 413 | otherAspectClasses = map show 414 | 415 | -- Notes are not included. 416 | noteClasses _s = [] 417 | 418 | -- Should we output a named anchor? 419 | -- Only if we are at the definition site now (@here@) 420 | -- and such a pretty named anchor exists (see 'defSiteAnchor'). 421 | hereAnchor :: Bool 422 | hereAnchor = here && isJust mDefSiteAnchor 423 | 424 | mDefinitionSite :: Maybe DefinitionSite 425 | mDefinitionSite = definitionSite mi 426 | 427 | -- Are we at the definition site now? 428 | here :: Bool 429 | here = maybe False defSiteHere mDefinitionSite 430 | 431 | mDefSiteAnchor :: Maybe String 432 | mDefSiteAnchor = maybe __IMPOSSIBLE__ defSiteAnchor mDefinitionSite 433 | 434 | link (DefinitionSite m defPos _here aName) = Attr.href $ stringValue $ 435 | -- If the definition site points to the top of a file, 436 | -- we drop the anchor part and just link to the file. 437 | applyUnless (defPos <= 1) 438 | (++ "#" ++ Network.URI.Encode.encode anchor) 439 | (maybe id () u $ Network.URI.Encode.encode $ modToFile m "") 440 | where 441 | u = Map.lookup m moduleToURL 442 | -- Use named anchors for external links as they should be more stable(?) 443 | anchor | Just a <- aName, Just u' <- u, u' /= "" = a 444 | | otherwise = show defPos 445 | --------------------------------------------------------------------------------