├── .gitignore ├── conference-paper.pdf ├── modelstructure.pdf ├── unifying-cartesian.pdf ├── agda ├── Makefile ├── glueing.agda ├── Data │ ├── nat.agda │ ├── id3.agda │ ├── inductive-fiber.agda │ ├── paths.agda │ ├── id.agda │ ├── products.agda │ ├── functions.agda │ └── id2.agda ├── everything.agda ├── pushout.agda ├── nonempty.agda ├── strictness-axioms.agda ├── unifying-summary.agda ├── interval.agda ├── README.md ├── equivscontrfib.agda ├── ABCFHL.agda ├── glueing │ ├── aligned.agda │ ├── strict.agda │ └── core.agda ├── wtypesred.agda ├── retract.agda ├── equivs.agda ├── trivialfib.agda ├── cofreplacement.agda ├── cofprop.agda ├── realignment.agda ├── hcomp-coe.agda ├── fiberwise-total.agda ├── univalence.agda ├── CCHM.agda ├── fibrations.agda ├── fibreplacement.agda └── prelude.agda └── README.md /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | -------------------------------------------------------------------------------- /conference-paper.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/mortberg/gen-cart/HEAD/conference-paper.pdf -------------------------------------------------------------------------------- /modelstructure.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/mortberg/gen-cart/HEAD/modelstructure.pdf -------------------------------------------------------------------------------- /unifying-cartesian.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/mortberg/gen-cart/HEAD/unifying-cartesian.pdf -------------------------------------------------------------------------------- /agda/Makefile: -------------------------------------------------------------------------------- 1 | check: $(wildcard **/*.agda) 2 | agda everything.agda 3 | 4 | clean:; rm -f *agdai */*agdai */*/*agdai */*/*/*agdai 5 | -------------------------------------------------------------------------------- /agda/glueing.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module glueing where 3 | 4 | open import glueing.core public 5 | open import glueing.strict public 6 | open import glueing.aligned public 7 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Generalized Cartesian Cubical Type Theory 2 | 3 | Agda code for a Cartesian cubical set model of univalent type theory 4 | based on a weakened version of ABCFHL fibrations. This generalizes the 5 | construction in https://github.com/dlicata335/cart-cube and 6 | https://github.com/IanOrton/cubical-topos-experiments/. 7 | 8 | 9 | This code is based on the formalization accompanying Ian Orton's 10 | Ph.D. thesis which can be found at: 11 | https://www.repository.cam.ac.uk/handle/1810/288558. We are grateful 12 | for Ian's permission to use his code. 13 | -------------------------------------------------------------------------------- /agda/Data/nat.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | The type of natural numbers is fibrant. 4 | 5 | -} 6 | module Data.nat where 7 | 8 | open import prelude 9 | open import interval 10 | open import cofprop 11 | open import fibrations 12 | open import hcomp-coe 13 | 14 | open import Agda.Builtin.Nat public 15 | using (zero; suc) 16 | renaming (Nat to ℕ) 17 | 18 | pred : ℕ → ℕ 19 | pred zero = zero 20 | pred (suc x) = x 21 | 22 | natDec : (x y : ℕ) → (x ≡ y) ⊎ ¬ (x ≡ y) 23 | natDec zero zero = inl refl 24 | natDec zero (suc y) = inr λ () 25 | natDec (suc x) zero = inr λ () 26 | natDec (suc x) (suc y) with natDec x y 27 | ... | inl p = inl (cong suc p) 28 | ... | inr p = inr (λ q → p (cong pred q)) 29 | 30 | isFibNat : {Γ : Set} → isFib {Γ = Γ} (λ _ → ℕ) 31 | isFibNat = constHFib→isFib ℕ (decEq→isHFib _ natDec) 32 | 33 | fibNat : Fib ℕ 34 | fibNat = _ , isFibNat 35 | -------------------------------------------------------------------------------- /agda/everything.agda: -------------------------------------------------------------------------------- 1 | module everything where 2 | 3 | open import ABCFHL 4 | open import CCHM 5 | open import cofprop 6 | open import cofreplacement 7 | open import Data.functions 8 | open import Data.id 9 | open import Data.id2 10 | open import Data.id3 11 | open import Data.inductive-fiber 12 | open import Data.nat 13 | open import Data.paths 14 | open import Data.products 15 | open import equivs 16 | open import equivscontrfib 17 | open import fibrations 18 | open import fibreplacement 19 | open import glueing.aligned 20 | open import glueing.core 21 | open import glueing.strict 22 | open import hcomp-coe 23 | open import interval 24 | open import nonempty 25 | open import prelude 26 | open import pushout 27 | open import realignment 28 | open import retract 29 | open import strictness-axioms 30 | open import unifying-summary 31 | open import univalence 32 | open import trivialfib 33 | open import wtypesred 34 | -------------------------------------------------------------------------------- /agda/Data/id3.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module Data.id3 {ℓ} {Γ : Set ℓ} (A : Γ → Set) where 3 | 4 | open import prelude 5 | open import Data.products 6 | open import interval 7 | open import cofprop 8 | open import fibrations 9 | open import fiberwise-total 10 | open import cofreplacement 11 | open import fibreplacement 12 | open import Data.products 13 | open import Data.inductive-fiber 14 | 15 | Δ : (x : Γ) → A x → A x × A x 16 | Δ x a = (a , a) 17 | 18 | Id3 : Σ x ∈ Γ , A x × A x → Set 19 | Id3 = uncurry (FR' Δ) 20 | 21 | refl3 : (x : Γ) (a : A x) → Id3 (x , (a , a)) 22 | refl3 x a = quot (Δ x) a .snd 23 | 24 | refl3IsTrivCofibration : (x : Γ) 25 | → triv-cofibration (λ a → ((a , a) , refl3 x a)) 26 | refl3IsTrivCofibration x = quotIsTrivCofibration (Δ x) 27 | 28 | Id3IsFib : isFib A → isFib Id3 29 | Id3IsFib α = FR'IsFib Δ α αα 30 | where 31 | αα = FibΣ {A = A} {B = λ {(x , _) → A x}} α (reindex A α fst) 32 | 33 | -------------------------------------------------------------------------------- /agda/pushout.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Axiomatization of pushouts. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | open import prelude 8 | 9 | module pushout where 10 | 11 | postulate 12 | Pushout : {ℓ m k : Level} {A : Set ℓ} {B : Set m} {C : Set k} 13 | (f : A → B) (g : A → C) → Set (ℓ ⊔ m ⊔ k) 14 | 15 | module _ {ℓ m k : Level} {A : Set ℓ} {B : Set m} {C : Set k} 16 | {f : A → B} {g : A → C} 17 | where 18 | postulate 19 | 20 | left : B → Pushout f g 21 | right : C → Pushout f g 22 | push : (a : A) → (left (f a) ≡ right (g a)) 23 | 24 | module _ {j : Level} (X : Pushout f g → Set j) 25 | (l : (b : B) → X (left b)) 26 | (r : (c : C) → X (right c)) 27 | (p : (a : A) → subst X (push a) (l (f a)) ≡ r (g a)) 28 | where 29 | postulate 30 | pushout-elim : (z : Pushout f g) → (X z) 31 | 32 | pushout-l-β : (b : B) → (pushout-elim (left b) ≡ l b) 33 | pushout-r-β : (c : C) → (pushout-elim (right c) ≡ r c) 34 | 35 | {-# REWRITE pushout-l-β #-} 36 | {-# REWRITE pushout-r-β #-} 37 | 38 | pushout-rec : {j : Level} {X : Set j} (l : B → X) (r : C → X) → 39 | (p : (a : A) → l (f a) ≡ r (g a)) → Pushout f g → X 40 | pushout-rec {X = X} l r p = 41 | pushout-elim (λ _ → X) l r 42 | (λ a → trans (p a) (substconst X (push a) (l (f a)))) 43 | -------------------------------------------------------------------------------- /agda/nonempty.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module nonempty where 3 | 4 | open import prelude 5 | open import interval 6 | open import cofprop 7 | open import fibrations 8 | open import hcomp-coe 9 | 10 | -- homogeneous composition for non-empty systems 11 | isRestrictedHFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 12 | isRestrictedHFib {Γ = Γ} A = (r : Int) (x : Γ) 13 | (φ : CofProp) (_ : ¬ ¬ [ φ ]) (f : [ φ ] → Int → A x) 14 | (a₀ : A x [ φ ↦ f ◆ r ]) 15 | → WComp r (λ _ → A x) φ f a₀ 16 | 17 | RestrictedHFib : ∀{ℓ}(Γ : Set ℓ) → Set (lsuc lzero ⊔ ℓ) 18 | RestrictedHFib Γ = Σ (Γ → Set) isRestrictedHFib 19 | 20 | postulate 21 | ~ : CofProp → CofProp 22 | explode : ∀ φ → [ φ ] → [ ~ φ ] → ∅ 23 | ¬¬LEM : ∀ φ → ¬ ¬ [ φ ∨ ~ φ ] 24 | 25 | isRestrictedHFib→isHFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) 26 | → isRestrictedHFib A → isHFib A 27 | isRestrictedHFib→isHFib A α r x φ f a₀ = 28 | record 29 | { comp = λ s → (fixed .comp s .fst , λ u → fixed .comp s .snd ∣ inl u ∣) 30 | ; cap = 31 | ( fixed .cap .fst 32 | , λ t u → fixed .cap .snd t ∣ inl u ∣ 33 | ) 34 | } 35 | where 36 | tube : [ φ ∨ ~ φ ] → Int → A x 37 | tube = 38 | ∨-rec φ (~ φ) 39 | f 40 | (λ _ _ → a₀ .fst) 41 | (λ u ~u → ∅-rec (explode φ u ~u)) 42 | 43 | base : A x [ φ ∨ ~ φ ↦ tube ◆ r ] 44 | base = 45 | ( a₀ .fst 46 | , ∨-elim φ (~ φ) _ 47 | (a₀ .snd) 48 | (λ _ → refl) 49 | (λ _ _ → uipImp) 50 | ) 51 | 52 | fixed = α r x (φ ∨ ~ φ) (¬¬LEM φ) tube base 53 | -------------------------------------------------------------------------------- /agda/strictness-axioms.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Axioms related to strictifying Glue types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module strictness-axioms where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | 14 | ---------------------------------------------------------------------- 15 | -- Strictifying 16 | ---------------------------------------------------------------------- 17 | postulate 18 | reIm : 19 | (φ : CofProp) 20 | (A : [ φ ] → Set) 21 | (B : Set) 22 | (m : (u : [ φ ]) → A u ≅ B) 23 | → ---------------------- 24 | Σ B' ∈ Set , Σ m' ∈ B' ≅ B , ((u : [ φ ]) → (A u , m u) ≡ (B' , m')) 25 | 26 | strictify : 27 | (φ : CofProp) 28 | (A : [ φ ] → Set) 29 | (B : Set) 30 | (m : (u : [ φ ]) → A u ≅ B) 31 | → ---------------------- 32 | Set 33 | strictify φ A B m = reIm φ A B m .fst 34 | 35 | isoB : 36 | (φ : CofProp) 37 | (A : [ φ ] → Set) 38 | (B : Set) 39 | (m : (u : [ φ ]) → A u ≅ B) 40 | → ---------------------- 41 | strictify φ A B m ≅ B 42 | isoB φ A B m = reIm φ A B m .snd .fst 43 | 44 | restrictsToA : 45 | (φ : CofProp) 46 | (A : [ φ ] → Set) 47 | (B : Set) 48 | (m : (u : [ φ ]) → A u ≅ B) 49 | (u : [ φ ]) 50 | → ---------------------- 51 | A u ≡ strictify φ A B m 52 | restrictsToA φ A B m u = cong fst (reIm φ A B m .snd .snd u) 53 | 54 | restrictsToM : 55 | (φ : CofProp) 56 | (A : [ φ ] → Set) 57 | (B : Set) 58 | (m : (u : [ φ ]) → A u ≅ B) 59 | (u : [ φ ]) 60 | → ---------------------- 61 | (A u , m u) ≡ (strictify φ A B m , isoB φ A B m) 62 | restrictsToM φ A B m u = reIm φ A B m .snd .snd u 63 | -------------------------------------------------------------------------------- /agda/unifying-summary.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | This file provides a stable interface to the formalization of the 4 | paper: 5 | 6 | Unifying Cubical Models of Univalent Type Theory 7 | 8 | by Evan Cavallo, Anders Mörtberg and Andrew Swan. 9 | 10 | PLEASE DO NOT RENAME THIS FILE - its name is referenced in the article 11 | about this formalization. 12 | 13 | -} 14 | module unifying-summary where 15 | 16 | -- | Section 2.1: The interval and Path types. 17 | -- (axioms 1 and 2) 18 | open import interval 19 | 20 | 21 | -- | Section 2.2: Cofibrant propositions (Φ is called CofProp) 22 | -- (axioms 3, 4, 5 and 7) 23 | open import cofprop 24 | 25 | 26 | -- | Section 2.3: Fibration structures (WComp and SComp) 27 | open import fibrations 28 | 29 | 30 | -- | Section 2.3.1: ABCFHL fibrations 31 | open import ABCFHL 32 | 33 | 34 | -- | Section 2.3.2: CCHM fibrations 35 | open import CCHM 36 | 37 | 38 | -- | Section 2.4: Fibration structures for basic type formers 39 | 40 | -- Σ-types: 41 | open import Data.products 42 | 43 | -- Π-types: 44 | open import Data.functions 45 | 46 | -- Path types: 47 | open import Data.paths 48 | 49 | -- Natural numbers 50 | open import Data.nat 51 | 52 | 53 | -- | Section 2.5: Glueing 54 | 55 | -- Definition 16 (Glueing) and Theorem 17 (Fibrant Glue types) 56 | open import glueing.core 57 | 58 | -- Axiom 6 59 | open import strictness-axioms 60 | 61 | -- Strict Glue types (SGlue) 62 | open import glueing.strict 63 | 64 | -- Alignment for SGlue 65 | open import glueing.aligned 66 | 67 | 68 | -- | Section 2.5.1: Univalence 69 | open import univalence 70 | 71 | 72 | -- | Section 3.2: Cofibration and trivial fibration awfs 73 | open import cofreplacement 74 | 75 | 76 | -- | Section 3.3: Trivial cofibration and fibration awfs 77 | open import fibreplacement 78 | 79 | 80 | -- | Section 4: Identity types and higher inductive types 81 | 82 | -- CCHM style identity types using a dominance on CofProp 83 | open import Data.id 84 | 85 | -- Identity types from cofibrant replacement 86 | open import Data.id2 87 | 88 | -- Identity types from fibrant replacement 89 | open import Data.id3 90 | -------------------------------------------------------------------------------- /agda/interval.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of the interval and Path types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module interval where 8 | 9 | open import prelude 10 | 11 | ---------------------------------------------------------------------- 12 | -- Interval 13 | ---------------------------------------------------------------------- 14 | postulate 15 | Int : Set 16 | O : Int 17 | I : Int 18 | ¬O≡I : ¬ (O ≡ I) 19 | cntd : (P : Int → Set) → ((i : Int) → P i ⊎ ¬ (P i)) 20 | → ((i : Int) → P i) ⊎ ((i : Int) → ¬ P i) 21 | 22 | O≠I : ∀ {ℓ} {A : Set ℓ} → O ≡ I → A 23 | O≠I x = ∅-rec (¬O≡I x) 24 | 25 | ---------------------------------------------------------------------- 26 | -- Dependently typed paths 27 | ---------------------------------------------------------------------- 28 | Π : ∀ {ℓ} → (Int → Set ℓ) → Set ℓ 29 | Π A = (i : Int) → A i 30 | 31 | Π′ : ∀ {ℓ} 32 | {A B : Int → Set ℓ} 33 | → --------------------------------- 34 | ((i : Int) → A i → B i) → Π A → Π B 35 | Π′ F p i = F i (p i) 36 | 37 | ---------------------------------------------------------------------- 38 | -- Path types 39 | ---------------------------------------------------------------------- 40 | 41 | record _~~_ {A : Int → Set}(a : A O)(a' : A I) : Set where 42 | constructor path 43 | field 44 | at : Π A 45 | atO : at O ≡ a 46 | atI : at I ≡ a' 47 | 48 | open _~~_ public 49 | 50 | _~_ : {A : Set}(a a' : A) → Set 51 | _~_ {A} a a' = _~~_ {A = λ _ → A} a a' 52 | 53 | refl~ : {A : Set}(a : A) → a ~ a 54 | refl~ a = path (λ _ → a) refl refl 55 | 56 | PathExt : {A : Set}{a a' : A}{p q : a ~ a'} → p .at ≡ q .at → p ≡ q 57 | PathExt {p = p} refl = cong₂ (path (p .at)) uipImp uipImp 58 | 59 | PathExtDep : {I : Set} {A : Set} {a₀ a₁ : I → A} 60 | {i i' : I} (r : i ≡ i') 61 | {p : a₀ i ~ a₁ i} {q : a₀ i' ~ a₁ i'} 62 | → p .at ≡ q .at 63 | → subst (λ i → a₀ i ~ a₁ i) r p ≡ q 64 | PathExtDep refl {p = p} refl = cong₂ (path (p .at)) uipImp uipImp 65 | 66 | eqToPath : {A : Set} {x y : A} → x ≡ y → x ~ y 67 | eqToPath refl = refl~ _ 68 | 69 | eqToPathAtO : {A : Set} {x y : A} (e : x ≡ y) (k : Int) → eqToPath e .at k ≡ x 70 | eqToPathAtO refl k = refl 71 | -------------------------------------------------------------------------------- /agda/Data/inductive-fiber.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module Data.inductive-fiber {ℓ} {Γ : Set ℓ} {A B : Γ → Set} (f : ∀ x → A x → B x) where 3 | 4 | open import prelude 5 | open import interval 6 | open import cofprop 7 | open import fibrations 8 | open import wtypesred 9 | open import hcomp-coe 10 | open import fiberwise-total 11 | open import retract 12 | open import cofreplacement 13 | open import fibreplacement 14 | open import Data.products 15 | open import Data.paths 16 | 17 | FR' : (x : Γ) → (b : B x) → Set 18 | FR' x = FR (f x) 19 | 20 | ΣFR'IsHFib : isFib B → isHFib (Σ' B (uncurry FR')) 21 | ΣFR'IsHFib β r x φ h σ₀ = 22 | FibΣ {Γ = Unit} 23 | (reindex B β (λ _ → x)) 24 | (reindex (FR' x) (FRIsFib (f x)) (λ {(_ , b) → b})) 25 | r _ 26 | φ h σ₀ 27 | 28 | module _ (α : isFib A) (β : isFib B) where 29 | 30 | eval : {x : Γ} → (b : B x) → FR' x b → A x 31 | eval {x} = curry 32 | (quotIsTrivCofibration (f x) .j _ (reindex A α (λ _ → x)) id) 33 | 34 | quotEval : {x : Γ} (b : B x) (c : FR' x b) 35 | → _~_ {A = Σ (B x) (FR' x)} (quot _ (eval b c)) (b , c) 36 | quotEval {x} = curry 37 | (quotIsTrivCofibration (f x) .j 38 | (λ {(b , c) → quot _ (eval b c) ~ (b , c)}) 39 | (reindex 40 | (Path (λ _ → Σ (B x) (FR' x))) 41 | (FibPath sigFib) 42 | (λ {(b , c) → tt , quot (f x) (eval b c) , (b , c)})) 43 | (λ a → 44 | eqToPath 45 | (trans 46 | (cong (λ emp → (f x a , ixsup (FRPoly _) (dom a) emp)) 47 | (funext λ _ → funext ∅-elim)) 48 | (cong (quot (f x)) 49 | (quotIsTrivCofibration (f x) .upper-tri 50 | _ (reindex A α (λ _ → x)) id a))))) 51 | where 52 | sigFib : isFib (λ _ → Σ (B x) (FR' x)) 53 | sigFib = 54 | FibΣ {B = λ {(_ , b) → FR' x b}} 55 | (reindex B β (λ _ → x)) 56 | (reindex (FR' x) (FRIsFib (f x)) (λ {(_ , b) → b})) 57 | 58 | ΣFR'IsFib : isFib (Σ' B (uncurry FR')) 59 | ΣFR'IsFib = 60 | homotopyRetractIsFib (Σ' B (uncurry FR')) A 61 | (ΣFR'IsHFib β) 62 | α 63 | (λ {x (b , c) → eval b c}) 64 | (λ x a → quot _ a) 65 | (λ {x (b , c) → quotEval b c}) 66 | 67 | FR'IsFib : isFib (uncurry FR') 68 | FR'IsFib = 69 | fiberwise+total→isFib B FR' β (λ x → FRIsFib (f x)) ΣFR'IsFib 70 | -------------------------------------------------------------------------------- /agda/README.md: -------------------------------------------------------------------------------- 1 | # Agda code for a generalized Cartesian Cubical Type Theory 2 | 3 | To typecheck everything run: 4 | ``` 5 | > make 6 | ``` 7 | 8 | This assumes that Agda version >=2.6.0 is installed. 9 | 10 | 11 | Contents 12 | -------- 13 | 14 | - ABCFHL: Equivalence between weak fibrations and ABCFHL fibrations 15 | when the diagonal on I is a cofibration. 16 | 17 | - CCHM: Equivalence with weak fibrations and CCHM fibrations in the 18 | presence of a connection algebra structure. 19 | 20 | - cofprop: Definition of the universe of cofibrant propositions and 21 | basic operations on these. 22 | 23 | - cofreplacement: Cofibrant replacement and definition of (trivial) 24 | cofibrations. 25 | 26 | - Data/functions: Fibrancy of Π-types. 27 | 28 | - Data/id: CCHM style identity types. 29 | 30 | - Data/id2: 31 | 32 | - Data/id3: 33 | 34 | - Data/inductive-fiber: 35 | 36 | - Data/nat: The type of natural numbers is fibrant. 37 | 38 | - Data/paths: Fibrancy of Path types. 39 | 40 | - Data/products: Fibrancy of Σ-types. 41 | 42 | - equivs: Definitions of contractible, extensible (SContr), fibers, 43 | equivalences and quasi-invertible maps. 44 | 45 | - equivscontrfib: Equivalent definition of isFib due to Dan Licata. 46 | 47 | - everything: File that imports all of the other files. 48 | 49 | - fiberwise-total: 50 | 51 | - fibrations: Definition of weak composition and fibrations. Proof 52 | that they are closed under isomorphism and that weak composition can 53 | be strictified. 54 | 55 | - fibreplacement: 56 | 57 | - glueing: Import the files related to Glue types (equivalence 58 | extension types). 59 | 60 | - glueing/aligned: Realigning strict Glue types. 61 | 62 | - glueing/core: Definition of (weak) Glue types and their (unaligned) 63 | fibrancy. 64 | 65 | - glueing/strict: Strict Glue types. 66 | 67 | - hcomp-coe: Definition of weak homogeneous composition and 68 | coercion. Proof that these imply weak composition. 69 | 70 | - interval: Definition of the interval and Path types. 71 | 72 | - prelude: Basics. 73 | 74 | - pushout: Axiomatization of pushouts. 75 | 76 | - realignment: Realign a fibration structure. 77 | 78 | - retract: 79 | 80 | - strictness-axioms: Axioms related to strictifying Glue types. 81 | 82 | - trivialfib: Definition of trivial fibrations and proof that they are 83 | actually fibrations. 84 | 85 | - univalence: Proof of univalence in the form of ua and uaβ. 86 | 87 | - wtypesred: W-types with reductions. 88 | -------------------------------------------------------------------------------- /agda/equivscontrfib.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Equivalent definition of isFib due to Dan Licata 4 | 5 | A : I → Set has weak Kan composition iff 6 | Π r:I, the map (λ f → f r) : (Π (x : I) → A x) → (A r) is an equivalence 7 | 8 | -} 9 | {-# OPTIONS --rewriting #-} 10 | module equivscontrfib where 11 | 12 | open import prelude 13 | open import interval 14 | open import cofprop 15 | open import fibrations 16 | open import equivs using (SContr ; Fiber ; FiberExt ; EquivSContr) 17 | 18 | isEquivSContrFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 19 | isEquivSContrFib {Γ = Γ} A = 20 | (r : Int) (p : Int → Γ) → EquivSContr (λ (f : (i : Int) → A (p i)) → f r) 21 | 22 | isEquivSContrFib→isFib : (A : Int → Set) (h : isEquivSContrFib A) → isFib A 23 | isEquivSContrFib→isFib A h r p φ f x₀ = record 24 | { comp = λ s → fib1 .fst s 25 | , λ u → appCong (cong fst (fib .snd u)) 26 | ; cap = fib1 .snd 27 | , λ k u → trans (cong (λ w → w .snd .at k) (fib .snd u)) (symm (eqToPathAtO (x₀ .snd u) k)) } 28 | where 29 | fib : Fiber (λ f → f r) (x₀ .fst) [ φ ↦ (λ x → (f x , eqToPath (x₀ .snd x))) ] 30 | fib = h r p (x₀ .fst) φ (λ x → (f x , eqToPath (x₀ .snd x))) 31 | 32 | fib1 : Fiber (λ f → f r) (x₀ .fst) 33 | fib1 = fib .fst 34 | 35 | isFib→isEquivSContrFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) (α : isFib A) → isEquivSContrFib A 36 | isFib→isEquivSContrFib A α r p a₀ φ t = 37 | ( ( (λ i → f+ .comp i .fst) 38 | , path 39 | (λ s → q+ s .fst) 40 | (trans 41 | (f+ .cap .fst .atO) 42 | (symm (q+ O .snd ∣ inr ∣ inl refl ∣ ∣))) 43 | (symm (q+ I .snd ∣ inr ∣ inr refl ∣ ∣)) 44 | ) 45 | , λ u → 46 | FiberExt 47 | (funext λ i → f+ .comp i .snd u) 48 | (funext λ s → q+ s .snd ∣ inl u ∣) 49 | ) 50 | where 51 | f : [ φ ] → Π (A ∘ p) 52 | f u = t u .fst 53 | 54 | q : (u : [ φ ]) → f u r ~ a₀ 55 | q u = t u .snd 56 | 57 | a₀Fix = 58 | strictifyFib A α I (λ _ → p r) φ (λ u → q u .at) (a₀ , λ u → q u .atI) 59 | 60 | f+ = α r p φ f 61 | ( a₀Fix .comp O (I≡IsCofProp O) .fst 62 | , λ u → 63 | trans 64 | (a₀Fix .comp O (I≡IsCofProp O) .snd u) 65 | (symm (q u .atO))) 66 | 67 | module _ (s : Int) where 68 | q+Tube : [ φ ∨ s ≈O ∨ s ≈I ] → Int → A (p r) 69 | q+Tube = 70 | ∨-rec φ (s ≈O ∨ s ≈I) 71 | (λ u _ → q u .at s) 72 | (OI-rec s 73 | (λ {refl j → f+ .cap .fst .at j}) 74 | (λ {refl _ → a₀})) 75 | (λ u → ∨-elimEq (s ≈O) (s ≈I) 76 | (λ {refl → funext λ j → 77 | trans (f+ .cap .snd j u) (t u .snd .atO)}) 78 | (λ {refl → funext λ j → t u .snd .atI})) 79 | 80 | q+Base : A (p r) [ φ ∨ s ≈O ∨ s ≈I ↦ q+Tube ◆ I ] 81 | q+Base = 82 | ( a₀Fix .comp s (I≡IsCofProp s) .fst 83 | , ∨-elimEq φ (s ≈O ∨ s ≈I) 84 | (λ u → a₀Fix .comp s (I≡IsCofProp s) .snd u) 85 | (∨-elimEq (s ≈O) (s ≈I) 86 | (λ {refl → f+ .cap .fst .atI}) 87 | λ {refl → symm (a₀Fix .cap (I≡IsCofProp I))}) 88 | ) 89 | 90 | q+ = α I (λ _ → p r) (φ ∨ s ≈O ∨ s ≈I) q+Tube q+Base .comp O 91 | -------------------------------------------------------------------------------- /agda/ABCFHL.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Proof that ABCFHL fibrations as defined in 4 | 5 | https://github.com/dlicata335/cart-cube 6 | 7 | are equivalent to weak fibrations when the diagonal is a cofibration. 8 | 9 | We also prove that ABCFHL fibrations are equivalent to the variation 10 | in 11 | 12 | https://github.com/riaqn/orton/blob/master/src/fibrations.agda 13 | 14 | This correspond to the "unbiased fibrations" of 15 | 16 | https://github.com/awodey/math/tree/master/QMS 17 | 18 | -} 19 | {-# OPTIONS --rewriting #-} 20 | module ABCFHL where 21 | 22 | open import prelude 23 | open import interval 24 | open import cofprop 25 | open import fibrations 26 | 27 | -- Assume that r≡s is a cofibration 28 | postulate 29 | diagCof : (r s : Int) → CofProp 30 | [diagCof] : (r s : Int) → [ diagCof r s ] ≡ (r ≡ s) 31 | 32 | {-# REWRITE [diagCof] #-} 33 | 34 | diagIsCofProp : (r s : Int) → isCofProp (r ≡ s) 35 | diagIsCofProp r s = (diagCof r s , id , id) 36 | 37 | -- ABCFHL fibrations 38 | record ABCFHLComp (r : Int) (A : Int → Set) (φ : CofProp) (f : [ φ ] → Π A) 39 | (x₀ : A r [ φ ↦ f ◆ r ]) : Set 40 | where 41 | field 42 | comp : (s : Int) → A s [ φ ↦ f ◆ s ] 43 | cap : comp r .fst ≡ x₀ .fst 44 | 45 | open ABCFHLComp public 46 | 47 | isABCFHLFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 48 | isABCFHLFib {Γ = Γ} A = ∀ r p φ f x₀ → ABCFHLComp r (A ∘ p) φ f x₀ 49 | 50 | isFib→isABCFHLFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isFib A → isABCFHLFib A 51 | isFib→isABCFHLFib A α r p ϕ f x₀ = record 52 | { comp = λ s → strictifyFib A α r p ϕ f x₀ .comp s (diagIsCofProp r s) 53 | ; cap = strictifyFib A α r p ϕ f x₀ .cap (diagIsCofProp r r) 54 | } 55 | 56 | isABCFHLFib→isFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isABCFHLFib A → isFib A 57 | isABCFHLFib→isFib A α r p ϕ f x₀ = record 58 | { comp = α r p ϕ f x₀ .comp 59 | ; cap = path (λ i → x₀ .fst) (symm (α r p ϕ f x₀ .cap)) refl 60 | , λ _ → x₀ .snd 61 | } 62 | 63 | -- Proof that ABCFHL fibrations are equivalent to the notion of fibration in 64 | -- https://github.com/riaqn/orton/blob/master/src/fibrations.agda 65 | private 66 | Comp : (Int → Set) → Set 67 | Comp A = (φ : CofProp) (f : [ φ ] → Π A) → (e₀ e₁ : Int) 68 | (h₀ : A e₀ [ φ ↦ f ◆ e₀ ]) 69 | → A e₁ [ φ ↦ f ◆ e₁ ] 70 | 71 | -- This definition of the cap condition is a bit messy to work with 72 | -- compared to ABCFHL fibrations as it involves equality in a Σ-type 73 | Reduce : {A : Int → Set} → (c : Comp A) → Set 74 | Reduce {A = A} c = (φ : CofProp) (f : [ φ ] → Π A) (e : Int) → 75 | (h : A e [ φ ↦ f ◆ e ]) → 76 | c φ f e e h ≡ h 77 | 78 | isFib' : ∀{a} {Γ : Set a} (A : Γ → Set) → Set a 79 | isFib' {a = a} {Γ = Γ} A = (p : Int → Γ) → Σ c ∈ Comp (A ∘ p) , Reduce c 80 | 81 | isABCFHLFib→isFib' : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isABCFHLFib A → isFib' A 82 | isABCFHLFib→isFib' A α p = (λ ϕ f r s x₀ → α r p ϕ f x₀ .comp s) 83 | , λ ϕ f r x₀ → Σext (α r p ϕ f x₀ .cap) (funext (λ _ → uipImp)) 84 | 85 | isFib'→isABCFHLFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isFib' A → isABCFHLFib A 86 | isFib'→isABCFHLFib A c r p ϕ f x₀ = record 87 | { comp = λ s → c p .fst ϕ f r s x₀ 88 | ; cap = cong fst (c p .snd ϕ f r x₀) 89 | } 90 | -------------------------------------------------------------------------------- /agda/glueing/aligned.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Realigning strict glue. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module glueing.aligned where 8 | 9 | open import glueing.core 10 | open import glueing.strict 11 | 12 | open import prelude 13 | open import interval 14 | open import cofprop 15 | open import fibrations 16 | open import equivs 17 | open import realignment 18 | 19 | ---------------------------------------------------------------------- 20 | -- Realigning strict glue 21 | ---------------------------------------------------------------------- 22 | 23 | FibSGlue : 24 | ∀{a}{Γ : Set a} 25 | (Φ : Γ → CofProp) 26 | {A : res Γ Φ → Set} 27 | {B : Γ → Set} 28 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 29 | (equiv : Equiv' f) 30 | → --------------- 31 | isFib A → isFib B → isFib (SGlue' Φ A B f) 32 | FibSGlue {a} {Γ} Φ {A} {B} f equiv α β = 33 | realign Φ (SGlue' Φ A B f) 34 | (subst isFib (SGlueStrictness' Φ f) α) 35 | (Misaligned.FibSGlue Φ f equiv α β) 36 | 37 | FibSGlueStrictness : 38 | {ℓ : Level} 39 | {Γ : Set ℓ} 40 | (Φ : Γ → CofProp) 41 | {A : res Γ Φ → Set} 42 | {B : Γ → Set} 43 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 44 | (equiv : Equiv' f) 45 | (α : isFib A) (β : isFib B) 46 | → --------------- 47 | subst isFib (SGlueStrictness' Φ f) α ≡ reindex (SGlue' Φ A B f) (FibSGlue Φ f equiv α β) fst 48 | FibSGlueStrictness Φ {A} {B} f equiv α β = 49 | symm 50 | (isRealigned Φ (SGlue' Φ A B f) 51 | (subst isFib (SGlueStrictness' Φ f) α) 52 | (Misaligned.FibSGlue Φ f equiv α β)) 53 | 54 | SGlueFib : 55 | ∀{a}{Γ : Set a} 56 | (Φ : Γ → CofProp) 57 | (Aα : Fib (res Γ Φ)) 58 | (Bβ : Fib Γ) 59 | (f : (xu : res Γ Φ) → Aα .fst xu → Bβ .fst (xu .fst)) 60 | (equiv : Equiv' f) 61 | → --------------- 62 | Fib Γ 63 | SGlueFib {a} {Γ} Φ (A , α) (B , β) f equiv = 64 | (SGlue' Φ A B f , FibSGlue Φ f equiv α β) 65 | 66 | SGlueFibStrictness : 67 | ∀{ℓ}{Γ : Set ℓ} 68 | (Φ : Γ → CofProp) 69 | (Aα : Fib (res Γ Φ)) 70 | (Bβ : Fib Γ) 71 | (f : (xu : res Γ Φ) → Aα .fst xu → Bβ .fst (xu .fst)) 72 | (equiv : Equiv' f) 73 | → Aα ≡ reindex' (SGlueFib Φ Aα Bβ f equiv) fst 74 | SGlueFibStrictness Φ (A , α) (B , β) f equiv = 75 | Σext (SGlueStrictness' Φ f) (FibSGlueStrictness Φ f equiv α β) 76 | 77 | reindexFibSGlue : 78 | ∀ {ℓ} {Δ Γ : Set ℓ} 79 | (Φ : Γ → CofProp) 80 | {A : res Γ Φ → Set} 81 | {B : Γ → Set} 82 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 83 | (equiv : Equiv' f) 84 | (α : isFib A) (β : isFib B) 85 | (ρ : Δ → Γ) 86 | → reindex (SGlue' Φ A B f) (FibSGlue Φ f equiv α β) ρ 87 | ≡ FibSGlue (Φ ∘ ρ) (f ∘ (ρ ×id)) (equiv ∘ (ρ ×id)) (reindex A α (ρ ×id)) (reindex B β ρ) 88 | reindexFibSGlue Φ {A} {B} f equiv α β ρ = 89 | trans 90 | (cong₂ (realign (Φ ∘ ρ) (SGlue' Φ A B f ∘ ρ)) 91 | (reindexSubst (ρ ×id) (SGlueStrictness' Φ f) (SGlueStrictness' (Φ ∘ ρ) (f ∘ (ρ ×id))) α) 92 | (Misaligned.reindexFibSGlue Φ f equiv α β ρ)) 93 | (reindexRealign Φ (SGlue' Φ A B f) 94 | (subst isFib (SGlueStrictness' Φ f) α) 95 | (Misaligned.FibSGlue Φ f equiv α β) 96 | ρ) 97 | where 98 | reindexSubst : ∀ {ℓ ℓ'} {Δ : Set ℓ} {Γ : Set ℓ'} {A A' : Γ → Set} 99 | (ρ : Δ → Γ)(P : A ≡ A') (Q : A ∘ ρ ≡ A' ∘ ρ) (α : isFib A) 100 | → reindex A' (subst isFib P α) ρ ≡ subst isFib Q (reindex A α ρ) 101 | reindexSubst ρ refl refl α = refl 102 | -------------------------------------------------------------------------------- /agda/wtypesred.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | W-types with reductions. 4 | 5 | We postulate cofibrant, non-indexed and indexed versions of W-types 6 | with reductions. These hold in any topos with n.n.o that satisfies 7 | WISC, and there is a more explicit description that does not use WISC 8 | for (possibly internal) presheaves when cofibrant propositions are 9 | locally decidable. 10 | 11 | -} 12 | {-# OPTIONS --rewriting #-} 13 | module wtypesred {ℓ} where 14 | 15 | open import prelude 16 | open import cofprop 17 | 18 | record Poly : Set (lsuc ℓ) where 19 | field 20 | Constr : Set ℓ 21 | Arity : Constr → Set ℓ 22 | Red : Constr → CofProp 23 | ev : (c : Constr) → [ Red c ] → (Arity c) 24 | 25 | module _ (P : Poly) where 26 | open Poly P 27 | 28 | postulate 29 | WR : Set ℓ 30 | 31 | sup : (c : Constr) → (Arity c → WR) → WR 32 | red : (c : Constr) → (α : Arity c → WR) → 33 | (x : [ Red c ]) → (sup c α ≡ α (ev c x)) 34 | 35 | WR-elim : ∀ {ℓ'} (X : WR → Set ℓ') 36 | (s : (c : Constr) (α : Arity c → WR) → ((a : Arity c) → X (α a)) → X (sup c α)) → 37 | (r : (c : Constr) (α : Arity c → WR) → (u : (a : Arity c) → X (α a)) → 38 | (x : [ Red c ]) → subst X (red c α x) (s c α u) ≡ u (ev c x)) 39 | (w : WR) → X w 40 | 41 | WR-elim-β : ∀ {ℓ'} (X : WR → Set ℓ') 42 | (s : (c : Constr) (α : Arity c → WR) → ((a : Arity c) → X (α a)) → X (sup c α)) 43 | (r : (c : Constr) (α : Arity c → WR) → (u : (a : Arity c) → X (α a)) → 44 | (x : [ Red c ]) → subst X (red c α x) (s c α u) ≡ u (ev c x)) 45 | (c : Constr) (α : Arity c → WR) (u : (a : Arity c) → X (α a)) → 46 | WR-elim X s r (sup c α) ≡ s c α (λ a → WR-elim X s r (α a)) 47 | 48 | record IxPoly (I : Set) : Set (lsuc ℓ) where 49 | field 50 | Constr : I → Set ℓ 51 | Arity : {i : I} → Constr i → I → Set ℓ 52 | Red : {i : I} → Constr i → CofProp 53 | ev : {i : I} → (c : Constr i) → [ Red c ] → (Arity c i) 54 | 55 | module _ {I : Set} (P : IxPoly I) where 56 | open IxPoly P 57 | 58 | postulate 59 | IxWR : I → Set ℓ 60 | 61 | ixsup : {i : I} (c : Constr i) → ((j : I) → Arity c j → IxWR j) → IxWR i 62 | ixred : {i : I} (c : Constr i) (α : (j : I) → Arity c j → IxWR j) → 63 | (x : [ Red c ]) → (ixsup c α ≡ α i (ev c x)) 64 | 65 | IxWR-elim : ∀ {ℓ'} (X : (i : I) → IxWR i → Set ℓ') 66 | (s : {i : I} (c : Constr i) (α : (j : I) → Arity c j → IxWR j) 67 | → ((j : I) (a : Arity c j) → X j (α j a)) → X i (ixsup c α)) 68 | (r : {i : I} (c : Constr i) (α : (j : I) → Arity c j → IxWR j) 69 | (u : (j : I) (a : Arity c j) → X j (α j a)) 70 | (x : [ Red c ]) → subst (X i) (ixred c α x) (s c α u) ≡ u i (ev c x)) 71 | (i : I) (w : IxWR i) → X i w 72 | 73 | IxWR-elim-β : ∀ {ℓ'} (X : (i : I) → IxWR i → Set ℓ') 74 | (s : {i : I} (c : Constr i) (α : (j : I) → Arity c j → IxWR j) 75 | → ((j : I) (a : Arity c j) → X j (α j a)) → X i (ixsup c α)) 76 | (r : {i : I} (c : Constr i) (α : (j : I) → Arity c j → IxWR j) 77 | (u : (j : I) (a : Arity c j) → X j (α j a)) 78 | (x : [ Red c ]) → subst (X i) (ixred c α x) (s c α u) ≡ u i (ev c x)) 79 | (i : I) (c : Constr i) (α : (j : I) → Arity c j → IxWR j) (u : (j : I) (a : Arity c j) → X j (α j a)) 80 | → IxWR-elim X s r i (ixsup c α) ≡ s c α (λ j a → IxWR-elim X s r j (α j a)) 81 | 82 | {-# REWRITE IxWR-elim-β #-} 83 | -------------------------------------------------------------------------------- /agda/retract.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module retract where 3 | 4 | open import prelude 5 | open import interval 6 | open import cofprop 7 | open import fibrations 8 | open import hcomp-coe 9 | 10 | homotopyRetractIsFib : ∀ {ℓ} {Γ : Set ℓ} 11 | (A B : Γ → Set) (α : isHFib A) (β : isFib B) 12 | (f : ∀ x → A x → B x) (g : ∀ x → B x → A x) 13 | (ret : ∀ x a → g x (f x a) ~ a) 14 | → isFib A 15 | homotopyRetractIsFib A B α β f g ret r p φ h (a₀ , ext) = 16 | record 17 | { comp = λ s → 18 | ( mainComp s O (gCompBase s) .comp I (≡IIsCofProp O) .fst 19 | , λ u → 20 | trans 21 | (mainComp s O (gCompBase s) .comp I (≡IIsCofProp O) .snd u) 22 | (symm (ret (p s) (h u s) .atI)) 23 | ) 24 | ; cap = 25 | ( path 26 | (λ t → mainCap t .comp I .fst) 27 | (symm (mainCap O .comp I .snd ∣ inr ∣ inl refl ∣ ∣)) 28 | (trans 29 | (trans 30 | (ret (p r) a₀ .atI) 31 | (mainComp r I (bRetBase I I) .cap (≡IIsCofProp I))) 32 | (symm (mainCap I .comp I .snd ∣ inr ∣ inr refl ∣ ∣))) 33 | , λ t u → mainCap t .comp I .snd ∣ inl u ∣ 34 | ) 35 | } 36 | where 37 | bTube : [ φ ] → (i : Int) → B (p i) 38 | bTube u i = f (p i) (h u i) 39 | 40 | bBase : B (p r) [ φ ↦ bTube ◆ r ] 41 | bBase = (f (p r) a₀ , λ u → cong (f (p r)) (ext u)) 42 | 43 | bComp = β r p φ bTube bBase 44 | 45 | module _ (s : Int) where 46 | 47 | compTube : [ φ ] → (j : Int) → A (p s) 48 | compTube u = ret (p s) (h u s) .at 49 | 50 | gCompBase : A (p s) [ φ ↦ compTube ◆ O ] 51 | gCompBase = 52 | ( g (p s) (bComp .comp s .fst) 53 | , λ u → 54 | trans 55 | (cong (g (p s)) (bComp .comp s .snd u)) 56 | (ret (p s) (h u s) .atO) 57 | ) 58 | 59 | mainComp : (i : Int) → _ 60 | mainComp i = strictifyHFib A α i (p s) φ compTube 61 | 62 | module _ (t : Int) where 63 | 64 | gCapBase : A (p r) [ φ ↦ compTube r ◆ O ] 65 | gCapBase = 66 | ( g (p r) (bComp .cap .fst .at t) 67 | , λ u → 68 | trans 69 | (cong (g (p r)) (bComp .cap .snd t u)) 70 | (ret (p r) (h u r) .atO) 71 | ) 72 | 73 | bRetBase : (k : Int) → A (p r) [ φ ↦ compTube r ◆ k ] 74 | bRetBase k = 75 | ( ret (p r) a₀ .at k 76 | , λ u → cong (λ a → ret (p r) a .at k) (ext u) 77 | ) 78 | 79 | capTube : [ φ ∨ t ≈O ∨ t ≈I ] → Int → A (p r) 80 | capTube = 81 | ∨-rec φ (t ≈O ∨ t ≈I) 82 | (λ u _ → h u r) 83 | (OI-rec t 84 | (λ {refl k → mainComp r O (gCompBase r) .comp I (≡IIsCofProp O) .fst}) 85 | (λ {refl k → mainComp r k (bRetBase k) .comp I (≡IIsCofProp k) .fst})) 86 | (λ u → ∨-elimEq (t ≈O) (t ≈I) 87 | (λ {refl → funext λ k → 88 | trans 89 | (mainComp r O (gCompBase r) .comp I (≡IIsCofProp O) .snd u) 90 | (symm (ret (p r) (h u r) .atI))}) 91 | λ {refl → funext λ k → 92 | trans 93 | (mainComp r k (bRetBase k) .comp I (≡IIsCofProp k) .snd u) 94 | (symm (ret (p r) (h u r) .atI))}) 95 | 96 | capBase : A (p r) [ φ ∨ t ≈O ∨ t ≈I ↦ capTube ◆ O ] 97 | capBase = 98 | ( mainComp r O gCapBase .comp I (≡IIsCofProp O) .fst 99 | , ∨-elimEq φ (t ≈O ∨ t ≈I) 100 | (λ u → 101 | trans 102 | (mainComp r O gCapBase .comp I (≡IIsCofProp O) .snd u) 103 | (symm (ret (p r) (h u r) .atI))) 104 | (OI-elim t _ 105 | (λ {refl → 106 | cong (λ base → mainComp r O base .comp I (≡IIsCofProp O) .fst) 107 | (Σext 108 | (cong (g (p r)) (symm (bComp .cap .fst .atO))) 109 | (funext λ _ → uipImp))}) 110 | λ {refl → 111 | cong (λ base → mainComp r O base .comp I (≡IIsCofProp O) .fst) 112 | (Σext 113 | (trans 114 | (cong (g (p r)) (symm (bComp .cap .fst .atI))) 115 | (ret (p r) a₀ .atO)) 116 | (funext λ _ → uipImp))}) 117 | ) 118 | 119 | mainCap = α O (p r) (φ ∨ t ≈O ∨ t ≈I) capTube capBase 120 | -------------------------------------------------------------------------------- /agda/Data/paths.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Fibrancy of Path types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module Data.paths where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | 14 | ---------------------------------------------------------------------- 15 | -- Path types 16 | ---------------------------------------------------------------------- 17 | 18 | enlargeSys : {A : Int → Set} (φ : CofProp) 19 | (f : (t : Int) → [ φ ] → A t) 20 | (a₀ : A O [ φ ↦ f O ]) (a₁ : A I [ φ ↦ f I ]) 21 | (t : Int) → [ φ ∨ t ≈O ∨ t ≈I ] → A t 22 | enlargeSys φ f (a₀ , ex₀) (a₁ , ex₁) t = 23 | ∨-rec φ (t ≈O ∨ t ≈I) 24 | (f t) 25 | (OI-rec t (λ {refl → a₀}) (λ {refl → a₁})) 26 | (λ u → 27 | ∨-elimEq (t ≈O) (t ≈I) 28 | (λ {refl → ex₀ u}) 29 | (λ {refl → ex₁ u})) 30 | 31 | enlargedExtends : {A : Int → Int → Set} (φ : CofProp) 32 | (f : (t : Int) → [ φ ∨ t ≈O ∨ t ≈I ] → (i : Int) → A t i) 33 | (r : Int) (a : (t : Int) → A t r) 34 | (pφ : (t : Int) → (λ u → f t ∣ inl u ∣ r) ↗ a t) 35 | (pO : f O ∣ inr ∣ inl refl ∣ ∣ r ≡ a O) 36 | (pI : f I ∣ inr ∣ inr refl ∣ ∣ r ≡ a I) 37 | (t : Int) → f t ◆ r ↗ a t 38 | enlargedExtends φ f r a pφ pO pI t = 39 | ∨-elimEq φ (t ≈O ∨ t ≈I) 40 | (pφ t) 41 | (∨-elimEq (t ≈O) (t ≈I) 42 | (λ {refl → pO}) 43 | (λ {refl → pI})) 44 | 45 | Path : ∀{ℓ}{Γ : Set ℓ}(A : Γ → Set) → Σ x ∈ Γ , A x × A x → Set 46 | Path A (x , (a , a')) = a ~ a' 47 | 48 | FibPathId : 49 | {A : Int → Set} 50 | (α : isFib A) 51 | → ----------- 52 | isFib (Path A) 53 | FibPathId {A} α r p φ f (q , ex) = 54 | record 55 | { comp = λ s → 56 | ( path 57 | (λ j → compA j .comp s .fst) 58 | (symm (compA O .comp s .snd ∣ inr ∣ inl refl ∣ ∣)) 59 | (symm (compA I .comp s .snd ∣ inr ∣ inr refl ∣ ∣)) 60 | , λ u → PathExt (funext λ j → compA j .comp s .snd ∣ inl u ∣) 61 | ) 62 | ; cap = 63 | ( path 64 | (λ k → 65 | path 66 | (λ j → compA j .cap .fst .at k) 67 | (symm (compA O .cap .snd k ∣ inr ∣ inl refl ∣ ∣)) 68 | (symm (compA I .cap .snd k ∣ inr ∣ inr refl ∣ ∣))) 69 | (PathExt (funext λ j → compA j .cap .fst .atO)) 70 | (PathExt (funext λ j → compA j .cap .fst .atI)) 71 | , λ k u → PathExt (funext λ j → compA j .cap .snd k ∣ inl u ∣) 72 | ) 73 | } 74 | where 75 | sysA : (j : Int) → [ φ ∨ j ≈O ∨ j ≈I ] → Π (λ i → A (p i .fst)) 76 | sysA = 77 | enlargeSys φ 78 | (λ j u i → f u i .at j) 79 | ((λ i → p i .snd .fst) , λ u → funext (λ i → f u i .atO)) 80 | ((λ i → p i .snd .snd) , λ u → funext (λ i → f u i .atI)) 81 | 82 | module _ (j : Int) where 83 | 84 | baseA : A (p r .fst) [ (φ ∨ j ≈O ∨ j ≈I) ↦ (sysA j) ◆ r ] 85 | baseA = 86 | ( q .at j 87 | , enlargedExtends φ sysA r (q .at) 88 | (λ j u → cong (λ z → z .at j) (ex u)) 89 | (symm (q .atO)) 90 | (symm (q .atI)) 91 | j 92 | ) 93 | 94 | compA : WComp r (λ i → A (p i .fst)) (φ ∨ j ≈O ∨ j ≈I) (sysA j) baseA 95 | compA = α r (λ i → p i .fst) (φ ∨ j ≈O ∨ j ≈I) (sysA j) baseA 96 | 97 | -- trick to get Agda to check reindexing stability quickly 98 | private 99 | abstract 100 | FibPathId' = FibPathId 101 | 102 | FibPath : 103 | ∀{ℓ}{Γ : Set ℓ} 104 | {A : Γ → Set} 105 | (α : isFib A) 106 | → ----------- 107 | isFib (Path A) 108 | FibPath {_} {Γ} {A} α e p = FibPathId' (reindex A α (fst ∘ p)) e (id× p) where 109 | id×_ : (p : Int → Σ Γ (λ x → A x × A x)) → Int → Σ Int (λ i → A (fst (p i)) × A (fst (p i))) 110 | (id× p) i = (i , snd (p i)) 111 | 112 | FibPath' : 113 | ∀{ℓ}{Γ : Set ℓ} 114 | (A : Fib Γ) 115 | → ----------- 116 | Fib (Σ x ∈ Γ , fst A x × fst A x) 117 | FibPath' (A , α) = (Path A , FibPath {A = A} α) 118 | 119 | ---------------------------------------------------------------------- 120 | -- Forming Path types is stable under reindexing 121 | ---------------------------------------------------------------------- 122 | reindexPath : 123 | ∀{ℓ ℓ'}{Δ : Set ℓ}{Γ : Set ℓ'} 124 | (A : Γ → Set) 125 | (α : isFib A) 126 | (ρ : Δ → Γ) 127 | → ---------------------- 128 | reindex (Path A) (FibPath α) (ρ ×id) ≡ FibPath (reindex A α ρ) 129 | reindexPath A α ρ = refl 130 | -------------------------------------------------------------------------------- /agda/equivs.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definitions of contractible, extensible (SContr), fibers, equivalences 4 | and quasi-invertible maps. 5 | 6 | -} 7 | {-# OPTIONS --rewriting #-} 8 | module equivs where 9 | 10 | open import prelude 11 | open import interval 12 | open import cofprop 13 | open import fibrations 14 | open import hcomp-coe 15 | 16 | ---------------------------------------------------------------------- 17 | -- Contr and Ext (now called SContr) 18 | ---------------------------------------------------------------------- 19 | 20 | Contr : Set → Set 21 | Contr A = Σ a₀ ∈ A , ((a : A) → a ~ a₀) 22 | 23 | Contr' : ∀ {ℓ} {Γ : Set ℓ} → (Γ → Set) → Set ℓ 24 | Contr' A = (x : _) → Contr (A x) 25 | 26 | SContr : ∀ {ℓ} (A : Set ℓ) → Set ℓ 27 | SContr A = (φ : CofProp) → (t : [ φ ] → A) → A [ φ ↦ t ] 28 | 29 | SContr→Contr : {A : Set} → SContr A → Contr A 30 | SContr→Contr h = x , λ y → path (λ i → α y i .fst) 31 | (symm (α y O .snd ∣ inl refl ∣)) 32 | (symm (α y I .snd ∣ inr refl ∣)) 33 | where 34 | x = h ⊥ O≠I .fst 35 | α = λ y i → h (i ≈O ∨ i ≈I) (∨-rec (i ≈O) (i ≈I) (λ _ → y) (λ _ → x) (λ {refl h → O≠I h})) 36 | 37 | -- TODO: We need to assume that A has hcomp to finish this 38 | -- Contr→SContr : (A : Set) → Contr A → SContr A 39 | -- Contr→SContr A h = λ φ t → h .fst , λ u → {!h .snd (t u)!} 40 | 41 | ---------------------------------------------------------------------- 42 | -- Equivalences and quasi-inverses 43 | ---------------------------------------------------------------------- 44 | 45 | Fiber : {A B : Set} (f : A → B) (b : B) → Set 46 | Fiber {A} f b = Σ a ∈ A , f a ~ b 47 | 48 | Equiv : {A B : Set} (f : A → B) → Set 49 | Equiv {B = B} f = (b : B) → Contr (Fiber f b) 50 | 51 | EquivSContr : {A B : Set} (f : A → B) → Set 52 | EquivSContr {B = B} f = (b : B) → SContr (Fiber f b) 53 | 54 | EquivSContr→Equiv : {A B : Set} {f : A → B} → EquivSContr f → Equiv f 55 | EquivSContr→Equiv h b = SContr→Contr (h b) 56 | 57 | Equiv' : ∀ {ℓ} {Γ : Set ℓ} {A B : Γ → Set} (f : (x : Γ) → A x → B x) → Set ℓ 58 | Equiv' {Γ = Γ} f = (x : Γ) → Equiv (f x) 59 | 60 | Qinv : {A B : Set} (f : A → B) → Set 61 | Qinv {A} {B} f = Σ g ∈ (B → A) , ((b : B) → f (g b) ~ b) × ((a : A) → g (f a) ~ a) 62 | 63 | FiberExt : {A B : Set} {f : A → B} {b : B} {x y : Fiber f b} 64 | → x .fst ≡ y .fst → x .snd .at ≡ y .snd .at → x ≡ y 65 | FiberExt refl refl = Σext refl (PathExt refl) 66 | 67 | FiberExtDep : {A B : Set} {f : A → B} {b b' : B} (p : b ≡ b') {x : Fiber f b} {y : Fiber f b'} 68 | → x .fst ≡ y .fst → x .snd .at ≡ y .snd .at → subst (Fiber f) p x ≡ y 69 | FiberExtDep refl = FiberExt 70 | 71 | fiberPathEq : {A B : Set} {f : A → B} {b : B} {x y : Fiber f b} 72 | → x ≡ y → ∀ k → x .snd .at k ≡ y .snd .at k 73 | fiberPathEq refl _ = refl 74 | 75 | fiberPathEqDep : {A B : Set} {f : A → B} {b b' : B} (p : b ≡ b') {x : Fiber f b} {y : Fiber f b'} 76 | → subst (Fiber f) p x ≡ y → ∀ k → x .snd .at k ≡ y .snd .at k 77 | fiberPathEqDep refl refl _ = refl 78 | 79 | fiberDomEqDep : {A B : Set} {f : A → B} {b b' : B} (p : b ≡ b') {x : Fiber f b} {y : Fiber f b'} 80 | → subst (Fiber f) p x ≡ y → x .fst ≡ y .fst 81 | fiberDomEqDep refl refl = refl 82 | 83 | -- The identity map on a fibration is an equivalence 84 | idEquiv : ∀ {ℓ} {Γ : Set ℓ} {A : Γ → Set} → isFib A → Equiv' {A = A} (λ _ → id) 85 | idEquiv {A = A} α x a .fst = (a , refl~ a) 86 | idEquiv {A = A} α x a .snd (a' , p) = 87 | path 88 | (λ i → 89 | ( q i .comp O (I≡IsCofProp O) .fst 90 | , path 91 | (λ j → q i .comp j (I≡IsCofProp j) .fst) 92 | refl 93 | (q i .cap (I≡IsCofProp I)) 94 | )) 95 | (FiberExt 96 | (trans (p .atO) (symm (q O .comp O (I≡IsCofProp O) .snd ∣ inl refl ∣))) 97 | (funext λ j → symm (q O .comp j (I≡IsCofProp j) .snd ∣ inl refl ∣))) 98 | (FiberExt 99 | (symm (q I .comp O (I≡IsCofProp O) .snd ∣ inr refl ∣)) 100 | (funext λ j → symm (q I .comp j (I≡IsCofProp j) .snd ∣ inr refl ∣))) 101 | where 102 | q : (i : Int) → _ 103 | q i = 104 | strictifyFib A α I (λ _ → x) (i ≈O ∨ i ≈I) 105 | (OI-rec i 106 | (λ {refl → p .at}) 107 | (λ {refl _ → a})) 108 | ( a 109 | , ∨-elimEq (i ≈O) (i ≈I) 110 | (λ {refl → p .atI}) 111 | (λ {refl → refl}) 112 | ) 113 | -------------------------------------------------------------------------------- /agda/trivialfib.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of trivial fibrations and proof that they are actually 4 | fibrations. 5 | 6 | -} 7 | {-# OPTIONS --rewriting #-} 8 | module trivialfib where 9 | 10 | open import prelude 11 | open import cofprop 12 | open import fibrations 13 | open import interval 14 | open import equivs 15 | 16 | O≡I-elim : ∀ {ℓ} {i : Int} {D : [ i ≈O ] → [ i ≈I ] → Set ℓ} (v : i ≡ O) (u : i ≡ I) → D v u 17 | O≡I-elim v u = O≠I (trans u (symm v)) 18 | 19 | path-contr : {A : Set} (ac : SContr A) (a a' : A) → SContr (a ~ a') 20 | path-contr {A} ac a a' φ u = 21 | path (λ i → fst (ac (φ' i) (u' i))) 22 | (symm (snd (ac (φ' O) (u' O)) ∣ inr ∣ inl refl ∣ ∣)) 23 | (symm (snd (ac (φ' I) (u' I)) ∣ inr ∣ inr refl ∣ ∣)) 24 | , λ x → PathExt (funext (λ i → snd (ac (φ' i) (u' i)) ∣ inl x ∣)) 25 | where 26 | φ' = λ i → φ ∨ (i ≈O) ∨ (i ≈I) 27 | u' : (i : Int) → [ φ' i ] → A 28 | u' i = ∨-rec φ ((i ≈O) ∨ (i ≈I)) (λ x → u x .at i) (OI-rec i (λ _ → a) (λ _ → a')) 29 | (λ w v → ∨-elim (i ≈O) (i ≈I) (λ v' → u w .at i ≡ OI-rec i (λ _ → a) (λ _ → a') v') 30 | (λ {refl → (u w) .atO}) (λ {refl → (u w) .atI}) O≡I-elim v) 31 | 32 | 33 | scontr-iso : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} (f : A → B) (g : B → A) 34 | (f-g : (b : B) → f (g b) ≡ b) (g-f : (a : A) → g (f a) ≡ a) → 35 | SContr A → SContr B 36 | scontr-iso f g f-g g-f acontr = 37 | λ φ u → (f (fst (acontr φ (g ∘ u)))) , 38 | λ x → trans (cong f (snd (acontr φ (g ∘ u)) x)) (symm (f-g (u x))) 39 | 40 | isTrivialFib : ∀ {ℓ ℓ'} {Γ : Set ℓ} (B : Γ → Set ℓ') → Set (ℓ ⊔ ℓ') 41 | isTrivialFib B = (γ : _) → SContr (B γ) 42 | 43 | SFiber : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Set ℓ'} → (A → B) → B → Set (ℓ ⊔ ℓ') 44 | SFiber f b = Σ _ (λ a → f a ≡ b) 45 | 46 | map-isTrivialFib : {A : Set} {B : Set} (f : A → B) → Set 47 | map-isTrivialFib f = isTrivialFib (SFiber f) 48 | 49 | TrivialFib-tomap : {Γ : Set} (A : Γ → Set) → isTrivialFib A → map-isTrivialFib fst 50 | TrivialFib-tomap {Γ} A tf γ = scontr-iso (λ a → (γ , a) , refl) (λ {((γ' , a'), p) → subst A p a'}) 51 | (λ {((γ' , a'), p) → lemma p a'}) (λ a → refl) (tf γ) 52 | where 53 | lemma : {γ γ' : Γ} (p : γ' ≡ γ) (a' : A γ') → (((γ , subst A p a') , refl) ≡ ((γ' , a') , p)) 54 | lemma refl a' = refl 55 | 56 | TrivialFib-∘ : {A B C : Set} (g : B → C) (f : A → B) → map-isTrivialFib g → map-isTrivialFib f → 57 | map-isTrivialFib (g ∘ f) 58 | TrivialFib-∘ g f gtf ftf c φ u = (a , (trans p (cong g q))) , (λ x → Σext (Σeq₁ (abdy x)) (uip _ _)) 59 | where 60 | u' : [ φ ] → SFiber g c 61 | u' x = f (fst (u x)) , snd (u x) 62 | 63 | bp0 = gtf c φ u' 64 | 65 | b = fst (fst bp0) 66 | p = snd (fst bp0) 67 | bbdy = snd bp0 68 | 69 | aq0 = ftf b φ (λ x → (fst (u x)) , Σeq₁ (bbdy x)) 70 | a = fst (fst aq0) 71 | q = snd (fst aq0) 72 | abdy = snd aq0 73 | 74 | 75 | SContrToPaths : {A : Int → Set} (tf : isTrivialFib A) (a : A O) (a' : A I) → a ~~ a' 76 | SContrToPaths {A} tf a a' = 77 | path p e0 e1 78 | where 79 | φ : Int → CofProp 80 | φ i = (i ≈O) ∨ (i ≈I) 81 | 82 | u : (i : Int) → [ φ i ] → A i 83 | u i x = OI-rec i ((λ p → subst A (symm p) a)) ((λ q → subst A (symm q) a')) x 84 | 85 | pc = λ i → tf i (φ i) (u i) 86 | p = λ i → fst (pc i) 87 | 88 | e0 : p O ≡ a 89 | e0 = symm (snd (pc O) ∣ inl refl ∣) 90 | 91 | e1 : p I ≡ a' 92 | e1 = symm (snd (pc I) ∣ inr refl ∣) 93 | 94 | TrivialFib-isFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isTrivialFib A → isFib A 95 | TrivialFib-isFib A tf r p φ f x₀ = 96 | record { comp = co ; cap = ca } 97 | where 98 | co = λ i → tf (p i) φ (λ x → f x i) 99 | 100 | φ' = λ i → φ ∨ (i ≈O) ∨ (i ≈I) 101 | u : (i : Int) → [ φ' i ] → A (p r) 102 | u i x = ∨-rec φ ((i ≈O) ∨ (i ≈I)) (λ x' → f x' r) 103 | endpoints 104 | (λ v w → ∨-elim (i ≈O) (i ≈I) (λ w' → f v r ≡ endpoints w') 105 | (λ p → snd (co r) v) 106 | (λ p → snd x₀ v) 107 | O≡I-elim w) 108 | x 109 | where 110 | endpoints = OI-rec i ((λ _ → fst (co r))) ((λ _ → fst x₀)) 111 | 112 | pth = λ i → fst (tf (p r) (φ' i) (u i)) 113 | bdy = λ i → snd (tf (p r) (φ' i) (u i)) 114 | 115 | ca = path pth (symm (bdy O ∣ inr ∣ inl refl ∣ ∣)) 116 | (symm (bdy I ∣ inr ∣ inr refl ∣ ∣)) , 117 | λ i x → bdy i ∣ inl x ∣ 118 | -------------------------------------------------------------------------------- /agda/Data/id.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | CCHM style identity types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module Data.id where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | open import Data.paths 14 | open import Data.products 15 | 16 | postulate 17 | _∧_ : (φ : CofProp) → ([ φ ] → CofProp) → CofProp 18 | [∧] : ∀ φ ψ → [ φ ∧ ψ ] ≡ Σ [ φ ] ([_] ∘ ψ) 19 | 20 | {-# REWRITE [∧] #-} 21 | 22 | cofExt : {φ ψ : CofProp} → ([ φ ] → [ ψ ]) → ([ ψ ] → [ φ ]) → φ ≡ ψ 23 | 24 | ---------------------------------------------------------------------- 25 | -- Id types 26 | ---------------------------------------------------------------------- 27 | 28 | Constancy : ∀{ℓ}{Γ : Set ℓ}(A : Γ → Set) 29 | → Σ (Σ Γ (λ x → A x × A x)) (Path A) → Set 30 | Constancy A ((x , (a , a')) , p) = Σ φ ∈ CofProp , ([ φ ] → (i : Int) → p .at i ≡ a) 31 | 32 | Id : ∀{ℓ}{Γ : Set ℓ}(A : Γ → Set) → Σ x ∈ Γ , A x × A x → Set 33 | Id {_} {Γ} A = Σ' (Path A) (Constancy A) 34 | 35 | pathToId : ∀{ℓ}{Γ : Set ℓ}{A : Γ → Set}{x : Γ}{a a' : A x} → Path A (x , a' , a') → Id A (x , a' , a') 36 | pathToId p = (p , ⊥ , ⊥→) 37 | 38 | ConstancyExt : ∀{ℓ}{Γ : Set ℓ}(A : Γ → Set) 39 | {x : Γ} {a a' : A x} 40 | {p p' : a ~ a'} (eq : p ≡ p') 41 | {u : Constancy A ((x , (a , a')) , p)} 42 | {v : Constancy A ((x , (a , a')) , p')} 43 | → u .fst ≡ v .fst 44 | → subst ((λ p → Constancy A ((x , a , a') , p))) eq u ≡ v 45 | ConstancyExt A {a = a} {p = p} {_} refl {(φ , _)} refl = 46 | Σext refl (funext λ _ → funext λ _ → uipImp) 47 | 48 | IdExt : 49 | ∀{ℓ}{Γ : Set ℓ} 50 | {A : Γ → Set} 51 | {x : Γ} 52 | {a a' : A x} 53 | {id id' : Id A (x , (a , a'))} 54 | → ------------ 55 | (id .fst .at ≡ id' .fst .at) → id .snd .fst ≡ id' .snd .fst → id ≡ id' 56 | IdExt {_} {Γ} {A} {x} {a} {a'} {(p , u)} refl refl = 57 | let pathEq = (PathExt refl) in 58 | Σext pathEq (ConstancyExt A pathEq refl) 59 | 60 | FibIdId : {A : Int → Set} → isFib A → isFib (Id A) 61 | FibIdId {A} α = FibΣ {B = Constancy A} (FibPath {A = A} α) β 62 | where 63 | β : isFib (Constancy A) 64 | β r p φ f ((ψ₀ , cst₀) , ex) = 65 | record 66 | { comp = λ s → 67 | ( ( φ ∧ (λ u → f u s .fst) 68 | , λ {(u , v) → f u s .snd v} 69 | ) 70 | , λ u → 71 | ConstancyExt A {p = snd (p s)} refl 72 | (cofExt 73 | (λ v → u , v) 74 | (λ {(u' , v) → subst (λ w → [ f w s .fst ]) (cofIsProp φ _ _) v})) 75 | ) 76 | ; cap = 77 | ( path 78 | (λ t → 79 | ( (φ ∧ (λ u → f u r .fst)) ∨ ((t ≈I) ∧ (λ _ → ψ₀)) 80 | , ∨-rec (φ ∧ (λ u → f u r .fst)) ((t ≈I) ∧ (λ _ → ψ₀)) 81 | (λ {(u , v) → f u r .snd v}) 82 | (λ {(refl , w) → cst₀ w}) 83 | (λ _ _ → funext λ _ → uipImp) 84 | ) 85 | ) 86 | (ConstancyExt A {p = snd (p r)} _ 87 | (cofExt 88 | (∨-rec (φ ∧ (λ u → f u r .fst)) (⊥ ∧ (λ _ → ψ₀)) 89 | id 90 | (O≠I ∘ fst) 91 | (λ _ _ → cofIsProp (φ ∧ (λ u → f u r .fst)) _ _)) 92 | (∣_∣ ∘ inl))) 93 | (ConstancyExt A {p = snd (p r)} _ 94 | (cofExt 95 | (∨-rec (φ ∧ (λ u → f u r .fst)) ((I ≈I) ∧ (λ _ → ψ₀)) 96 | (λ {(u , v) → subst ([_] ∘ fst) (ex u) v}) 97 | snd 98 | (λ _ _ → cofIsProp ψ₀ _ _)) 99 | (λ w → ∣ inr (refl , w) ∣))) 100 | , λ t u → 101 | ConstancyExt A {p = snd (p r)} _ 102 | (cofExt 103 | (λ v → ∣ inl (u , v) ∣) 104 | (∨-rec (φ ∧ (λ u → f u r .fst)) ((t ≈I) ∧ (λ _ → ψ₀)) 105 | (λ {(u' , v) → subst (λ x → [ f x r .fst ]) (cofIsProp φ _ _) v}) 106 | (λ {(refl , w) → subst ([_] ∘ fst) (symm (ex u)) w}) 107 | (λ _ _ → cofIsProp (f u r .fst) _ _))) 108 | ) 109 | } 110 | 111 | FibId : ∀{ℓ}{Γ : Set ℓ}{A : Γ → Set} → isFib A → isFib (Id A) 112 | FibId {_} {Γ} {A} α e p = FibIdId (reindex A α (fst ∘ p)) e (id× p) where 113 | id×_ : (p : Int → Σ Γ (λ x → A x × A x)) → Int → Σ Int (λ i → A (fst (p i)) × A (fst (p i))) 114 | (id× p) i = (i , snd (p i)) 115 | 116 | ---------------------------------------------------------------------- 117 | -- Forming Id types is stable under reindexing 118 | ---------------------------------------------------------------------- 119 | reindexId : 120 | ∀{ℓ ℓ'}{Δ : Set ℓ}{Γ : Set ℓ'} 121 | (A : Γ → Set) 122 | (α : isFib A) 123 | (ρ : Δ → Γ) 124 | → ---------------------- 125 | reindex (Id A) (FibId α) (ρ ×id) ≡ FibId (reindex A α ρ) 126 | reindexId A α ρ = refl 127 | 128 | 129 | -- EC : TODO J 130 | -------------------------------------------------------------------------------- /agda/cofreplacement.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Cofibrant replacement and definition of (trivial) cofibrations. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module cofreplacement where 8 | 9 | open import prelude 10 | open import cofprop 11 | open import fibrations 12 | open import interval 13 | open import trivialfib 14 | open import wtypesred 15 | open import equivs 16 | 17 | crpoly : Set → Poly 18 | crpoly A = record { Constr = A ⊎ CofProp 19 | ; Arity = λ {(inl a) → ∅ ; (inr φ) → [ φ ]} 20 | ; Red = λ {(inl a) → (O ≈I) ; (inr φ) → φ } 21 | ; ev = λ {(inl a) → O≠I ; (inr φ) → λ x → x} } 22 | 23 | 24 | CR : Set → Set 25 | CR A = WR (crpoly A) 26 | 27 | incl : {A : Set} → A → CR A 28 | incl a = sup _ (inl a) ∅-elim 29 | 30 | fill : {A : Set} (φ : CofProp) → (u : [ φ ] → CR A) → CR A 31 | fill φ u = sup _ (inr φ) u 32 | 33 | CR-red : {A : Set} {φ : CofProp} (u : [ φ ] → CR A) → 34 | (x : [ φ ]) → (fill φ u) ≡ u x 35 | CR-red u x = red _ (inr _) u x 36 | 37 | CR-elim : ∀ {ℓ'} {A : Set} (B : CR A → Set ℓ') 38 | (b₀ : (a : A) → (B (incl a))) 39 | (h : (φ : CofProp) → (u : [ φ ] → CR A) → 40 | (f : (x : [ φ ]) → B (u x)) → B (fill φ u)) → 41 | (p : (φ : CofProp) → (u : [ φ ] → CR A) → 42 | (f : (x : [ φ ]) → B (u x)) → 43 | (x : [ φ ]) → (subst B (CR-red u x) (h φ u f)) ≡ f x) → 44 | (z : CR A) → (B z) 45 | CR-elim B b₀ h p = WR-elim _ B 46 | (λ { (inl a) → λ _ _ → subst B (cong (sup _ (inl a)) (funext ∅-elim)) (b₀ a) 47 | ; (inr φ) → h φ}) 48 | λ { (inl a) → λ α u x → O≠I x 49 | ; (inr φ) → p φ} 50 | 51 | CR-β : ∀ {ℓ'} {A : Set} (B : CR A → Set ℓ') 52 | {b₀ : (a : A) → (B (incl a))} 53 | {h : (φ : CofProp) → (u : [ φ ] → CR A) → 54 | (f : (x : [ φ ]) → B (u x)) → B (fill φ u)} → 55 | {p : (φ : CofProp) → (u : [ φ ] → CR A) → 56 | (f : (x : [ φ ]) → B (u x)) → 57 | (x : [ φ ]) → (subst B (CR-red u x) (h φ u f)) ≡ f x} → 58 | (a : A) → (CR-elim B b₀ h p (incl a) ≡ b₀ a) 59 | CR-β B {b₀ = b₀} a = trans (cong (λ p → subst B p (b₀ a)) (uip ((cong (sup _ (inl a)) (funext ∅-elim))) refl)) (WR-elim-β _ B _ _ _ _ ∅-elim) 60 | 61 | 62 | CR-iscontr : {A : Set} → SContr (CR A) 63 | CR-iscontr φ u = fill φ u , λ x → symm (CR-red u x) 64 | 65 | CR-isfib : ∀ {ℓ} (Γ : Set ℓ) (A : Γ → Set) → isFib (CR ∘ A) 66 | CR-isfib Γ A = TrivialFib-isFib (CR ∘ A) (λ a → CR-iscontr) 67 | 68 | record cofibration {A B : Set} (f : A → B) : Set₁ where 69 | field 70 | j : (X : B → Set) (xfib : isTrivialFib X) (x₀ : (a : A) → X (f a)) → (b : B) → X b 71 | upper-tri : (X : B → Set) (xfib : isTrivialFib X) (x₀ : (a : A) → X (f a)) → (a : A) → 72 | j X xfib x₀ (f a) ≡ x₀ a 73 | 74 | record triv-cofibration {A B : Set} (f : A → B) : Set₁ where 75 | field 76 | j : (X : B → Set) (xfib : isFib X) (x₀ : (a : A) → X (f a)) → (b : B) → X b 77 | upper-tri : (X : B → Set) (xfib : isFib X) (x₀ : (a : A) → X (f a)) → (a : A) → 78 | j X xfib x₀ (f a) ≡ x₀ a 79 | 80 | module cof-replace {A B : Set} (f : A → B) where 81 | M₀ : B → Set 82 | M₀ b = CR (SFiber f b) 83 | 84 | M₀-isTrivialFib : isTrivialFib M₀ 85 | M₀-isTrivialFib b = CR-iscontr 86 | 87 | M : Set 88 | M = Σ B M₀ 89 | 90 | L₀ : (a : A) → (M₀ (f a)) 91 | L₀ a = incl (a , refl) 92 | 93 | L : A → M 94 | L a = f a , L₀ a 95 | 96 | L-cof : cofibration L 97 | L-cof = 98 | record { j = j 99 | ; upper-tri = ut } 100 | where 101 | module _ (X : M → Set) (xtf : isTrivialFib X) (x₀ : (a : A) → X (L a)) where 102 | X' : (b : B) → CR (SFiber f b) → Set 103 | X' b z = X (b , z) 104 | 105 | lemma1 : (a : A) (b : B) (p : f a ≡ b) → (f a , incl (a , refl)) ≡ (b , incl (a , p)) 106 | lemma1 a .(f a) refl = refl 107 | 108 | j : (m : M) → X m 109 | j (b , z) = CR-elim (X' b) (λ {(a , p) → subst X (lemma1 a b p) (x₀ a)}) 110 | (λ φ u g → fst (h φ u g)) 111 | (λ φ u g x → lemma2 (CR-red u x) (g x) (fst (h φ u g)) (snd (h φ u g) x)) 112 | z 113 | where 114 | lemma2 : {x y : CR (SFiber f b)} (q : y ≡ x) (z : X' b x) (w : X' b y) → 115 | (subst (X' b) (symm q) z ≡ w) → (subst (X' b) q w ≡ z) 116 | lemma2 refl x y = symm 117 | 118 | h : (φ : CofProp) (u : [ φ ] → CR (SFiber f b)) → 119 | (s : (x : [ φ ]) → X' b (u x)) → 120 | X' b (fill φ u) [ φ ↦ (λ x → subst (λ c' → X' b c') (symm (CR-red u x)) (s x)) ] 121 | h φ u s = xtf (b , fill φ u) φ _ 122 | 123 | ut : (a : A) → j (L a) ≡ x₀ a 124 | ut a = CR-β (X' (f a)) (a , refl) 125 | 126 | 127 | R : M → B 128 | R = fst 129 | 130 | R-tf : map-isTrivialFib R 131 | R-tf = TrivialFib-tomap _ M₀-isTrivialFib 132 | 133 | 134 | -- TODO: make things universe polymorphic so that this works properly 135 | record factorization {A B X : Set} (f : A → B) (P : (A → X) → Set₁) (Q : (X → B) → Set) : Set₁ where 136 | field 137 | g : A → X 138 | hg : P g 139 | h : X → B 140 | hh : Q h 141 | triangle : (a : A) → h (g a) ≡ f a 142 | 143 | open cof-replace 144 | 145 | c-tf-factorization : {A B : Set} (f : A → B) → factorization f cofibration map-isTrivialFib 146 | c-tf-factorization f = record { g = L f ; hg = L-cof f ; h = R f ; hh = R-tf f ; triangle = λ a → refl } 147 | -------------------------------------------------------------------------------- /agda/cofprop.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of the universe of cofibrant propositions and basic 4 | operations on these. 5 | 6 | -} 7 | {-# OPTIONS --rewriting #-} 8 | 9 | module cofprop where 10 | 11 | open import prelude 12 | open import interval 13 | 14 | infixr 4 _∨_ 15 | infix 5 _↗_ 16 | 17 | ---------------------------------------------------------------------- 18 | -- Cofibrant propositions 19 | ---------------------------------------------------------------------- 20 | 21 | postulate 22 | CofProp : Set 23 | [_] : CofProp → Set 24 | 25 | _≈O _≈I : (i : Int) → CofProp 26 | [≈O] : ∀ i → [ i ≈O ] ≡ (i ≡ O) 27 | [≈I] : ∀ i → [ i ≈I ] ≡ (i ≡ I) 28 | 29 | _∨_ : CofProp → CofProp → CofProp 30 | [∨] : ∀ φ ψ → [ φ ∨ ψ ] ≡ ∥ [ φ ] ⊎ [ ψ ] ∥ 31 | 32 | ∀I : (Int → CofProp) → CofProp 33 | [∀I] : ∀ φ → [ ∀I φ ] ≡ ((i : Int) → [ φ i ]) 34 | 35 | {-# REWRITE [≈O] [≈I] [∨] [∀I] #-} 36 | 37 | cofIsProp : (φ : CofProp) → (u v : [ φ ]) → u ≡ v 38 | 39 | ⊥ : CofProp 40 | ⊥ = O ≈I 41 | 42 | ⊥→ : ∀ {ℓ} {A : Set ℓ} → [ ⊥ ] → A 43 | ⊥→ x = O≠I x 44 | 45 | ---------------------------------------------------------------------- 46 | -- Cofibrant-partial function classifier 47 | ---------------------------------------------------------------------- 48 | 49 | _↗_ : ∀ {ℓ} {A : Set ℓ} {ϕ : Set} → (ϕ → A) → A → Set ℓ 50 | f ↗ x = (u : _) → f u ≡ x 51 | 52 | _◆_ : ∀ {ℓ ℓ'} {A : Set ℓ} {B : Int → Set ℓ'} 53 | → (A → (i : Int) → B i) → (i : Int) → A → B i 54 | (f ◆ i) a = f a i 55 | 56 | _[_↦_] : ∀ {ℓ} (A : Set ℓ) (φ : CofProp) → ([ φ ] → A) → Set ℓ 57 | A [ φ ↦ f ] = Σ a ∈ A , f ↗ a 58 | 59 | ---------------------------------------------------------------------- 60 | -- Restricting a context by a cofibrant propositions 61 | ---------------------------------------------------------------------- 62 | res : ∀ {ℓ} (Γ : Set ℓ) (Φ : Γ → CofProp) → Set ℓ 63 | res Γ Φ = Σ x ∈ Γ , [ Φ x ] 64 | 65 | ---------------------------------------------------------------------- 66 | -- Property of being a cofibration 67 | ---------------------------------------------------------------------- 68 | isCofProp : ∀ {ℓ} (A : Set ℓ) → Set ℓ 69 | isCofProp A = Σ φ ∈ CofProp , ([ φ ] → A) × (A → [ φ ]) 70 | 71 | ≡OIsCofProp : (i : Int) → isCofProp (i ≡ O) 72 | ≡OIsCofProp i = i ≈O , id , id 73 | 74 | ≡IIsCofProp : (i : Int) → isCofProp (i ≡ I) 75 | ≡IIsCofProp i = i ≈I , id , id 76 | 77 | O≡IsCofProp : (i : Int) → isCofProp (O ≡ i) 78 | O≡IsCofProp i = i ≈O , symm , symm 79 | 80 | I≡IsCofProp : (i : Int) → isCofProp (I ≡ i) 81 | I≡IsCofProp i = i ≈I , symm , symm 82 | 83 | isCofProp-∅ : isCofProp ∅ 84 | isCofProp-∅ = ⊥ , ⊥→ , ∅-rec 85 | 86 | ---------------------------------------------------------------------- 87 | -- Compatible partial functions 88 | ---------------------------------------------------------------------- 89 | □ : ∀ {ℓ} → Set ℓ → Set ℓ 90 | □ A = Σ φ ∈ CofProp , ([ φ ] → A) 91 | 92 | _⌣_ : ∀ {ℓ} {A : Set ℓ} → □ A → □ A → Set ℓ 93 | (φ , f) ⌣ (ψ , g) = (u : [ φ ]) (v : [ ψ ]) → f u ≡ g v 94 | 95 | ∨-rec : ∀ {ℓ} 96 | (φ ψ : CofProp) 97 | {A : Set ℓ} 98 | (f : [ φ ] → A) 99 | (g : [ ψ ] → A) 100 | (p : (φ , f) ⌣ (ψ , g)) 101 | → --------------------------- 102 | [ φ ∨ ψ ] → A 103 | ∨-rec φ ψ {A = A} f g p = ∥∥-rec _ h q where 104 | 105 | h : [ φ ] ⊎ [ ψ ] → A 106 | h (inl u) = f u 107 | h (inr v) = g v 108 | 109 | q : (z z' : [ φ ] ⊎ [ ψ ]) → h z ≡ h z' 110 | q (inl _) (inl _) = cong f (cofIsProp φ _ _) 111 | q (inl u) (inr v) = p u v 112 | q (inr v) (inl u) = symm (p u v) 113 | q (inr _) (inr _) = cong g (cofIsProp ψ _ _) 114 | 115 | OI-rec : ∀ {ℓ} 116 | (r : Int) 117 | {A : Set ℓ} 118 | → ([ r ≈O ] → A) 119 | → ([ r ≈I ] → A) 120 | → --------------------------- 121 | [ r ≈O ∨ r ≈I ] → A 122 | OI-rec r f g = 123 | ∨-rec (r ≈O) (r ≈I) f g (λ {refl r≡I → O≠I r≡I}) 124 | 125 | ∨-elim : ∀ {ℓ} 126 | (φ ψ : CofProp) 127 | (P : [ φ ∨ ψ ] → Set ℓ) 128 | (f : (u : [ φ ]) → P ∣ inl u ∣) 129 | (g : (v : [ ψ ]) → P ∣ inr v ∣) 130 | (p : (u : [ φ ]) (v : [ ψ ]) → subst P (trunc _ _) (f u) ≡ g v) 131 | → --------------------------- 132 | (w : [ φ ∨ ψ ]) → P w 133 | ∨-elim φ ψ P f g p = ∥∥-elim _ h q 134 | where 135 | h : (z : [ φ ] ⊎ [ ψ ]) → P ∣ z ∣ 136 | h (inl u) = f u 137 | h (inr v) = g v 138 | 139 | q : (z z' : [ φ ] ⊎ [ ψ ]) → subst P (trunc _ _) (h z) ≡ h z' 140 | q (inl u) (inl u') = 141 | trans 142 | (congdep f (cofIsProp φ u u')) 143 | (trans 144 | (appCong (symm (subst-cong-assoc P (∣_∣ ∘ inl) (cofIsProp φ u u')))) 145 | (cong (λ r → subst P r (f u)) 146 | (uip (trunc ∣ inl u ∣ ∣ inl u' ∣) (cong (∣_∣ ∘ inl) (cofIsProp φ u u'))))) 147 | q (inl u) (inr v) = p u v 148 | q (inr v) (inl u) = 149 | adjustSubstEq P 150 | refl (trunc ∣ inl u ∣ ∣ inr v ∣) 151 | (trunc ∣ inr v ∣ ∣ inl u ∣) refl 152 | (symm (p u v)) 153 | q (inr v) (inr v') = 154 | trans 155 | (congdep g (cofIsProp ψ v v')) 156 | (trans 157 | (appCong (symm (subst-cong-assoc P (∣_∣ ∘ inr) (cofIsProp ψ v v')))) 158 | (cong (λ r → subst P r (g v)) 159 | (uip (trunc ∣ inr v ∣ ∣ inr v' ∣) (cong (∣_∣ ∘ inr) (cofIsProp ψ v v'))))) 160 | 161 | OI-elim : ∀ {ℓ} 162 | (r : Int) 163 | (A : [ r ≈O ∨ r ≈I ] → Set ℓ) 164 | → ((rO : [ r ≈O ]) → A ∣ inl rO ∣) 165 | → ((rI : [ r ≈I ]) → A ∣ inr rI ∣) 166 | → --------------------------- 167 | (rOI : [ r ≈O ∨ r ≈I ]) → A rOI 168 | OI-elim r A f g = 169 | ∨-elim (r ≈O) (r ≈I) A f g (λ {refl r≡I → O≠I r≡I}) 170 | 171 | ∨-elimEq : ∀ {ℓ} 172 | (φ ψ : CofProp) {A : Set ℓ} 173 | {f g : [ φ ∨ ψ ] → A} 174 | → ((u : [ φ ]) → f ∣ inl u ∣ ≡ g ∣ inl u ∣) 175 | → ((v : [ ψ ]) → f ∣ inr v ∣ ≡ g ∣ inr v ∣) 176 | → --------------------------- 177 | (w : [ φ ∨ ψ ]) → f w ≡ g w 178 | ∨-elimEq φ ψ p q = 179 | ∨-elim φ ψ _ p q (λ _ _ → uipImp) 180 | -------------------------------------------------------------------------------- /agda/realignment.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Realign a fibration structure. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module realignment where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | 14 | module Realign {ℓ}{Γ : Set ℓ} 15 | (Φ : Γ → CofProp) (A : Γ → Set) 16 | (β : isFib {Γ = res Γ Φ} (A ∘ fst)) (α : isFib A) 17 | (r : Int) (p : Int → Γ) (ψ : CofProp) (f : [ ψ ] → Π (A ∘ p)) 18 | (a₀ : A (p r) [ ψ ↦ f ◆ r ]) 19 | where 20 | wB : [ ∀I (Φ ∘ p) ] → WComp r (A ∘ p) ψ f a₀ 21 | wB u = β r (λ i → p i , u i) ψ f a₀ 22 | 23 | f' : [ ψ ∨ ∀I (Φ ∘ p) ] → (i : Int) → A (p i) 24 | f' = 25 | ∨-rec ψ (∀I (Φ ∘ p)) 26 | f 27 | (λ u i → wB u .comp i .fst) 28 | (λ v u → funext λ i → wB u .comp i .snd v) 29 | 30 | a₀FixTube : [ ψ ∨ ∀I (Φ ∘ p) ] → Int → A (p r) 31 | a₀FixTube = 32 | ∨-rec ψ (∀I (Φ ∘ p)) 33 | (λ v _ → f v r) 34 | (λ u j → wB u .cap .fst .at j) 35 | (λ v u → funext λ j → wB u .cap .snd j v) 36 | 37 | a₀FixBase : A (p r) [ ψ ∨ ∀I (Φ ∘ p) ↦ a₀FixTube ◆ I ] 38 | a₀FixBase = 39 | ( a₀ .fst 40 | , ∨-elimEq ψ (∀I (Φ ∘ p)) 41 | (a₀ .snd) 42 | (λ u → wB u .cap .fst .atI) 43 | ) 44 | 45 | a₀Fix = strictifyFib A α I (λ _ → p r) (ψ ∨ ∀I (Φ ∘ p)) a₀FixTube a₀FixBase 46 | 47 | a₀' : A (p r) [ ψ ∨ ∀I (Φ ∘ p) ↦ f' ◆ r ] 48 | a₀' = 49 | ( a₀Fix .comp O (I≡IsCofProp O) .fst 50 | , ∨-elimEq ψ (∀I (Φ ∘ p)) 51 | (λ v → a₀Fix .comp O (I≡IsCofProp O) .snd ∣ inl v ∣) 52 | (λ u → 53 | trans 54 | (a₀Fix .comp O (I≡IsCofProp O) .snd ∣ inr u ∣) 55 | (symm (wB u .cap .fst .atO))) 56 | ) 57 | 58 | wA = α r p (ψ ∨ ∀I (Φ ∘ p)) f' a₀' 59 | 60 | module _ (t : Int) where 61 | 62 | capTube : [ ψ ∨ ∀I (Φ ∘ p) ∨ t ≈O ∨ t ≈I ] → Int → A (p r) 63 | capTube = 64 | ∨-rec ψ (∀I (Φ ∘ p) ∨ t ≈O ∨ t ≈I) 65 | (λ v _ → f v r) 66 | (∨-rec (∀I (Φ ∘ p)) (t ≈O ∨ t ≈I) 67 | (λ u j → wB u .cap .fst .at t) 68 | (OI-rec t 69 | (λ {refl j → wA .cap .fst .at j}) 70 | (λ {refl j → a₀ .fst})) 71 | (λ u → ∨-elimEq (t ≈O) (t ≈I) 72 | (λ {refl → funext λ j → trans (wA .cap .snd j ∣ inr u ∣) (wB u .cap .fst .atO)}) 73 | (λ {refl → funext λ j → wB u .cap .fst .atI}))) 74 | (λ v → ∨-elimEq (∀I (Φ ∘ p)) (t ≈O ∨ t ≈I) 75 | (λ u → funext λ _ → wB u .cap .snd t v) 76 | (∨-elimEq (t ≈O) (t ≈I) 77 | (λ {refl → funext λ j → wA .cap .snd j ∣ inl v ∣}) 78 | (λ {refl → funext λ _ → a₀ .snd v}))) 79 | 80 | capBase : A (p r) [ ψ ∨ ∀I (Φ ∘ p) ∨ t ≈O ∨ t ≈I ↦ capTube ◆ I ] 81 | capBase = 82 | ( a₀Fix .comp t (I≡IsCofProp t) .fst 83 | , ∨-elimEq ψ (∀I (Φ ∘ p) ∨ t ≈O ∨ t ≈I) 84 | (λ v → a₀Fix .comp t (I≡IsCofProp t) .snd ∣ inl v ∣) 85 | (∨-elimEq (∀I (Φ ∘ p)) (t ≈O ∨ t ≈I) 86 | (λ u → a₀Fix .comp t (I≡IsCofProp t) .snd ∣ inr u ∣) 87 | (∨-elimEq (t ≈O) (t ≈I) 88 | (λ {refl → wA .cap .fst .atI}) 89 | (λ {refl → symm (a₀Fix .cap (I≡IsCofProp I))}))) 90 | ) 91 | 92 | capComp = α I (λ _ → p r) (ψ ∨ ∀I (Φ ∘ p) ∨ t ≈O ∨ t ≈I) capTube capBase .comp O 93 | 94 | abstract 95 | realignId : 96 | (Φ : Int → CofProp) 97 | (A : Int → Set) 98 | (β : isFib {Γ = res Int Φ} (A ∘ fst)) 99 | (α : isFib A) 100 | → --------------- 101 | isFib A 102 | realignId Φ A β α r p ψ f a₀ = 103 | record 104 | { comp = λ s → 105 | ( wA .comp s .fst 106 | , λ v → wA .comp s .snd ∣ inl v ∣ 107 | ) 108 | ; cap = 109 | ( path 110 | (λ t → capComp t .fst) 111 | (trans 112 | (wA .cap .fst .atO) 113 | (symm (capComp O .snd ∣ inr ∣ inr ∣ inl refl ∣ ∣ ∣))) 114 | (symm (capComp I .snd ∣ inr ∣ inr ∣ inr refl ∣ ∣ ∣)) 115 | , (λ t v → capComp t .snd ∣ inl v ∣) 116 | ) 117 | } 118 | where 119 | open Realign Φ A β α r p ψ f a₀ 120 | 121 | isRealignedId : 122 | (Φ : Int → CofProp) 123 | (A : Int → Set) 124 | (β : isFib {Γ = res Int Φ} (A ∘ fst)) 125 | (α : isFib A) 126 | → --------------- 127 | reindex A (realignId Φ A β α) fst ≡ β 128 | isRealignedId Φ A β α = 129 | fibExt {A = A ∘ fst} λ r p ψ f a₀ → 130 | let 131 | open Realign Φ A β α r (fst ∘ p) ψ f a₀ 132 | in 133 | ( (λ s → symm (wA .comp s .snd ∣ inr (λ i → p i .snd) ∣)) 134 | , λ t → symm (capComp t .snd ∣ inr ∣ inl (λ i → p i .snd) ∣ ∣) 135 | ) 136 | 137 | realign : 138 | ∀{ℓ}{Γ : Set ℓ} 139 | (Φ : Γ → CofProp) 140 | (A : Γ → Set) 141 | (β : isFib {Γ = res Γ Φ} (A ∘ fst)) 142 | (α : isFib A) 143 | → --------------- 144 | isFib A 145 | realign Φ A β α r p = 146 | realignId (Φ ∘ p) 147 | (A ∘ p) 148 | (reindex (A ∘ fst) β (p ×id)) 149 | (reindex A α p) 150 | r 151 | id 152 | 153 | isRealigned : 154 | ∀{ℓ}{Γ : Set ℓ} 155 | (Φ : Γ → CofProp) 156 | (A : Γ → Set) 157 | (β : isFib {Γ = res Γ Φ} (A ∘ fst)) 158 | (α : isFib A) 159 | → --------------- 160 | reindex A (realign Φ A β α) fst ≡ β 161 | isRealigned {ℓ} {Γ} Φ A β α = 162 | funext λ r → funext λ pu → 163 | let 164 | p = fst ∘ pu 165 | in 166 | appCong {x = λ i → (i , pu i .snd)} 167 | (appCong {x = r} 168 | (isRealignedId 169 | (Φ ∘ p) 170 | (A ∘ p) 171 | (reindex (A ∘ fst) β (p ×id)) 172 | (reindex A α p))) 173 | 174 | reindexRealign : 175 | ∀{ℓ ℓ'} {Δ : Set ℓ} {Γ : Set ℓ'} 176 | (Φ : Γ → CofProp) 177 | (A : Γ → Set) 178 | (β : isFib {Γ = res Γ Φ} (A ∘ fst)) 179 | (α : isFib A) 180 | (ρ : Δ → Γ) 181 | → --------------- 182 | reindex A (realign Φ A β α) ρ 183 | ≡ realign (Φ ∘ ρ) (A ∘ ρ) (reindex (A ∘ fst) β (ρ ×id)) (reindex A α ρ) 184 | reindexRealign Φ A β α ρ = refl 185 | -------------------------------------------------------------------------------- /agda/Data/products.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Fibrancy of Σ-types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module Data.products where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | open import Data.paths 14 | 15 | ---------------------------------------------------------------------- 16 | -- Dependent products 17 | ---------------------------------------------------------------------- 18 | 19 | Σ' : ∀{a}{Γ : Set a}(A : Γ → Set)(B : (Σ x ∈ Γ , A x) → Set) → Γ → Set 20 | Σ' A B x = Σ a ∈ A x , B (x , a) 21 | 22 | FibΣid : 23 | {A : Int → Set} 24 | {B : (Σ x ∈ Int , A x) → Set} 25 | (α : isFib A) 26 | (β : isFib B) 27 | → ----------- 28 | isFib (Σ' A B) 29 | FibΣid {A} {B} α β r p φ f ((a₀ , b₀) , ex₀) = 30 | record 31 | { comp = λ s → 32 | ( (wA .comp s .fst , wB .comp s .fst) 33 | , λ u → Σext (wA .comp s .snd u) (wB .comp s .snd u) 34 | ) 35 | ; cap = 36 | ( path 37 | (λ t → (wA .cap .fst .at t , c t .fst)) 38 | (Σext (wA .cap .fst .atO) 39 | (trans 40 | (wB .cap .fst .atO) 41 | (adjustSubstEq (λ a → B (p r , a)) 42 | refl (symm (wA .cap .fst .atO)) 43 | (wA .cap .fst .atO) refl 44 | (symm (c O .snd ∣ inr ∣ inl refl ∣ ∣))))) 45 | (Σext (wA .cap .fst .atI) 46 | (adjustSubstEq (λ a → B (p r , a)) 47 | refl (symm (wA .cap .fst .atI)) 48 | (wA .cap .fst .atI) refl 49 | (symm (c I .snd ∣ inr ∣ inr refl ∣ ∣)))) 50 | , λ t u → 51 | Σext (wA .cap .snd t u) (c t .snd ∣ inl u ∣) 52 | ) 53 | } 54 | where 55 | fA : [ φ ] → Π (A ∘ p) 56 | fA u i = f u i .fst 57 | 58 | wA = α r p φ fA (a₀ , λ u → Σeq₁ (ex₀ u)) 59 | 60 | pcomp : Int → (Σ x ∈ Int , A x) 61 | pcomp i = (p i , wA .comp i .fst) 62 | 63 | pcap : Int → (Σ x ∈ Int , A x) 64 | pcap i = (p r , wA .cap .fst .at i) 65 | 66 | fB : [ φ ] → Π (B ∘ pcomp) 67 | fB u i = subst (λ a → B (p i , a)) (wA .comp i .snd u) (snd (f u i)) 68 | 69 | b₀FixTube : [ φ ] → (i : Int) → B (pcap i) 70 | b₀FixTube u i = subst (λ a → B (p r , a)) (wA .cap .snd i u) (snd (f u r)) 71 | 72 | b₀FixBase : B (pcap I) [ φ ↦ b₀FixTube ◆ I ] 73 | b₀FixBase = 74 | ( subst (λ a → B (p r , a)) (symm (wA .cap .fst .atI)) b₀ 75 | , λ u → 76 | adjustSubstEq (λ a → B (p r , a)) 77 | (Σeq₁ (ex₀ u)) refl 78 | (wA .cap .snd I u) (symm (wA .cap .fst .atI)) 79 | (Σeq₂ (ex₀ u)) 80 | ) 81 | 82 | b₀Fix : SComp I (B ∘ pcap) φ b₀FixTube b₀FixBase 83 | b₀Fix = 84 | strictifyFib B β I pcap φ 85 | b₀FixTube 86 | b₀FixBase 87 | 88 | b₀Fixed : B (pcomp r) [ φ ↦ fB ◆ r ] 89 | b₀Fixed = 90 | ( subst (λ a → B (p r , a)) (wA .cap .fst .atO) 91 | (b₀Fix .comp O (I≡IsCofProp O) .fst) 92 | , λ u → 93 | adjustSubstEq (λ a → B (p r , a)) 94 | (wA .cap .snd O u) refl 95 | (wA .comp r .snd u) (wA .cap .fst .atO) 96 | (b₀Fix .comp O (I≡IsCofProp O) .snd u) 97 | ) 98 | 99 | wB : WComp r (B ∘ pcomp) φ fB b₀Fixed 100 | wB = β r pcomp φ fB b₀Fixed 101 | 102 | -- defining the system for the final composition 103 | 104 | csysO : Int → B (pcap O) 105 | csysO j = 106 | subst (λ a → B (p r , a)) (symm (wA .cap .fst .atO)) (wB .cap .fst .at j) 107 | 108 | csysI : Int → B (pcap I) 109 | csysI j = subst (λ a → B (p r , a)) (symm (wA .cap .fst .atI)) b₀ 110 | 111 | module _ (u : [ φ ]) where 112 | 113 | csysOCompat : (j : Int) → b₀FixTube u O ≡ csysO j 114 | csysOCompat j = 115 | adjustSubstEq (λ a → B (p r , a)) 116 | (wA .comp r .snd u) refl 117 | (wA .cap .snd O u) (symm (wA .cap .fst .atO)) 118 | (wB .cap .snd j u) 119 | 120 | csysICompat : (j : Int) → b₀FixTube u I ≡ csysI j 121 | csysICompat j = 122 | adjustSubstEq (λ a → B (p r , a)) 123 | (Σeq₁ (ex₀ u)) refl 124 | (wA .cap .snd I u) (symm (wA .cap .fst .atI)) 125 | (Σeq₂ (ex₀ u)) 126 | 127 | csys : (t : Int) → [ φ ∨ t ≈O ∨ t ≈I ] → Int → B (pcap t) 128 | csys = 129 | enlargeSys φ 130 | (λ t u _ → b₀FixTube u t) 131 | (csysO , λ u → funext (csysOCompat u)) 132 | (csysI , λ u → funext (csysICompat u)) 133 | 134 | c : (t : Int) → B (pcap t) [ φ ∨ t ≈O ∨ t ≈I ↦ csys t ◆ O ] 135 | c t = 136 | β I (λ _ → pcap t) (φ ∨ t ≈O ∨ t ≈I) (csys t) 137 | ( b₀Fix .comp t (I≡IsCofProp t) .fst 138 | , enlargedExtends φ csys I (λ t → b₀Fix .comp t (I≡IsCofProp t) .fst) 139 | (λ t → b₀Fix .comp t (I≡IsCofProp t) .snd) 140 | (adjustSubstEq (λ a → B (p r , a)) 141 | refl (wA .cap .fst .atO) 142 | (symm (wA .cap .fst .atO)) refl 143 | (wB .cap .fst .atI)) 144 | (adjustSubstEq (λ a → B (p r , a)) 145 | (symm (wA .cap .fst .atI)) refl 146 | (symm (wA .cap .fst .atI)) refl 147 | (symm (b₀Fix .cap (I≡IsCofProp I)))) 148 | t 149 | ) 150 | .comp O 151 | 152 | -- trick to get Agda to check reindexing stability quickly 153 | private 154 | abstract 155 | FibΣidAbstract = FibΣid 156 | 157 | FibΣ : ∀ {ℓ} 158 | {Γ : Set ℓ} 159 | {A : Γ → Set} 160 | {B : (Σ x ∈ Γ , A x) → Set} 161 | (α : isFib A) 162 | (β : isFib B) 163 | → ----------- 164 | isFib (Σ' A B) 165 | FibΣ {Γ = Γ} {A} {B} α β r p = 166 | FibΣidAbstract (reindex A α p) (reindex B β (p ×id)) r id 167 | 168 | FibΣ' : 169 | {Γ : Set} 170 | (A : Fib Γ) 171 | (B : Fib (Σ x ∈ Γ , fst A x)) 172 | → ----------- 173 | Fib Γ 174 | FibΣ' (A , α) (B , β) = Σ' A B , FibΣ {A = A} {B = B} α β 175 | 176 | ---------------------------------------------------------------------- 177 | -- Forming Σ-types is stable under reindexing 178 | ---------------------------------------------------------------------- 179 | reindexΣ : 180 | {Δ Γ : Set} 181 | (A : Γ → Set) 182 | (B : Σ Γ A → Set) 183 | (α : isFib A) 184 | (β : isFib B) 185 | (ρ : Δ → Γ) 186 | → ---------------------- 187 | reindex (Σ' A B) (FibΣ {B = B} α β) ρ ≡ FibΣ {B = B ∘ (ρ ×id)} (reindex A α ρ) (reindex B β (ρ ×id)) 188 | reindexΣ A B α β ρ = refl 189 | -------------------------------------------------------------------------------- /agda/hcomp-coe.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of weak homogeneous composition and coercion. Proof that 4 | these imply weak composition. 5 | 6 | -} 7 | {-# OPTIONS --rewriting #-} 8 | module hcomp-coe where 9 | 10 | open import prelude 11 | open import interval 12 | open import cofprop 13 | open import fibrations 14 | 15 | isHFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 16 | isHFib {Γ = Γ} A = ∀ r x φ f a₀ → WComp r (λ _ → A x) φ f a₀ 17 | 18 | record WCoe (r : Int) (A : Int → Set) (a₀ : A r) : Set 19 | where 20 | field 21 | comp : (s : Int) → A s 22 | cap : comp r ~ a₀ 23 | 24 | open WCoe public 25 | 26 | isWCoeFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 27 | isWCoeFib A = ∀ r p a₀ → WCoe r (A ∘ p) a₀ 28 | 29 | ---------------------------------------------------------------------- 30 | -- Strict homogeneous composition 31 | ---------------------------------------------------------------------- 32 | 33 | isSHFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 34 | isSHFib {Γ = Γ} A = ∀ r x φ f a₀ → SComp r (λ _ → A x) φ f a₀ 35 | 36 | strictifyHFib : ∀{a}{Γ : Set a}(A : Γ → Set) → isHFib A → isSHFib A 37 | strictifyHFib A α r y φ f x₀ = 38 | strictifyFib {Γ = Unit} 39 | (λ _ → A y) 40 | (λ r _ φ f x₀ → α r _ φ f x₀) 41 | r _ φ f x₀ 42 | 43 | ------------------------------------------------------------------------- 44 | -- Trivial direction: We can derive homogeneous composition + coercion 45 | -- from weak composition 46 | ------------------------------------------------------------------------- 47 | 48 | isFib→isHFib : ∀ {ℓ} {Γ : Set ℓ} {A : Γ → Set} → isFib A → isHFib A 49 | isFib→isHFib α r x φ f a₀ = α r (λ _ → x) φ f a₀ 50 | 51 | isFib→isWCoeFib : ∀ {ℓ} {Γ : Set ℓ} {A : Γ → Set} → isFib A → isWCoeFib A 52 | isFib→isWCoeFib α r p a₀ = 53 | let f = α r p ⊥ ⊥→ (a₀ , λ u → ⊥→ u) 54 | in record { comp = λ s → f .comp s .fst ; cap = f .cap .fst } 55 | 56 | ------------------------------------------------------------------------------ 57 | -- We can derive a fibration structure from homogeneous composition + coercion 58 | ------------------------------------------------------------------------------ 59 | 60 | H+WCoe→Fib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isHFib A → isWCoeFib A → isFib A 61 | H+WCoe→Fib A η κ r p φ f a₀ = record 62 | { comp = λ s → 63 | ( fixedComp s .fst 64 | , λ u → trans (fixedComp s .snd u) (symm (κ s p (f u s) .cap .atI)) 65 | ) 66 | ; cap = 67 | ( path 68 | (λ t → fixedCap t .fst) 69 | (symm (fixedCap O .snd ∣ inr ∣ inl refl ∣ ∣)) 70 | (trans 71 | (trans 72 | (κ r p (a₀ .fst) .cap .atI) 73 | (d I .cap (≡IIsCofProp I))) 74 | (symm (fixedCap I .snd ∣ inr ∣ inr refl ∣ ∣))) 75 | , λ t u → fixedCap t .snd ∣ inl u ∣ 76 | ) 77 | } 78 | where 79 | strictη = strictifyHFib A η 80 | 81 | module _ (k : Int) where 82 | 83 | capTube : [ φ ] → Int → A (p k) 84 | capTube u j = κ k p (f u k) .cap .at j 85 | 86 | module _ (s : Int) where 87 | 88 | compTube : [ φ ] → Int → A (p s) 89 | compTube u i = κ i p (f u i) .comp s 90 | 91 | compBase : A (p s) [ φ ↦ compTube ◆ r ] 92 | compBase = 93 | ( κ r p (a₀ .fst) .comp s 94 | , λ u → cong (λ a → κ r p a .comp s) (a₀ .snd u) 95 | ) 96 | 97 | h = η r (p s) φ compTube compBase 98 | 99 | fixedBase : A (p s) [ φ ↦ capTube s ◆ O ] 100 | fixedBase = 101 | ( h .comp s .fst 102 | , λ u → trans (h .comp s .snd u) (κ s p (f u s) .cap .atO) 103 | ) 104 | 105 | fixedComp = strictη O (p s) φ (capTube s) fixedBase .comp I (≡IIsCofProp O) 106 | 107 | module _ (k : Int) where 108 | 109 | dBase : A (p r) [ φ ↦ capTube r ◆ k ] 110 | dBase = 111 | ( κ r p (a₀ .fst) .cap .at k 112 | , λ u → cong (λ a → κ r p a .cap .at k) (a₀ .snd u) 113 | ) 114 | 115 | d = strictη k (p r) φ (capTube r) dBase 116 | 117 | module _ (t : Int) where 118 | 119 | cBase : A (p r) [ φ ↦ capTube r ◆ O ] 120 | cBase = 121 | ( h r .cap .fst .at t 122 | , λ u → trans (h r .cap .snd t u) (κ r p (f u r) .cap .atO) 123 | ) 124 | 125 | c = strictη O (p r) φ (capTube r) cBase .comp I (≡IIsCofProp O) 126 | 127 | fixedCapTube : [ φ ∨ t ≈O ∨ t ≈I ] → Int → A (p r) 128 | fixedCapTube = 129 | ∨-rec φ (t ≈O ∨ t ≈I) 130 | (λ u _ → f u r) 131 | (OI-rec t 132 | (λ {refl k → fixedComp r .fst}) 133 | (λ {refl k → d k .comp I (≡IIsCofProp k) .fst})) 134 | (λ u → ∨-elimEq (t ≈O) (t ≈I) 135 | (λ {refl → funext λ k → 136 | trans (fixedComp r .snd u) (symm (κ r p (f u r) .cap .atI))}) 137 | (λ {refl → funext λ k → 138 | trans (d k .comp I (≡IIsCofProp k) .snd u) (symm (κ r p (f u r) .cap .atI))})) 139 | 140 | fixedCapBase : A (p r) [ φ ∨ t ≈O ∨ t ≈I ↦ fixedCapTube ◆ O ] 141 | fixedCapBase = 142 | ( c .fst 143 | , ∨-elimEq φ (t ≈O ∨ t ≈I) 144 | (λ u → trans (c .snd u) (symm (κ r p (f u r) .cap .atI))) 145 | (∨-elimEq (t ≈O) (t ≈I) 146 | (λ {refl → 147 | cong (λ a → strictη O (p r) φ (capTube r) a .comp I (≡IIsCofProp O) .fst) 148 | (Σext (symm (h r .cap .fst .atO)) (funext λ _ → uipImp))}) 149 | (λ {refl → 150 | cong (λ a → strictη O (p r) φ (capTube r) a .comp I (≡IIsCofProp O) .fst) 151 | (Σext 152 | (trans (symm (h r .cap .fst .atI)) (κ r p (a₀ .fst) .cap .atO)) 153 | (funext λ _ → uipImp))})) 154 | ) 155 | 156 | fixedCap = η O (p r) (φ ∨ t ≈O ∨ t ≈I) fixedCapTube fixedCapBase .comp I 157 | 158 | 159 | ------------------------------ 160 | -- More lemmas about isHFib -- 161 | ------------------------------ 162 | 163 | -- A constant type with homogeneous fibration structure is fibrant 164 | constHFib→isFib : ∀ {ℓ} {Γ : Set ℓ} (A : Set) → isHFib {Γ = Γ} (λ _ → A) → isFib {Γ = Γ} (λ _ → A) 165 | constHFib→isFib A h r p φ f x₀ = record 166 | { comp = λ s → h r (p r) φ f x₀ .comp s 167 | ; cap = h r (p r) φ f x₀ .cap } 168 | 169 | -- Types with decidable equality has a homogeneous fibration structure 170 | decEq→isHFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) (dec : {γ : Γ} (x y : A γ) → (x ≡ y) ⊎ (x ≡ y → ∅)) → isHFib A 171 | decEq→isHFib {Γ = Γ} A dec r x φ f a₀ = record 172 | { comp = λ s → a₀ .fst , λ u → trans (a₀ .snd u) (constant (f u) s r) 173 | ; cap = refl~ _ , λ _ → a₀ .snd } 174 | where 175 | constant : ∀ {γ : Γ} (t : Int → A γ) (r s : Int) → t r ≡ t s 176 | constant t r s with cntd (λ i → t r ≡ t i) (λ i → dec (t r) (t i)) 177 | constant t r s | inl f = f s 178 | constant t r s | inr e = ∅-rec (e r refl) 179 | -------------------------------------------------------------------------------- /agda/glueing/strict.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Strict Glue types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module glueing.strict where 8 | 9 | open import glueing.core 10 | 11 | open import prelude 12 | open import interval 13 | open import cofprop 14 | open import fibrations 15 | open import equivs 16 | open import Data.paths 17 | open import Data.products 18 | open import strictness-axioms 19 | 20 | ---------------------------------------------------------------------- 21 | -- Strict glueing 22 | ---------------------------------------------------------------------- 23 | includeA : 24 | (φ : CofProp) 25 | {A : [ φ ] → Set} 26 | {B : Set} 27 | (f : (u : [ φ ]) → A u → B) 28 | → --------------- 29 | (u : [ φ ]) → A u → Glue φ A B f 30 | includeA φ {A} {B} f u a = 31 | glue 32 | (λ v → moveA v a) 33 | (f u a) 34 | (λ v → cong (λ{(u , a) → f u a}) (symm (moveMove v))) 35 | where 36 | moveA : (v : [ φ ]) → A u → A v 37 | moveA v = subst A (cofIsProp φ _ _) 38 | moveMove : (v : [ φ ]) → (u , a) ≡ (v , moveA v a) 39 | moveMove v = Σext (cofIsProp φ _ _) refl 40 | 41 | includeAIso : 42 | (φ : CofProp) 43 | {A : [ φ ] → Set} 44 | {B : Set} 45 | (w : (u : [ φ ]) → A u → B) 46 | → --------------- 47 | (u : [ φ ]) → A u ≅ Glue φ A B w 48 | includeAIso φ {A} {B} w u = iso 49 | where 50 | prfIr : {a : A u} → subst A (cofIsProp φ u u) a ≡ a 51 | prfIr {a} = cong (λ p → subst A p a) (uip (cofIsProp φ u u) refl) 52 | 53 | iso : A u ≅ Glue φ A B w 54 | iso .to a = includeA φ w u a 55 | iso .from (glue a _ _) = a u 56 | iso .inv₁ = funext (λ a → prfIr) 57 | iso .inv₂ = funext fg≡id where 58 | parEq : {a a' : (u : [ φ ]) → A u} → a u ≡ a' u → a ≡ a' 59 | parEq {a} {a'} eq = funext (λ v → subst (λ v → a v ≡ a' v) (cofIsProp φ u v) eq) 60 | 61 | fg≡id : (gl : Glue φ A B w) → (includeA φ w u (gl .dom u)) ≡ gl 62 | fg≡id gl = 63 | glueExt w (includeA φ w u (gl .dom u)) gl (parEq prfIr) (gl .match u) 64 | 65 | SGlue : 66 | (φ : CofProp) 67 | (A : [ φ ] → Set) 68 | (B : Set) 69 | (f : (u : [ φ ]) → A u → B) 70 | → --------------- 71 | Set 72 | SGlue φ A B f = strictify φ A (Glue φ A B f) (includeAIso φ f) 73 | 74 | strictifyGlueIso : 75 | (φ : CofProp) 76 | {A : [ φ ] → Set} 77 | {B : Set} 78 | (f : (u : [ φ ]) → A u → B) 79 | → --------------- 80 | SGlue φ A B f ≅ Glue φ A B f 81 | strictifyGlueIso φ {A} {B} f = isoB φ A (Glue φ A B f) (includeAIso φ f) 82 | 83 | sglue : 84 | {φ : CofProp} 85 | {A : [ φ ] → Set} 86 | {B : Set} 87 | (f : (u : [ φ ]) → A u → B) 88 | → --------------- 89 | (u : [ φ ]) → A u → SGlue φ A B f 90 | sglue {φ} f u = strictifyGlueIso φ f .from ∘ includeA φ f u 91 | 92 | sunglue : 93 | {φ : CofProp} 94 | {A : [ φ ] → Set} 95 | {B : Set} 96 | (f : (u : [ φ ]) → A u → B) 97 | → --------------- 98 | SGlue φ A B f → B 99 | sunglue {φ} f = cod ∘ strictifyGlueIso φ f .to 100 | 101 | SGlueStrictness : 102 | (φ : CofProp) 103 | {A : [ φ ] → Set} 104 | {B : Set} 105 | (f : (u : [ φ ]) → A u → B) 106 | (u : [ φ ]) 107 | → --------------- 108 | A u ≡ SGlue φ A B f 109 | SGlueStrictness φ {A} {B} f u = 110 | restrictsToA φ A (Glue φ A B f) (includeAIso φ f) u 111 | 112 | sunglue-boundary : 113 | (φ : CofProp) 114 | {A : [ φ ] → Set} 115 | {B : Set} 116 | (f : (u : [ φ ]) → A u → B) 117 | (u : [ φ ]) (a : A u) 118 | → sunglue f (coe (SGlueStrictness φ f u) a) ≡ f u a 119 | sunglue-boundary φ {A} {B} f u a = 120 | symm 121 | (appdep 122 | (restrictsToA φ A (Glue φ A B f) (includeAIso φ f) u) 123 | (trans 124 | (congdep (λ p → cod ∘ p .snd .to) (restrictsToM φ A (Glue φ A B f) (includeAIso φ f) u)) 125 | (symm 126 | (appCong 127 | (subst-cong-assoc (λ D → D → B) fst 128 | (restrictsToM φ A (Glue φ A B f) (includeAIso φ f) u))))) 129 | refl) 130 | where 131 | appdep : ∀ {ℓ ℓ' ℓ''} {A : Set ℓ} {B : A → Set ℓ'} {C : Set ℓ''} 132 | {a a' : A} (p : a ≡ a') {f : B a → C} {f' : B a' → C} 133 | (q : subst (λ a → B a → C) p f ≡ f') 134 | {b : B a} {b' : B a'} (r : subst B p b ≡ b') 135 | → f b ≡ f' b' 136 | appdep refl refl refl = refl 137 | 138 | ---------------------------------------------------------------------- 139 | -- Indexed versions 140 | ---------------------------------------------------------------------- 141 | 142 | SGlue' : 143 | ∀{a}{Γ : Set a} 144 | (Φ : Γ → CofProp) 145 | (A : res Γ Φ → Set) 146 | (B : Γ → Set) 147 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 148 | → --------------- 149 | Γ → Set 150 | SGlue' Φ A B f x = SGlue (Φ x) (λ u → A (x , u)) (B x) (λ u → f (x , u)) 151 | 152 | strictifyGlueIso' : 153 | ∀{a}{Γ : Set a} 154 | (Φ : Γ → CofProp) 155 | {A : res Γ Φ → Set} 156 | {B : Γ → Set} 157 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 158 | → --------------- 159 | SGlue' Φ A B f ≅' Glue' Φ A B f 160 | strictifyGlueIso' Φ {A} {B} f x = strictifyGlueIso (Φ x) (λ u → f (x , u)) 161 | 162 | SGlueStrictness' : 163 | ∀{a}{Γ : Set a} 164 | (Φ : Γ → CofProp) 165 | {A : res Γ Φ → Set} 166 | {B : Γ → Set} 167 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 168 | → --------------- 169 | A ≡ SGlue' Φ A B f ∘ fst 170 | SGlueStrictness' Φ {A} {B} f = 171 | funext λ {(x , u) → SGlueStrictness (Φ x) (λ u' → f (x , u')) u} 172 | 173 | module Misaligned where 174 | 175 | abstract 176 | FibSGlueId : 177 | (Φ : Int → CofProp) 178 | {A : res Int Φ → Set} 179 | {B : Int → Set} 180 | (f : (xu : res Int Φ) → A xu → B (xu .fst)) 181 | (equiv : Equiv' f) 182 | → --------------- 183 | isFib A → isFib B → isFib (SGlue' Φ A B f) 184 | FibSGlueId Φ {A} {B} f equiv α β = 185 | FibIso (SGlue' Φ A B f) (Glue' Φ A B f) (strictifyGlueIso' Φ f) (FibGlueId Φ f equiv α β) 186 | 187 | FibSGlue : 188 | ∀{a}{Γ : Set a} 189 | (Φ : Γ → CofProp) 190 | {A : res Γ Φ → Set} 191 | {B : Γ → Set} 192 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 193 | (equiv : Equiv' f) 194 | → --------------- 195 | isFib A → isFib B → isFib (SGlue' Φ A B f) 196 | FibSGlue {a} {Γ} Φ {A} {B} f equiv α β r p = 197 | FibSGlueId (Φ ∘ p) 198 | (f ∘ (p ×id)) 199 | (equiv ∘ (p ×id)) 200 | (reindex A α (p ×id)) 201 | (reindex B β p) 202 | r 203 | id 204 | 205 | reindexFibSGlue : 206 | ∀ {ℓ} {Δ Γ : Set ℓ} 207 | (Φ : Γ → CofProp) 208 | {A : res Γ Φ → Set} 209 | {B : Γ → Set} 210 | (f : (xu : res Γ Φ) → A xu → B (xu .fst)) 211 | (equiv : Equiv' f) 212 | (α : isFib A) (β : isFib B) 213 | (ρ : Δ → Γ) 214 | → reindex (SGlue' Φ A B f) (FibSGlue Φ f equiv α β) ρ 215 | ≡ FibSGlue (Φ ∘ ρ) (f ∘ (ρ ×id)) (equiv ∘ (ρ ×id)) (reindex A α (ρ ×id)) (reindex B β ρ) 216 | reindexFibSGlue Φ f equiv α β ρ = refl 217 | -------------------------------------------------------------------------------- /agda/fiberwise-total.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module fiberwise-total where 3 | 4 | open import prelude 5 | open import interval 6 | open import cofprop 7 | open import fibrations 8 | open import wtypesred 9 | open import Data.products 10 | 11 | fiberwise+total→isFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) (B : (x : Γ) → A x → Set) 12 | (α : isFib A) 13 | (fβ : ∀ x → isFib (B x)) 14 | (tβ : isFib (Σ' A (uncurry B))) 15 | → isFib (uncurry B) 16 | fiberwise+total→isFib A B α fβ tβ r pa φ f (b₀ , ext) = record 17 | { comp = λ s → 18 | ( subst (B (p s)) 19 | (symm (aFix I .comp s .snd ∣ inr ∣ inr refl ∣ ∣)) 20 | (actual s .comp I (O≡IsCofProp I) .fst) 21 | , λ u → 22 | adjustSubstEq (B (p s)) 23 | (aFix I .comp s .snd ∣ inl u ∣) refl 24 | refl (symm (aFix I .comp s .snd ∣ inr ∣ inr refl ∣ ∣)) 25 | (actual s .comp I (O≡IsCofProp I) .snd u) 26 | ) 27 | ; cap = 28 | ( path 29 | (λ t → 30 | subst (B (p r)) 31 | (symm (aFix I .cap .snd t ∣ inr ∣ inr refl ∣ ∣)) 32 | (capPath t .fst)) 33 | (adjustSubstEq (B (p r)) 34 | refl (symm (aFix I .cap .fst .atO)) 35 | (symm (aFix I .cap .snd O ∣ inr ∣ inr refl ∣ ∣)) 36 | (symm (aFix I .comp r .snd ∣ inr ∣ inr refl ∣ ∣)) 37 | (symm (capPath O .snd ∣ inr ∣ inl refl ∣ ∣))) 38 | (trans 39 | (trans 40 | (congdep snd (total .cap .fst .atI)) 41 | (appCong (symm (subst-cong-assoc (B (p r)) fst (total .cap .fst .atI))))) 42 | (adjustSubstEq (B (p r)) 43 | refl (symm (aFix I .cap .fst .atI)) 44 | (symm (aFix I .cap .snd I ∣ inr ∣ inr refl ∣ ∣)) (cong fst (total .cap .fst .atI)) 45 | (symm (capPath I .snd ∣ inr ∣ inr refl ∣ ∣)))) 46 | , λ t u → 47 | adjustSubstEq (B (p r)) 48 | (aFix I .cap .snd t ∣ inl u ∣) refl 49 | refl (symm (aFix I .cap .snd t ∣ inr ∣ inr refl ∣ ∣)) 50 | (capPath t .snd ∣ inl u ∣) 51 | ) 52 | } 53 | where 54 | p = fst ∘ pa 55 | a = snd ∘ pa 56 | 57 | total = tβ r p φ 58 | (λ u i → (a i , f u i)) 59 | ((a r , b₀) , λ u → cong (_,_ _) (ext u)) 60 | 61 | module _ (j : Int) where 62 | 63 | aFixTube : [ φ ∨ j ≈O ∨ j ≈I ] → Π (A ∘ p) 64 | aFixTube = 65 | ∨-rec φ (j ≈O ∨ j ≈I) 66 | (λ _ i → a i) 67 | (OI-rec j 68 | (λ _ i → total .comp i .fst .fst) 69 | (λ _ i → a i)) 70 | (λ u → ∨-elimEq (j ≈O) (j ≈I) 71 | (λ {refl → funext λ i → cong fst (total .comp i .snd u)}) 72 | (λ _ → refl)) 73 | 74 | aFixBase : A (p r) [ φ ∨ j ≈O ∨ j ≈I ↦ aFixTube ◆ r ] 75 | aFixBase = 76 | ( total .cap .fst .at j .fst 77 | , ∨-elimEq φ (j ≈O ∨ j ≈I) 78 | (λ u → cong fst (total .cap .snd j u)) 79 | (∨-elimEq (j ≈O) (j ≈I) 80 | (λ {refl → cong fst (symm (total .cap .fst .atO))}) 81 | (λ {refl → cong fst (symm (total .cap .fst .atI))})) 82 | ) 83 | 84 | aFix = α r p (φ ∨ j ≈O ∨ j ≈I) aFixTube aFixBase 85 | 86 | module _ (s : Int) where 87 | 88 | actualTube : [ φ ] → (j : Int) → B (p s) (aFix j .comp s .fst) 89 | actualTube u j = 90 | subst (B (p s)) 91 | (aFix j .comp s .snd ∣ inl u ∣) 92 | (f u s) 93 | 94 | actualBase : B (p s) (aFix O .comp s .fst) [ φ ↦ actualTube ◆ O ] 95 | actualBase = 96 | ( subst (B (p s)) 97 | (aFix O .comp s .snd ∣ inr ∣ inl refl ∣ ∣) 98 | (total .comp s .fst .snd) 99 | , λ u → 100 | adjustSubstEq (B (p s)) 101 | (cong fst (total .comp s .snd u)) refl 102 | (aFix O .comp s .snd ∣ inl u ∣) (aFix O .comp s .snd ∣ inr ∣ inl refl ∣ ∣) 103 | (trans 104 | (congdep snd (total .comp s .snd u)) 105 | (appCong (symm (subst-cong-assoc (B (p s)) fst (total .comp s .snd u))))) 106 | ) 107 | 108 | actual = 109 | strictifyFib (B (p s)) (fβ (p s)) 110 | O 111 | (λ j → aFix j .comp s .fst) 112 | φ 113 | actualTube 114 | actualBase 115 | 116 | module _ (t : Int) where 117 | 118 | capTube : [ φ ∨ t ≈O ∨ t ≈I ] → (j : Int) → B (p r) (aFix j .cap .fst .at t) 119 | capTube = 120 | ∨-rec φ (t ≈O ∨ t ≈I) 121 | (λ u j → 122 | subst (B (p r)) 123 | (aFix j .cap .snd t ∣ inl u ∣) 124 | (f u r)) 125 | (OI-rec t 126 | (λ {refl j → 127 | subst (B (p r)) 128 | (symm (aFix j .cap .fst .atO)) 129 | (actual r .comp j (O≡IsCofProp j) .fst)}) 130 | (λ {refl j → 131 | subst (B (p r)) 132 | (symm (aFix j .cap .fst .atI)) 133 | (total .cap .fst .at j .snd)})) 134 | (λ u → ∨-elimEq (t ≈O) (t ≈I) 135 | (λ {refl → funext λ j → 136 | adjustSubstEq (B (p r)) 137 | (aFix j .comp r .snd ∣ inl u ∣) refl 138 | (aFix j .cap .snd O ∣ inl u ∣) (symm (aFix j .cap .fst .atO)) 139 | (actual r .comp j (O≡IsCofProp j) .snd u)}) 140 | (λ {refl → funext λ j → 141 | adjustSubstEq (B (p r)) 142 | (cong fst (total .cap .snd j u)) refl 143 | (aFix j .cap .snd I ∣ inl u ∣) (symm (aFix j .cap .fst .atI)) 144 | (trans 145 | (congdep snd (total .cap .snd j u)) 146 | (appCong (symm (subst-cong-assoc (B (p r)) fst (total .cap .snd j u)))))})) 147 | 148 | capBase : B (p r) (aFix O .cap .fst .at t) [ φ ∨ t ≈O ∨ t ≈I ↦ capTube ◆ O ] 149 | capBase = 150 | ( subst (B (p r)) 151 | (aFix O .cap .snd t ∣ inr ∣ inl refl ∣ ∣) 152 | (total .comp r .fst .snd) 153 | , ∨-elimEq φ (t ≈O ∨ t ≈I) 154 | (λ u → 155 | adjustSubstEq (B (p r)) 156 | (cong fst (total .comp r .snd u)) refl 157 | (aFix O .cap .snd t ∣ inl u ∣) (aFix O .cap .snd t ∣ inr ∣ inl refl ∣ ∣) 158 | (trans 159 | (congdep snd (total .comp r .snd u)) 160 | (appCong (symm (subst-cong-assoc (B (p r)) fst (total .comp r .snd u)))))) 161 | (∨-elimEq (t ≈O) (t ≈I) 162 | (λ {refl → 163 | adjustSubstEq (B (p r)) 164 | refl (aFix O .comp r .snd ∣ inr ∣ inl refl ∣ ∣) 165 | (symm (aFix O .cap .fst .atO)) (aFix O .cap .snd O ∣ inr ∣ inl refl ∣ ∣) 166 | (actual r .cap (O≡IsCofProp O))}) 167 | λ {refl → 168 | adjustSubstEq (B (p r)) 169 | (cong fst (total .cap .fst .atO)) refl 170 | (symm (aFix O .cap .fst .atI)) (aFix O .cap .snd I ∣ inr ∣ inl refl ∣ ∣) 171 | (trans 172 | (congdep snd (total .cap .fst .atO)) 173 | (appCong (symm (subst-cong-assoc (B (p r)) fst (total .cap .fst .atO)))))}) 174 | ) 175 | 176 | capPath = 177 | fβ (p r) 178 | O 179 | (λ j → aFix j .cap .fst .at t) 180 | (φ ∨ t ≈O ∨ t ≈I) 181 | capTube 182 | capBase 183 | .comp I 184 | -------------------------------------------------------------------------------- /agda/Data/functions.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Fibrancy of Π-types. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module Data.functions where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | open import Data.paths 14 | 15 | ---------------------------------------------------------------------- 16 | -- Dependent functions 17 | ---------------------------------------------------------------------- 18 | Π' : ∀{a}{Γ : Set a}(A : Γ → Set)(B : (Σ x ∈ Γ , A x) → Set) → Γ → Set 19 | Π' A B x = (a : A x) → B (x , a) 20 | 21 | FibΠid : 22 | {A : Int → Set} 23 | {B : (Σ x ∈ Int , A x) → Set} 24 | (α : isFib A) 25 | (β : isFib B) 26 | → ----------- 27 | isFib (Π' A B) 28 | FibΠid {A} {B} α β r p φ f (g , ex) = 29 | record 30 | { comp = λ s → 31 | ( (λ a → 32 | let open Comp s a 33 | in 34 | subst (λ a → B (p s , a)) 35 | (wA .cap .fst .atI) 36 | (fix O compUnfixed .comp I (≡IIsCofProp O) .fst) 37 | ) 38 | , λ u → funext λ a → 39 | let open Comp s a 40 | in 41 | adjustSubstEq (λ a → B (p s , a)) 42 | (symm (wA .cap .fst .atI)) refl 43 | refl (wA .cap .fst .atI) 44 | (trans 45 | (fix O compUnfixed .comp I (≡IIsCofProp O) .snd u) 46 | (congdep (f u s) (symm (wA .cap .fst .atI)))) 47 | ) 48 | ; cap = 49 | ( path 50 | (λ t a → fixCap a t .fst) 51 | (funext λ a → 52 | let open Comp r a in 53 | trans 54 | (cong 55 | (λ b → subst (λ a → B (p r , a)) (wA .cap .fst .atI) (fix O b .comp I (≡IIsCofProp O) .fst)) 56 | (Σext 57 | (cong (subst (λ a → B (p r , a)) (symm (wA .cap .fst .atO))) 58 | (wB .cap .fst .atO)) 59 | (funext λ _ → uipImp))) 60 | (symm (fixCap a O .snd ∣ inr ∣ inl refl ∣ ∣))) 61 | (funext λ a → 62 | symm (fixCap a I .snd ∣ inr ∣ inr refl ∣ ∣)) 63 | , λ t u → funext λ a → 64 | fixCap a t .snd ∣ inl u ∣ 65 | ) 66 | } 67 | where 68 | module Comp (s : Int) (a : A (p s)) where 69 | 70 | wA = α s p ⊥ O≠I (a , λ u → O≠I u) 71 | 72 | wBTube : [ φ ] → (i : Int) → B (p i , wA .comp i .fst) 73 | wBTube u i = f u i (wA .comp i .fst) 74 | 75 | wBBase : B (p r , wA .comp r .fst) [ φ ↦ wBTube ◆ r ] 76 | wBBase = 77 | ( g (wA .comp r .fst) 78 | , λ u → cong (λ h → h (wA .comp r .fst)) (ex u) 79 | ) 80 | 81 | wB = β r (λ i → (p i , wA .comp i .fst)) φ wBTube wBBase 82 | 83 | fixTube : [ φ ] → (j : Int) → B (p s , wA .cap .fst .at j) 84 | fixTube u j = f u s (wA .cap .fst .at j) 85 | 86 | fix : ∀ k y → SComp k (λ j → B (p s , wA .cap .fst .at j)) φ fixTube y 87 | fix k = strictifyFib B β k (λ j → (p s , wA .cap .fst .at j)) φ fixTube 88 | 89 | compUnfixed : B (p s , wA .cap .fst .at O) [ φ ↦ fixTube ◆ O ] 90 | compUnfixed = 91 | ( subst (λ a → B (p s , a)) (symm (wA .cap .fst .atO)) (wB .comp s .fst) 92 | , λ u → 93 | adjustSubstEq (λ a → B (p s , a)) 94 | (wA .cap .fst .atO) refl 95 | refl (symm (wA .cap .fst .atO)) 96 | (trans (wB .comp s .snd u) (congdep (f u s) (wA .cap .fst .atO))) 97 | ) 98 | 99 | module _ (a : A (p r)) where 100 | 101 | open Comp r a 102 | 103 | capUnfixedO : Int → B (p r , wA .cap .fst .at O) [ φ ↦ fixTube ◆ O ] 104 | capUnfixedO j = 105 | ( subst (λ a → B (p r , a)) (symm (wA .cap .fst .atO)) (wB .cap .fst .at j) 106 | , λ u → 107 | adjustSubstEq (λ a → B (p r , a)) 108 | (wA .cap .fst .atO) refl 109 | refl (symm (wA .cap .fst .atO)) 110 | (trans (wB .cap .snd j u) (congdep (f u r) (wA .cap .fst .atO))) 111 | ) 112 | 113 | fixCapTube : (t : Int) → [ φ ∨ t ≈O ∨ t ≈I ] → Int → B (p r , a) 114 | fixCapTube = 115 | enlargeSys φ 116 | (λ t u _ → f u r a) 117 | ( (λ j → 118 | subst (λ a → B (p r , a)) 119 | (wA .cap .fst .atI) 120 | (fix O (capUnfixedO j) .comp I (≡IIsCofProp O) .fst)) 121 | , λ u → funext λ j → 122 | adjustSubstEq (λ a → B (p r , a)) 123 | (symm (wA .cap .fst .atI)) refl 124 | refl (wA .cap .fst .atI) 125 | (trans 126 | (fix O (capUnfixedO j) .comp I (≡IIsCofProp O) .snd u) 127 | (congdep (f u r) (symm (wA .cap .fst .atI)))) 128 | ) 129 | ( (λ _ → g a) 130 | , λ u → funext λ _ → cong (λ h → h a) (ex u) 131 | ) 132 | 133 | capUnfixedBase : (t : Int) → B (p r , wA .cap .fst .at t) [ φ ↦ fixTube ◆ t ] 134 | capUnfixedBase t = 135 | ( g (wA .cap .fst .at t) 136 | , λ u → cong (λ h → h (wA .cap .fst .at t)) (ex u) 137 | ) 138 | 139 | fixCapBase : (t : Int) 140 | → B (p r , a) [ φ ∨ t ≈O ∨ t ≈I ↦ fixCapTube t ◆ I ] 141 | fixCapBase t = 142 | ( elem t 143 | , enlargedExtends φ fixCapTube I elem 144 | (λ t u → 145 | adjustSubstEq (λ a → B (p r , a)) 146 | (symm (wA .cap .fst .atI)) refl 147 | refl (wA .cap .fst .atI) 148 | (trans 149 | (fix t (capUnfixedBase t) .comp I (≡IIsCofProp t) .snd u) 150 | (congdep (f u r) (symm (wA .cap .fst .atI))))) 151 | (cong 152 | (λ b → subst (λ a → B (p r , a)) (wA .cap .fst .atI) (fix O b .comp I (≡IIsCofProp O) .fst)) 153 | (Σext 154 | (adjustSubstEq (λ a → B (p r , a)) 155 | refl (wA .cap .fst .atO) 156 | (symm (wA .cap .fst .atO)) refl 157 | (trans (symm (congdep g (wA .cap .fst .atO))) (wB .cap .fst .atI))) 158 | (funext λ _ → uipImp))) 159 | (adjustSubstEq (λ a → B (p r , a)) 160 | (symm (wA .cap .fst .atI)) refl 161 | refl (wA .cap .fst .atI) 162 | (trans 163 | (symm (fix I (capUnfixedBase I) .cap (≡IIsCofProp I))) 164 | (congdep g (symm (wA .cap .fst .atI))))) 165 | t 166 | ) 167 | where 168 | elem : (t : Int) → B (p r , a) 169 | elem t = 170 | subst (λ a → B (p r , a)) 171 | (wA .cap .fst .atI) 172 | (fix t (capUnfixedBase t) .comp I (≡IIsCofProp t) .fst) 173 | 174 | fixCap : (t : Int) → B (p r , a) [ φ ∨ t ≈O ∨ t ≈I ↦ fixCapTube t ◆ O ] 175 | fixCap t = 176 | β I (λ _ → (p r , a)) (φ ∨ t ≈O ∨ t ≈I) (fixCapTube t) (fixCapBase t) .comp O 177 | 178 | -- trick to get Agda to check reindexing stability quickly 179 | private 180 | abstract 181 | FibΠid' = FibΠid 182 | 183 | FibΠ : 184 | ∀{a}{Γ : Set a} 185 | {A : Γ → Set} 186 | {B : (Σ x ∈ Γ , A x) → Set} 187 | (α : isFib A) 188 | (β : isFib B) 189 | → ----------- 190 | isFib (Π' A B) 191 | FibΠ {_} {Γ} {A} {B} α β e p = FibΠid' (reindex A α p) (reindex B β (p ×id)) e id 192 | 193 | FibΠ' : 194 | {Γ : Set} 195 | (A : Fib Γ) 196 | (B : Fib (Σ x ∈ Γ , fst A x)) 197 | → ----------- 198 | Fib Γ 199 | FibΠ' (A , α) (B , β) = (Π' A B , FibΠ {A = A} {B = B} α β) 200 | 201 | ---------------------------------------------------------------------- 202 | -- Forming Π-types is stable under reindexing 203 | ---------------------------------------------------------------------- 204 | reindexΠ : 205 | ∀{a b}{Δ : Set a}{Γ : Set b} 206 | (A : Γ → Set) 207 | (B : Σ Γ A → Set) 208 | (α : isFib A) 209 | (β : isFib B) 210 | (ρ : Δ → Γ) 211 | → ---------------------- 212 | reindex (Π' A B) (FibΠ {B = B} α β) ρ ≡ FibΠ {B = B ∘ (ρ ×id)} (reindex A α ρ) (reindex B β (ρ ×id)) 213 | reindexΠ A B α β ρ = refl 214 | -------------------------------------------------------------------------------- /agda/univalence.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Proof of univalence in the form of ua and uaβ. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module univalence where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | open import Data.functions 14 | open import Data.paths 15 | open import equivs 16 | open import glueing 17 | 18 | record PathU {ℓ} {Γ : Set ℓ} (Aα Bβ : Fib Γ) : Set (lsuc lzero ⊔ ℓ) where 19 | constructor pathU 20 | field 21 | line : Fib (Γ × Int) 22 | atO : reindex' line (λ x → x , O) ≡ Aα 23 | atI : reindex' line (λ x → x , I) ≡ Bβ 24 | 25 | open PathU public 26 | 27 | record PathD {ℓ} {Γ : Set ℓ} (Aα Bβ : Fib Γ) (P : PathU Aα Bβ) (x : Γ) (a : Aα .fst x) (b : Bβ .fst x) 28 | : Set (lsuc lzero ⊔ ℓ) where 29 | constructor pathD 30 | field 31 | line : ∀ i → P .line .fst (x , i) 32 | atO : subst (λ F → F .fst x) (P .atO) (line O) ≡ a 33 | atI : subst (λ F → F .fst x) (P .atI) (line I) ≡ b 34 | 35 | open PathD public 36 | 37 | module _ {ℓ} {Γ : Set ℓ} {A B : Fib Γ} (P : PathU A B) (x : Γ) (a : A .fst x) where 38 | 39 | private 40 | H = P .line .fst 41 | η = P .line .snd 42 | 43 | com = 44 | strictifyFib H η O (λ i → x , i) ⊥ (∅-rec ∘ O≠I) 45 | ( subst (λ F → F .fst x) (symm (P .atO)) a 46 | , λ u → O≠I u 47 | ) 48 | 49 | abstract 50 | coerce : B .fst x 51 | coerce = subst (λ F → F .fst x) (P .atI) (com .comp I (O≡IsCofProp I) .fst) 52 | 53 | coerceFill : PathD A B P x a coerce 54 | coerceFill = 55 | pathD 56 | (λ i → com .comp i (O≡IsCofProp i) .fst) 57 | (adjustSubstEq (λ F → F .fst x) 58 | refl (symm (P .atO)) 59 | (P .atO) refl 60 | (com .cap (O≡IsCofProp O))) 61 | refl 62 | 63 | private 64 | ∈OI : ∀ {ℓ} {Γ : Set ℓ} → Γ × Int → CofProp 65 | ∈OI (_ , i) = (i ≈O) ∨ (i ≈I) 66 | 67 | extractOI : (i : Int) → [ i ≈O ∨ i ≈I ] → i ≡ O ⊎ i ≡ I 68 | extractOI i = OI-rec i inl inr 69 | 70 | module _ (f : Int → Int) (inOI : ∀ i → f i ≡ O ⊎ f i ≡ I) where 71 | 72 | private 73 | caseO : ∀ i → f i ≡ O ⊎ ¬ (f i ≡ O) 74 | caseO i with inOI i 75 | caseO i | inl fi≡O = inl fi≡O 76 | caseO i | inr fi≡I = inr λ fi≡O → O≠I (trans fi≡I (symm fi≡O)) 77 | 78 | discriminate : (∀ i → f i ≡ O) ⊎ (∀ i → f i ≡ I) 79 | discriminate with cntd _ caseO 80 | discriminate | inl f≡O = inl f≡O 81 | discriminate | inr f≠O = inr (λ i → rem i (inOI i)) 82 | where 83 | rem : ∀ i → f i ≡ O ⊎ f i ≡ I → f i ≡ I 84 | rem i (inl fi≡O) = ∅-rec (f≠O i fi≡O) 85 | rem i (inr fi≡I) = fi≡I 86 | 87 | module _ {ℓ} {Γ : Set ℓ} (Aα Bβ : Fib Γ) 88 | (f : (x : Γ) → Aα .fst x → Bβ .fst x) (fEquiv : Equiv' f) 89 | where 90 | 91 | private 92 | T : res (Γ × Int) ∈OI → Set 93 | T ((x , i) , u) = 94 | OI-rec i 95 | (λ {refl → Aα .fst x}) 96 | (λ {refl → Bβ .fst x}) 97 | u 98 | 99 | Options : (p : Int → res (Γ × Int) ∈OI) → Set ℓ 100 | Options p = 101 | ((λ i → (p i .fst .fst , O) , ∣ inl refl ∣)) ≡ p 102 | ⊎ ((λ i → (p i .fst .fst , I) , ∣ inr refl ∣)) ≡ p 103 | 104 | extendedDiscr : (p : Int → res (Γ × Int) ∈OI) → Options p 105 | extendedDiscr p with discriminate (snd ∘ fst ∘ p) (λ i → extractOI _ (p i .snd)) 106 | extendedDiscr p | inl pO = 107 | inl (funext λ i → Σext (Σext refl (symm (pO i))) (trunc _ _)) 108 | extendedDiscr p | inr pI = 109 | inr (funext λ i → Σext (Σext refl (symm (pI i))) (trunc _ _)) 110 | 111 | 112 | τHelper : (r : Int) (p : Int → res (Γ × Int) ∈OI) 113 | → Options p → ∀ φ f x₀ → WComp r (λ x → T (p x)) φ f x₀ 114 | τHelper r p (inl pO) = 115 | subst (λ p → ∀ φ f x₀ → WComp r (T ∘ p) φ f x₀) pO (Aα .snd r (fst ∘ fst ∘ p)) 116 | τHelper r p (inr pI) = 117 | subst (λ p → ∀ φ f x₀ → WComp r (T ∘ p) φ f x₀) pI (Bβ .snd r (fst ∘ fst ∘ p)) 118 | 119 | τ : isFib T 120 | τ r p = τHelper r p (extendedDiscr _) 121 | 122 | τOHelper : (r : Int) (p : Int → Γ) (op : Options (λ i → (p i , O) , ∣ inl refl ∣)) 123 | → τHelper r (λ i → (p i , O) , ∣ inl refl ∣) op ≡ Aα .snd r p 124 | τOHelper r p (inl pO) = 125 | cong 126 | (λ q → subst (λ p → ∀ φ f x₀ → WComp r (T ∘ p) φ f x₀) q (Aα .snd r p)) 127 | (uip pO refl) 128 | τOHelper r p (inr pI) = 129 | O≠I (symm (cong (λ f → f O .fst .snd) pI)) 130 | 131 | τO : reindex T τ (λ x → (x , O) , ∣ inl refl ∣) ≡ Aα .snd 132 | τO = funext λ r → funext λ p → τOHelper r p (extendedDiscr _) 133 | 134 | τIHelper : (r : Int) (p : Int → Γ) (op : Options (λ i → (p i , I) , ∣ inr refl ∣)) 135 | → τHelper r (λ i → (p i , I) , ∣ inr refl ∣) op ≡ Bβ .snd r p 136 | τIHelper r p (inl pO) = 137 | O≠I (cong (λ f → f I .fst .snd) pO) 138 | τIHelper r p (inr pI) = 139 | cong 140 | (λ q → subst (λ p → ∀ φ f x₀ → WComp r (T ∘ p) φ f x₀) q (Bβ .snd r p)) 141 | (uip pI refl) 142 | 143 | τI : reindex T τ (λ x → (x , I) , ∣ inr refl ∣) ≡ Bβ .snd 144 | τI = funext λ r → funext λ p → τIHelper r p (extendedDiscr _) 145 | 146 | g : (xiu : res (Γ × Int) ∈OI) → T xiu → Bβ .fst (xiu .fst .fst) 147 | g ((x , i) , u) = 148 | ∨-elim (i ≈O) (i ≈I) 149 | (λ u → T ((x , i) , u) → Bβ .fst x) 150 | (λ {refl → f x}) 151 | (λ {refl → id}) 152 | (λ {refl i≡I → O≠I i≡I}) 153 | u 154 | 155 | gEquiv : Equiv' g 156 | gEquiv ((x , i) , u) = 157 | ∨-elim (i ≈O) (i ≈I) 158 | (λ u → Equiv (g ((x , i) , u))) 159 | (λ {refl → fEquiv x}) 160 | (λ {refl → idEquiv (Bβ .snd) x}) 161 | (λ {refl i≡I → O≠I i≡I}) 162 | u 163 | 164 | ua : PathU Aα Bβ 165 | ua = 166 | pathU 167 | (SGlueFib ∈OI (T , τ) (reindex' Bβ fst) g gEquiv) 168 | (trans 169 | (Σext refl τO) 170 | (cong (λ Hη → reindex' Hη (λ x → ((x , O) , ∣ inl refl ∣))) 171 | (symm (SGlueFibStrictness ∈OI (T , τ) (reindex' Bβ fst) g gEquiv)))) 172 | (trans 173 | (Σext refl τI) 174 | (cong (λ Hη → reindex' Hη (λ x → ((x , I) , ∣ inr refl ∣))) 175 | (symm (SGlueFibStrictness ∈OI (T , τ) (reindex' Bβ fst) g gEquiv)))) 176 | 177 | uaβ : (x : Γ) (a : Aα .fst x) → f x a ~ coerce ua x a 178 | uaβ x a = 179 | path 180 | (λ k → sunglue (λ u → g ((x , k) , u)) (coerceFill ua x a .line k)) 181 | (trans 182 | (sunglue-boundary (∈OI (x , O)) (λ u → g ((x , O) , u)) ∣ inl refl ∣ a) 183 | (cong (sunglue (λ u → g ((x , O) , u))) lemO)) 184 | (trans 185 | (sunglue-boundary (∈OI (x , I)) (λ u → g ((x , I) , u)) ∣ inr refl ∣ (coerce ua x a)) 186 | (cong (sunglue (λ u → g ((x , I) , u))) lemI)) 187 | where 188 | lemO : 189 | coerceFill ua x a .line O 190 | ≡ coe (SGlueStrictness (∈OI (x , O)) (λ u → g ((x , O) , u)) ∣ inl refl ∣) a 191 | lemO = 192 | adjustSubstEq id 193 | (cong (λ F → F .fst x) (ua .atO)) refl 194 | refl (SGlueStrictness (∈OI (x , O)) (λ u → g ((x , O) , u)) ∣ inl refl ∣) 195 | (trans 196 | (coerceFill ua x a .atO) 197 | (symm (appCong (subst-cong-assoc id (λ F → F .fst x) (ua .atO))))) 198 | 199 | lemI : 200 | coerceFill ua x a .line I 201 | ≡ coe (SGlueStrictness (∈OI (x , I)) (λ u → g ((x , I) , u)) ∣ inr refl ∣) (coerce ua x a) 202 | lemI = 203 | adjustSubstEq id 204 | (cong (λ F → F .fst x) (ua .atI)) refl 205 | refl (SGlueStrictness (∈OI (x , I)) (λ u → g ((x , I) , u)) ∣ inr refl ∣) 206 | (trans 207 | (coerceFill ua x a .atI) 208 | (symm (appCong (subst-cong-assoc id (λ F → F .fst x) (ua .atI))))) 209 | -------------------------------------------------------------------------------- /agda/CCHM.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Proof that CCHM fibrations defined in 4 | 5 | https://arxiv.org/abs/1611.02108 6 | 7 | are equivalent to weak fibrations in the presence of a connection 8 | algebra à la Orton-Pitts: 9 | 10 | https://lmcs.episciences.org/5028/pdf 11 | https://github.com/IanOrton/cubical-topos-experiments 12 | 13 | We also prove that CCHM composition implies filling and squeeze 14 | operations. 15 | 16 | -} 17 | {-# OPTIONS --rewriting #-} 18 | module CCHM where 19 | 20 | -- Use Bool to represent the endpoints O and I 21 | open import Agda.Builtin.Bool renaming ( Bool to OI ; false to O' ; true to I' ) 22 | 23 | open import prelude 24 | open import interval 25 | open import cofprop 26 | open import fibrations 27 | 28 | ! : OI → OI 29 | ! O' = I' 30 | ! I' = O' 31 | 32 | ⟨_⟩ : OI → Int 33 | ⟨ O' ⟩ = O 34 | ⟨ I' ⟩ = I 35 | 36 | ---------------------------------------------------------------------- 37 | -- CCHM Fibrations 38 | ---------------------------------------------------------------------- 39 | 40 | Comp : OI → (Int → Set) → Set 41 | Comp e A = (φ : CofProp) (f : [ φ ] → Π A) → 42 | A ⟨ e ⟩ [ φ ↦ f ◆ ⟨ e ⟩ ] → 43 | A ⟨ ! e ⟩ [ φ ↦ f ◆ ⟨ ! e ⟩ ] 44 | 45 | isCCHMFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 46 | isCCHMFib {Γ = Γ} A = (e : OI) (p : Int → Γ) → Comp e (A ∘ p) 47 | 48 | isFib→isCCHMFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isFib A → isCCHMFib A 49 | isFib→isCCHMFib A α e p φ f x₀ = α ⟨ e ⟩ p φ f x₀ .comp ⟨ ! e ⟩ 50 | 51 | 52 | -- We now assume that we have a connection algebra structure 53 | postulate 54 | min max : Int → Int → Int 55 | 56 | minOi=O : (i : Int) → min O i ≡ O 57 | minIi=i : (i : Int) → min I i ≡ i 58 | miniO=O : (i : Int) → min i O ≡ O 59 | miniI=i : (i : Int) → min i I ≡ i 60 | 61 | {-# REWRITE minOi=O minIi=i miniO=O miniI=i #-} 62 | 63 | maxOi=i : (i : Int) → max O i ≡ i 64 | maxIi=I : (i : Int) → max I i ≡ I 65 | maxiO=i : (i : Int) → max i O ≡ i 66 | maxiI=I : (i : Int) → max i I ≡ I 67 | 68 | {-# REWRITE maxOi=i maxIi=I maxiO=i maxiI=I #-} 69 | 70 | ---------------------------------------------------------------------- 71 | -- Deriving weak composition for a CCHM fibration 72 | ---------------------------------------------------------------------- 73 | 74 | isCCHMFib→isFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isCCHMFib A → isFib A 75 | isCCHMFib→isFib {Γ = Γ} A α r p φ f x₀ = record 76 | { comp = λ s → (fil s .fst , λ u → fil s .snd ∣ inl u ∣) 77 | ; cap = 78 | ( path 79 | (λ t → fil' t .fst) 80 | (cong 81 | (λ {(f , ex) → α O' (λ i → p (min i r)) (φ ∨ ⊥) f (sqz O .fst , ex) .fst}) 82 | (Σext 83 | (funext 84 | (∨-elimEq φ ⊥ (λ _ → refl) (λ v → ⊥→ v))) 85 | (funext λ _ → uipImp))) 86 | (symm (fil' I .snd ∣ inr refl ∣)) 87 | , λ t u → fil' t .snd ∣ inl u ∣ 88 | ) 89 | } 90 | where 91 | module _ (t : Int) where 92 | 93 | sqzTube : [ φ ∨ t ≈I ] → (i : Int) → A (p (min (max t i) r)) 94 | sqzTube = 95 | ∨-rec φ (t ≈I) 96 | (λ u i → f u (min (max t i) r)) 97 | (λ {refl i → x₀ .fst}) 98 | (λ {u refl → funext λ i → x₀ .snd u}) 99 | 100 | x₀' : A (p r) [ φ ∨ t ≈I ↦ sqzTube ◆ I ] 101 | x₀' = x₀ .fst , ∨-elimEq φ (t ≈I) (x₀ .snd) (λ {refl → refl}) 102 | 103 | sqz = α I' (λ i → p (min (max t i) r)) (φ ∨ t ≈I) sqzTube x₀' 104 | 105 | module _ (s : Int) where 106 | 107 | filTube : [ φ ∨ ⊥ ] → (i : Int) → A (p (min i s)) 108 | filTube = 109 | ∨-rec φ ⊥ 110 | (λ u i → f u (min i s)) 111 | O≠I 112 | (λ _ O≡I → O≠I O≡I) 113 | 114 | filBase : A (p O) [ φ ∨ ⊥ ↦ filTube ◆ O ] 115 | filBase = 116 | ( sqz O .fst 117 | , ∨-elimEq φ ⊥ 118 | (λ u → sqz O .snd ∣ inl u ∣) 119 | (λ v → O≠I v) 120 | ) 121 | 122 | fil = α O' (λ i → p (min i s)) (φ ∨ ⊥) filTube filBase 123 | 124 | module _ (t : Int) where 125 | 126 | fil' = α O' (λ i → p (min (max t i) r)) (φ ∨ t ≈I) (sqzTube t) (sqz t) 127 | 128 | ---------------------------------------------------------------------- 129 | -- Deriving fill and squeeze operations from CCHM composition 130 | ---------------------------------------------------------------------- 131 | 132 | !! : ∀ e → ! (! e) ≡ e 133 | !! true = refl 134 | !! false = refl 135 | 136 | _≈OI_ : Int → OI → CofProp 137 | r ≈OI O' = r ≈O 138 | r ≈OI I' = r ≈I 139 | 140 | [≈OI] : ∀ i e → [ i ≈OI e ] ≡ (i ≡ ⟨ e ⟩) 141 | [≈OI] i false = refl 142 | [≈OI] i true = refl 143 | 144 | {-# REWRITE !! [≈OI] #-} 145 | 146 | cnx : OI → Int → Int → Int 147 | cnx O' = min 148 | cnx I' = max 149 | 150 | cnxeei=e : (e : OI)(i : Int) → cnx e ⟨ e ⟩ i ≡ ⟨ e ⟩ 151 | cnxeei=e O' i = refl 152 | cnxeei=e I' i = refl 153 | 154 | cnxeie=e : (e : OI)(i : Int) → cnx e i ⟨ e ⟩ ≡ ⟨ e ⟩ 155 | cnxeie=e O' i = refl 156 | cnxeie=e I' i = refl 157 | 158 | cnx!eei=i : (e : OI)(i : Int) → cnx (! e) ⟨ e ⟩ i ≡ i 159 | cnx!eei=i O' i = refl 160 | cnx!eei=i I' i = refl 161 | 162 | cnx!eie=i : (e : OI)(i : Int) → cnx (! e) i ⟨ e ⟩ ≡ i 163 | cnx!eie=i O' i = refl 164 | cnx!eie=i I' i = refl 165 | 166 | cnxe!ei=i : (e : OI)(i : Int) → cnx e ⟨ ! e ⟩ i ≡ i 167 | cnxe!ei=i O' i = refl 168 | cnxe!ei=i I' i = refl 169 | 170 | cnxei!e=i : (e : OI)(i : Int) → cnx e i ⟨ ! e ⟩ ≡ i 171 | cnxei!e=i O' i = refl 172 | cnxei!e=i I' i = refl 173 | 174 | {-# REWRITE cnxeei=e cnxeie=e cnx!eei=i cnx!eie=i cnxe!ei=i cnxei!e=i #-} 175 | 176 | record Fill (e : OI) (A : Int → Set) (φ : CofProp) (f : [ φ ] → Π A) 177 | (x₀ : A ⟨ e ⟩ [ φ ↦ f ◆ ⟨ e ⟩ ]) : Set where 178 | field 179 | fill : (r : Int) → A r [ φ ↦ f ◆ r ] 180 | cap : fill ⟨ e ⟩ .fst ≡ x₀ .fst 181 | 182 | open Fill public 183 | 184 | hasFill : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 185 | hasFill {Γ = Γ} A = (e : OI) (p : Int → Γ) (φ : CofProp) (f : [ φ ] → Π (A ∘ p)) 186 | (x₀ : A (p ⟨ e ⟩) [ φ ↦ f ◆ ⟨ e ⟩ ]) → Fill e (A ∘ p) φ f x₀ 187 | 188 | isCCHMFib→Fill : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isCCHMFib A → hasFill A 189 | isCCHMFib→Fill A α e p φ f x₀ = record 190 | { fill = λ r → (a r .fst , λ u → a r .snd ∣ inl u ∣) 191 | ; cap = symm (a ⟨ e ⟩ .snd ∣ inr refl ∣) 192 | } 193 | where 194 | module _ (r : Int) where 195 | φ∨r≡e : CofProp 196 | φ∨r≡e = φ ∨ (r ≈OI e) 197 | f' : [ φ∨r≡e ] → (i : Int) → A (p (cnx e i r)) 198 | f' = ∨-rec φ (r ≈OI e) 199 | (λ x i → f x (cnx e i r)) 200 | (λ {refl → λ _ → x₀ .fst}) 201 | (λ u → λ {refl → funext λ x → x₀ .snd u}) 202 | a = α e (λ i → p (cnx e i r)) φ∨r≡e f' 203 | (x₀ .fst , ∨-elimEq φ (r ≈OI e) (snd x₀) (λ {refl → refl})) 204 | 205 | record Squeeze (r : Int) (A : Int → Set) (φ : CofProp) (f : [ φ ] → Π A) 206 | (x₀ : A r [ φ ↦ f ◆ r ]) : Set where 207 | field 208 | squeeze : (e : OI) → A ⟨ e ⟩ [ φ ↦ f ◆ ⟨ e ⟩ ] 209 | cap : (e : OI) (p : ⟨ e ⟩ ≡ r) → subst A p (squeeze e .fst) ≡ x₀ .fst 210 | 211 | open Squeeze public 212 | 213 | hasSqueeze : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 214 | hasSqueeze {Γ = Γ} A = (r : Int) (p : Int → Γ) (φ : CofProp) (f : [ φ ] → Π (A ∘ p)) 215 | (x₀ : A (p r) [ φ ↦ f ◆ r ]) → Squeeze r (A ∘ p) φ f x₀ 216 | 217 | isCCHMFib→Squeeze : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isCCHMFib A → hasSqueeze A 218 | isCCHMFib→Squeeze A α r p φ f x₀ = record 219 | { squeeze = λ e → (a e .fst , λ u → a e .snd ∣ inl u ∣) 220 | ; cap = λ {e refl → symm (a e .snd ∣ inr refl ∣)} 221 | } 222 | where 223 | module _ (e : OI) where 224 | f' : [ φ ∨ (r ≈OI e) ] → (i : Int) → A (p (cnx e i r)) 225 | f' = ∨-rec φ (r ≈OI e) 226 | (λ x i → f x (cnx e i r)) 227 | (λ {refl → λ _ → x₀ .fst}) 228 | (λ u → λ {refl → funext λ x → x₀ .snd u}) 229 | a = α (! e) (λ i → p (cnx e i r)) (φ ∨ (r ≈OI e)) f' 230 | (x₀ .fst , ∨-elimEq φ (r ≈OI e) (snd x₀) (λ {refl → refl})) 231 | -------------------------------------------------------------------------------- /agda/Data/id2.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Identity types constructed using cofibrant replacement (factorisation 4 | of a map as a cofibration followed by trivial fibration). 5 | 6 | In this version the cofibrant replacement is constructed by assuming W 7 | types with reductions rather than ax7 and extensionality for cofibrant 8 | propositions. 9 | 10 | -} 11 | 12 | 13 | {-# OPTIONS --rewriting #-} 14 | module Data.id2 (Γ : Set) (A : Γ → Set) where 15 | 16 | open import prelude 17 | open import cofreplacement 18 | open import Data.products 19 | open import interval 20 | open import cofprop 21 | open import fibrations 22 | open import trivialfib 23 | open import Data.paths 24 | 25 | Path' : Γ → Set 26 | Path' γ = Int → A γ 27 | 28 | r-path : {γ : Γ} → A γ → Path' γ 29 | r-path a = λ _ → a 30 | 31 | eO-path : {γ : Γ} → Path' γ → A γ 32 | eO-path p = p O 33 | 34 | {- Endpoint maps are both trivial fibrations (we only show eO) -} 35 | eO-path-tf : {γ : Γ} → (isFib A) → map-isTrivialFib (eO-path {γ}) 36 | eO-path-tf {γ} afib a φ u = 37 | ((λ i → fst (comp sc i (O≡IsCofProp i))) , 38 | cap sc (O≡IsCofProp O)) , 39 | λ x → Σext (funext (λ i → snd (comp sc i (O≡IsCofProp i)) x)) (uip _ _) 40 | where 41 | sc = strictifyFib A afib O (λ _ → γ) φ (λ x → fst (u x)) (a , (λ x → snd (u x))) 42 | 43 | {- Id2' has the wrong type to work as identity type, but is more convenient to work with. -} 44 | Id2' : Γ → Set 45 | Id2' γ = cof-replace.M (r-path {γ}) 46 | 47 | {- A bunch more utility functions. -} 48 | {- This one is a variant of the reflexivity map. -} 49 | r' : {γ : Γ} → A γ → Id2' γ 50 | r' = cof-replace.L r-path 51 | 52 | {- Endpoint projection. -} 53 | eO' : {γ : Γ} → Id2' γ → A γ 54 | eO' = eO-path ∘ cof-replace.R r-path 55 | 56 | eO-tf : {γ : Γ} (afib : isFib A) → map-isTrivialFib (eO' {γ}) 57 | eO-tf afib = TrivialFib-∘ eO-path (cof-replace.R r-path) (eO-path-tf afib) (cof-replace.R-tf r-path) 58 | 59 | r-cof : {γ : Γ} → cofibration (r' {γ}) 60 | r-cof = cof-replace.L-cof r-path 61 | 62 | Path-fmap : {B C : Set} {b b' : B} (f : B → C) → (b ~ b') → (f b ~ f b') 63 | Path-fmap f (path p bO bI) = path (f ∘ p) (cong f bO) (cong f bI) 64 | 65 | private 66 | extract-pth : (x : Σ (Σ γ ∈ Γ , A γ × A γ) (Path A)) → Path' (fst (fst x)) 67 | extract-pth ((γ , (a , a')) , pth) = at pth 68 | 69 | {- Actual identity type (with correct type) is Id2 here. -} 70 | Id2 : Σ γ ∈ Γ , (A γ × A γ) → Set 71 | Id2 = Σ' (Path A) (λ x → cof-replace.M₀ r-path (extract-pth x)) 72 | 73 | {- Reflexivity map -} 74 | r : (γ : Γ) → (a : A γ) → Id2 (γ , (a , a)) 75 | r γ a = (refl~ a) , cof-replace.L₀ r-path a 76 | 77 | {- The identity type is a fibration over Γ, A γ, A γ. -} 78 | Id2-isFib : (afib : isFib A) → isFib Id2 79 | Id2-isFib afib = FibΣ {Γ = Σ γ ∈ Γ , (A γ × A γ)} {A = Path A} 80 | {B = λ x → cof-replace.M₀ r-path (extract-pth x)} (FibPath afib) 81 | (TrivialFib-isFib (λ x → cof-replace.M₀ r-path (extract-pth x)) 82 | (λ x → cof-replace.M₀-isTrivialFib r-path (extract-pth x))) 83 | 84 | 85 | module _ (afib : isFib A) (γ : Γ) where 86 | open cofibration (r-cof {γ}) 87 | 88 | id-path' : (z : Id2' γ) → (r' (eO' z) , refl) ~ (z , refl) 89 | id-path' = 90 | j (λ z → (r' (eO' z) , refl) ~ (z , refl)) 91 | (λ z' → path-contr (eO-tf afib (eO' z')) _ _) 92 | λ a → refl~ _ 93 | 94 | 95 | id-path : (z : Id2' γ) → r' (eO' z) ~ z 96 | id-path z = Path-fmap fst (id-path' z) 97 | 98 | id-refl' : (a : A γ) → (id-path' (r' a) ≡ refl~ _) 99 | id-refl' a = upper-tri (((λ z → (r' (eO' z) , refl) ~ (z , refl)))) 100 | (λ z' → path-contr (eO-tf afib (eO' z')) _ _) 101 | (λ a₁ → refl~ _) a 102 | 103 | id-refl : (a : A γ) → (id-path (r' a) ≡ refl~ _) 104 | id-refl a = cong (Path-fmap fst) (id-refl' a) 105 | 106 | 107 | {- Leibniz exponential from endpoint inclusion -} 108 | δ₀⇒ : {ℓ' : Level} {Γ' : Set ℓ'} (C : Γ' → Set) → 109 | Σ (Int → Γ') (λ p → C (p O)) → Set 110 | δ₀⇒ C (p , c₀) = Σ ((i : Int) → C (p i)) λ p' → p' O ≡ c₀ 111 | 112 | {- Leibniz exponential with an endpoint sends fibrations to trivial fibrations. -} 113 | path-lift-tf : {ℓ' : Level} {Γ' : Set ℓ'} (C : Γ' → Set) (cfib : isFib C) → isTrivialFib (δ₀⇒ C) 114 | path-lift-tf C cfib (p , c₀) φ u = 115 | ((λ i → fst (comp sc i (O≡IsCofProp i))) 116 | , cap sc (O≡IsCofProp O)) 117 | , (λ x → Σext (funext (λ i → snd (comp sc i (O≡IsCofProp i)) x)) (uip _ _)) 118 | where 119 | u' : [ φ ] → (i : Int) → C (p i) 120 | u' = fst ∘ u 121 | 122 | sc = strictifyFib C cfib O p φ u' (c₀ , (λ x → snd (u x))) 123 | 124 | 125 | {- If a map is both a cofibration and strong deformation retract, then it is a 126 | trivial cofibration. 127 | -} 128 | sdr-cof-to-tc : {A B : Set} (r : A → B) (s : B → A) (s-r : (a : A) → s (r a) ≡ a) 129 | (r-s : (b : B) → r (s b) ~ b) → 130 | ((a : A) → r-s (r a) ≡ path (λ _ → r a) (cong r (symm (s-r a))) refl) → 131 | (cofibration r) → (triv-cofibration r) 132 | sdr-cof-to-tc {A} {B} r s s-r r-s str rcof = 133 | record { j = j 134 | ; upper-tri = ut } 135 | where 136 | module _ (X : B → Set) (xfib : isFib X) (x₀ : (a : A) → X (r a)) where 137 | str' : (a : A) → (i : Int) → (r-s (r a) .at i ≡ r a) 138 | str' a = subst (λ pth → (i : Int) → pth .at i ≡ r a) (symm (str a)) λ _ → refl 139 | 140 | X' : B → Set 141 | X' b = δ₀⇒ X ((at (r-s b)) , (subst X (symm (atO (r-s b))) (x₀ (s b)))) 142 | 143 | X'-tf : isTrivialFib X' 144 | X'-tf = λ b φ u → path-lift-tf X xfib (((at (r-s b)) , (subst X (symm (atO (r-s b))) (x₀ (s b))))) φ u 145 | 146 | {- NB: Both lemmas use axiom K. -} 147 | lemma1 : {a a' : A} (q : a' ≡ a) {b : B} (l : b ≡ r a) (p : b ≡ r a') → 148 | (subst X (symm l) (x₀ a) ≡ (subst X (symm p) (x₀ a'))) 149 | lemma1 refl refl refl = refl 150 | 151 | lemma2 : {b b' : B} (p q : b ≡ b') → (x : X b') → subst X p (subst X (symm q) x) ≡ x 152 | lemma2 refl refl x = refl 153 | 154 | x₀' : (a : A) → X' (r a) 155 | x₀' a = (λ i → subst X (symm (str' a i)) (x₀ a)) 156 | , lemma1 (s-r a) (str' a O) (atO (r-s (r a))) 157 | 158 | j = λ b → subst X (atI (r-s b)) (fst (cofibration.j rcof X' X'-tf x₀' b) I) 159 | 160 | ut = λ a → trans (lemma2 (atI (r-s (r a))) (str' a I) (x₀ a)) 161 | (cong (λ x → subst X (atI (r-s (r a))) (fst x I)) (cofibration.upper-tri rcof X' X'-tf x₀' a)) 162 | 163 | 164 | 165 | 166 | module _ {γ : Γ} (afib : isFib A) (C : Id2' γ → Set) (cfib : isFib C) 167 | (c₀ : (a : A γ) → C (r' a)) where 168 | 169 | open triv-cofibration (sdr-cof-to-tc r' eO' (λ a → refl) 170 | (λ b → id-path afib γ b) 171 | (id-refl afib γ) (cof-replace.L-cof r-path)) 172 | 173 | J' : (z : Id2' γ) → C z 174 | J' = j C cfib c₀ 175 | 176 | 177 | J-comp' : (a : A γ) → J' (r' a) ≡ c₀ a 178 | J-comp' = upper-tri C cfib c₀ 179 | 180 | 181 | {- J rule formatted for the actual identity type - Id2 -} 182 | module _ (afib : isFib A) (C : Σ (Σ γ ∈ Γ , A γ × A γ) Id2 → Set) 183 | (cfib : isFib C) (c₀ : (γ : Γ) → (a : A γ) → C ((γ , (a , a)) , (r γ a))) where 184 | 185 | C' : (γ : Γ) → Id2' γ → Set 186 | C' γ (p , z) = C ((γ , ((p O) , (p I))) , (path p refl refl) , z) 187 | 188 | c₀' : (γ : Γ) → (a : A γ) → C' γ (r' a) 189 | c₀' γ a = c₀ γ a 190 | 191 | cfib' : (γ : Γ) → isFib (C' γ) 192 | cfib' γ = reindex C cfib (λ {(p , z) → (γ , (p O) , (p I)) , (path p refl refl) , z}) 193 | 194 | J : (z : Σ (Σ γ ∈ Γ , A γ × A γ) Id2) → C z 195 | J ((γ , a , a') , (pth , z)) = 196 | subst C (symm (lemma (at pth) (atO pth) (atI pth) z)) 197 | (J' afib (C' γ) (cfib' γ) (c₀' γ) ((at pth) , z)) 198 | where 199 | lemma : {b b' : A γ} (α : Int → A γ) (p : α O ≡ b) (q : α I ≡ b') → 200 | (z' : _) → 201 | ((γ , b , b') , (path α p q , z')) ≡ ((γ , α O , α I) , (path α refl refl , z')) 202 | lemma α refl refl z' = refl 203 | -------------------------------------------------------------------------------- /agda/fibrations.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of weak composition and fibrations. Proof that they are 4 | closed under isomorphism and that weak composition can be strictified. 5 | 6 | -} 7 | {-# OPTIONS --rewriting #-} 8 | module fibrations where 9 | 10 | open import prelude 11 | open import interval 12 | open import cofprop 13 | 14 | ---------------------------------------------------------------------- 15 | -- Weak composition structure 16 | ---------------------------------------------------------------------- 17 | 18 | record WComp (r : Int) (A : Int → Set) (φ : CofProp) (f : [ φ ] → Π A) 19 | (x₀ : A r [ φ ↦ f ◆ r ]) : Set 20 | where 21 | field 22 | comp : (s : Int) → A s [ φ ↦ f ◆ s ] 23 | cap : Σ p ∈ comp r .fst ~ x₀ .fst , ((k : Int) → f ◆ r ↗ p .at k) 24 | 25 | open WComp public 26 | 27 | ---------------------------------------------------------------------- 28 | -- Fibrations 29 | ---------------------------------------------------------------------- 30 | isFib : ∀{a}{Γ : Set a}(A : Γ → Set) → Set a 31 | isFib {Γ = Γ} A = ∀ r p φ f x₀ → WComp r (A ∘ p) φ f x₀ 32 | 33 | Fib : ∀{a}(Γ : Set a) → Set (lsuc lzero ⊔ a) 34 | Fib {a} Γ = Σ (Γ → Set) isFib 35 | 36 | ---------------------------------------------------------------------- 37 | -- Fibrations can be reindexed 38 | ---------------------------------------------------------------------- 39 | reindex : ∀{a a'}{Δ : Set a}{Γ : Set a'}(A : Γ → Set)(α : isFib A)(ρ : Δ → Γ) → isFib (A ∘ ρ) 40 | reindex A α ρ r p = α r (ρ ∘ p) 41 | 42 | reindex' : ∀{a a'}{Δ : Set a}{Γ : Set a'}(Aα : Fib Γ)(ρ : Δ → Γ) → Fib Δ 43 | reindex' (A , α) ρ = (A ∘ ρ , reindex A α ρ) 44 | 45 | ---------------------------------------------------------------------- 46 | -- Reindexing is functorial 47 | ---------------------------------------------------------------------- 48 | reindexAlongId : ∀{a}{Γ : Set a}{A : Γ → Set}{α : isFib A} → α ≡ reindex A α id 49 | reindexAlongId = refl 50 | 51 | reindexWComp : 52 | ∀{a b c}{Γ₁ : Set a}{Γ₂ : Set b}{Γ₃ : Set c}{A : Γ₃ → Set}{α : isFib A} 53 | (f : Γ₁ → Γ₂)(g : Γ₂ → Γ₃) 54 | → ---------------------- 55 | reindex A α (g ∘ f) ≡ reindex (A ∘ g) (reindex A α g) f 56 | reindexWComp g f = refl 57 | 58 | reindexAlongId' : ∀{a}{Γ : Set a}{A : Fib Γ} → reindex' A id ≡ A 59 | reindexAlongId' = refl 60 | 61 | reindexWComp' : 62 | ∀{a b c}{Γ₁ : Set a}{Γ₂ : Set b}{Γ₃ : Set c} 63 | {A : Fib Γ₃} 64 | (f : Γ₁ → Γ₂)(g : Γ₂ → Γ₃) 65 | → ---------------------- 66 | reindex' A (g ∘ f) ≡ reindex' (reindex' A g) f 67 | reindexWComp' g f = refl 68 | 69 | ---------------------------------------------------------------------- 70 | -- An extensionality principle for fibration structures 71 | ---------------------------------------------------------------------- 72 | fibExt : ∀{ℓ}{Γ : Set ℓ}{A : Γ → Set}{α α' : isFib A} → 73 | ((r : Int) (p : Int → Γ) (φ : CofProp) (f : [ φ ] → Π (A ∘ p)) (a₀ : A (p r) [ φ ↦ f ◆ r ]) 74 | → ((s : Int) → α r p φ f a₀ .comp s .fst ≡ α' r p φ f a₀ .comp s .fst) 75 | × ((t : Int) → α r p φ f a₀ .cap .fst .at t ≡ α' r p φ f a₀ .cap .fst .at t)) 76 | → α ≡ α' 77 | fibExt {Γ = Γ} {A} {α} {α'} ext = 78 | funext λ r → funext λ p → funext λ φ → funext λ f → funext λ a₀ → full r p φ f a₀ 79 | where 80 | module _ (r : Int) (p : Int → Γ) (φ : CofProp) (f : [ φ ] → Π (A ∘ p)) 81 | (a₀ : A (p r) [ φ ↦ f ◆ r ]) 82 | where 83 | pairEqs : {w w' : WComp r (A ∘ p) φ f a₀} 84 | (q : w .comp ≡ w' .comp) 85 | → subst (λ c → Σ p ∈ c r .fst ~ a₀ .fst , ((k : Int) → f ◆ r ↗ p .at k)) q (w .cap) ≡ w' .cap 86 | → w ≡ w' 87 | pairEqs refl refl = refl 88 | 89 | compEq : α r p φ f a₀ .comp ≡ α' r p φ f a₀ .comp 90 | compEq = funext λ s → Σext (ext r p φ f a₀ .fst s) (funext λ _ → uipImp) 91 | 92 | full : α r p φ f a₀ ≡ α' r p φ f a₀ 93 | full = 94 | pairEqs compEq 95 | (Σextdep compEq 96 | (PathExtDep compEq 97 | (funext λ t → ext r p φ f a₀ .snd t)) 98 | (funext λ _ → funext λ _ → uipImp)) 99 | 100 | ---------------------------------------------------------------------- 101 | -- Terminal object is fibrant 102 | ---------------------------------------------------------------------- 103 | FibUnit : {Γ : Set} → isFib(λ(_ : Γ) → Unit) 104 | FibUnit r p φ f (unit , _) .comp s = (unit , λ _ → refl) 105 | FibUnit r p φ f (unit , _) .cap = (refl~ unit , λ _ _ → refl) 106 | 107 | ---------------------------------------------------------------------- 108 | -- Initial object is fibrant 109 | ---------------------------------------------------------------------- 110 | Fib∅ : {Γ : Set} → isFib(λ(_ : Γ) → ∅) 111 | Fib∅ r p φ f (() , _) 112 | 113 | ---------------------------------------------------------------------- 114 | -- Fibrations are closed under isomorphism 115 | ---------------------------------------------------------------------- 116 | record _≅_ (A B : Set) : Set where 117 | field 118 | to : A → B 119 | from : B → A 120 | inv₁ : from ∘ to ≡ id 121 | inv₂ : to ∘ from ≡ id 122 | 123 | open _≅_ public 124 | 125 | _≅'_ : ∀{a}{Γ : Set a}(A B : Γ → Set) → Set a 126 | _≅'_ {_} {Γ} A B = (x : Γ) → A x ≅ B x 127 | 128 | FibIso : ∀{a}{Γ : Set a}(A B : Γ → Set) → (A ≅' B) → isFib B → isFib A 129 | FibIso A B iso β r p φ f (a₀ , ex) = 130 | record 131 | { comp = λ s → 132 | ( iso (p s) .from (inB .comp s .fst) 133 | , λ u → 134 | trans 135 | (cong (iso (p s) .from) (inB .comp s .snd u)) 136 | (symm (appCong (iso (p s) .inv₁))) 137 | ) 138 | ; cap = 139 | ( path 140 | (λ t → iso (p r) .from (inB .cap .fst .at t)) 141 | (cong (iso (p r) .from) (inB .cap .fst .atO)) 142 | (trans 143 | (appCong (iso (p r) .inv₁)) 144 | (cong (iso (p r) .from) (inB .cap .fst .atI))) 145 | , λ t u → 146 | trans 147 | (cong (iso (p r) .from) (inB .cap .snd t u)) 148 | (symm (appCong (iso (p r) .inv₁))) 149 | ) 150 | } 151 | where 152 | tube : [ φ ] → Π (B ∘ p) 153 | tube u i = iso (p i) .to (f u i) 154 | 155 | base : B (p r) [ φ ↦ tube ◆ r ] 156 | base = (iso (p r) .to a₀ , λ u → cong (iso (p r) .to) (ex u)) 157 | 158 | inB = β r p φ tube base 159 | 160 | ---------------------------------------------------------------------- 161 | -- Strict composition 162 | ---------------------------------------------------------------------- 163 | record SComp (r : Int) (A : Int → Set) (φ : CofProp) (f : [ φ ] → Π A) 164 | (x₀ : A r [ φ ↦ f ◆ r ]) : Set 165 | where 166 | field 167 | comp : (s : Int) → isCofProp (r ≡ s) → A s [ φ ↦ f ◆ s ] 168 | cap : (c : isCofProp (r ≡ r)) → comp r c .fst ≡ x₀ .fst 169 | 170 | open SComp public 171 | 172 | isSFib : ∀{ℓ} {Γ : Set ℓ} (A : Γ → Set) → Set ℓ 173 | isSFib {Γ = Γ} A = ∀ r p φ f x₀ → SComp r (A ∘ p) φ f x₀ 174 | 175 | strictifyFib : ∀ {ℓ} {Γ : Set ℓ} (A : Γ → Set) → isFib A → isSFib A 176 | strictifyFib A α r p φ f x₀ = 177 | record 178 | { comp = λ s c → 179 | (strong s c .fst , λ v → strong s c .snd ∣ inl v ∣) 180 | ; cap = λ c → 181 | trans 182 | (atI (α r p φ f x₀ .cap .fst)) 183 | (trans 184 | (cong (λ q → fix r c q I) (uip (c .snd .fst (c .snd .snd refl)) refl)) 185 | (symm (strong r c .snd ∣ inr (c .snd .snd refl) ∣))) 186 | } 187 | where 188 | module _ (s : Int) (c : isCofProp (r ≡ s)) 189 | where 190 | 191 | r≡s = c .fst 192 | 193 | weak : A (p s) [ φ ↦ f ◆ s ] 194 | weak = α r p φ f x₀ .comp s 195 | 196 | f' : [ φ ] → (k : Int) → A (p s) 197 | f' v k = f v s 198 | 199 | fix : r ≡ s → Int → A (p s) 200 | fix eq k = subst (A ∘ p) eq (α r p φ f x₀ .cap .fst .at k) 201 | 202 | fixed-f' : [ φ ∨ r≡s ] → Int → A (p s) 203 | fixed-f' = 204 | ∨-rec φ r≡s f' (fix ∘ c .snd .fst) 205 | (λ u eq → funext (λ k → 206 | trans (cong (subst (A ∘ p) (c .snd .fst eq)) (α r p φ f x₀ .cap .snd k u)) 207 | (symm (substfam (c .snd .fst eq) (f u))))) 208 | 209 | weak' : Σ x₁ ∈ (A (p s)) , (fixed-f' ◆ O ↗ x₁) 210 | weak' = 211 | ( weak .fst 212 | , ∨-elimEq φ r≡s 213 | (weak .snd) 214 | (λ eq → 215 | trans 216 | (substfam (c .snd .fst eq) (λ k → α r p φ f x₀ .comp k .fst)) 217 | (cong (subst (A ∘ p) (c .snd .fst eq)) (α r p φ f x₀ .cap .fst .atO)) 218 | ) 219 | ) 220 | 221 | strong : Σ x₂ ∈ (A (p s)) , (fixed-f' ◆ I ↗ x₂) 222 | strong = α O (λ _ → p s) (φ ∨ r≡s) fixed-f' weak' .comp I 223 | -------------------------------------------------------------------------------- /agda/fibreplacement.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module fibreplacement where 3 | 4 | open import prelude 5 | open import interval 6 | open import cofprop 7 | open import fibrations 8 | open import wtypesred 9 | open import cofreplacement 10 | 11 | open triv-cofibration public 12 | 13 | module _ {A B : Set} (f : A → B) where 14 | 15 | data FRConstr : B → Set where 16 | dom : (a : A) → FRConstr (f a) 17 | fcomp : (r : Int) (p : Int → B) (φ : CofProp) (s t : Int) → ∥ t ≡ O ⊎ r ≡ s ∥ → FRConstr (p s) 18 | 19 | data FcompArity (r : Int) (p : Int → B) (φ : CofProp) : B → Set where 20 | fcompArg : (i : Int) → ∥ [ φ ] ⊎ (r ≡ i) ∥ → FcompArity r p φ (p i) 21 | 22 | FRPoly : IxPoly B 23 | FRPoly = record 24 | { Constr = FRConstr 25 | ; Arity = λ 26 | { (dom a) _ → ∅ 27 | ; (fcomp r p φ s t w) → FcompArity r p φ 28 | } 29 | ; Red = λ 30 | { (dom a) → ⊥ 31 | ; (fcomp r p φ s t w) → φ ∨ t ≈I 32 | } 33 | ; ev = λ {b} → λ 34 | { (dom a) → O≠I 35 | ; (fcomp r p φ s t w) → 36 | let 37 | t≡ICase : t ≡ I → ∥ t ≡ O ⊎ r ≡ s ∥ → FcompArity r p φ (p s) 38 | t≡ICase t≡I = 39 | ∥⊎∥-rec uip uip 40 | (λ {refl → O≠I t≡I}) 41 | (λ {refl → fcompArg r ∣ inr refl ∣}) 42 | (λ {refl refl → O≠I t≡I}) 43 | in 44 | ∨-rec φ (t ≈I) 45 | (λ u → fcompArg s ∣ inl u ∣) 46 | (λ t≡I → t≡ICase t≡I w) 47 | (λ u t≡I → 48 | ∥∥-elim (λ w → fcompArg s ∣ inl u ∣ ≡ t≡ICase t≡I w) 49 | (λ 50 | { (inl refl) → O≠I t≡I 51 | ; (inr refl) → cong (fcompArg s) (trunc _ _) 52 | }) 53 | (λ _ _ → uipImp) 54 | w) 55 | } 56 | } 57 | 58 | FR : B → Set 59 | FR b = IxWR FRPoly b 60 | 61 | FRIsFib : isFib FR 62 | FRIsFib r p φ h c₀ = 63 | record 64 | { comp = λ s → 65 | ( joined s O ∣ inl refl ∣ 66 | , λ u → 67 | symm (ixred FRPoly (fcomp r p φ s O ∣ inl refl ∣) _ ∣ inl u ∣) 68 | ) 69 | ; cap = 70 | ( path 71 | (λ t → joined r t ∣ inr refl ∣) 72 | (cong (joined r O) (trunc _ _)) 73 | (ixred FRPoly (fcomp r p φ r I ∣ inr refl ∣) _ ∣ inr refl ∣) 74 | , λ t u → symm (ixred FRPoly (fcomp r p φ r t ∣ inr refl ∣) _ ∣ inl u ∣) 75 | ) 76 | 77 | } 78 | where 79 | box : (i : Int) → ∥ [ φ ] ⊎ (r ≡ i) ∥ → IxWR FRPoly (p i) 80 | box i = ∥∥-rec _ 81 | (λ 82 | { (inl u) → h u i 83 | ; (inr refl) → c₀ .fst 84 | }) 85 | (λ 86 | { (inl u) (inl u') → cong (λ y → h y i) (cofIsProp φ u u') 87 | ; (inl u) (inr refl) → c₀ .snd u 88 | ; (inr refl) (inl u') → symm (c₀ .snd u') 89 | ; (inr refl) (inr refl) → refl 90 | }) 91 | 92 | joined : (s t : Int) → ∥ t ≡ O ⊎ r ≡ s ∥ → IxWR FRPoly (p s) 93 | joined s t ⊔ = 94 | ixsup FRPoly (fcomp r p φ s t ⊔) (λ {._ (fcompArg i w) → box i w}) 95 | 96 | quot : A → Σ B FR 97 | quot a = (_ , ixsup FRPoly (dom a) (λ _ → ∅-rec)) 98 | 99 | quotIsTrivCofibration : triv-cofibration quot 100 | quotIsTrivCofibration = record 101 | { j = λ X χ quotCase → uncurry 102 | (IxWR-elim _ _ 103 | (λ 104 | { (dom a) v _ → 105 | subst 106 | (λ v' → X (f a , ixsup _ (dom a) v')) 107 | (funext λ _ → funext ∅-elim) 108 | (quotCase a) 109 | ; (fcomp r p φ s t w) box boxrec → 110 | joined X χ r p φ box boxrec s t w 111 | }) 112 | (λ 113 | { (dom a) _ _ O≡I → O≠I O≡I 114 | ; (fcomp r p φ s t w) box boxrec → 115 | boundary X χ r p φ box boxrec s t w 116 | })) 117 | ; upper-tri = λ X χ quotCase a → 118 | cong 119 | (λ p → subst (λ v' → X (f a , ixsup _ (dom a) v')) p (quotCase a)) 120 | (uip (funext λ _ → funext ∅-elim) refl) 121 | } 122 | where 123 | module _ (X : Σ B FR → Set) (χ : isFib X) 124 | (r : Int) (p : Int → B) (φ : CofProp) 125 | (box : (b : B) → FcompArity r p φ b → FR b) 126 | (boxrec : (b : B) (w' : FcompArity r p φ b) → X (b , box b w')) 127 | where 128 | 129 | fcompTerm : (s t : Int) → ∥ t ≡ O ⊎ r ≡ s ∥ → _ 130 | fcompTerm s t w = ixsup _ (fcomp r p φ s t w) box 131 | 132 | fixTube : [ φ ] → (j : Int) → X (p r , fcompTerm r j ∣ inr refl ∣) 133 | fixTube u j = 134 | subst (curry X (p r)) 135 | (symm (ixred _ (fcomp r p φ r j ∣ inr refl ∣) box ∣ inl u ∣)) 136 | (boxrec (p r) (fcompArg r ∣ inl u ∣)) 137 | 138 | fixBase : _ [ φ ↦ fixTube ◆ I ] 139 | fixBase = 140 | ( subst (curry X (p r)) 141 | (symm (ixred _ (fcomp r p φ r I ∣ inr refl ∣) box ∣ inr refl ∣)) 142 | (boxrec (p r) (fcompArg r ∣ inr refl ∣)) 143 | , λ u → 144 | adjustSubstEq (curry X (p r)) 145 | (cong (λ v → box (p r) (fcompArg r v)) (trunc _ _)) refl 146 | (symm (ixred _ (fcomp r p φ r I ∣ inr refl ∣) box ∣ inl u ∣)) 147 | (symm (ixred _ (fcomp r p φ r I ∣ inr refl ∣) box ∣ inr refl ∣)) 148 | (trans 149 | (congdep (λ v → boxrec (p r) (fcompArg r v)) (trunc ∣ inl u ∣ ∣ inr refl ∣)) 150 | (symm 151 | (appCong 152 | (subst-cong-assoc (curry X (p r)) (λ v → box (p r) (fcompArg r v)) 153 | (trunc ∣ inl u ∣ ∣ inr refl ∣))))) 154 | ) 155 | 156 | fix = 157 | strictifyFib X χ I 158 | (λ j → p r , ixsup _ (fcomp r p φ r j ∣ inr refl ∣) box) 159 | φ 160 | fixTube 161 | fixBase 162 | 163 | compTube : [ φ ] → (i : Int) → X (p i , fcompTerm i O ∣ inl refl ∣) 164 | compTube u i = 165 | subst (curry X (p i)) 166 | (symm (ixred _ (fcomp r p φ i O ∣ inl refl ∣) box ∣ inl u ∣)) 167 | (boxrec (p i) (fcompArg i ∣ inl u ∣)) 168 | 169 | compBase : _ [ φ ↦ compTube ◆ r ] 170 | compBase = 171 | ( subst (curry X (p r)) 172 | (cong (fcompTerm r O) (trunc _ _)) 173 | (fix .comp O (I≡IsCofProp O) .fst) 174 | , λ u → 175 | trans 176 | (cong (subst (curry X (p r)) (cong (fcompTerm r O) (trunc _ _))) 177 | (fix .comp O (I≡IsCofProp O) .snd u)) 178 | (trans 179 | (substTrans (curry X (p r)) 180 | (cong (fcompTerm r O) (trunc _ _)) 181 | (symm (ixred _ (fcomp r p φ r O ∣ inr refl ∣) box ∣ inl u ∣))) 182 | (adjustSubstEq (curry X (p r)) 183 | refl refl 184 | (symm (ixred _ (fcomp r p φ r O ∣ inl refl ∣) box ∣ inl u ∣)) 185 | (trans 186 | (cong (fcompTerm r O) (trunc ∣ inr refl ∣ ∣ inl refl ∣)) 187 | (symm (ixred _ (fcomp r p φ r O ∣ inr refl ∣) box ∣ inl u ∣))) 188 | refl)) 189 | ) 190 | 191 | compMain = 192 | χ r (λ i → p i , fcompTerm i O ∣ inl refl ∣) 193 | φ 194 | compTube 195 | compBase 196 | 197 | module _ (t : Int) where 198 | 199 | capTube : [ φ ∨ t ≈O ∨ t ≈I ] → (i : Int) → X (p r , fcompTerm r t ∣ inr refl ∣) 200 | capTube = 201 | ∨-rec φ (t ≈O ∨ t ≈I) 202 | (λ u _ → fixTube u t) 203 | (OI-rec t 204 | (λ {refl j → 205 | subst (curry X (p r)) 206 | (cong (fcompTerm r O) (trunc _ _)) 207 | (compMain .cap .fst .at j)}) 208 | (λ {refl j → fixBase .fst})) 209 | (λ u → OI-elim t _ 210 | (λ {refl → funext λ j → 211 | adjustSubstEq (curry X (p r)) 212 | (symm (ixred _ (fcomp r p φ r O ∣ inl refl ∣) box ∣ inl u ∣)) 213 | refl 214 | (symm (ixred _ (fcomp r p φ r O ∣ inr refl ∣) box ∣ inl u ∣)) 215 | (cong (fcompTerm r O) (trunc _ _)) 216 | (compMain .cap .snd j u)}) 217 | (λ {refl → funext λ j → fixBase .snd u})) 218 | 219 | capBase : _ [ φ ∨ t ≈O ∨ t ≈I ↦ capTube ◆ I ] 220 | capBase = 221 | ( fix .comp t (I≡IsCofProp t) .fst 222 | , ∨-elimEq φ (t ≈O ∨ t ≈I) 223 | (fix .comp t (I≡IsCofProp t) .snd) 224 | (∨-elimEq (t ≈O) (t ≈I) 225 | (λ {refl → 226 | adjustSubstEq (curry X (p r)) 227 | refl (cong (fcompTerm r O) (trunc _ _)) 228 | (cong (fcompTerm r O) (trunc _ _)) refl 229 | (compMain .cap .fst .atI)}) 230 | (λ {refl → symm (fix .cap (I≡IsCofProp I))})) 231 | ) 232 | 233 | capMain = 234 | χ I (λ _ → p r , fcompTerm r t ∣ inr refl ∣) 235 | (φ ∨ t ≈O ∨ t ≈I) 236 | capTube 237 | capBase 238 | .comp O 239 | 240 | joined : (s t : Int) (w : ∥ t ≡ O ⊎ r ≡ s ∥) 241 | → X (p s , ixsup _ (fcomp r p φ s t w) box) 242 | joined s t = 243 | ∥⊎∥-elim uip uip _ 244 | (λ {refl → compMain .comp s .fst}) 245 | (λ {refl → capMain t .fst }) 246 | (λ {refl refl → 247 | trans 248 | (capMain O .snd ∣ inr ∣ inl refl ∣ ∣) 249 | (trans 250 | (appCong 251 | (subst-cong-assoc (curry X (p r)) 252 | (fcompTerm r O) 253 | (trunc _ _))) 254 | (cong (subst (curry X (p r) ∘ fcompTerm r O) (trunc _ _)) 255 | (symm (compMain .cap .fst .atO))))}) 256 | 257 | boundary : (s t : Int) 258 | (w : ∥ t ≡ O ⊎ r ≡ s ∥) (y : [ φ ∨ t ≈I ]) 259 | → subst (curry X (p s)) 260 | (ixred _ (fcomp r p φ s t w) box y) 261 | (joined s t w) 262 | ≡ boxrec (p s) (IxPoly.ev FRPoly (fcomp r p φ s t w) y) 263 | boundary s t = 264 | ∥⊎∥-elim uip uip _ 265 | (λ {refl → ∥⊎∥-elim (cofIsProp φ) uip _ 266 | (λ u → 267 | adjustSubstEq (curry X (p s)) 268 | refl 269 | (symm (ixred _ (fcomp r p φ s O ∣ inl refl ∣) box ∣ inl u ∣)) 270 | (ixred _ (fcomp r p φ s O ∣ inl refl ∣) box ∣ inl u ∣) 271 | refl 272 | (symm (compMain .comp s .snd u))) 273 | (λ O≡I → O≠I O≡I) 274 | (λ _ _ → uipImp)}) 275 | (λ {refl → ∥⊎∥-elim (cofIsProp φ) uip _ 276 | (λ u → 277 | adjustSubstEq (curry X (p s)) 278 | refl 279 | (symm (ixred _ (fcomp r p φ r t ∣ inr refl ∣) box ∣ inl u ∣)) 280 | (ixred _ (fcomp r p φ r t ∣ inr refl ∣) box ∣ inl u ∣) 281 | refl 282 | (symm (capMain t .snd ∣ inl u ∣))) 283 | (λ {refl → 284 | adjustSubstEq (curry X (p s)) 285 | refl 286 | (symm (ixred _ (fcomp r p φ r I ∣ inr refl ∣) box ∣ inr refl ∣)) 287 | (ixred _ (fcomp r p φ r I ∣ inr refl ∣) box ∣ inr refl ∣) 288 | refl 289 | (symm (capMain I .snd ∣ inr ∣ inr refl ∣ ∣))}) 290 | (λ _ _ → uipImp)}) 291 | (λ _ _ → funext λ y → uipImp) 292 | -------------------------------------------------------------------------------- /agda/glueing/core.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Definition of (weak) Glue and their (unaligned) fibrancy. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module glueing.core where 8 | 9 | open import prelude 10 | open import interval 11 | open import cofprop 12 | open import fibrations 13 | open import equivs 14 | open import Data.paths 15 | open import Data.products 16 | 17 | ---------------------------------------------------------------------- 18 | -- Glueing 19 | ---------------------------------------------------------------------- 20 | record Glue (Φ : CofProp) 21 | (T : [ Φ ] → Set) 22 | (A : Set) 23 | (f : (u : [ Φ ]) → T u → A) : Set 24 | where 25 | constructor glue 26 | field 27 | dom : (u : [ Φ ]) → T u 28 | cod : A 29 | match : (u : [ Φ ]) → f u (dom u) ≡ cod 30 | 31 | open Glue public 32 | 33 | Glue' : 34 | ∀{a}{Γ : Set a} 35 | (Φ : Γ → CofProp) 36 | (T : res Γ Φ → Set) 37 | (A : Γ → Set) 38 | (f : (xu : res Γ Φ) → T xu → A (xu .fst)) 39 | → --------------- 40 | Γ → Set 41 | Glue' Φ T A f x = Glue (Φ x) (λ u → T (x , u)) (A x) (λ u → f (x , u)) 42 | 43 | glueExt : 44 | {Φ : CofProp} 45 | {T : [ Φ ] → Set} 46 | {A : Set} 47 | (f : (u : [ Φ ]) → T u → A) 48 | (g g' : Glue Φ T A f) 49 | (p : g .dom ≡ g' .dom) 50 | (q : g .cod ≡ g' .cod) 51 | → --------------- 52 | g ≡ g' 53 | glueExt _ (glue t a ft≡a) (glue _ _ ft≡a') refl refl = 54 | cong (glue t a) (funext λ u → uip (ft≡a u) (ft≡a' u)) 55 | 56 | FibGlueId : 57 | (Φ : Int → CofProp) 58 | {T : res Int Φ → Set} 59 | {A : Int → Set} 60 | (f : (xu : res Int Φ) → T xu → A (xu .fst)) 61 | (e : Equiv' f) 62 | → --------------- 63 | isFib T → isFib A → isFib (Glue' Φ T A f) 64 | FibGlueId Φ {T} {A} f e τ α r p ψ q (glue t₀ a₀ ft₀↗a₀ , ext) = 65 | record 66 | { comp = λ s → 67 | ( glue 68 | (λ us → R s us .comp O (I≡IsCofProp O) .fst .fst) 69 | (wA' s .comp O (I≡IsCofProp O) .fst) 70 | (λ us → 71 | trans 72 | (wA' s .comp O (I≡IsCofProp O) .snd ∣ inl us ∣) 73 | (symm (R s us .comp O (I≡IsCofProp O) .fst .snd .atO))) 74 | , λ v → 75 | glueExt (λ u → f (p s , u)) _ _ 76 | (funext λ us → 77 | cong fst 78 | (trans 79 | (R s us .comp O (I≡IsCofProp O) .snd v) 80 | (symm (C₂ s us (RFiber s us v) .atO)))) 81 | (wA' s .comp O (I≡IsCofProp O) .snd ∣ inr v ∣) 82 | ) 83 | ; cap = 84 | ( path 85 | (λ t → 86 | glue 87 | (λ ur → R̲ t ur .fst .fst) 88 | (w̲A' t .fst) 89 | (λ ur → 90 | trans 91 | (w̲A' t .snd ∣ inl ur ∣) 92 | (symm (R̲ t ur .fst .snd .atO)))) 93 | (glueExt (λ u → f (p r , u)) _ _ 94 | (funext λ ur → 95 | trans 96 | (symm (fiberDomEqDep (wAr≡dO O ur refl) refl)) 97 | (symm (cong fst (R̲ O ur .snd ∣ inl refl ∣)))) 98 | (symm (w̲A' O .snd ∣ inr ∣ inl refl ∣ ∣))) 99 | (glueExt (λ u → f (p r , u)) _ _ 100 | (funext λ ur → 101 | trans 102 | (cong fst (C̲₂ I ur (R̲Fiber I ur ∣ inl refl ∣) .atO)) 103 | (symm (cong fst (R̲ I ur .snd ∣ inr ∣ inl refl ∣ ∣)))) 104 | (trans 105 | (symm (d I .snd ∣ inr ∣ inl refl ∣ ∣)) 106 | (symm (w̲A' I .snd ∣ inr ∣ inr ∣ inl refl ∣ ∣ ∣)))) 107 | , λ t v → 108 | glueExt (λ u → f (p r , u)) _ _ 109 | (funext λ ur → 110 | trans 111 | (cong fst (R̲ t ur .snd ∣ inr ∣ inr v ∣ ∣)) 112 | (symm (cong fst (C̲₂ t ur (R̲Fiber t ur ∣ inr v ∣) .atO)))) 113 | (trans 114 | (w̲A' t .snd ∣ inr ∣ inr ∣ inr v ∣ ∣ ∣) 115 | (trans 116 | (d t .snd ∣ inr ∣ inr v ∣ ∣) 117 | (cong cod (ext v)))) 118 | ) 119 | } 120 | 121 | where 122 | -- Step 4 123 | 124 | wATube : [ ψ ] → (i : Int) → A (p i) 125 | wATube v i = q v i .cod 126 | 127 | wABase : A (p r) [ ψ ↦ wATube ◆ r ] 128 | wABase = (a₀ , λ v → cong cod (ext v)) 129 | 130 | wA = α r p ψ wATube wABase 131 | 132 | -- Step 5 133 | 134 | module _ (s : Int) where 135 | 136 | module _ (us : [ Φ (p s) ]) where 137 | 138 | C₁ = e (p s , us) (wA .comp s .fst) .fst 139 | C₂ = e (p s , us) (wA .comp s .fst) .snd 140 | 141 | -- Step 6 142 | 143 | σ : isFib (Fiber (f (p s , us))) 144 | σ = 145 | FibΣ {B = λ{ (a , t) → f (p s , us) t ~ a}} 146 | (reindex T τ (λ _ → p s , us)) 147 | (reindex (Path A) (FibPath α) (λ{ (a , t) → (p s , f (p s , us) t , a)})) 148 | 149 | 150 | RFiber : [ ψ ] → Fiber (f (p s , us)) (wA .comp s .fst) 151 | RFiber v = 152 | ( q v s .dom us 153 | , path 154 | (λ _ → wA .comp s .fst) 155 | (trans 156 | (symm (q v s .match us)) 157 | (symm (wA .comp s .snd v))) 158 | refl 159 | ) 160 | 161 | RTube : [ ψ ] → Int → Fiber (f (p s , us)) (wA .comp s .fst) 162 | RTube uv k = C₂ (RFiber uv) .at k 163 | 164 | RBase : Fiber (f (p s , us)) (wA .comp s .fst) [ ψ ↦ RTube ◆ I ] 165 | RBase = ( C₁ , λ uv → C₂ (RFiber uv) .atI) 166 | 167 | R = strictifyFib (Fiber (f (p s , us))) σ I (λ _ → wA .comp s .fst) ψ RTube RBase 168 | 169 | -- Step 7 170 | 171 | wA'Tube : [ Φ (p s) ∨ ψ ] → Int → A (p s) 172 | wA'Tube = 173 | ∨-rec (Φ (p s)) ψ 174 | (λ us k → R us .comp O (I≡IsCofProp O) .fst .snd .at k) 175 | (λ v _ → q v s .cod) 176 | (λ us v → funext λ k → 177 | trans 178 | (symm (wA .comp s .snd v)) 179 | (cong (λ fib → fib .snd .at k) 180 | (trans 181 | (C₂ us (RFiber us v) .atO) 182 | (symm (R us .comp O (I≡IsCofProp O) .snd v))))) 183 | 184 | wA'Base : A (p s) [ Φ (p s) ∨ ψ ↦ wA'Tube ◆ I ] 185 | wA'Base = 186 | ( wA .comp s .fst 187 | , ∨-elimEq (Φ (p s)) ψ 188 | (λ us → R us .comp O (I≡IsCofProp O) .fst .snd .atI) 189 | (λ v → wA .comp s .snd v) 190 | ) 191 | 192 | wA' = strictifyFib A α I (λ _ → p s) (Φ (p s) ∨ ψ) wA'Tube wA'Base 193 | 194 | -- Step 4̲ 195 | 196 | module _ (t : Int) where 197 | 198 | dTube : [ t ≈O ∨ t ≈I ∨ ψ ] → Int → A (p r) 199 | dTube = 200 | ∨-rec (t ≈O) (t ≈I ∨ ψ) 201 | (λ t≡O j → wA .cap .fst .at j) 202 | (λ _ _ → a₀) 203 | (λ {refl → ∨-elimEq (t ≈I) ψ 204 | (λ t≡I → O≠I t≡I) 205 | (λ v → funext λ j → 206 | trans 207 | (cong cod (ext v)) 208 | (symm (wA .cap .snd j v)))}) 209 | 210 | dBase : A (p r) [ t ≈O ∨ t ≈I ∨ ψ ↦ dTube ◆ I ] 211 | dBase = 212 | ( a₀ 213 | , ∨-elimEq (t ≈O) (t ≈I ∨ ψ) 214 | (λ {refl → wA .cap .fst .atI}) 215 | (λ _ → refl) 216 | ) 217 | 218 | d = α I (λ _ → p r) (t ≈O ∨ t ≈I ∨ ψ) dTube dBase .comp O 219 | 220 | -- Step 5̲ 221 | 222 | module _ (ur : [ Φ (p r) ]) where 223 | 224 | C̲₁ = e (p r , ur) (d .fst) .fst 225 | C̲₂ = e (p r , ur) (d .fst) .snd 226 | 227 | -- Step 6̲ 228 | 229 | R̲Fiber : [ t ≈I ∨ ψ ] → Fiber (f (p r , ur)) (d .fst) 230 | R̲Fiber = 231 | ∨-rec (t ≈I) (ψ) 232 | (λ {refl → 233 | ( t₀ ur 234 | , path 235 | (λ _ → d .fst) 236 | (trans 237 | (symm (ft₀↗a₀ ur)) 238 | (symm (d .snd ∣ inr ∣ inl refl ∣ ∣))) 239 | refl 240 | )}) 241 | (λ v → 242 | ( q v r .dom ur 243 | , path 244 | (λ _ → d .fst) 245 | (trans 246 | (trans 247 | (symm (q v r .match ur)) 248 | (symm (cong cod (ext v)))) 249 | (symm (d .snd ∣ inr ∣ inr v ∣ ∣))) 250 | refl 251 | )) 252 | (λ {refl v → 253 | FiberExt 254 | (symm (cong (λ g → g .dom ur) (ext v))) 255 | refl}) 256 | 257 | wAr≡dO : t ≡ O → wA .comp r .fst ≡ d .fst 258 | wAr≡dO t≡O = trans (d .snd ∣ inl t≡O ∣) (symm (wA .cap .fst .atO)) 259 | 260 | dO≡wAr : t ≡ O → d .fst ≡ wA .comp r .fst 261 | dO≡wAr t≡O = trans (wA .cap .fst .atO) (symm (d .snd ∣ inl t≡O ∣)) 262 | 263 | R̲FiberAtO : (t≡O : t ≡ O) (v : [ ψ ]) 264 | → subst (Fiber (f (p r , ur))) (wAr≡dO t≡O) (RFiber r ur v) ≡ R̲Fiber ∣ inr v ∣ 265 | R̲FiberAtO refl v = 266 | FiberExtDep (wAr≡dO refl) refl (funext λ _ → wAr≡dO refl) 267 | 268 | R̲Tube : [ t ≈O ∨ t ≈I ∨ ψ ] → Int → Fiber (f (p r , ur)) (d .fst) 269 | R̲Tube = 270 | ∨-rec (t ≈O) (t ≈I ∨ ψ) 271 | (λ t≡O k → 272 | subst (Fiber (f (p r , ur))) 273 | (wAr≡dO t≡O) 274 | (R r ur .comp k (I≡IsCofProp k) .fst)) 275 | (λ w k → C̲₂ (R̲Fiber w) .at k) 276 | (λ {refl → 277 | ∨-elimEq (t ≈I) ψ 278 | (λ t≡I → O≠I t≡I) 279 | (λ v → funext λ k → 280 | trans 281 | (congdep₂ 282 | (λ b fib → e (p r , ur) b .snd fib .at k) 283 | (wAr≡dO refl) 284 | (FiberExtDep (wAr≡dO refl) refl (funext λ _ → wAr≡dO refl))) 285 | (cong (subst (Fiber (f (p r , ur))) (wAr≡dO refl)) 286 | (symm (R r ur .comp k (I≡IsCofProp k) .snd v))))}) 287 | 288 | R̲Base : _ [ t ≈O ∨ t ≈I ∨ ψ ↦ R̲Tube ◆ I ] 289 | R̲Base = 290 | ( C̲₁ 291 | , ∨-elimEq (t ≈O) (t ≈I ∨ ψ) 292 | (λ {refl → 293 | trans 294 | (congdep (λ b → e (p r , ur) b .fst) (wAr≡dO refl)) 295 | (cong (subst (Fiber (f (p r , ur))) (wAr≡dO refl)) 296 | (R r ur .cap (I≡IsCofProp I)))}) 297 | (λ Iuv → C̲₂ (R̲Fiber Iuv) .atI) 298 | ) 299 | 300 | R̲ = σ r ur I (λ _ → d .fst) (t ≈O ∨ t ≈I ∨ ψ) R̲Tube R̲Base .comp O 301 | 302 | w̲A'Tube : [ Φ (p r) ∨ t ≈O ∨ t ≈I ∨ ψ ] → Int → A (p r) 303 | w̲A'Tube = 304 | ∨-rec (Φ (p r)) (t ≈O ∨ t ≈I ∨ ψ) 305 | (λ ur k → R̲ ur .fst .snd .at k) 306 | (∨-rec (t ≈O) (t ≈I ∨ ψ) 307 | (λ {refl k → wA' r .comp k (I≡IsCofProp k) .fst}) 308 | (λ _ _ → d .fst) 309 | (λ {refl → 310 | ∨-elimEq (t ≈I) ψ 311 | (λ t≡I → O≠I t≡I) 312 | (λ v → funext λ k → 313 | trans 314 | (d .snd ∣ inr ∣ inr v ∣ ∣) 315 | (trans 316 | (cong cod (ext v)) 317 | (symm (wA' r .comp k (I≡IsCofProp k) .snd ∣ inr v ∣))))})) 318 | (λ ur → ∨-elimEq (t ≈O) (t ≈I ∨ ψ) 319 | (λ {refl → funext λ k → 320 | trans 321 | (wA' r .comp k (I≡IsCofProp k) .snd ∣ inl ur ∣) 322 | (trans 323 | (symm (fiberPathEqDep (wAr≡dO ur refl) refl k)) 324 | (symm (fiberPathEq (R̲ ur .snd ∣ inl refl ∣) k)))}) 325 | (∨-elimEq (t ≈I) ψ 326 | (λ {refl → funext λ k → 327 | fiberPathEq 328 | (trans 329 | (C̲₂ ur (R̲Fiber ur ∣ inl refl ∣) .atO) 330 | (symm (R̲ ur .snd ∣ inr ∣ inl refl ∣ ∣))) 331 | k}) 332 | (λ v → funext λ k → 333 | fiberPathEq 334 | (trans 335 | (C̲₂ ur (R̲Fiber ur ∣ inr v ∣) .atO) 336 | (symm (R̲ ur .snd ∣ inr ∣ inr v ∣ ∣))) 337 | k))) 338 | 339 | w̲A'Base : A (p r) [ Φ (p r) ∨ t ≈O ∨ t ≈I ∨ ψ ↦ w̲A'Tube ◆ I ] 340 | w̲A'Base = 341 | ( d .fst 342 | , ∨-elimEq (Φ (p r)) (t ≈O ∨ t ≈I ∨ ψ) 343 | (λ ur → R̲ ur .fst .snd .atI) 344 | (∨-elimEq (t ≈O) (t ≈I ∨ ψ) 345 | (λ {refl → 346 | trans 347 | (d .snd ∣ inl refl ∣) 348 | (trans 349 | (symm (wA .cap .fst .atO)) 350 | (wA' r .cap (I≡IsCofProp I)))}) 351 | (λ _ → refl)) 352 | ) 353 | 354 | w̲A' = α I (λ _ → p r) (Φ (p r) ∨ t ≈O ∨ t ≈I ∨ ψ) w̲A'Tube w̲A'Base .comp O 355 | -------------------------------------------------------------------------------- /agda/prelude.agda: -------------------------------------------------------------------------------- 1 | {- 2 | 3 | Basics. 4 | 5 | -} 6 | {-# OPTIONS --rewriting #-} 7 | module prelude where 8 | 9 | open import Agda.Primitive public 10 | 11 | infix 1 Σ 12 | infixr 3 _,_ _×_ _⊎_ 13 | infixr 5 _∘_ 14 | 15 | ---------------------------------------------------------------------- 16 | -- Identity function 17 | ---------------------------------------------------------------------- 18 | id : ∀{a}{A : Set a} → A → A 19 | id x = x 20 | 21 | ---------------------------------------------------------------------- 22 | -- Composition 23 | ---------------------------------------------------------------------- 24 | _∘_ : 25 | {ℓ m n : Level} 26 | {A : Set ℓ} 27 | {B : Set m} 28 | {C : B → Set n} 29 | (g : (b : B) → C b) 30 | (f : A → B) 31 | → ------------- 32 | (a : A) → C (f a) 33 | (g ∘ f) x = g (f x) 34 | 35 | ---------------------------------------------------------------------- 36 | -- Propositional equality 37 | ---------------------------------------------------------------------- 38 | open import Agda.Builtin.Equality public 39 | 40 | {-# BUILTIN REWRITE _≡_ #-} 41 | 42 | trans : 43 | {ℓ : Level} 44 | {A : Set ℓ} 45 | {x y z : A} 46 | (q : y ≡ z) 47 | (p : x ≡ y) 48 | → --------- 49 | x ≡ z 50 | trans q refl = q 51 | 52 | symm : 53 | {ℓ : Level} 54 | {A : Set ℓ} 55 | {x y : A} 56 | (p : x ≡ y) 57 | → --------- 58 | y ≡ x 59 | symm refl = refl 60 | 61 | cong : 62 | {ℓ ℓ' : Level} 63 | {A : Set ℓ} 64 | {B : Set ℓ'} 65 | (f : A → B) 66 | {x y : A} 67 | (p : x ≡ y) 68 | → ----------- 69 | f x ≡ f y 70 | cong _ refl = refl 71 | 72 | cong₂ : 73 | {ℓ ℓ' : Level} 74 | {A A' : Set ℓ} 75 | {B : Set ℓ'} 76 | (f : A → A' → B) 77 | {x y : A} 78 | {x' y' : A'} 79 | (p : x ≡ y) 80 | (q : x' ≡ y') 81 | → -------------- 82 | f x x' ≡ f y y' 83 | cong₂ _ refl refl = refl 84 | 85 | subst : 86 | {ℓ ℓ' : Level} 87 | {A : Set ℓ} 88 | (B : A → Set ℓ') 89 | {x y : A} 90 | (p : x ≡ y) 91 | → -------------- 92 | B x → B y 93 | subst _ refl = λ b → b 94 | 95 | congdep : 96 | {ℓ ℓ' : Level} 97 | {A : Set ℓ} 98 | {B : A → Set ℓ'} 99 | (f : (a : A) → B a) 100 | {x y : A} 101 | (p : x ≡ y) 102 | → ----------- 103 | subst B p (f x) ≡ f y 104 | congdep _ refl = refl 105 | 106 | congdep₂ : 107 | {ℓ ℓ' ℓ'' : Level} 108 | {A : Set ℓ} 109 | {B : A → Set ℓ'} 110 | {C : A → Set ℓ''} 111 | (f : (a : A) → B a → C a) 112 | {x y : A} 113 | (p : x ≡ y) 114 | {z : B x} {w : B y} 115 | (q : subst B p z ≡ w) 116 | → subst C p (f x z) ≡ f y w 117 | congdep₂ _ refl refl = refl 118 | 119 | substconst : 120 | {ℓ ℓ' : Level} 121 | {A : Set ℓ} 122 | (B : Set ℓ') 123 | {x y : A} 124 | (p : x ≡ y) 125 | (b : B) 126 | → ------------------------ 127 | (subst (λ _ → B) p) b ≡ b 128 | substconst _ refl _ = refl 129 | 130 | substfam : 131 | {ℓ ℓ' : Level} 132 | {A : Set ℓ} 133 | {B : A → Set ℓ'} 134 | {x y : A} 135 | (p : x ≡ y) 136 | (b : (a : A) → B a) 137 | → ------------------------ 138 | subst B p (b x) ≡ b y 139 | substfam refl b = refl 140 | 141 | subst₂ : 142 | {ℓ ℓ' : Level} 143 | {A A' : Set ℓ} 144 | (B : A → A' → Set ℓ') 145 | {x y : A} 146 | {x' y' : A'} 147 | (p : x ≡ y) 148 | (p' : x' ≡ y') 149 | → ------------------ 150 | B x x' → B y y' 151 | subst₂ _ refl refl = λ b → b 152 | 153 | subst₂dep : 154 | {ℓ ℓ' ℓ'' : Level} 155 | {A : Set ℓ} (A' : A → Set ℓ') 156 | (B : (a : A) → A' a → Set ℓ'') 157 | {x y : A} 158 | {x' : A' x} {y' : A' y} 159 | (p : x ≡ y) 160 | (p' : subst A' p x' ≡ y') 161 | → ------------------ 162 | B x x' → B y y' 163 | subst₂dep _ _ refl refl = λ b → b 164 | 165 | subst-cong-assoc : 166 | {l l' l'' : Level} 167 | {A : Set l} 168 | {B : Set l'} 169 | (C : B → Set l'') 170 | (f : A → B) 171 | {x y : A} 172 | (p : x ≡ y) 173 | → ------------------------------------------ 174 | subst (λ x → C (f x)) p ≡ subst C (cong f p) 175 | subst-cong-assoc _ _ refl = refl 176 | 177 | substTrans : 178 | {ℓ ℓ' : Level} 179 | {A : Set ℓ} 180 | (B : A → Set ℓ') 181 | {x y z : A} 182 | (q : y ≡ z) (p : x ≡ y) 183 | {b : B x} 184 | → ------------------------------------------ 185 | subst B (trans q p) b ≡ subst B q (subst B p b) 186 | substTrans B refl refl = refl 187 | 188 | uip : 189 | {ℓ : Level} 190 | {A : Set ℓ} 191 | {x y : A} 192 | (p q : x ≡ y) 193 | → ----------- 194 | p ≡ q 195 | uip refl refl = refl 196 | 197 | uipImp : 198 | {ℓ : Level} 199 | {A : Set ℓ} 200 | {x y : A} 201 | {p q : x ≡ y} 202 | → ----------- 203 | p ≡ q 204 | uipImp {p = refl} {q = refl} = refl 205 | 206 | appCong : 207 | {ℓ ℓ' : Level} 208 | {A : Set ℓ} 209 | {B : A → Set ℓ'} 210 | {f g : (a : A) → B a} 211 | {x : A} 212 | (p : f ≡ g) 213 | → ----------- 214 | f x ≡ g x 215 | appCong refl = refl 216 | 217 | adjustSubstEq : 218 | {ℓ ℓ' : Level} 219 | {A : Set ℓ} 220 | (B : A → Set ℓ') 221 | {x y z w : A} 222 | (p : x ≡ z) (p' : y ≡ z) 223 | (q : x ≡ w) (q' : y ≡ w) 224 | {b : B x} {b' : B y} 225 | → subst B p b ≡ subst B p' b' 226 | → subst B q b ≡ subst B q' b' 227 | adjustSubstEq B refl refl refl refl = id 228 | 229 | ---------------------------------------------------------------------- 230 | -- Type coercion 231 | ---------------------------------------------------------------------- 232 | coe : ∀ {ℓ} {A B : Set ℓ} → A ≡ B → A → B 233 | coe = subst id 234 | 235 | ---------------------------------------------------------------------- 236 | -- Empty type 237 | ---------------------------------------------------------------------- 238 | data ∅ : Set where 239 | 240 | ∅-elim : 241 | {ℓ : Level} 242 | {A : ∅ → Set ℓ} 243 | → --------- 244 | (v : ∅) → A v 245 | ∅-elim () 246 | 247 | ∅-rec : 248 | {ℓ : Level} 249 | {A : Set ℓ} 250 | → --------- 251 | ∅ → A 252 | ∅-rec () 253 | 254 | ¬_ : ∀ {ℓ} → Set ℓ → Set ℓ 255 | ¬ A = A → ∅ 256 | 257 | ---------------------------------------------------------------------- 258 | -- One-element type 259 | ---------------------------------------------------------------------- 260 | open import Agda.Builtin.Unit renaming (⊤ to Unit) public 261 | 262 | ---------------------------------------------------------------------- 263 | -- Booleans 264 | ---------------------------------------------------------------------- 265 | open import Agda.Builtin.Bool renaming (Bool to 𝔹) public 266 | 267 | ---------------------------------------------------------------------- 268 | -- Natural Numbers 269 | ---------------------------------------------------------------------- 270 | open import Agda.Builtin.Nat renaming (Nat to ℕ) public 271 | 272 | ---------------------------------------------------------------------- 273 | -- Disjoint union 274 | ---------------------------------------------------------------------- 275 | data _⊎_ {ℓ m : Level}(A : Set ℓ)(B : Set m) : Set (ℓ ⊔ m) where 276 | inl : A → A ⊎ B 277 | inr : B → A ⊎ B 278 | 279 | ---------------------------------------------------------------------- 280 | -- Propositional truncation 281 | ---------------------------------------------------------------------- 282 | 283 | postulate 284 | ∥_∥ : ∀ {ℓ} → Set ℓ → Set ℓ 285 | 286 | module _ {ℓ : Level} {A : Set ℓ} where 287 | postulate 288 | ∣_∣ : A → ∥ A ∥ 289 | 290 | trunc : (t u : ∥ A ∥) → t ≡ u 291 | 292 | ∥∥-rec : ∀ {ℓ'} 293 | (P : Set ℓ') 294 | (f : A → P) 295 | (p : ∀ a b → f a ≡ f b) 296 | → --------------- 297 | ∥ A ∥ → P 298 | 299 | ∥∥-rec-β : ∀ {ℓ'} (P : Set ℓ') f p → (a : A) 300 | → ∥∥-rec P f p ∣ a ∣ ≡ f a 301 | 302 | ∥∥-elim : ∀ {ℓ'} 303 | (P : ∥ A ∥ → Set ℓ') 304 | (f : (a : A) → P ∣ a ∣) 305 | (p : ∀ a b → subst P (trunc ∣ a ∣ ∣ b ∣) (f a) ≡ f b) 306 | → --------------- 307 | (t : ∥ A ∥) → P t 308 | 309 | ∥∥-elim-β : ∀ {ℓ'} (P : ∥ A ∥ → Set ℓ') f p → (a : A) 310 | → ∥∥-elim P f p ∣ a ∣ ≡ f a 311 | 312 | {-# REWRITE ∥∥-rec-β ∥∥-elim-β #-} 313 | 314 | ---------------------------------------------------------------------- 315 | -- Truncated disjunction 316 | ---------------------------------------------------------------------- 317 | 318 | ∥⊎∥-rec : ∀ {ℓ ℓ' ℓ''} 319 | {A : Set ℓ} {B : Set ℓ'} {C : Set ℓ''} 320 | (pA : (a a' : A) → a ≡ a') 321 | (pB : (b b' : B) → b ≡ b') 322 | (f : A → C) 323 | (g : B → C) 324 | (p : ∀ a b → f a ≡ g b) 325 | → ∥ A ⊎ B ∥ → C 326 | ∥⊎∥-rec pA pB f g p = 327 | ∥∥-rec _ 328 | (λ {(inl a) → f a ; (inr b) → g b}) 329 | (λ 330 | { (inl a) (inl a') → cong f (pA a a') 331 | ; (inl a) (inr b') → p a b' 332 | ; (inr b) (inl a') → symm (p a' b) 333 | ; (inr b) (inr b') → cong g (pB b b') 334 | }) 335 | 336 | ∥⊎∥-elim : ∀ {ℓ ℓ' ℓ''} 337 | {A : Set ℓ} {B : Set ℓ'} 338 | (pA : (a a' : A) → a ≡ a') 339 | (pB : (b b' : B) → b ≡ b') 340 | (C : ∥ A ⊎ B ∥ → Set ℓ'') 341 | (f : (a : A) → C ∣ inl a ∣) 342 | (g : (b : B) → C ∣ inr b ∣) 343 | (p : ∀ a b → subst C (trunc _ _) (f a) ≡ g b) 344 | → (t : ∥ A ⊎ B ∥) → C t 345 | ∥⊎∥-elim pA pB C f g p = 346 | ∥∥-elim _ 347 | (λ {(inl a) → f a ; (inr b) → g b}) 348 | (λ 349 | { (inl a) (inl a') → 350 | trans 351 | (trans 352 | (congdep f (pA a a')) 353 | (symm (appCong (subst-cong-assoc C (∣_∣ ∘ inl) (pA a a'))))) 354 | (cong (λ q → subst C q (f a)) (uip (trunc _ _) (cong (∣_∣ ∘ inl) (pA a a')))) 355 | ; (inl a) (inr b') → p a b' 356 | ; (inr b) (inl a') → 357 | adjustSubstEq C 358 | refl (trunc ∣ inl a' ∣ ∣ inr b ∣) 359 | (trunc ∣ inr b ∣ ∣ inl a' ∣) refl 360 | (symm (p a' b)) 361 | ; (inr b) (inr b') → 362 | trans 363 | (trans 364 | (congdep g (pB b b')) 365 | (symm (appCong (subst-cong-assoc C (∣_∣ ∘ inr) (pB b b'))))) 366 | (cong (λ q → subst C q (g b)) (uip (trunc _ _) (cong (∣_∣ ∘ inr) (pB b b')))) 367 | }) 368 | 369 | ---------------------------------------------------------------------- 370 | -- Dependent products 371 | ---------------------------------------------------------------------- 372 | record Σ {ℓ m : Level} (A : Set ℓ) (B : A → Set m) : Set (ℓ ⊔ m) where 373 | constructor _,_ 374 | field 375 | fst : A 376 | snd : B fst 377 | 378 | open Σ public 379 | 380 | _&_ : ∀ {ℓ m : Level} {A : Set ℓ} {B : A → Set m} (a : A) → B a → Σ A B 381 | _&_ = _,_ 382 | 383 | syntax Σ A (λ x → B) = Σ x ∈ A , B 384 | 385 | _×_ : {ℓ m : Level} → Set ℓ → Set m → Set (ℓ ⊔ m) 386 | A × B = Σ A (λ _ → B) 387 | 388 | _×'_ : {ℓ ℓ' m m' : Level}{A : Set ℓ}{A' : Set ℓ'}{B : Set m}{B' : Set m'} 389 | → (A → A') → (B → B') → A × B → A' × B' 390 | (f ×' g) (a , b) = f a , g b 391 | 392 | _×id : {ℓ ℓ' m : Level}{A : Set ℓ}{A' : Set ℓ'}{B : A' → Set m} 393 | (f : A → A') → Σ A (B ∘ f) → Σ A' B 394 | (f ×id) (a , b) = (f a , b) 395 | 396 | Σext : 397 | {ℓ m : Level} 398 | {A : Set ℓ} 399 | {B : A → Set m} 400 | {x x' : A} 401 | {y : B x} 402 | {y' : B x'} 403 | (p : x ≡ x') 404 | (q : subst B p y ≡ y') 405 | → -------------------- 406 | (x , y) ≡ (x' , y') 407 | Σext refl refl = refl 408 | 409 | Σextdep : 410 | {ℓ m n : Level} 411 | {A : Set ℓ} 412 | {B : A → Set m} 413 | {C : (a : A) → B a → Set n} 414 | {x x' : A} 415 | (p : x ≡ x') 416 | {y : B x} {y' : B x'} 417 | (q : subst B p y ≡ y') 418 | {z : C x y} {z' : C x' y'} 419 | (r : subst₂dep B C p q z ≡ z') 420 | → subst (λ x → Σ (B x) (C x)) p (y , z) ≡ (y' , z') 421 | Σextdep refl refl refl = refl 422 | 423 | Σeq₁ : 424 | {ℓ ℓ' : Level} 425 | {A : Set ℓ} 426 | {B : A → Set ℓ'} 427 | {x y : Σ A B} 428 | (p : x ≡ y) 429 | → -------------- 430 | fst x ≡ fst y 431 | Σeq₁ refl = refl 432 | 433 | Σeq₂ : 434 | {ℓ ℓ' : Level} 435 | {A : Set ℓ} 436 | {B : A → Set ℓ'} 437 | {x y : Σ A B} 438 | (p : x ≡ y) 439 | → -------------- 440 | subst B (Σeq₁ p) (snd x) ≡ snd y 441 | Σeq₂ refl = refl 442 | 443 | uncurry : ∀ {ℓ ℓ' ℓ''} {A : Set ℓ} {B : A → Set ℓ'} {C : (a : A) → B a → Set ℓ''} 444 | → (∀ a b → C a b) 445 | → ((p : Σ A B) → C (p .fst) (p .snd)) 446 | uncurry f (a , b) = f a b 447 | 448 | curry : ∀ {ℓ ℓ' ℓ''} {A : Set ℓ} {B : A → Set ℓ'} {C : (a : A) → B a → Set ℓ''} 449 | → ((p : Σ A B) → C (p .fst) (p .snd)) 450 | → (∀ a b → C a b) 451 | curry f a b = f (a , b) 452 | 453 | ---------------------------------------------------------------------- 454 | -- Function extensionality 455 | ---------------------------------------------------------------------- 456 | postulate 457 | funext : 458 | {ℓ ℓ' : Level} 459 | {A : Set ℓ} 460 | {B : A → Set ℓ'} 461 | {f g : (x : A) → B x} 462 | → ---------------------------- 463 | ((x : A) → f x ≡ g x) → f ≡ g 464 | --------------------------------------------------------------------------------