├── .gitignore ├── LICENSE.md ├── Map ├── H-equivalence.agda ├── H-fiber.agda ├── Injection.agda ├── Surjection.agda └── WeakEquivalence.agda ├── Path.agda ├── Path ├── HigherOrder.agda ├── Lemmas.agda ├── Omega2-abelian.agda ├── Sum.agda └── Symbol.agda ├── Prelude.agda ├── README.agda ├── README.md ├── Space ├── Bool │ └── Initial.agda ├── Fin │ └── Lemmas.agda ├── Flower.agda ├── Flower │ └── Omega1.agda ├── FreeGroup.agda ├── Integer.agda ├── Interval.agda ├── List │ └── Lemmas.agda ├── Nat │ └── Lemmas.agda ├── Sphere.agda ├── Sphere │ ├── HopfJunior.agda │ ├── Omega1.agda │ └── TwoPoints.agda └── Torus.agda ├── Univalence.agda └── Univalence └── Lemmas.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | *.swp 3 | Test.agda 4 | Test 5 | -------------------------------------------------------------------------------- /LICENSE.md: -------------------------------------------------------------------------------- 1 | Please note that a significant portion of code was copied, modified 2 | and merged from Nils Anders Danielsson' library and Agda standard 3 | library. All code is released under the same BSD3 license. 4 | 5 | # Copyright Owners 6 | 7 | - My own code:\ 8 | Copyright (c) 2012 Favonia 9 | 10 | - [Nils' library](http://www.cse.chalmers.se/~nad/repos/equality/)\ 11 | Copyright (c) 2011-2012 Nils Anders Danielsson 12 | 13 | - [Agda stardard library](http://www.cse.chalmers.se/~nad/repos/lib/)\ 14 | Copyright (c) 2007-2012 Nils Anders Danielsson, Ulf Norell, Shin-Cheng 15 | Mu, Samuel Bronson, Dan Doel, Patrik Jansson, Liang-Ting Chen, 16 | Jean-Philippe Bernardy, Andrés Sicard-Ramírez, Nicolas Pouillard, 17 | Darin Morrison, Peter Berry, Daniel Brown, Simon Foster, Dominique 18 | Devriese, Andreas Abel, Alcatel-Lucent, Eric Mertens, Joachim 19 | Breitner, Liyang Hu, Noam Zeilberger, Érdi Gergő, Stevan Andjelkovic 20 | 21 | # BSD3 License Text 22 | 23 | ~~~~ 24 | Permission is hereby granted, free of charge, to any person obtaining a 25 | copy of this software and associated documentation files (the 26 | "Software"), to deal in the Software without restriction, including 27 | without limitation the rights to use, copy, modify, merge, publish, 28 | distribute, sublicense, and/or sell copies of the Software, and to 29 | permit persons to whom the Software is furnished to do so, subject to 30 | the following conditions: 31 | 32 | The above copyright notice and this permission notice shall be included 33 | in all copies or substantial portions of the Software. 34 | 35 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS 36 | OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF 37 | MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND 38 | NONINFRINGEMENT. IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE 39 | LIABLE FOR ANY CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER IN AN ACTION 40 | OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION 41 | WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. 42 | ~~~~ 43 | -------------------------------------------------------------------------------- /Map/H-equivalence.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Homotopy equivalence 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | module Map.H-equivalence where 11 | 12 | open import Prelude as P hiding (id) renaming (_∘_ to _⊚_) 13 | open import Path 14 | 15 | open import Map.Injection using (Injective; _↣_) 16 | open import Map.Surjection as Surjection using (_↠_; module _↠_) 17 | 18 | ------------------------------------------------------------------------ 19 | -- Homotopy equivalence 20 | 21 | infix 0 _↔_ 22 | 23 | record _↔_ {f t} (From : Set f) (To : Set t) : Set (f ⊔ t) where 24 | field 25 | surjection : From ↠ To 26 | 27 | open _↠_ surjection 28 | 29 | field 30 | left-inverse-of : ∀ x → from (to x) ≡ x 31 | 32 | injective : Injective to 33 | injective {x = x} {y = y} to-x≡to-y = 34 | x ≡⟨ sym (left-inverse-of x) ⟩ 35 | from (to x) ≡⟨ cong from to-x≡to-y ⟩ 36 | from (to y) ≡⟨ left-inverse-of y ⟩∎ 37 | y ∎ 38 | 39 | injection : From ↣ To 40 | injection = record 41 | { to = to 42 | ; injective = injective 43 | } 44 | 45 | -- A lemma. 46 | 47 | to-from : ∀ {x y} → to x ≡ y → from y ≡ x 48 | to-from {x} {y} to-x≡y = 49 | from y ≡⟨ cong from $ sym to-x≡y ⟩ 50 | from (to x) ≡⟨ left-inverse-of x ⟩∎ 51 | x ∎ 52 | 53 | open _↠_ surjection public 54 | 55 | ------------------------------------------------------------------------ 56 | -- Equivalence 57 | 58 | -- _↔_ is an equivalence relation. 59 | 60 | id : ∀ {a} {A : Set a} → A ↔ A 61 | id = record 62 | { surjection = Surjection.id 63 | ; left-inverse-of = refl 64 | } 65 | 66 | inverse : ∀ {a b} {A : Set a} {B : Set b} → A ↔ B → B ↔ A 67 | inverse A↔B = record 68 | { surjection = record 69 | { from = to 70 | ; to = from 71 | ; right-inverse-of = left-inverse-of 72 | } 73 | ; left-inverse-of = right-inverse-of 74 | } where open _↔_ A↔B 75 | 76 | infixr 9 _∘_ 77 | 78 | _∘_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 79 | B ↔ C → A ↔ B → A ↔ C 80 | f ∘ g = record 81 | { surjection = Surjection._∘_ (surjection f) (surjection g) 82 | ; left-inverse-of = from∘to 83 | } 84 | where 85 | open _↔_ 86 | 87 | abstract 88 | from∘to : ∀ x → from g (from f (to f (to g x))) ≡ x 89 | from∘to = λ x → 90 | from g (from f (to f (to g x))) ≡⟨ cong (from g) (left-inverse-of f (to g x)) ⟩ 91 | from g (to g x) ≡⟨ left-inverse-of g x ⟩∎ 92 | x ∎ 93 | 94 | -- "Equational" reasoning combinators. 95 | 96 | infix 0 finally-↔ 97 | infixr 0 _↔⟨_⟩_ 98 | 99 | _↔⟨_⟩_ : ∀ {a b c} (A : Set a) {B : Set b} {C : Set c} → 100 | A ↔ B → B ↔ C → A ↔ C 101 | _ ↔⟨ A↔B ⟩ B↔C = B↔C ∘ A↔B 102 | 103 | finally-↔ : ∀ {a b} (A : Set a) (B : Set b) → A ↔ B → A ↔ B 104 | finally-↔ _ _ A↔B = A↔B 105 | 106 | syntax finally-↔ A B A↔B = A ↔⟨ A↔B ⟩□ B □ 107 | -------------------------------------------------------------------------------- /Map/H-fiber.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Homotopy Fibers 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | -- Partly based on Voevodsky's work on univalent foundations. 11 | 12 | module Map.H-fiber where 13 | 14 | open import Prelude 15 | open import Path 16 | 17 | open import Map.Surjection hiding (id; _∘_) 18 | open import Map.H-equivalence hiding (id; _∘_) 19 | 20 | -- The homotopy fiber of y under f is denoted by f ⁻¹ y. 21 | 22 | infix 5 _⁻¹_ 23 | 24 | _⁻¹_ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → B → Set (a ⊔ b) 25 | f ⁻¹ y = ∃ λ x → f x ≡ y 26 | 27 | -- This is subject to changes 28 | id⁻¹-contractible : ∀ {ℓ} {A : Set ℓ} (x : A) → Contractible (id ⁻¹ x) 29 | id⁻¹-contractible y = (y , refl y) , λ {(_ , p) → elim″ (λ {x} p → (y , refl y) ≡ (x , p)) (refl _) p} 30 | 31 | postulate 32 | hequiv⁻¹-contractible : 33 | ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} (A↔B : A ↔ B) → 34 | let open _↔_ A↔B in ∀ y → Contractible (to ⁻¹ y) 35 | {- 36 | hequiv⁻¹-contractible A↔B y = 37 | (from y , right-inverse-of y) , 38 | λ {(x , to-x≡y) → elim′ (λ {x} p → (from y , right-inverse-of y) ≡ (x , p)) (refl _) to-x≡y} 39 | where 40 | open _↔_ A↔B 41 | 42 | lemma : ∀ x → (from (to x) , right-inverse-of (to x)) ≡ (x , refl (to x)) 43 | -} 44 | -------------------------------------------------------------------------------- /Map/Injection.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Injections 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | module Map.Injection where 11 | 12 | open import Prelude as P hiding (id) renaming (_∘_ to _⊚_) 13 | open import Path 14 | 15 | ------------------------------------------------------------------------ 16 | -- Injections 17 | 18 | -- The property of being injective. 19 | 20 | Injective : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Set (a ⊔ b) 21 | Injective f = ∀ {x y} → f x ≡ f y → x ≡ y 22 | 23 | infix 0 _↣_ 24 | 25 | -- Injections. 26 | 27 | record _↣_ {f t} (From : Set f) (To : Set t) : Set (f ⊔ t) where 28 | field 29 | to : From → To 30 | injective : Injective to 31 | 32 | ------------------------------------------------------------------------ 33 | -- Preorder 34 | 35 | -- _↣_ is a preorder. 36 | 37 | id : ∀ {a} {A : Set a} → A ↣ A 38 | id = record 39 | { to = P.id 40 | ; injective = P.id 41 | } 42 | 43 | infixr 9 _∘_ 44 | 45 | _∘_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 46 | B ↣ C → A ↣ B → A ↣ C 47 | f ∘ g = record 48 | { to = to′ 49 | ; injective = injective′ 50 | } 51 | where 52 | open _↣_ 53 | 54 | to′ = to f ⊚ to g 55 | 56 | abstract 57 | injective′ : Injective to′ 58 | injective′ = injective g ⊚ injective f 59 | 60 | -- "Equational" reasoning combinators. 61 | 62 | infix 0 finally-↣ 63 | infixr 0 _↣⟨_⟩_ 64 | 65 | _↣⟨_⟩_ : ∀ {a b c} (A : Set a) {B : Set b} {C : Set c} → 66 | A ↣ B → B ↣ C → A ↣ C 67 | _ ↣⟨ A↣B ⟩ B↣C = B↣C ∘ A↣B 68 | 69 | finally-↣ : ∀ {a b} (A : Set a) (B : Set b) → A ↣ B → A ↣ B 70 | finally-↣ _ _ A↣B = A↣B 71 | 72 | syntax finally-↣ A B A↣B = A ↣⟨ A↣B ⟩□ B □ 73 | -------------------------------------------------------------------------------- /Map/Surjection.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Surjections 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | module Map.Surjection where 11 | 12 | open import Prelude as P hiding (id) renaming (_∘_ to _⊚_) 13 | open import Path 14 | 15 | ------------------------------------------------------------------------ 16 | -- Surjections 17 | 18 | infix 0 _↠_ 19 | 20 | -- Surjections. 21 | 22 | record _↠_ {f t} (From : Set f) (To : Set t) : Set (f ⊔ t) where 23 | field 24 | to : From → To 25 | from : To → From 26 | right-inverse-of : ∀ x → to (from x) ≡ x 27 | 28 | -- A lemma. 29 | 30 | from-to : ∀ {x y} → from x ≡ y → to y ≡ x 31 | from-to {x} {y} from-x≡y = 32 | to y ≡⟨ cong to $ sym from-x≡y ⟩ 33 | to (from x) ≡⟨ right-inverse-of x ⟩∎ 34 | x ∎ 35 | 36 | ------------------------------------------------------------------------ 37 | -- Preorder 38 | 39 | -- _↠_ is a preorder. 40 | 41 | id : ∀ {a} {A : Set a} → A ↠ A 42 | id = record 43 | { to = P.id 44 | ; from = P.id 45 | ; right-inverse-of = refl 46 | } 47 | 48 | infixr 9 _∘_ 49 | 50 | _∘_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 51 | B ↠ C → A ↠ B → A ↠ C 52 | f ∘ g = record 53 | { to = to f ⊚ to g 54 | ; from = from g ⊚ from f 55 | ; right-inverse-of = to∘from 56 | } 57 | where 58 | open _↠_ 59 | 60 | abstract 61 | to∘from : ∀ x → to f (to g (from g (from f x))) ≡ x 62 | to∘from = λ x → 63 | to f (to g (from g (from f x))) ≡⟨ cong (to f) (right-inverse-of g (from f x)) ⟩ 64 | to f (from f x) ≡⟨ right-inverse-of f x ⟩∎ 65 | x ∎ 66 | 67 | -- "Equational" reasoning combinators. 68 | 69 | infix 0 finally-↠ 70 | infixr 0 _↠⟨_⟩_ 71 | 72 | _↠⟨_⟩_ : ∀ {a b c} (A : Set a) {B : Set b} {C : Set c} → 73 | A ↠ B → B ↠ C → A ↠ C 74 | _ ↠⟨ A↠B ⟩ B↠C = B↠C ∘ A↠B 75 | 76 | finally-↠ : ∀ {a b} (A : Set a) (B : Set b) → A ↠ B → A ↠ B 77 | finally-↠ _ _ A↠B = A↠B 78 | 79 | syntax finally-↠ A B A↠B = A ↠⟨ A↠B ⟩□ B □ 80 | 81 | -------------------------------------------------------------------------------- /Map/WeakEquivalence.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Weak equivalences 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | -- Partly based on Voevodsky's work on univalent foundations. 11 | 12 | module Map.WeakEquivalence where 13 | 14 | open import Prelude as P hiding (id) renaming (_∘_ to _⊚_) 15 | open import Path 16 | 17 | open import Map.Surjection using (_↠_; module _↠_) 18 | open import Map.H-equivalence as H-equiv hiding (id; _∘_; inverse) 19 | open import Map.H-fiber as H-fiber 20 | 21 | ------------------------------------------------------------------------ 22 | -- Is-weak-equivalence 23 | 24 | -- A function f is a weak equivalence if all h-fibers under f are 25 | -- contractible. 26 | 27 | Is-weak-equivalence : ∀ {a b} {A : Set a} {B : Set b} → 28 | (A → B) → Set (a ⊔ b) 29 | Is-weak-equivalence f = ∀ y → Contractible (f ⁻¹ y) 30 | 31 | ------------------------------------------------------------------------ 32 | -- _≈_ 33 | 34 | -- Weak equivalences. 35 | 36 | infix 4 _≈_ 37 | 38 | record _≈_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where 39 | constructor weq 40 | field 41 | to : A → B 42 | is-weak-equivalence : Is-weak-equivalence to 43 | 44 | -- Weakly equivalent sets are isomorphic. 45 | 46 | from : B → A 47 | from = proj₁ ⊚ proj₁ ⊚ is-weak-equivalence 48 | 49 | abstract 50 | right-inverse-of : ∀ x → to (from x) ≡ x 51 | right-inverse-of = proj₂ ⊚ proj₁ ⊚ is-weak-equivalence 52 | 53 | left-inverse-of : ∀ x → from (to x) ≡ x 54 | left-inverse-of = λ x → 55 | cong (proj₁ {B = λ x′ → to x′ ≡ to x}) $ 56 | proj₂ (is-weak-equivalence (to x)) (x , refl (to x)) 57 | 58 | hequivalence : A ↔ B 59 | hequivalence = record 60 | { surjection = record 61 | { to = to 62 | ; from = from 63 | ; right-inverse-of = right-inverse-of 64 | } 65 | ; left-inverse-of = left-inverse-of 66 | } 67 | 68 | open _↔_ hequivalence public 69 | hiding (from; to; right-inverse-of; left-inverse-of) 70 | 71 | abstract 72 | 73 | -- All homotopy fibers of an element under the weak equivalence are 74 | -- equal. 75 | 76 | irrelevance : ∀ y (p : to ⁻¹ y) → (from y , right-inverse-of y) ≡ p 77 | irrelevance = proj₂ ⊚ is-weak-equivalence 78 | 79 | -- The two proofs left-inverse-of and right-inverse-of are 80 | -- related. (I.e., zigzag identity for adjoint equivalences.) 81 | 82 | left-right-lemma : 83 | ∀ x → cong to (left-inverse-of x) ≡ right-inverse-of (to x) 84 | left-right-lemma x = 85 | lemma₁ to (left-inverse-of x) _ (lemma₂ (irrelevance (to x) (x , refl (to x)))) 86 | where 87 | lemma₁ : {x y : A} (f : A → B) (p : x ≡ y) (q : f x ≡ f y) → 88 | refl (f y) ≡ trans (cong f (sym p)) q → 89 | cong f p ≡ q 90 | lemma₁ f = elim 91 | (λ {x y} p → ∀ q → refl (f y) ≡ trans (cong f (sym p)) q → 92 | cong f p ≡ q) 93 | (λ x q hyp → 94 | cong f (refl x) ≡⟨ refl _ ⟩ 95 | refl (f x) ≡⟨ hyp ⟩∎ 96 | -- trans (cong f (sym (refl x))) q ≡⟨ refl _ ⟩ 97 | -- trans (cong f (refl x)) q ≡⟨ refl _ ⟩ 98 | -- trans (refl (f x)) q ≡⟨ refl _ ⟩∎ 99 | q ∎) 100 | 101 | lemma₂ : ∀ {f : A → B} {y} {f⁻¹y₁ f⁻¹y₂ : f ⁻¹ y} 102 | (p : f⁻¹y₁ ≡ f⁻¹y₂) → 103 | proj₂ f⁻¹y₂ ≡ 104 | trans (cong f (sym (cong (proj₁ {B = λ x → f x ≡ y}) p))) 105 | (proj₂ f⁻¹y₁) 106 | lemma₂ {f} {y} = 107 | let pr = proj₁ {B = λ x → f x ≡ y} in 108 | elim {A = f ⁻¹ y} 109 | (λ {f⁻¹y₁ f⁻¹y₂} p → 110 | proj₂ f⁻¹y₂ ≡ 111 | trans (cong f (sym (cong pr p))) (proj₂ f⁻¹y₁)) 112 | (λ f⁻¹y → 113 | proj₂ f⁻¹y ≡⟨ refl _ ⟩ 114 | trans (refl (f (proj₁ f⁻¹y))) (proj₂ f⁻¹y) ≡⟨ refl _ ⟩ 115 | trans (cong f (refl (proj₁ f⁻¹y))) (proj₂ f⁻¹y) ≡⟨ refl _ ⟩ 116 | trans (cong f (sym (refl (proj₁ f⁻¹y)))) (proj₂ f⁻¹y) ≡⟨ refl _ ⟩∎ 117 | trans (cong f (sym (cong pr (refl f⁻¹y)))) (proj₂ f⁻¹y) ∎) 118 | 119 | -- Homotopy equivalences are weak equivalences. 120 | 121 | ↔⇒≈ : ∀ {a b} {A : Set a} {B : Set b} → A ↔ B → A ≈ B 122 | ↔⇒≈ A↔B = record 123 | { to = _↔_.to A↔B 124 | ; is-weak-equivalence = H-fiber.hequiv⁻¹-contractible A↔B 125 | } 126 | 127 | ------------------------------------------------------------------------ 128 | -- Equivalence 129 | 130 | -- Weak equivalences are equivalence relations. 131 | 132 | -- This is subject to changes 133 | id : ∀ {a} {A : Set a} → A ≈ A 134 | -- id = ↔⇒≈ H-equiv.id 135 | id = record 136 | { to = P.id 137 | ; is-weak-equivalence = H-fiber.id⁻¹-contractible 138 | } 139 | 140 | inverse : ∀ {a b} {A : Set a} {B : Set b} → A ≈ B → B ≈ A 141 | inverse = ↔⇒≈ ⊚ H-equiv.inverse ⊚ _≈_.hequivalence 142 | 143 | infixr 9 _∘_ 144 | 145 | _∘_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 146 | B ≈ C → A ≈ B → A ≈ C 147 | f ∘ g = ↔⇒≈ $ H-equiv._∘_ (_≈_.hequivalence f) (_≈_.hequivalence g) 148 | 149 | -- Equational reasoning combinators. 150 | 151 | infixr 0 _≈⟨_⟩_ 152 | infix 0 finally-≈ 153 | 154 | _≈⟨_⟩_ : ∀ {a b c} (A : Set a) {B : Set b} {C : Set c} → 155 | A ≈ B → B ≈ C → A ≈ C 156 | _ ≈⟨ A≈B ⟩ B≈C = B≈C ∘ A≈B 157 | 158 | finally-≈ : ∀ {a b} (A : Set a) (B : Set b) → A ≈ B → A ≈ B 159 | finally-≈ _ _ A≈B = A≈B 160 | 161 | syntax finally-≈ A B A≈B = A ≈⟨ A≈B ⟩□ B □ 162 | -------------------------------------------------------------------------------- /Path.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Path 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | module Path where 11 | 12 | open import Prelude 13 | 14 | ------------------------------------------------------------------------ 15 | -- Formation and introduction 16 | 17 | private 18 | data Path′ {ℓ} {A : Set ℓ} (x : A) : A → Set ℓ where 19 | refl′ : Path′ x x 20 | 21 | infix 4 _≡_ 22 | 23 | _≡_ : ∀ {ℓ} {A : Set ℓ} (x : A) → A → Set ℓ 24 | _≡_ = Path′ 25 | 26 | Path : ∀ {ℓ} {A : Set ℓ} (x : A) → A → Set ℓ 27 | Path = Path′ 28 | 29 | refl : ∀ {ℓ} {A : Set ℓ} (x : A) → x ≡ x 30 | refl _ = refl′ 31 | 32 | ------------------------------------------------------------------------ 33 | -- Elimination and computation 34 | 35 | -- I think dependent pattern matching is fine, because it seems that 36 | -- we can construct another equality with its destructor exposed and 37 | -- show two equalities are equivalent. However the destructor is still 38 | -- hidden so that people need to write elim explicitly (at least for 39 | -- this datatype). 40 | 41 | elim : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (P : {x y : A} → x ≡ y → Set ℓ₂) → 42 | (∀ x → P (refl x)) → 43 | ∀ {x y} (x≡y : x ≡ y) → P x≡y 44 | elim P r refl′ = r _ 45 | 46 | ------------------------------------------------------------------------ 47 | -- Alternative elimination/computation rules 48 | 49 | elim′ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {x : A} (P : {y : A} → x ≡ y → Set ℓ₂) → 50 | P (refl x) → ∀ {y} (x≡y : x ≡ y) → P x≡y 51 | -- elim′ P r refl′ = r 52 | elim′ {ℓ₂ = ℓ₂} {A = A} P r p = 53 | elim 54 | (λ {x y} p → (P : {y : A} → x ≡ y → Set ℓ₂) → P (refl x) → P p) 55 | (λ x P r → r) 56 | p P r 57 | 58 | elim″ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {y : A} (P : {x : A} → x ≡ y → Set ℓ₂) → 59 | P (refl y) → ∀ {x} (x≡y : x ≡ y) → P x≡y 60 | -- elim′ P r refl′ = r 61 | elim″ {ℓ₂ = ℓ₂} {A = A} P r p = 62 | elim 63 | (λ {x y} p → (P : {x : A} → x ≡ y → Set ℓ₂) → P (refl y) → P p) 64 | (λ x P r → r) 65 | p P r 66 | 67 | ------------------------------------------------------------------------ 68 | -- Congruence (respect or map) and substitutivity (tranport) 69 | 70 | cong : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} 71 | (f : A → B) {x y : A} → x ≡ y → f x ≡ f y 72 | cong f = elim (λ {u v} _ → f u ≡ f v) (λ x → refl (f x)) 73 | 74 | subst : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (P : A → Set ℓ₂) {x y : A} → 75 | x ≡ y → P x → P y 76 | subst P = elim (λ {u v} _ → P u → P v) (λ x p → p) 77 | 78 | ------------------------------------------------------------------------ 79 | -- Transitivity and symmetry 80 | 81 | -- Here we makes "trans (refl _) p" definitionally equal to "p". 82 | -- The reason is that we usually need to deal with "trans (refl (trans ...))" 83 | -- in a complex proof of equivalence between paths. 84 | -- (This is different from the intension of Nils' original code.) 85 | 86 | trans : ∀ {ℓ} {A : Set ℓ} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 87 | trans {x = x} {y} {z} x≡y = 88 | elim 89 | (λ {x y} x≡y → y ≡ z → x ≡ z) 90 | (λ _ → id) 91 | x≡y 92 | 93 | sym : ∀ {ℓ} {A : Set ℓ} {x y : A} → x ≡ y → y ≡ x 94 | sym {x = x} x≡y = subst (λ z → x ≡ z → z ≡ x) x≡y id x≡y 95 | 96 | ------------------------------------------------------------------------ 97 | -- A must-have for beautiful proofs 98 | 99 | infix 0 finally 100 | infixr 0 _≡⟨_⟩_ 101 | 102 | _≡⟨_⟩_ : ∀ {a} {A : Set a} x {y z : A} → x ≡ y → y ≡ z → x ≡ z 103 | _ ≡⟨ x≡y ⟩ y≡z = trans x≡y y≡z 104 | 105 | finally : ∀ {a} {A : Set a} (x y : A) → x ≡ y → x ≡ y 106 | finally _ _ x≡y = x≡y 107 | 108 | syntax finally x y x≡y = x ≡⟨ x≡y ⟩∎ y ∎ 109 | 110 | ------------------------------------------------------------------------ 111 | -- Some terminologies 112 | 113 | -- A type is contractible if it is inhabited and all elements are 114 | -- equal. 115 | Contractible : ∀ {ℓ} → Set ℓ → Set ℓ 116 | Contractible A = ∃ λ (x : A) → ∀ y → x ≡ y 117 | 118 | -- Singleton x is a set which contains all elements which are equal 119 | -- to x. 120 | Singleton : ∀ {ℓ} {A : Set ℓ} → A → Set ℓ 121 | Singleton x = ∃ λ y → y ≡ x 122 | 123 | -- Extensionality for (non-dependent) functions of a certain type. 124 | Extensionality : ∀ {ℓ₁ ℓ₂} → Set ℓ₁ → Set ℓ₂ → Set (ℓ₁ ⊔ ℓ₂) 125 | Extensionality A B = {f g : A → B} → (∀ x → f x ≡ g x) → f ≡ g 126 | 127 | -- Extensionality for dependent functions of a certain type. 128 | Extensionality[dep] : ∀ {ℓ₁ ℓ₂} (A : Set ℓ₁) → (A → Set ℓ₂) → Set (ℓ₁ ⊔ ℓ₂) 129 | Extensionality[dep] A B = {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g 130 | -------------------------------------------------------------------------------- /Path/HigherOrder.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Higher-order paths and loops 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | module Path.HigherOrder where 8 | 9 | open import Prelude 10 | open import Path 11 | open import Path.Lemmas 12 | open import Space.Nat.Lemmas 13 | 14 | -- TODO Generalized cong[dep]-const lemma 15 | 16 | ------------------------------------------------------------------------ 17 | -- Higher-order loop spaces 18 | 19 | Ω : ∀ {ℓ} n {A : Set ℓ} → A → Set ℓ 20 | private Ω-refl : ∀ {ℓ} n {A : Set ℓ} (base : A) → Ω n base 21 | 22 | Ω 0 {A} _ = A 23 | Ω (suc n) {A} base = Ω-refl n base ≡ Ω-refl n base 24 | 25 | Ω-refl 0 base = base 26 | Ω-refl (suc n) base = refl (Ω-refl n base) 27 | 28 | ------------------------------------------------------------------------ 29 | -- Higher-order path spaces 30 | 31 | -- Endpoints[d] 0 = tt 32 | -- Endpoints[d] 1 = (tt, x₁ , y₁) 33 | -- Endpoints[d] 2 = ((tt, x₁ , y₂) , x₂ , y₂) 34 | 35 | -- This definition is perfect for induction but probably not user-friendly 36 | Endpoints⇑ : ∀ {ℓ} n → Set ℓ → Set ℓ 37 | Path⇑ : ∀ {ℓ} n {A : Set ℓ} → Endpoints⇑ n A → Set ℓ 38 | 39 | Endpoints⇑ {ℓ} 0 A = ↑ ℓ ⊤ 40 | Endpoints⇑ (suc n) A = Σ (Endpoints⇑ n A) (λ c → Path⇑ n c × Path⇑ n c) 41 | 42 | Path⇑ 0 {A} _ = A 43 | Path⇑ (suc n) (_ , x , y) = x ≡ y 44 | 45 | -- This definition is terrible for induction, but seems nicer to use? 46 | Endpoints′⇑ : ∀ {ℓ} n → Set ℓ → Set ℓ 47 | Endpoints′⇑ {ℓ} 0 A = ↑ ℓ ⊤ 48 | Endpoints′⇑ (suc n) A = Σ A (λ x → Σ A ( λ y → Endpoints′⇑ n (x ≡ y))) 49 | 50 | -- A helper function to convert the above definition 51 | endpoints′⇒endpoints⇑ : ∀ {ℓ} n m {A : Set ℓ} (eA : Endpoints⇑ n A) → Endpoints′⇑ m (Path⇑ n eA) → Endpoints⇑ (m + n) A 52 | endpoints′⇒endpoints⇑ n 0 c _ = c 53 | endpoints′⇒endpoints⇑ n (suc m) {A} c (x , y , p) = 54 | subst (λ n → Endpoints⇑ n A) (n+suc m n) $ endpoints′⇒endpoints⇑ (suc n) m (_ , (x , y)) p 55 | 56 | -- A section that parametrized by a high-order path 57 | -- Endpoints[d] 0 = tt 58 | -- Endpoints[d] 1 = (tt, bx₁ , by₁) 59 | -- Endpoints[d] 2 = ((tt, bx₁ , by₂) , bx₂ , by₂) 60 | Endpoints[dep]⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A → Set ℓ₂) → Endpoints⇑ n A → Set ℓ₂ 61 | Path[dep]⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A → Set ℓ₂) {eA : Endpoints⇑ n A} → Path⇑ n eA → Endpoints[dep]⇑ n B eA → Set ℓ₂ 62 | 63 | Endpoints[dep]⇑ {ℓ₁} {ℓ₂} 0 _ _ = ↑ ℓ₂ ⊤ 64 | Endpoints[dep]⇑ (suc n) B (eA , x , y) = Σ (Endpoints[dep]⇑ n B eA) 65 | (λ c → Path[dep]⇑ n B x c × Path[dep]⇑ n B y c) 66 | 67 | Path[dep]⇑ 0 B a _ = B a 68 | Path[dep]⇑ (suc n) B a (eB , bx , by) = subst (λ x → Path[dep]⇑ n B x eB) a bx ≡ by 69 | 70 | ------------------------------------------------------------------------ 71 | -- cong (resp, map) for higher-order path 72 | 73 | -- Higher-order cong(ruence) that takes high-order paths 74 | -- Reuse Path[dep] to avoid all sorts of problems in Space.Sphere 75 | cong-endpoints[dep]⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A → Set ℓ₂) (f : (x : A) → B x) 76 | (eA : Endpoints⇑ n A) → Endpoints[dep]⇑ n B eA 77 | cong[dep]⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A → Set ℓ₂) (f : (x : A) → B x) 78 | (eA : Endpoints⇑ n A) (pA : Path⇑ n eA) → Path[dep]⇑ n B pA (cong-endpoints[dep]⇑ n B f eA) 79 | 80 | cong-endpoints[dep]⇑ 0 B f _ = lift tt 81 | cong-endpoints[dep]⇑ (suc n) B f (eA , x , y) = (cong-endpoints[dep]⇑ n B f eA , 82 | cong[dep]⇑ n B f eA x , 83 | cong[dep]⇑ n B f eA y) 84 | 85 | cong[dep]⇑ 0 B f t p = f p 86 | cong[dep]⇑ (suc n) B f (eA , (x , y)) p = cong[dep] (λ x → Path[dep]⇑ n B x (cong-endpoints[dep]⇑ n B f eA)) (cong[dep]⇑ n B f eA) p 87 | 88 | -- Higher-order non-dependent cong(ruence) that takes high-order paths 89 | cong-endpoints⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) 90 | (eA : Endpoints⇑ n A) → Endpoints⇑ n B 91 | cong⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) 92 | (eA : Endpoints⇑ n A) (pA : Path⇑ n eA) → Path⇑ n (cong-endpoints⇑ n f eA) 93 | 94 | cong-endpoints⇑ 0 f _ = lift tt 95 | cong-endpoints⇑ (suc n) f (eA , x , y) = (cong-endpoints⇑ n f eA , 96 | cong⇑ n f eA x , 97 | cong⇑ n f eA y) 98 | 99 | cong⇑ 0 f t p = f p 100 | cong⇑ (suc n) f (eA , (x , y)) p = cong (cong⇑ n f eA) p 101 | 102 | -- Generalized cong[dep]-const 103 | 104 | cong-endpoints[dep]-const⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) (eA : Endpoints⇑ n A) (pA : Path⇑ n eA) → 105 | Path[dep]⇑ n (const B) pA (cong-endpoints[dep]⇑ n (const B) f eA) ≡ Path⇑ n (cong-endpoints⇑ n f eA) 106 | 107 | cong[dep]-const⇑ : ∀ {ℓ₁ ℓ₂} n {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) (eA : Endpoints⇑ n A) (pA : Path⇑ n eA) → 108 | subst id (cong-endpoints[dep]-const⇑ n f eA pA) (cong[dep]⇑ n (const B) f eA pA) ≡ cong⇑ n f eA pA 109 | 110 | cong-endpoints[dep]-const⇑ 0 _ _ _ = refl _ 111 | cong-endpoints[dep]-const⇑ (suc n) {B = B} f (eA , x , y) pA = elim 112 | (λ {x y} pA → 113 | Path[dep]⇑ (suc n) (const B) pA (cong-endpoints[dep]⇑ (suc n) (const B) f (eA , x , y)) ≡ 114 | Path⇑ (suc n) (cong-endpoints⇑ (suc n) f (eA , x , y))) 115 | (λ x → cong-≡ (cong-endpoints[dep]-const⇑ n f eA x) (cong[dep]-const⇑ n f eA x) (cong[dep]-const⇑ n f eA x)) 116 | pA 117 | 118 | cong[dep]-const⇑ 0 _ _ _ = refl _ 119 | cong[dep]-const⇑ (suc n) {B = B} f (eA , x , y) pA = elim 120 | (λ {x y} pA → 121 | subst id (cong-endpoints[dep]-const⇑ (suc n) f (eA , x , y) pA) (cong[dep]⇑ (suc n) (const B) f (eA , x , y) pA) ≡ 122 | cong⇑ (suc n) f (eA , x , y) pA) 123 | (λ x → subst-cong-≡ (cong-endpoints[dep]-const⇑ n f eA x) (cong[dep]-const⇑ n f eA x)) 124 | pA 125 | -------------------------------------------------------------------------------- /Path/Lemmas.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Lemmas for equality 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | open import Prelude 11 | open import Path 12 | 13 | module Path.Lemmas where 14 | 15 | ------------------------------------------------------------------------ 16 | -- Tons of lemmas to handle paths 17 | 18 | trans-reflʳ : ∀ {ℓ} {A : Set ℓ} {x y : A} (x≡y : x ≡ y) → 19 | trans x≡y (refl y) ≡ x≡y 20 | trans-reflʳ = 21 | elim 22 | (λ {x y} x≡y → trans x≡y (refl y) ≡ x≡y) 23 | (λ x → refl _) 24 | 25 | {- 26 | -- definitional 27 | trans-reflˡ : ∀ {ℓ} {A : Set ℓ} {x y : A} (x≡y : x ≡ y) → 28 | trans (refl x) x≡y ≡ x≡y 29 | -} 30 | 31 | sym-sym : ∀ {ℓ} {A : Set ℓ} {x y : A} (x≡y : x ≡ y) → 32 | sym (sym x≡y) ≡ x≡y 33 | sym-sym = elim (λ {u v} u≡v → sym (sym u≡v) ≡ u≡v) 34 | (λ x → refl _) 35 | 36 | sym-trans : ∀ {a} {A : Set a} {x y z : A} 37 | (x≡y : x ≡ y) (y≡z : y ≡ z) → 38 | sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y) 39 | sym-trans {a} = 40 | elim (λ x≡y → ∀ y≡z → sym (trans x≡y y≡z) ≡ trans (sym y≡z) (sym x≡y)) 41 | (λ y y≡z → sym (trans (refl y) y≡z) ≡⟨ refl _ ⟩ 42 | sym y≡z ≡⟨ sym $ trans-reflʳ _ ⟩∎ 43 | trans (sym y≡z) (sym (refl y)) ∎) 44 | 45 | cong₂ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} 46 | (f : A → B → C) {x y : A} {u v : B} → 47 | x ≡ y → u ≡ v → f x u ≡ f y v 48 | cong₂ f {x} {y} {u} {v} x≡y u≡v = 49 | f x u ≡⟨ cong (flip f u) x≡y ⟩ 50 | f y u ≡⟨ cong (f y) u≡v ⟩∎ 51 | f y v ∎ 52 | 53 | cong₂′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} 54 | (f : A → B → C) {x y : A} {u v : B} → 55 | x ≡ y → u ≡ v → f x u ≡ f y v 56 | cong₂′ f {x} {y} {u} {v} x≡y u≡v = 57 | f x u ≡⟨ cong (f x) u≡v ⟩ 58 | f x v ≡⟨ cong (flip f v) x≡y ⟩∎ 59 | f y v ∎ 60 | 61 | cong-id : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) → 62 | cong id x≡y ≡ x≡y 63 | cong-id = elim (λ {x y} x≡y → cong id x≡y ≡ x≡y) 64 | (λ x → cong id (refl x) ≡⟨ refl _ ⟩∎ 65 | refl x ∎) 66 | 67 | cong-cong : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} {x y : A} 68 | (f : B → C) (g : A → B) (x≡y : x ≡ y) → 69 | cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y 70 | cong-cong f g = 71 | elim 72 | (λ x≡y → cong f (cong g x≡y) ≡ cong (f ∘ g) x≡y) 73 | (λ x → refl _) 74 | 75 | cong-sym : ∀ {a b} {A : Set a} {B : Set b} {x y : A} 76 | (f : A → B) (x≡y : x ≡ y) → 77 | cong f (sym x≡y) ≡ sym (cong f x≡y) 78 | cong-sym f = elim (λ x≡y → cong f (sym x≡y) ≡ sym (cong f x≡y)) 79 | (λ x → refl _) 80 | 81 | cong-trans : ∀ {a b} {A : Set a} {B : Set b} {x y z : A} 82 | (f : A → B) (x≡y : x ≡ y) (y≡z : y ≡ z) → 83 | cong f (trans x≡y y≡z) ≡ trans (cong f x≡y) (cong f y≡z) 84 | cong-trans f = 85 | elim (λ x≡y → ∀ y≡z → cong f (trans x≡y y≡z) ≡ 86 | trans (cong f x≡y) (cong f y≡z)) 87 | (λ y y≡z → refl _) 88 | 89 | trans-assoc : ∀ {a} {A : Set a} {x y z u : A} 90 | (x≡y : x ≡ y) (y≡z : y ≡ z) (z≡u : z ≡ u) → 91 | trans (trans x≡y y≡z) z≡u ≡ trans x≡y (trans y≡z z≡u) 92 | trans-assoc = 93 | elim (λ x≡y → ∀ y≡z z≡u → trans (trans x≡y y≡z) z≡u ≡ 94 | trans x≡y (trans y≡z z≡u)) 95 | (λ y y≡z z≡u → refl _) 96 | 97 | -- Derived² lemmas 98 | 99 | trans-symˡ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) → 100 | trans (sym x≡y) x≡y ≡ refl y 101 | trans-symˡ = 102 | elim 103 | (λ {x y} (x≡y : x ≡ y) → trans (sym x≡y) x≡y ≡ refl y) 104 | (λ x → refl _) 105 | 106 | trans-symʳ : ∀ {a} {A : Set a} {x y : A} (x≡y : x ≡ y) → 107 | trans x≡y (sym x≡y) ≡ refl x 108 | trans-symʳ = 109 | elim 110 | (λ {x y} p → trans p (sym p) ≡ refl x) 111 | (λ x → refl _) 112 | 113 | cong[dep] : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) (f : (x : A) → B x) 114 | {x y : A} (p : x ≡ y) → subst B p (f x) ≡ (f y) 115 | cong[dep] B f = 116 | elim 117 | (λ {x y} p → subst B p (f x) ≡ (f y) ) 118 | (λ x → refl _) 119 | 120 | cong[dep]₂ : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : A → Set ℓ₃) 121 | (f : (x : A) → B x → C x) {a₁ a₂ : A} (p : a₁ ≡ a₂) {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → 122 | subst C p (f a₁ b₁) ≡ f a₂ b₂ 123 | cong[dep]₂ B C f = 124 | elim 125 | (λ {a₁ a₂} p → {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → subst C p (f a₁ b₁) ≡ f a₂ b₂) 126 | (λ a q → cong (f a) q) 127 | 128 | cong-const : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} (b : B) {x y : A} (p : x ≡ y) → cong (const b) p ≡ refl b 129 | cong-const {B = B} b p = 130 | elim 131 | (λ {_ _} p → cong (const b) p ≡ refl b ) 132 | (λ x → refl _) 133 | p 134 | 135 | subst-trans : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {x y : A} 136 | (p₁ : x ≡ y) {z : A} (p₂ : y ≡ z) (b : B x) → 137 | subst B (trans p₁ p₂) b ≡ subst B p₂ (subst B p₁ b) 138 | subst-trans {A = A} B = 139 | elim 140 | (λ {x y} p₁ → {z : A} (p₂ : y ≡ z) (b : B x) → 141 | subst B (trans p₁ p₂) b ≡ subst B p₂ (subst B p₁ b)) 142 | (λ x p₂ b → refl _) 143 | 144 | -- TODO a better interface? 145 | subst-sym : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) 146 | {x y : A} (p : x ≡ y) (by : B y) (bx : B x) → 147 | subst B p bx ≡ by → subst B (sym p) by ≡ bx 148 | subst-sym {A = A} B {x} {y} p by bx q = 149 | subst B (sym p) by ≡⟨ cong (subst B (sym p)) $ sym q ⟩ 150 | subst B (sym p) (subst B p bx) ≡⟨ sym $ subst-trans B p (sym p) bx ⟩ 151 | subst B (trans p (sym p)) bx ≡⟨ cong (λ p → subst B p bx) $ trans-symʳ p ⟩∎ 152 | bx ∎ 153 | 154 | subst-const : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} {x y : A} (p : x ≡ y) (b : B) → subst (const B) p b ≡ b 155 | subst-const {B = B} p b = 156 | elim 157 | (λ {_ _} p → subst (const B) p b ≡ b ) 158 | (λ x → refl _) 159 | p 160 | 161 | subst-cong : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} {B : Set ℓ₂} (C : B → Set ℓ₃) 162 | (f : A → B) {x y : A} (p : x ≡ y) (c : C (f x)) → 163 | subst C (cong f p) c ≡ subst (C ∘ f) p c 164 | subst-cong C f = 165 | elim 166 | (λ {x y} p → (c : C (f x)) → subst C (cong f p) c ≡ subst (C ∘ f) p c ) 167 | (λ x c → refl _) 168 | 169 | subst-app : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : A → Set ℓ₃) 170 | {x y : A} (p : x ≡ y) (f : B x → C x) {bx : B x} {by : B y} (q : subst B p bx ≡ by) → 171 | subst C p (f bx) ≡ subst (λ x → B x → C x) p f by 172 | subst-app {A = A} B C = 173 | elim 174 | (λ {x y} p → (f : B x → C x) {bx : B x} {by : B y} (q : subst B p bx ≡ by) → 175 | subst C p (f bx) ≡ (subst (λ x → B x → C x) p f) by) 176 | (λ x f → elim 177 | (λ {bx by} q → subst C (refl x) (f bx) ≡ (subst (λ x → B x → C x) (refl x) f) by) 178 | (λ b → refl _)) 179 | 180 | -- TODO a better interface? 181 | subst-func : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : A → Set ℓ₃) 182 | {x y : A} (p : x ≡ y) (f : B x → C x) {bx : B x} {by : B y} (q : subst B p bx ≡ by) → 183 | subst (λ x → B x → C x) p f by ≡ subst C p (f bx) 184 | subst-func B C p f q = sym $ subst-app B C p f q 185 | 186 | subst-path[idʳ] : ∀ {ℓ₁} {A : Set ℓ₁} (f : A → A) 187 | {x y : A} (p : x ≡ y) (q : f x ≡ x) → 188 | subst (λ x → f x ≡ x) p q ≡ trans (sym (cong f p)) (trans q p) 189 | subst-path[idʳ] f = 190 | elim 191 | (λ {x y} (p : x ≡ y) → (q : f x ≡ x) → 192 | subst (λ x → f x ≡ x) p q ≡ trans (sym (cong f p)) (trans q p)) 193 | (λ x q → 194 | subst (λ x → f x ≡ x) (refl _) q ≡⟨ refl _ ⟩ 195 | q ≡⟨ sym $ trans-reflʳ q ⟩∎ 196 | trans (sym (cong f (refl _))) (trans q (refl _)) ∎) 197 | 198 | subst-path[simp₂] : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} (f g : A → B) 199 | {x y : A} (p : x ≡ y) (q : f x ≡ g x) → 200 | subst (λ x → f x ≡ g x) p q ≡ trans (sym (cong f p)) (trans q (cong g p)) 201 | subst-path[simp₂] f g = 202 | elim 203 | (λ {x y} (p : x ≡ y) → (q : f x ≡ g x) → 204 | subst (λ x → f x ≡ g x) p q ≡ trans (sym (cong f p)) (trans q (cong g p))) 205 | (λ x q → 206 | subst (λ x → f x ≡ g x) (refl _) q ≡⟨ refl _ ⟩ 207 | q ≡⟨ sym $ trans-reflʳ q ⟩∎ 208 | trans (sym (cong f (refl _))) (trans q (cong g (refl _))) ∎) 209 | 210 | subst-path[constʳ] : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) (g′ : B) 211 | {x y : A} (p : x ≡ y) (q : f x ≡ g′) → 212 | subst (λ x → f x ≡ g′) p q ≡ trans (sym (cong f p)) q 213 | subst-path[constʳ] f g′ = 214 | elim 215 | (λ {x y} (p : x ≡ y) → (q : f x ≡ g′) → 216 | subst (λ x → f x ≡ g′) p q ≡ trans (sym (cong f p)) q) 217 | (λ x q → refl _) 218 | 219 | cong-≡ : ∀ {ℓ₁} {A B : Set ℓ₁} (up : A ≡ B) 220 | {a₁ : A} {b₁ : B} (p₁ : subst id up a₁ ≡ b₁) 221 | {a₂ : A} {b₂ : B} (p₂ : subst id up a₂ ≡ b₂) → 222 | (a₁ ≡ a₂) ≡ (b₁ ≡ b₂) 223 | cong-≡ = elim 224 | (λ {A B} (up : A ≡ B) → 225 | {a₁ : A} {b₁ : B} (p₁ : subst id up a₁ ≡ b₁) 226 | {a₂ : A} {b₂ : B} (p₂ : subst id up a₂ ≡ b₂) → 227 | (a₁ ≡ a₂) ≡ (b₁ ≡ b₂)) 228 | (λ A p₁ p₂ → cong₂ (_≡_) p₁ p₂) 229 | 230 | subst-cong-≡ : ∀ {ℓ₁} {A B : Set ℓ₁} (up : A ≡ B) 231 | {a : A} {b : B} (p : subst id up a ≡ b) → 232 | subst id (cong-≡ up p p) (refl a) ≡ refl b 233 | subst-cong-≡ = elim 234 | (λ {A B} up → 235 | {a : A} {b : B} (p : subst id up a ≡ b) → 236 | subst id (cong-≡ up p p) (refl a) ≡ refl b) 237 | (λ x → elim 238 | (λ {a b} p → 239 | subst id (cong-≡ (refl x) p p) (refl a) ≡ refl b) 240 | (λ ab → refl _)) 241 | 242 | -- Derived³ lemmas 243 | 244 | cong[dep]-const : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} (f : A → B) 245 | {x y : A} (p : x ≡ y) → cong[dep] (const B) f p ≡ (trans (subst-const p (f x)) (cong f p)) 246 | cong[dep]-const {B = B} f = 247 | elim 248 | (λ {x y} p → cong[dep] (const B) f p ≡ (trans (subst-const p (f x)) (cong f p)) ) 249 | (λ x → 250 | cong[dep] (const B) f (refl x) ≡⟨ refl _ ⟩ 251 | refl (f x) ≡⟨ refl _ ⟩ 252 | subst-const (refl x) (f x) ≡⟨ sym $ trans-reflʳ _ ⟩∎ 253 | trans (subst-const (refl x) (f x)) (cong f (refl _))∎) 254 | 255 | trans-sym-trans : ∀ {ℓ} {A : Set ℓ} {x y z : A} (p : y ≡ x) (q : x ≡ z) → trans (sym p) (trans p q) ≡ q 256 | trans-sym-trans p q = 257 | trans (sym p) (trans p q) ≡⟨ sym $ trans-assoc (sym p) _ _ ⟩ 258 | trans (trans (sym p) p) q ≡⟨ cong (λ p → trans p q) $ trans-symˡ p ⟩∎ 259 | q ∎ 260 | 261 | trans-transʳ-sym : ∀ {ℓ} {A : Set ℓ} {x y z : A} (p : x ≡ y) (q : x ≡ z) → trans p (trans (sym p) q) ≡ q 262 | trans-transʳ-sym p q = 263 | trans p (trans (sym p) q) ≡⟨ sym $ trans-assoc p _ _ ⟩ 264 | trans (trans p (sym p)) q ≡⟨ cong (λ p → trans p q) $ trans-symʳ p ⟩∎ 265 | q ∎ 266 | 267 | trans-trans-symʳ : ∀ {ℓ} {A : Set ℓ} {x y z : A} (q : x ≡ y) (p : z ≡ y) → trans (trans q (sym p)) p ≡ q 268 | trans-trans-symʳ q p = 269 | trans (trans q (sym p)) p ≡⟨ trans-assoc q _ _ ⟩ 270 | trans q (trans (sym p) p) ≡⟨ cong (trans q) $ trans-symˡ p ⟩ 271 | trans q (refl _) ≡⟨ trans-reflʳ _ ⟩∎ 272 | q ∎ 273 | 274 | trans-trans-symʳʳ : ∀ {ℓ} {A : Set ℓ} {x y z : A} (q : x ≡ y) (p : y ≡ z) → trans (trans q p) (sym p) ≡ q 275 | trans-trans-symʳʳ q p = 276 | trans (trans q p) (sym p) ≡⟨ trans-assoc q _ _ ⟩ 277 | trans q (trans p (sym p)) ≡⟨ cong (trans q) $ trans-symʳ p ⟩ 278 | trans q (refl _) ≡⟨ trans-reflʳ _ ⟩∎ 279 | q ∎ 280 | 281 | -------------------------------------------------------------------------------- /Path/Omega2-abelian.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Ω₂(A) is abelian 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | -- Ω₂(A) for any type A is abelian. 10 | 11 | -- Credits: 12 | -- 13 | -- * The long (unfinished) proof is from 14 | -- (1) Dan Licata's post 15 | -- http://homotopytypetheory.org/2011/03/26/higher-fundamental-groups-are-abelian/ 16 | -- (2) Wikipedia on the Eckmann-Hilton argument 17 | -- http://en.wikipedia.org/wiki/Eckmann%E2%80%93Hilton_argument 18 | -- 19 | -- * I found a much shorter proof while following the above approach. 20 | 21 | module Path.Omega2-abelian {ℓ} {A : Set ℓ} (base : A) where 22 | 23 | open import Prelude hiding (_∘_) 24 | open import Path 25 | open import Path.Lemmas 26 | open import Path.HigherOrder 27 | 28 | private 29 | lemma₁ : ∀ (p : Ω 2 base) → cong (λ p′ → trans p′ (refl base)) p ≡ p 30 | lemma₁ p = elim″ (λ {x} p → cong (λ p → trans p (refl _)) p ≡ trans (trans-reflʳ x) p) (refl _) p 31 | 32 | lemma₂ : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} {B : Set ℓ₂} {C : Set ℓ₃} 33 | (f : A → B → C) {x y : A} {u v : B} 34 | (p : x ≡ y) (q : u ≡ v) → cong₂ f p q ≡ cong₂′ f p q 35 | lemma₂ f p q = elim (λ {_ _} p → cong₂ f p q ≡ cong₂′ f p q) (λ _ → sym $ trans-reflʳ _) p 36 | 37 | abelian : ∀ (p q : Ω 2 base) → trans p q ≡ trans q p 38 | abelian p q = 39 | trans p q ≡⟨ cong (trans p) $ sym $ cong-id q ⟩ 40 | trans p (cong id q) ≡⟨ cong (flip trans (cong id q)) $ sym $ lemma₁ p ⟩ 41 | trans (cong (λ p′ → trans p′ (refl base)) p) (cong id q) ≡⟨ lemma₂ trans p q ⟩ 42 | trans (cong id q) (cong (λ p′ → trans p′ (refl base)) p) ≡⟨ cong (trans (cong id q)) $ lemma₁ p ⟩ 43 | trans (cong id q) p ≡⟨ cong (flip trans p) $ cong-id q ⟩∎ 44 | trans q p ∎ 45 | 46 | ------------------------------------------------------------------------ 47 | -- The following are way too complex for the abelianess theorem, 48 | -- but might be useful in the future. 49 | 50 | private 51 | record IsUnital {ℓ} (A : Set ℓ) (_∘_ : A → A → A) (unit : A) : Set ℓ where 52 | field 53 | unitʳ : ∀ (x : A) → x ∘ unit ≡ x 54 | unitˡ : ∀ (x : A) → unit ∘ x ≡ x 55 | -- Monoid is unnecessary. Unital is enough. 56 | -- assoc : ∀ (x y z : A) → (x ∘ y) ∘ z ≡ x ∘ (y ∘ z) 57 | 58 | module EckmannHilton {ℓ} 59 | {A : Set ℓ} {_∘_ : A → A → A} {_∙_ : A → A → A} {unit : A} 60 | (M₁′ : IsUnital A _∘_ unit) (M₂′ : IsUnital A _∙_ unit) 61 | (interchange : ∀ (x y z w : A) → (x ∙ y) ∘ (z ∙ w) ≡ (x ∘ z) ∙ (y ∘ w)) where 62 | 63 | private 64 | module M₁ where 65 | open IsUnital M₁′ public 66 | module M₂ where 67 | open IsUnital M₂′ public 68 | 69 | same : ∀ (x y : A) → x ∘ y ≡ x ∙ y 70 | same x y = 71 | x ∘ y ≡⟨ cong (λ z → z ∘ y) $ sym $ M₂.unitʳ x ⟩ 72 | (x ∙ unit) ∘ y ≡⟨ cong (λ z → (x ∙ unit) ∘ z) $ sym $ M₂.unitˡ y ⟩ 73 | (x ∙ unit) ∘ (unit ∙ y) ≡⟨ interchange x unit unit y ⟩ 74 | (x ∘ unit) ∙ (unit ∘ y) ≡⟨ cong (λ z → z ∙ (unit ∘ y)) $ M₁.unitʳ x ⟩ 75 | x ∙ (unit ∘ y) ≡⟨ cong (λ z → x ∙ z) $ M₁.unitˡ y ⟩∎ 76 | x ∙ y ∎ 77 | 78 | abelian′ : ∀ (x y : A) → x ∘ y ≡ y ∘ x 79 | abelian′ x y = 80 | x ∘ y ≡⟨ same x y ⟩ 81 | x ∙ y ≡⟨ cong (λ z → z ∙ y) $ sym $ M₁.unitˡ x ⟩ 82 | (unit ∘ x) ∙ y ≡⟨ cong (λ z → (unit ∘ x) ∙ z) $ sym $ M₁.unitʳ y ⟩ 83 | (unit ∘ x) ∙ (y ∘ unit) ≡⟨ sym $ interchange unit y x unit ⟩ 84 | (unit ∙ y) ∘ (x ∙ unit) ≡⟨ cong (λ z → (unit ∙ y) ∘ z) $ M₂.unitʳ x ⟩ 85 | (unit ∙ y) ∘ x ≡⟨ cong (λ z → z ∘ x) $ M₂.unitˡ y ⟩∎ 86 | y ∘ x ∎ 87 | 88 | _∘_ : ∀ {ℓ} {B : Set ℓ} {x y z : B} → x ≡ y → y ≡ z → x ≡ z 89 | _∘_ = trans 90 | _∙_ : ∀ {x y z w : base ≡ base} → x ≡ y → z ≡ w → trans x z ≡ trans y w 91 | _∙_ = cong₂ trans 92 | unit = refl (refl base) 93 | 94 | trans-unital : IsUnital (Ω 2 base) _∘_ unit 95 | trans-unital = record 96 | { unitʳ = trans-reflʳ 97 | ; unitˡ = λ _ → refl _ 98 | --; assoc = trans-assoc 99 | } 100 | 101 | cong₂-trans-unital : IsUnital (Ω 2 base) _∙_ unit 102 | cong₂-trans-unital = record 103 | { unitʳ = unitʳ 104 | ; unitˡ = unitˡ 105 | --; assoc = assoc 106 | } where 107 | unitʳ : ∀ p → p ∙ unit ≡ p 108 | unitʳ p = elim″ (λ {x} p → p ∙ unit ≡ (trans-reflʳ x) ∘ p) (refl _) p 109 | 110 | unitˡ : ∀ p → unit ∙ p ≡ p 111 | unitˡ p = elim″ (λ {_} p → unit ∙ p ≡ p) (refl _) p 112 | {- 113 | assoc : ∀ p q r → (p ∙ q) ∙ r ≡ p ∙ (q ∙ r) 114 | assoc p q r = elim″ 115 | (λ {x} p → (p ∙ q) ∙ r ≡ trans-assoc x (refl base) (refl base) ∘ (p ∙ (q ∙ r))) 116 | ((unit ∙ q) ∙ r ≡⟨ cong (λ pq → pq ∙ r) $ unitˡ q ⟩ 117 | q ∙ r ≡⟨ sym $ unitˡ (q ∙ r) ⟩∎ 118 | unit ∙ (q ∙ r) ∎) 119 | p 120 | -} 121 | 122 | -- The effort to prove this is more than enough to prove abelianness: 123 | -- interchange : ∀ (x y z w : Ω₂A) → (x ∙ y) ∘ (z ∙ w) ≡ (x ∘ z) ∙ (y ∘ w) 124 | -------------------------------------------------------------------------------- /Path/Sum.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Paths in Σ types 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | -- The public interface includes 10 | -- * A homotopy equivalence between (a₁ , b₁) ≡ (a₂ , b₂) 11 | -- and Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂) 12 | 13 | module Path.Sum where 14 | 15 | open import Prelude 16 | open import Path 17 | open import Path.Lemmas 18 | open import Map.H-equivalence hiding (id; _∘_; inverse) 19 | 20 | ------------------------------------------------------------------------ 21 | -- A homotopy equivalence between ((a₁ , b₁) ≡ (a₂ , b₂)) and 22 | -- Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂) 23 | 24 | -- Compose equalities in a Σ type 25 | 26 | Σ≡⇒≡Σ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} → 27 | Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂) → (a₁ , b₁) ≡ (a₂ , b₂) 28 | Σ≡⇒≡Σ B {a₁} {a₂} {b₁} {b₂} (a≡a , b≡b) = elim 29 | (λ {a₁ a₂} a≡a → ∀ {b₁ : B a₁} {b₂ : B a₂} → subst B a≡a b₁ ≡ b₂ → (a₁ , b₁) ≡ (a₂ , b₂)) 30 | (λ a → λ {b₁} {b₂} b≡b → 31 | (a , b₁) ≡⟨ refl _ ⟩ 32 | (a , subst B (refl a) b₁) ≡⟨ cong (λ b → (a , b)) b≡b ⟩∎ 33 | (a , b₂) ∎) 34 | a≡a b≡b 35 | 36 | -- Decompose equalities for a Σ type 37 | 38 | ≡Σ⇒Σ≡ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} → 39 | (a₁ , b₁) ≡ (a₂ , b₂) → Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂) 40 | ≡Σ⇒Σ≡ {A = A} B {b₁ = b₁} {b₂ = b₂} s≡s = (a≡a , b≡b) 41 | where 42 | a≡a = cong proj₁ s≡s 43 | b≡b = 44 | subst B a≡a b₁ ≡⟨ subst-cong B proj₁ s≡s b₁ ⟩ 45 | subst (B ∘ proj₁) s≡s b₁ ≡⟨ cong[dep] (B ∘ proj₁) proj₂ s≡s ⟩∎ 46 | b₂ ∎ 47 | 48 | -- Composition and decomposition form a homotopy equivalence! 49 | 50 | ≡Σ↔Σ≡ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} → 51 | ((a₁ , b₁) ≡ (a₂ , b₂)) ↔ Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂) 52 | ≡Σ↔Σ≡ B {a₁} {a₂} {b₁} {b₂} = 53 | record 54 | { surjection = record 55 | { to = ≡Σ⇒Σ≡ B 56 | ; from = Σ≡⇒≡Σ B 57 | ; right-inverse-of = right-inverse-of 58 | } 59 | ; left-inverse-of = left-inverse-of 60 | } 61 | where 62 | left-inverse-of : (p : (a₁ , b₁) ≡ (a₂ , b₂)) → Σ≡⇒≡Σ B (≡Σ⇒Σ≡ B p) ≡ p 63 | left-inverse-of p = 64 | elim 65 | (λ {s₁ s₂} p → Σ≡⇒≡Σ B (≡Σ⇒Σ≡ B p) ≡ p) 66 | (λ s → refl _) 67 | p 68 | 69 | right-inverse-of : (s : Σ (a₁ ≡ a₂) (λ p → subst B p b₁ ≡ b₂)) → ≡Σ⇒Σ≡ B (Σ≡⇒≡Σ B s) ≡ s 70 | right-inverse-of (pa , pb) = 71 | elim 72 | (λ {a₁ a₂} pa → {b₁ : B a₁} {b₂ : B a₂} (pb : subst B pa b₁ ≡ b₂) → ≡Σ⇒Σ≡ B (Σ≡⇒≡Σ B (pa , pb)) ≡ (pa , pb)) 73 | (λ a → 74 | elim 75 | (λ {b₁ : B a} {b₂ : B a} (pb : b₁ ≡ b₂) → ≡Σ⇒Σ≡ B (Σ≡⇒≡Σ B (refl a , pb)) ≡ (refl a , pb)) 76 | (λ b → refl _)) 77 | pa pb 78 | 79 | ------------------------------------------------------------------------ 80 | -- A homotopy equivalence between ((a₁ , b₁) ≡ (a₂ , b₂)) and 81 | -- (a₁ ≡ a₂) × (b₁ ≡ b₂). I.e., non-dependent version 82 | 83 | -- Compose equalities in a Σ type 84 | 85 | ×≡⇒≡× : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} {a₁ a₂ : A} {b₁ b₂ : B} → 86 | (a₁ ≡ a₂) × (b₁ ≡ b₂) → (a₁ , b₁) ≡ (a₂ , b₂) 87 | ×≡⇒≡× {a₁} {a₂} {b₁} {b₂} (a≡a , b≡b) = elim 88 | (λ {a₁ a₂} a≡a → ∀ {b₁ b₂} → b₁ ≡ b₂ → (a₁ , b₁) ≡ (a₂ , b₂)) 89 | (λ a → λ {b₁} {b₂} b≡b → 90 | (a , b₁) ≡⟨ cong (λ b → (a , b)) b≡b ⟩∎ 91 | (a , b₂) ∎) 92 | a≡a b≡b 93 | 94 | -- Decompose equalities for a Σ type 95 | 96 | ≡×⇒×≡ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} {a₁ a₂ : A} {b₁ b₂ : B} → 97 | (a₁ , b₁) ≡ (a₂ , b₂) → (a₁ ≡ a₂) × (b₁ ≡ b₂) 98 | ≡×⇒×≡ {b₁ = b₁} {b₂ = b₂} s≡s = (a≡a , b≡b) 99 | where 100 | a≡a = cong proj₁ s≡s 101 | b≡b = cong proj₂ s≡s 102 | 103 | -- Composition and decomposition form a homotopy equivalence! 104 | 105 | ≡×↔×≡ : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : Set ℓ₂) {a₁ a₂ : A} {b₁ b₂ : B} → 106 | ((a₁ , b₁) ≡ (a₂ , b₂)) ↔ (a₁ ≡ a₂) × (b₁ ≡ b₂) 107 | ≡×↔×≡ B {a₁} {a₂} {b₁} {b₂} = 108 | record 109 | { surjection = record 110 | { to = ≡×⇒×≡ 111 | ; from = ×≡⇒≡× 112 | ; right-inverse-of = right-inverse-of 113 | } 114 | ; left-inverse-of = left-inverse-of 115 | } 116 | where 117 | left-inverse-of : (p : (a₁ , b₁) ≡ (a₂ , b₂)) → ×≡⇒≡× (≡×⇒×≡ p) ≡ p 118 | left-inverse-of p = 119 | elim 120 | (λ {s₁ s₂} p → ×≡⇒≡× (≡×⇒×≡ p) ≡ p) 121 | (λ s → refl _) 122 | p 123 | 124 | right-inverse-of : (s : (a₁ ≡ a₂) × (b₁ ≡ b₂)) → ≡×⇒×≡ (×≡⇒≡× s) ≡ s 125 | right-inverse-of (pa , pb) = 126 | elim 127 | (λ {a₁ a₂} pa → ∀ {b₁ b₂} (pb : b₁ ≡ b₂) → ≡×⇒×≡ (×≡⇒≡× (pa , pb)) ≡ (pa , pb)) 128 | (λ a → 129 | elim 130 | (λ {b₁ b₂} (pb : b₁ ≡ b₂) → ≡×⇒×≡ (×≡⇒≡× (refl a , pb)) ≡ (refl a , pb)) 131 | (λ b → refl _)) 132 | pa pb 133 | 134 | ------------------------------------------------------------------------ 135 | -- Other lemmas for paths in Σ 136 | 137 | private 138 | Cong-Σ≡⇒≡Σ : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : Set ℓ₃) 139 | (f : (x : A) → B x → C) {a₁ a₂ : A} (p : a₁ ≡ a₂) {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → 140 | Set ℓ₃ 141 | Cong-Σ≡⇒≡Σ B C f {a₁} {a₂} p {b₁} {b₂} q = 142 | cong (uncurry f) (Σ≡⇒≡Σ B (p , q)) ≡ 143 | (f a₁ b₁ ≡⟨ sym $ subst-const p (f a₁ b₁) ⟩ 144 | subst (const C) p (f a₁ b₁) ≡⟨ subst-app B (const C) p (f a₁) q ⟩ 145 | subst (λ x → B x → C) p (f a₁) b₂ ≡⟨ cong (λ f′ → f′ b₂) $ cong[dep] (λ x → B x → C) f p ⟩∎ 146 | f a₂ b₂ ∎) 147 | 148 | cong-Σ≡⇒≡Σ : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : Set ℓ₃) 149 | (f : (x : A) → B x → C) {a₁ a₂ : A} (p : a₁ ≡ a₂) {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → 150 | Cong-Σ≡⇒≡Σ B C f p q 151 | 152 | cong-Σ≡⇒≡Σ B C f = 153 | elim 154 | (λ {a₁ a₂} p → {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → Cong-Σ≡⇒≡Σ B C f p q) 155 | (λ a → 156 | elim 157 | (λ {b₁ b₂} q → Cong-Σ≡⇒≡Σ B C f (refl a) q) 158 | (λ b → refl _)) 159 | 160 | private 161 | Subst-Σfunc : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : Σ A B → Set ℓ₃) 162 | {a₁ a₂ : A} (p : a₁ ≡ a₂) (f₁ : (b₁ : B a₁) → C (a₁ , b₁)) 163 | {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → Set ℓ₃ 164 | Subst-Σfunc B C p f₁ {b₁} {b₂} q = 165 | subst (λ a → (b : B a) → C (a , b)) p f₁ b₂ ≡ 166 | subst C (Σ≡⇒≡Σ B (p , q)) (f₁ b₁) 167 | 168 | subst-Σfunc : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : Σ A B → Set ℓ₃) 169 | {a₁ a₂ : A} (p : a₁ ≡ a₂) (f₁ : (b₁ : B a₁) → C (a₁ , b₁)) 170 | {b₁ : B a₁} {b₂ : B a₂} (q : subst B p b₁ ≡ b₂) → Subst-Σfunc B C p f₁ q 171 | subst-Σfunc B C = 172 | elim 173 | (λ {a₁ a₂} p → ∀ f₁ {b₁ : B a₁} {b₂ : B a₂} q → Subst-Σfunc B C p f₁ q) 174 | (λ a f₁ → elim 175 | (λ {b₁} {b₂} q → Subst-Σfunc B C (refl a) f₁ q) 176 | (λ b → refl _)) 177 | 178 | -------------------------------------------------------------------------------- /Path/Symbol.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Symbols for paths 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Path.Symbol where 10 | 11 | open import Path 12 | 13 | !_ : ∀ {ℓ} {A : Set ℓ} {x y : A} → x ≡ y → y ≡ x 14 | !_ = sym 15 | 16 | _◇_ : ∀ {ℓ} {A : Set ℓ} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 17 | _◇_ = trans 18 | 19 | infixr 6 _◇_ 20 | -------------------------------------------------------------------------------- /Prelude.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- A small prelude 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | -- Note that parts of Agda's standard library make use of the K rule. 11 | 12 | module Prelude where 13 | 14 | ------------------------------------------------------------------------ 15 | -- Support for universe polymorphism 16 | 17 | -- Universe levels. 18 | 19 | infixl 6 _⊔_ 20 | 21 | postulate 22 | Level : Set 23 | lzero : Level 24 | lsuc : Level → Level 25 | _⊔_ : Level → Level → Level 26 | 27 | {-# BUILTIN LEVEL Level #-} 28 | {-# BUILTIN LEVELZERO lzero #-} 29 | {-# BUILTIN LEVELSUC lsuc #-} 30 | {-# BUILTIN LEVELMAX _⊔_ #-} 31 | 32 | -- Lifting. 33 | 34 | record ↑ {a} ℓ (A : Set a) : Set (a ⊔ ℓ) where 35 | constructor lift 36 | field lower : A 37 | 38 | open ↑ public 39 | 40 | ------------------------------------------------------------------------ 41 | -- Some finite types 42 | 43 | -- The empty type. 44 | 45 | data ⊥ {ℓ} : Set ℓ where 46 | 47 | ⊥-elim : ∀ {w ℓ} {Whatever : Set w} → ⊥ {ℓ = ℓ} → Whatever 48 | ⊥-elim () 49 | 50 | -- Negation. 51 | 52 | infix 3 ¬_ 53 | 54 | ¬_ : ∀ {ℓ} → Set ℓ → Set ℓ 55 | ¬ P = P → ⊥ {ℓ = lzero} 56 | 57 | -- The unit type. 58 | 59 | record ⊤ : Set where 60 | constructor tt 61 | 62 | -- Booleans. 63 | 64 | data Bool : Set where 65 | true false : Bool 66 | 67 | {-# BUILTIN BOOL Bool #-} 68 | {-# BUILTIN TRUE true #-} 69 | {-# BUILTIN FALSE false #-} 70 | 71 | -- Conditional. 72 | 73 | if_then_else_ : ∀ {a} {A : Set a} → Bool → A → A → A 74 | if true then t else f = t 75 | if false then t else f = f 76 | 77 | -- Not. 78 | 79 | not : Bool → Bool 80 | not b = if b then false else true 81 | 82 | ------------------------------------------------------------------------ 83 | -- Natural numbers 84 | 85 | data ℕ : Set where 86 | zero : ℕ 87 | suc : (n : ℕ) → ℕ 88 | 89 | -- Support for natural number literals. 90 | 91 | {-# BUILTIN NATURAL ℕ #-} 92 | {-# BUILTIN ZERO zero #-} 93 | {-# BUILTIN SUC suc #-} 94 | 95 | -- Dependent eliminator. 96 | 97 | ℕ-rec : ∀ {p} {P : ℕ → Set p} → 98 | P 0 → (∀ n → P n → P (suc n)) → ∀ n → P n 99 | ℕ-rec z s zero = z 100 | ℕ-rec z s (suc n) = s n (ℕ-rec z s n) 101 | 102 | -- Addition. 103 | 104 | infixl 6 _+_ 105 | 106 | _+_ : ℕ → ℕ → ℕ 107 | zero + n = n 108 | suc m + n = suc (m + n) 109 | 110 | ------------------------------------------------------------------------ 111 | -- Simple combinators working solely on and with functions 112 | 113 | infixr 9 _∘_ 114 | infixr 0 _$_ 115 | 116 | -- The identity function. 117 | 118 | id : ∀ {a} {A : Set a} → A → A 119 | id x = x 120 | 121 | -- Composition. 122 | 123 | _∘_ : ∀ {a b c} 124 | {A : Set a} {B : A → Set b} {C : {x : A} → B x → Set c} → 125 | (∀ {x} (y : B x) → C y) → (g : (x : A) → B x) → 126 | ((x : A) → C (g x)) 127 | f ∘ g = λ x → f (g x) 128 | 129 | -- Application. 130 | 131 | _$_ : ∀ {a b} {A : Set a} {B : A → Set b} → 132 | ((x : A) → B x) → ((x : A) → B x) 133 | f $ x = f x 134 | 135 | -- Constant functions. 136 | 137 | const : ∀ {a b} {A : Set a} {B : Set b} → A → (B → A) 138 | const x = λ _ → x 139 | 140 | -- Flips the first two arguments. 141 | 142 | flip : ∀ {a b c} {A : Set a} {B : Set b} {C : A → B → Set c} → 143 | ((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y) 144 | flip f = λ x y → f y x 145 | 146 | ------------------------------------------------------------------------ 147 | -- Σ-types 148 | 149 | infixr 4 _,_ 150 | infixr 2 _×_ 151 | 152 | record Σ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where 153 | constructor _,_ 154 | field 155 | proj₁ : A 156 | proj₂ : B proj₁ 157 | 158 | open Σ public 159 | 160 | -- A variant where the first argument is implicit. 161 | 162 | ∃ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b) 163 | ∃ = Σ _ 164 | 165 | -- Binary products. 166 | 167 | _×_ : ∀ {a b} (A : Set a) (B : Set b) → Set (a ⊔ b) 168 | A × B = Σ A (const B) 169 | 170 | -- A map function. 171 | 172 | Σ-map : ∀ {a b p q} 173 | {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} → 174 | (f : A → B) → (∀ {x} → P x → Q (f x)) → 175 | Σ A P → Σ B Q 176 | Σ-map f g = λ p → (f (proj₁ p) , g (proj₂ p)) 177 | 178 | -- Curry and uncurry. 179 | 180 | curry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} → 181 | ((p : Σ A B) → C p) → 182 | ((x : A) (y : B x) → C (x , y)) 183 | curry f x y = f (x , y) 184 | 185 | uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} → 186 | ((x : A) (y : B x) → C (x , y)) → 187 | ((p : Σ A B) → C p) 188 | uncurry f (x , y) = f x y 189 | 190 | ------------------------------------------------------------------------ 191 | -- W-types 192 | 193 | data W {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where 194 | sup : (x : A) (f : B x → W A B) → W A B 195 | 196 | -- Projections. 197 | 198 | head : ∀ {a b} {A : Set a} {B : A → Set b} → 199 | W A B → A 200 | head (sup x f) = x 201 | 202 | tail : ∀ {a b} {A : Set a} {B : A → Set b} → 203 | (x : W A B) → B (head x) → W A B 204 | tail (sup x f) = f 205 | 206 | -- If B is always inhabited, then W A B is empty. 207 | 208 | abstract 209 | 210 | inhabited⇒W-empty : ∀ {a b} {A : Set a} {B : A → Set b} → 211 | (∀ x → B x) → ¬ W A B 212 | inhabited⇒W-empty b (sup x f) = inhabited⇒W-empty b (f (b x)) 213 | 214 | ------------------------------------------------------------------------ 215 | -- Support for coinduction 216 | 217 | infix 1000 ♯_ 218 | 219 | postulate 220 | ∞ : ∀ {a} (A : Set a) → Set a 221 | ♯_ : ∀ {a} {A : Set a} → A → ∞ A 222 | ♭ : ∀ {a} {A : Set a} → ∞ A → A 223 | 224 | {-# BUILTIN INFINITY ∞ #-} 225 | {-# BUILTIN SHARP ♯_ #-} 226 | {-# BUILTIN FLAT ♭ #-} 227 | 228 | ------------------------------------------------------------------------ 229 | -- M-types 230 | 231 | data M {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where 232 | dns : (x : A) (f : B x → ∞ (M A B)) → M A B 233 | 234 | -- Projections. 235 | 236 | pɐǝɥ : ∀ {a b} {A : Set a} {B : A → Set b} → 237 | M A B → A 238 | pɐǝɥ (dns x f) = x 239 | 240 | lıɐʇ : ∀ {a b} {A : Set a} {B : A → Set b} → 241 | (x : M A B) → B (pɐǝɥ x) → M A B 242 | lıɐʇ (dns x f) y = ♭ (f y) 243 | 244 | ------------------------------------------------------------------------ 245 | -- Binary sums 246 | 247 | infixr 1 _⊎_ 248 | 249 | data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where 250 | inj₁ : (x : A) → A ⊎ B 251 | inj₂ : (y : B) → A ⊎ B 252 | 253 | -- Eliminator for binary sums. 254 | 255 | [_,_] : ∀ {a b c} {A : Set a} {B : Set b} {C : A ⊎ B → Set c} → 256 | ((x : A) → C (inj₁ x)) → ((x : B) → C (inj₂ x)) → 257 | ((x : A ⊎ B) → C x) 258 | [ f , g ] (inj₁ x) = f x 259 | [ f , g ] (inj₂ y) = g y 260 | 261 | -- A map function. 262 | 263 | ⊎-map : ∀ {a₁ a₂ b₁ b₂} 264 | {A₁ : Set a₁} {A₂ : Set a₂} {B₁ : Set b₁} {B₂ : Set b₂} → 265 | (A₁ → A₂) → (B₁ → B₂) → A₁ ⊎ B₁ → A₂ ⊎ B₂ 266 | ⊎-map f g = [ inj₁ ∘ f , inj₂ ∘ g ] 267 | 268 | -- A special case of binary sums: decided predicates. 269 | 270 | Dec : ∀ {p} → Set p → Set p 271 | Dec P = P ⊎ ¬ P 272 | 273 | -- Decidable relations. 274 | 275 | Decidable : ∀ {a b ℓ} {A : Set a} {B : Set b} → 276 | (A → B → Set ℓ) → Set (a ⊔ b ⊔ ℓ) 277 | Decidable _∼_ = ∀ x y → Dec (x ∼ y) 278 | 279 | ------------------------------------------------------------------------ 280 | -- Lists 281 | 282 | infixr 5 _∷_ 283 | 284 | data List {a} (A : Set a) : Set a where 285 | [] : List A 286 | _∷_ : (x : A) (xs : List A) → List A 287 | 288 | {-# BUILTIN LIST List #-} 289 | {-# BUILTIN NIL [] #-} 290 | {-# BUILTIN CONS _∷_ #-} 291 | 292 | [_] : ∀ {ℓ} {A : Set ℓ} → A → List A 293 | [_] a = a ∷ [] 294 | 295 | -- Right fold. 296 | 297 | foldr : ∀ {a b} {A : Set a} {B : Set b} → 298 | (A → B → B) → B → List A → B 299 | foldr _⊕_ ε [] = ε 300 | foldr _⊕_ ε (x ∷ xs) = x ⊕ foldr _⊕_ ε xs 301 | 302 | -- The length of a list. 303 | 304 | length : ∀ {a} {A : Set a} → List A → ℕ 305 | length = foldr (const suc) 0 306 | 307 | -- Appends two lists. 308 | 309 | infixr 5 _++_ 310 | 311 | _++_ : ∀ {a} {A : Set a} → List A → List A → List A 312 | xs ++ ys = foldr _∷_ ys xs 313 | 314 | -- Maps a function over a list. 315 | 316 | map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List A → List B 317 | map f = foldr (λ x ys → f x ∷ ys) [] 318 | 319 | -- Concatenates a list of lists. 320 | 321 | concat : ∀ {a} {A : Set a} → List (List A) → List A 322 | concat = foldr _++_ [] 323 | 324 | -- The list monad's bind operation. 325 | 326 | infixl 5 _>>=_ 327 | 328 | _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → 329 | List A → (A → List B) → List B 330 | xs >>= f = concat (map f xs) 331 | 332 | -- A filter function. 333 | 334 | filter : ∀ {a} {A : Set a} → (A → Bool) → List A → List A 335 | filter p = foldr (λ x xs → if p x then x ∷ xs else xs) [] 336 | 337 | ------------------------------------------------------------------------ 338 | -- Finite sets 339 | 340 | Fin : ℕ → Set 341 | Fin zero = ⊥ 342 | Fin (suc n) = ⊤ ⊎ Fin n 343 | 344 | -- A lookup function. 345 | 346 | lookup : ∀ {a} {A : Set a} (xs : List A) → Fin (length xs) → A 347 | lookup [] () 348 | lookup (x ∷ xs) (inj₁ tt) = x 349 | lookup (x ∷ xs) (inj₂ i) = lookup xs i 350 | -------------------------------------------------------------------------------- /README.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- My Code for Homotopy Type Theory 3 | -- 4 | -- Favonia 5 | ------------------------------------------------------------------------ 6 | 7 | -- Copyright (c) 2012 Favonia 8 | 9 | -- A large portion of code was copied from Nils Anders Danielsson' 10 | -- library released under BSD-like license 11 | -- http://www.cse.chalmers.se/~nad/repos/equality/ 12 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 13 | 14 | -- See LICENSE for detailed copyright notice. 15 | 16 | {-# OPTIONS --without-K #-} 17 | 18 | module README where 19 | 20 | ------------------------------------------------------------------------ 21 | -- Common definitions 22 | 23 | import Prelude 24 | 25 | ------------------------------------------------------------------------ 26 | -- Maps, continuous functions between spaces 27 | 28 | -- Homotopy equivalence 29 | import Map.H-equivalence 30 | 31 | -- Injections 32 | import Map.Injection 33 | 34 | -- Surjections 35 | import Map.Surjection 36 | 37 | -- Homotopy Fiber 38 | import Map.H-fiber 39 | 40 | -- Weak equivalent 41 | import Map.WeakEquivalence 42 | 43 | ------------------------------------------------------------------------ 44 | -- Paths (propositional equalities in type theories) 45 | 46 | -- The definition of paths and trans, subst, cong 47 | import Path 48 | 49 | -- Some really basic lemmas for equivalence of paths 50 | import Path.Lemmas 51 | 52 | -- Fancy Unicode symbols for writing incomprehensible proofs 53 | -- (Subjecting to changes.) 54 | import Path.Symbol 55 | 56 | -- Higher-order paths and loops 57 | import Path.HigherOrder 58 | 59 | -- A short proof that Ω₂(A) is abelian for any space A 60 | import Path.Omega2-abelian 61 | 62 | -- Tools to compose/decompose paths in Σ-type 63 | import Path.Sum 64 | 65 | -- Tools to manipulate paths in Π-type (extensionality) 66 | --import Path.Prod 67 | 68 | -- Definition of H-level and some basic lemmas 69 | --import Path.H-level 70 | 71 | ------------------------------------------------------------------------ 72 | -- Space 73 | 74 | -- Kristina's theorem: hom is contractable iff we have a dependent 75 | -- eliminator. 76 | -- (Only the interesting direction.) 77 | import Space.Bool.Initial 78 | 79 | -- Basic facts about Fin 80 | import Space.Fin.Lemmas 81 | 82 | -- Definition of flowers 83 | import Space.Flower 84 | 85 | -- A proof that Ω₁(Flower) is FreeGroup (currently broken) 86 | -- import Space.Flower.Omega1 87 | 88 | -- Definition of free groups (currently broken) 89 | -- import Space.FreeGroup 90 | 91 | -- Definition of integers 92 | import Space.Integer 93 | 94 | -- Definition of intervals 95 | import Space.Interval 96 | 97 | -- Basic facts about Fin 98 | import Space.List.Lemmas 99 | 100 | -- Some basic facts about Nat 101 | -- (Definition of Nat is in the Prelude) 102 | import Space.Nat.Lemmas 103 | 104 | -- Definition of spheres (base + loop) 105 | import Space.Sphere 106 | 107 | -- Alternative definition of spheres (two-point) 108 | import Space.Sphere.TwoPoints 109 | 110 | -- Definition of the Hopf junior (S₀ ↪ S₁ → S₁) 111 | -- and a proof that the total space is indeed S₁ 112 | import Space.Sphere.HopfJunior 113 | 114 | -- A proof that Ω₁(S₁) is ℤ 115 | import Space.Sphere.Omega1 116 | 117 | -- Definition of torus 118 | import Space.Torus 119 | 120 | ------------------------------------------------------------------------ 121 | -- The Univalence axiom 122 | 123 | -- Definition of the Univalence axiom 124 | import Univalence 125 | 126 | -- A proof that the Univalence axiom implies extensionality for functions 127 | -- Might be moved to Path.Prod later 128 | --import Univalence.Extensionality 129 | 130 | -- Some basic lemmas implied by the Univalence axiom 131 | import Univalence.Lemmas 132 | 133 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Introduction 2 | 3 | This repository hosts my Agda code for homotopy type theory. 4 | Homotopy type theory is a young theory 5 | seeking the connections between (abstract) homotopy theory, 6 | type theory and category theory. 7 | For more information about this theory 8 | I recommend [this blog](http://homotopytypetheory.org). 9 | 10 | # Acknowledgment 11 | 12 | This material is partially based upon work supported by the 13 | National Science Foundation under Grant Number 1116703. Any 14 | opinions, findings, and conclusions or recommendations expressed 15 | in this material are those of the author(s) and do not necessarily 16 | reflect the views of the National Science Foundation. 17 | 18 | # Copyright 19 | 20 | See `LICENSE.md` for more information. 21 | Please consider using open-source licenses in your derived work 22 | to maximize the availability and usefulness. 23 | 24 | # Documentation 25 | 26 | For an overview of the code, please take a look at `README.agda`. 27 | -------------------------------------------------------------------------------- /Space/Bool/Initial.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Alternative axioms for Boolean 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- Two equivalent ways to axiomatize Boolean 8 | -- with extensionality for functions: 9 | -- (1) non-dependent elim + being inital 10 | -- (2) dependent elim 11 | 12 | {-# OPTIONS --without-K #-} 13 | 14 | -- Credits: 15 | -- * Kristina Sojakova stated the theorem and gave a proof in 2011 Nov 16 | -- but I was too inexperienced to understand it back then. 17 | -- * The direction from "homotopy" to "dependent elim" is done by 18 | -- myself (Favonia) as an Agda exercise. The other direction 19 | -- should not be hard with extensionalty for functions... 20 | 21 | module Space.Bool.Initial where 22 | 23 | open import Prelude hiding (Bool; true; false; if_then_else_) 24 | open import Path 25 | open import Path.Lemmas 26 | open import Path.Sum 27 | 28 | private 29 | subst-× : ∀ {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} (B : A → Set ℓ₂) (C : A → Set ℓ₃) 30 | {x y : A} (p : x ≡ y) (bx : B x) (cx : C x) → 31 | subst (λ x → (B x) × (C x)) p (bx , cx) ≡ (subst B p bx , subst C p cx) 32 | subst-× B C = 33 | elim 34 | (λ {x y} p → (bx : B x) (cx : C x) → 35 | subst (λ x → (B x) × (C x)) p (bx , cx) ≡ (subst B p bx , subst C p cx)) 36 | (λ x bx cx → refl _) 37 | 38 | record Bool-axiom {ℓ} : Set (lsuc ℓ) where 39 | field 40 | {- form -} 41 | Bool : Set ℓ 42 | {- intro -} 43 | true false : Bool 44 | {- elim -} 45 | if : ∀ (C : Bool → Set ℓ) (b : Bool) → C true → C false → C b 46 | {- conv -} 47 | if-true : ∀ (C : Bool → Set ℓ) (et : C true) (ef : C false) → if C true et ef ≡ et 48 | if-false : ∀ (C : Bool → Set ℓ) (et : C true) (ef : C false) → if C false et ef ≡ ef 49 | 50 | record Bool-axiom′ {ℓ} : Set (lsuc ℓ) where 51 | field 52 | {- form -} 53 | Bool : Set ℓ 54 | {- intro -} 55 | true false : Bool 56 | {- elim -} 57 | if : ∀ {C : Set ℓ} → Bool → C → C → C 58 | {- conv -} 59 | if-true : ∀ {C : Set ℓ} (et ef : C) → if true et ef ≡ et 60 | if-false : ∀ {C : Set ℓ} (et ef : C) → if false et ef ≡ ef 61 | 62 | hom : ∀ {C : Set ℓ} {et ef : C} (h : Bool → C) → 63 | (vt : h true ≡ et) (vf : h false ≡ ef) → 64 | _≡_ {A = Σ (Bool → C) (λ h′ → (h′ true ≡ et) × (h′ false ≡ ef))} 65 | (h , (vt , vf )) 66 | ((λ b → if b et ef) , (if-true et ef , if-false et ef)) 67 | 68 | private 69 | -- A module to extact information out of the hom axiom 70 | module HomModule {C : Set ℓ} {et ef : C} (h : Bool → C) (vt′ : h true ≡ et) (vf′ : h false ≡ ef) where 71 | vt = vt′ 72 | vf = vf′ 73 | 74 | private 75 | lemma-hom = ≡Σ⇒Σ≡ (λ h′ → (h′ true ≡ et) × (h′ false ≡ ef)) $ hom h vt vf 76 | 77 | path-h = proj₁ lemma-hom 78 | 79 | private 80 | lemma-true-false = ≡×⇒×≡ $ 81 | (subst (λ h′ → (h′ true ≡ et)) path-h vt , subst (λ h′ → (h′ false ≡ ef)) path-h vf) 82 | ≡⟨ sym $ subst-× (λ h′ → (h′ true ≡ et)) (λ h′ → (h′ false ≡ ef)) path-h vt vf ⟩ 83 | subst (λ h′ → (h′ true ≡ et) × (h′ false ≡ ef)) path-h (vt , vf) 84 | ≡⟨ proj₂ lemma-hom ⟩∎ 85 | (if-true et ef , if-false et ef) 86 | ∎ 87 | 88 | abstract 89 | path-true : subst (λ h → (h true ≡ et)) path-h vt ≡ if-true et ef 90 | path-true = proj₁ lemma-true-false 91 | path-false : subst (λ h → (h false ≡ ef)) path-h vf ≡ if-false et ef 92 | path-false = proj₂ lemma-true-false 93 | 94 | sym-cong-true-h : trans (sym (cong (λ f → f true) path-h)) vt ≡ if-true et ef 95 | sym-cong-true-h = 96 | trans (sym (cong (λ f → f true) path-h)) vt 97 | ≡⟨ sym $ subst-path[constʳ] (λ f → f true) et path-h vt ⟩ 98 | subst (λ h → (h true ≡ et)) path-h vt 99 | ≡⟨ path-true ⟩∎ 100 | if-true et ef 101 | ∎ 102 | 103 | sym-cong-false-h : trans (sym (cong (λ f → f false) path-h)) vf ≡ if-false et ef 104 | sym-cong-false-h = 105 | trans (sym (cong (λ f → f false) path-h)) vf 106 | ≡⟨ sym $ subst-path[constʳ] (λ f → f false) ef path-h vf ⟩ 107 | subst (λ h → (h false ≡ ef)) path-h vf 108 | ≡⟨ path-false ⟩∎ 109 | if-false et ef 110 | ∎ 111 | 112 | -- A module to provide a nice dependent "if" interface 113 | module IfModule (C : Bool → Set ℓ) (et : C true) (ef : C false) where 114 | private 115 | et′ : Σ Bool C 116 | et′ = (true , et) 117 | ef′ : Σ Bool C 118 | ef′ = (false , ef) 119 | 120 | module ID = HomModule id (refl true) (refl false) 121 | 122 | module Proj₁ = HomModule 123 | (λ b → proj₁ (if b et′ ef′)) 124 | (cong proj₁ (if-true et′ ef′)) 125 | (cong proj₁ (if-false et′ ef′)) 126 | 127 | if′ : ∀ (b : Bool) → C b 128 | if′ b = subst (λ f → C (f b)) (trans Proj₁.path-h (sym ID.path-h)) $ proj₂ (if b et′ ef′) 129 | 130 | abstract 131 | if′-true : if′ true ≡ et 132 | if′-true = 133 | subst (λ f → C (f true)) (trans Proj₁.path-h (sym ID.path-h)) (proj₂ (if true et′ ef′)) 134 | ≡⟨ sym $ subst-cong C (λ f → f true) (trans Proj₁.path-h (sym ID.path-h)) _ ⟩ 135 | subst C (cong (λ f → f true) (trans Proj₁.path-h (sym ID.path-h))) (proj₂ (if true et′ ef′)) 136 | ≡⟨ cong (λ p → subst C p $ proj₂ (if true et′ ef′)) $ cong-trans (λ f → f true) Proj₁.path-h (sym ID.path-h) ⟩ 137 | subst C (trans (cong (λ f → f true) Proj₁.path-h) (cong (λ f → f true) (sym ID.path-h))) (proj₂ (if true et′ ef′)) 138 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f true) Proj₁.path-h) p) $ proj₂ (if true et′ ef′)) $ cong-sym (λ f → f true) ID.path-h ⟩ 139 | subst C (trans (cong (λ f → f true) Proj₁.path-h) (sym $ cong (λ f → f true) ID.path-h)) (proj₂ (if true et′ ef′)) 140 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f true) Proj₁.path-h) p) $ proj₂ (if true et′ ef′)) $ 141 | sym $ trans-reflʳ (sym $ cong (λ f → f true) ID.path-h) ⟩ 142 | subst C (trans (cong (λ f → f true) Proj₁.path-h) (trans (sym $ cong (λ f → f true) ID.path-h) (refl true))) (proj₂ (if true et′ ef′)) 143 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f true) Proj₁.path-h) p) $ proj₂ (if true et′ ef′)) $ ID.sym-cong-true-h ⟩ 144 | subst C (trans (cong (λ f → f true) Proj₁.path-h) (if-true true false)) (proj₂ (if true et′ ef′)) 145 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f true) Proj₁.path-h) p) $ proj₂ (if true et′ ef′)) $ sym $ Proj₁.sym-cong-true-h ⟩ 146 | subst C (trans (cong (λ f → f true) Proj₁.path-h) (trans (sym (cong (λ f → f true) Proj₁.path-h)) Proj₁.vt)) (proj₂ (if true et′ ef′)) 147 | ≡⟨ cong (λ p → subst C p $ proj₂ (if true et′ ef′)) $ trans-transʳ-sym (cong (λ f → f true) Proj₁.path-h) _ ⟩ 148 | subst C Proj₁.vt (proj₂ (if true et′ ef′)) 149 | ≡⟨ subst-cong C proj₁ (if-true et′ ef′) $ proj₂ (if true et′ ef′) ⟩ 150 | subst (C ∘ proj₁) (if-true et′ ef′) (proj₂ (if true et′ ef′)) 151 | ≡⟨ cong[dep] (λ s → C (proj₁ s)) proj₂ (if-true et′ ef′) ⟩∎ 152 | proj₂ et′ 153 | ∎ 154 | 155 | if′-false : if′ false ≡ ef 156 | if′-false = 157 | subst (λ f → C (f false)) (trans Proj₁.path-h (sym ID.path-h)) (proj₂ (if false et′ ef′)) 158 | ≡⟨ sym $ subst-cong C (λ f → f false) (trans Proj₁.path-h (sym ID.path-h)) _ ⟩ 159 | subst C (cong (λ f → f false) (trans Proj₁.path-h (sym ID.path-h))) (proj₂ (if false et′ ef′)) 160 | ≡⟨ cong (λ p → subst C p $ proj₂ (if false et′ ef′)) $ cong-trans (λ f → f false) Proj₁.path-h (sym ID.path-h) ⟩ 161 | subst C (trans (cong (λ f → f false) Proj₁.path-h) (cong (λ f → f false) (sym ID.path-h))) (proj₂ (if false et′ ef′)) 162 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f false) Proj₁.path-h) p) $ proj₂ (if false et′ ef′)) $ cong-sym (λ f → f false) ID.path-h ⟩ 163 | subst C (trans (cong (λ f → f false) Proj₁.path-h) (sym $ cong (λ f → f false) ID.path-h)) (proj₂ (if false et′ ef′)) 164 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f false) Proj₁.path-h) p) $ proj₂ (if false et′ ef′)) $ 165 | sym $ trans-reflʳ (sym $ cong (λ f → f false) ID.path-h) ⟩ 166 | subst C (trans (cong (λ f → f false) Proj₁.path-h) (trans (sym $ cong (λ f → f false) ID.path-h) (refl false))) (proj₂ (if false et′ ef′)) 167 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f false) Proj₁.path-h) p) $ proj₂ (if false et′ ef′)) $ ID.sym-cong-false-h ⟩ 168 | subst C (trans (cong (λ f → f false) Proj₁.path-h) (if-false true false)) (proj₂ (if false et′ ef′)) 169 | ≡⟨ cong (λ p → subst C (trans (cong (λ f → f false) Proj₁.path-h) p) $ proj₂ (if false et′ ef′)) $ sym $ Proj₁.sym-cong-false-h ⟩ 170 | subst C (trans (cong (λ f → f false) Proj₁.path-h) (trans (sym (cong (λ f → f false) Proj₁.path-h)) Proj₁.vf)) (proj₂ (if false et′ ef′)) 171 | ≡⟨ cong (λ p → subst C p $ proj₂ (if false et′ ef′)) $ trans-transʳ-sym (cong (λ f → f false) Proj₁.path-h) _ ⟩ 172 | subst C Proj₁.vf (proj₂ (if false et′ ef′)) 173 | ≡⟨ subst-cong C proj₁ (if-false et′ ef′) $ proj₂ (if false et′ ef′) ⟩ 174 | subst (C ∘ proj₁) (if-false et′ ef′) (proj₂ (if false et′ ef′)) 175 | ≡⟨ cong[dep] (λ s → C (proj₁ s)) proj₂ (if-false et′ ef′) ⟩∎ 176 | proj₂ ef′ 177 | ∎ 178 | 179 | if′ : ∀ (C : Bool → Set ℓ) (b : Bool) → C true → C false → C b 180 | if′ C b et ef = IfModule.if′ C et ef b 181 | 182 | if′-true : ∀ (C : Bool → Set ℓ) (et : C true) (ef : C false) → if′ C true et ef ≡ et 183 | if′-true C et ef = IfModule.if′-true C et ef 184 | 185 | if′-false : ∀ (C : Bool → Set ℓ) (et : C true) (ef : C false) → if′ C false et ef ≡ ef 186 | if′-false C et ef = IfModule.if′-false C et ef 187 | 188 | -------------------------------------------------------------------------------- /Space/Fin/Lemmas.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Lemmas for Fin 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Space.Fin.Lemmas where 10 | 11 | open import Prelude 12 | open import Path 13 | 14 | decide-≡ : ∀ {n} (x y : Fin n) → Dec (x ≡ y) 15 | decide-≡ {zero} () () 16 | decide-≡ {suc _} (inj₁ tt) (inj₁ tt) = inj₁ $ refl _ 17 | decide-≡ {suc _} (inj₂ x) (inj₂ y) with decide-≡ x y 18 | ... | inj₁ eq = inj₁ $ cong inj₂ eq 19 | ... | inj₂ neq = inj₂ $ neq ∘ cong λ{(inj₂ x) → x; (inj₁ _) → x} 20 | decide-≡ {suc _} (inj₂ _) (inj₁ _) = inj₂ λ eq → subst (λ{(inj₂ _) → ⊤; (inj₁ _) → ⊥}) eq tt 21 | decide-≡ {suc _} (inj₁ _) (inj₂ _) = inj₂ λ eq → subst (λ{(inj₁ _) → ⊤; (inj₂ _) → ⊥}) eq tt 22 | 23 | data _<>_ : ∀ {n} (x y : Fin n) → Set where 24 | 0<>s : ∀ {n} {x : Fin n} → inj₁ tt <> inj₂ x 25 | s<>0 : ∀ {n} {x : Fin n} → inj₂ x <> inj₁ tt 26 | s<>s : ∀ {n} {x y : Fin n} → x <> y → inj₂ x <> inj₂ y 27 | 28 | decide-≡' : ∀ {n} (x y : Fin n) → x ≡ y ⊎ x <> y 29 | decide-≡' {zero} () () 30 | decide-≡' {suc _} (inj₁ tt) (inj₁ tt) = inj₁ $ refl _ 31 | decide-≡' {suc _} (inj₂ _) (inj₁ _) = inj₂ s<>0 32 | decide-≡' {suc _} (inj₁ _) (inj₂ _) = inj₂ 0<>s 33 | decide-≡' {suc _} (inj₂ x) (inj₂ y) with decide-≡' x y 34 | ... | inj₁ eq = inj₁ $ cong inj₂ eq 35 | ... | inj₂ neq = inj₂ $ s<>s neq 36 | 37 | <>-sym : ∀ {n} {x y : Fin n} → x <> y → y <> x 38 | <>-sym 0<>s = s<>0 39 | <>-sym s<>0 = 0<>s 40 | <>-sym (s<>s neq) = s<>s $ <>-sym neq 41 | -------------------------------------------------------------------------------- /Space/Flower.agda: -------------------------------------------------------------------------------- 1 | module Space.Flower where 2 | 3 | open import Prelude 4 | 5 | open import Path 6 | open import Path.Lemmas 7 | 8 | private 9 | data Flower′ (n : ℕ) : Set where 10 | core′ : Flower′ n 11 | 12 | Flower : ℕ → Set 13 | Flower n = Flower′ n 14 | 15 | core : ∀ n → Flower n 16 | core n = core′ 17 | 18 | -- Dependent version 19 | 20 | postulate 21 | petal : ∀ {n} → Fin n → core n ≡ core n 22 | 23 | Flower-elim : ∀ {ℓ} {n} (P : Flower n → Set ℓ) 24 | (pcore : P (core n)) 25 | (ppetal : ∀ i → subst P (petal i) pcore ≡ pcore) → 26 | ∀ f → P f 27 | Flower-elim _ pcore _ core′ = pcore 28 | 29 | postulate 30 | Flower-elim-petal : ∀ {ℓ} {n} (P : Flower n → Set ℓ) 31 | (pcore : P (core n)) 32 | (ppetal : ∀ i → subst P (petal i) pcore ≡ pcore) → 33 | ∀ i → cong[dep] P (Flower-elim P pcore ppetal) (petal i) ≡ ppetal i 34 | 35 | -- Non-dependent version 36 | 37 | Flower-elim[simp] : ∀ {ℓ} {n} {P : Set ℓ} 38 | (pcore : P) (ppetal : ∀ (i : Fin n) → pcore ≡ pcore) → 39 | ∀ f → P 40 | Flower-elim[simp] {n = n} {P = P} pcore ppetal = 41 | Flower-elim (const P) pcore (λ i → trans (boring-petal i) (ppetal i)) 42 | where 43 | boring-petal : ∀ (i : Fin n) → subst (const P) (petal i) pcore ≡ pcore 44 | boring-petal i = subst-const (petal i) pcore 45 | 46 | -- This propositional equality is derivable from the dependent version. 47 | Flower-elim[simp]-petal : ∀ {ℓ} {n} {P : Set ℓ} 48 | (pcore : P) 49 | (ppetal : ∀ (i : Fin n) → pcore ≡ pcore) → 50 | ∀ i → cong (Flower-elim[simp] pcore ppetal) (petal i) ≡ ppetal i 51 | Flower-elim[simp]-petal {n = n} {P = P} pcore ppetal i = 52 | cong pflower (petal i) 53 | ≡⟨ sym $ trans-sym-trans (boring-petal i) _ ⟩ 54 | trans (sym $ boring-petal i) (trans (boring-petal i) $ cong pflower $ petal i) 55 | ≡⟨ cong (trans $ sym $ boring-petal i) $ sym $ cong[dep]-const pflower (petal i) ⟩ 56 | trans (sym $ boring-petal i) (cong[dep] (const P) pflower $ petal i) 57 | ≡⟨ cong (trans $ sym $ boring-petal i) $ 58 | Flower-elim-petal (const P) pcore (λ i → trans (boring-petal i) (ppetal i)) i ⟩ 59 | trans (sym $ boring-petal i) (trans (boring-petal i) $ ppetal i) 60 | ≡⟨ trans-sym-trans (boring-petal i) _ ⟩∎ 61 | ppetal i 62 | ∎ 63 | where 64 | boring-petal : ∀ (i : Fin n) → subst (const P) (petal i) pcore ≡ pcore 65 | boring-petal i = subst-const (petal i) pcore 66 | 67 | pflower = Flower-elim[simp] pcore ppetal 68 | -------------------------------------------------------------------------------- /Space/Flower/Omega1.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Paths in flowers 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | open import Univalence 8 | 9 | module Space.Flower.Omega1 10 | (univ : ∀ {ℓ} (A B : Set ℓ) → Univalence-axiom A B) where 11 | 12 | open import Prelude renaming (_∘_ to _⊙_) 13 | open import Path 14 | open import Path.Lemmas 15 | open import Path.Sum 16 | open import Path.HigherOrder 17 | open import Map.H-equivalence hiding (_∘_; id) 18 | open import Map.WeakEquivalence as Weak hiding (_∘_; id) 19 | 20 | import Univalence.Lemmas; open Univalence.Lemmas univ 21 | 22 | import Space.FreeGroup; open Space.FreeGroup univ 23 | open import Space.Flower 24 | 25 | ------------------------------------------------------------------------ 26 | -- From free groups to loops 27 | 28 | private 29 | word⇒petal : ∀ {n} → Word n → core n ≡ core n 30 | word⇒petal [] = refl _ 31 | word⇒petal (gen i ∷ xs) = trans (word⇒petal xs) (petal i) 32 | word⇒petal (inv i ∷ xs) = trans (word⇒petal xs) (sym $ petal i) 33 | 34 | free-group⇒petal : ∀ {n} → FreeGroup n → core n ≡ core n 35 | free-group⇒petal (xs , n) = word⇒petal xs 36 | 37 | ------------------------------------------------------------------------ 38 | -- From loops to free groups 39 | 40 | private 41 | cons-gen : ∀ {n} (i : Fin n) → FreeGroup n → FreeGroup n 42 | cons-gen i = cons (gen i) 43 | 44 | cons-inv : ∀ {n} (i : Fin n) → FreeGroup n → FreeGroup n 45 | cons-inv i = cons (inv i) 46 | 47 | cons-gen-inv : ∀ {n} i (f : FreeGroup n) → cons-gen i (cons-inv i f) ≡ f 48 | cons-gen-inv i = cons₂-flipʳ (gen i) 49 | 50 | cons-inv-gen : ∀ {n} i (f : FreeGroup n) → cons-inv i (cons-gen i f) ≡ f 51 | cons-inv-gen i = cons₂-flipʳ (inv i) 52 | 53 | -- The non-trivial homotopies equivalence between the free group 54 | cons-gen-↔ : ∀ {n} → Fin n → FreeGroup n ↔ FreeGroup n 55 | cons-gen-↔ i = record 56 | { surjection = record 57 | { to = cons-gen i 58 | ; from = cons-inv i 59 | ; right-inverse-of = cons-gen-inv i 60 | } 61 | ; left-inverse-of = cons-inv-gen i 62 | } 63 | 64 | cons-gen-≈ : ∀ {n} → Fin n → FreeGroup n ≈ FreeGroup n 65 | cons-gen-≈ = ↔⇒≈ ⊙ cons-gen-↔ 66 | 67 | cons-gen-≡ : ∀ {n} → Fin n → FreeGroup n ≡ FreeGroup n 68 | cons-gen-≡ = ≈⇒≡ ⊙ cons-gen-≈ 69 | 70 | -- The standard trick to build a universal covering 71 | C : ∀ {n} → Flower n → Set 72 | C {n} = Flower-elim[simp] (FreeGroup n) cons-gen-≡ 73 | 74 | abstract 75 | subst-C-petal : ∀ {n} i (f : FreeGroup n) → subst C (petal i) f ≡ cons-gen i f 76 | subst-C-petal {n} i f = 77 | subst C (petal i) f 78 | ≡⟨ sym $ subst-cong id C (petal i) f ⟩ 79 | subst id (cong C (petal i)) f 80 | ≡⟨ cong (λ p → subst id p f) $ Flower-elim[simp]-petal (FreeGroup n) cons-gen-≡ i ⟩ 81 | subst id (cons-gen-≡ i) f 82 | ≡⟨ subst-id-univ (cons-gen-≡ i) f ⟩ 83 | _≈_.to (≡⇒≈ $ cons-gen-≡ i) f 84 | ≡⟨ cong (λ weq → _≈_.to weq f) $ right-inverse-of $ cons-gen-≈ i ⟩∎ 85 | cons-gen i f 86 | ∎ 87 | where 88 | open _≈_ (≡≈≈ (FreeGroup n) (FreeGroup n)) 89 | 90 | subst-C-petal-cons-inv : ∀ {n} i (f : FreeGroup n) → subst C (petal i) (cons-inv i f) ≡ f 91 | subst-C-petal-cons-inv {n} i f = 92 | subst C (petal i) (cons-inv i f) 93 | ≡⟨ subst-C-petal i (cons-inv i f) ⟩ 94 | cons-gen i (cons-inv i f) 95 | ≡⟨ cons-gen-inv i f ⟩∎ 96 | f 97 | ∎ 98 | 99 | subst-C-sym-petal : ∀ {n} i (f : FreeGroup n) → subst C (sym $ petal i) f ≡ cons-inv i f 100 | subst-C-sym-petal i f = subst-sym C (petal i) f (cons-inv i f) $ subst-C-petal-cons-inv i f 101 | 102 | -- The counting function 103 | petal⇒free-group : ∀ {n} → core n ≡ core n → FreeGroup n 104 | petal⇒free-group p = subst C p unit 105 | 106 | ------------------------------------------------------------------------ 107 | -- Homotopy equivalence 108 | 109 | private 110 | private 111 | -- The right inverse 112 | right-inverse-of′ : ∀ {n} (xs : Word n) → Normalized xs → 113 | proj₁ (petal⇒free-group (word⇒petal xs)) ≡ xs 114 | right-inverse-of′ [] ∅ = refl _ 115 | right-inverse-of′ (gen i ∷ []) ✭ = 116 | proj₁ (subst C (petal i) unit) 117 | ≡⟨ cong proj₁ $ subst-C-petal i unit ⟩∎ 118 | gen i ∷ [] 119 | ∎ 120 | right-inverse-of′ (inv i ∷ []) ✭ = 121 | proj₁ (subst C (sym $ petal i) unit) 122 | ≡⟨ cong proj₁ $ subst-C-sym-petal i unit ⟩∎ 123 | inv i ∷ [] 124 | ∎ 125 | right-inverse-of′ (gen i ∷ x ∷ xs) (s ▸ n) = 126 | proj₁ (subst C (trans (word⇒petal $ x ∷ xs) (petal i)) unit) 127 | ≡⟨ cong proj₁ $ subst-trans C (word⇒petal $ x ∷ xs) (petal i) unit ⟩ 128 | proj₁ (subst C (petal i) $ subst C (word⇒petal $ x ∷ xs) unit) 129 | ≡⟨ cong (proj₁ ⊙ subst C (petal i)) $ word-≡⇒free-group-≡ $ right-inverse-of′ (x ∷ xs) n ⟩ 130 | proj₁ (subst C (petal i) (x ∷ xs , n)) 131 | ≡⟨ cong proj₁ $ subst-C-petal i (x ∷ xs , n) ⟩ 132 | proj₁ (cons-gen i (x ∷ xs , n)) 133 | ≡⟨ cong proj₁ $ cons-stable s n ⟩∎ 134 | gen i ∷ x ∷ xs 135 | ∎ 136 | right-inverse-of′ (inv i ∷ x ∷ xs) (s ▸ n) = 137 | proj₁ (subst C (trans (word⇒petal $ x ∷ xs) (sym $ petal i)) unit) 138 | ≡⟨ cong proj₁ $ subst-trans C (word⇒petal $ x ∷ xs) (sym $ petal i) unit ⟩ 139 | proj₁ (subst C (sym $ petal i) $ subst C (word⇒petal $ x ∷ xs) unit) 140 | ≡⟨ cong (proj₁ ⊙ subst C (sym $ petal i)) $ word-≡⇒free-group-≡ $ right-inverse-of′ (x ∷ xs) n ⟩ 141 | proj₁ (subst C (sym $ petal i) (x ∷ xs , n)) 142 | ≡⟨ cong proj₁ $ subst-C-sym-petal i (x ∷ xs , n) ⟩ 143 | proj₁ (cons-inv i (x ∷ xs , n)) 144 | ≡⟨ cong proj₁ $ cons-stable s n ⟩∎ 145 | inv i ∷ x ∷ xs 146 | ∎ 147 | 148 | right-inverse-of : ∀ {n} (f : FreeGroup n) → petal⇒free-group (free-group⇒petal f) ≡ f 149 | right-inverse-of (xs , n) = word-≡⇒free-group-≡ $ right-inverse-of′ xs n 150 | 151 | -- Intuition: 152 | -- 153 | -- J rule works on paths, not loops... 154 | -- Let's free one of the end points! 155 | 156 | free-group⇒petal-cons-inv : ∀ {n} i (f : FreeGroup n) → 157 | free-group⇒petal (cons-inv i f) ≡ trans (free-group⇒petal f) (sym $ petal i) 158 | free-group⇒petal-cons-inv i ([] , ∅) = refl _ 159 | free-group⇒petal-cons-inv i (x ∷ xs , n) with decideStable (inv i) x 160 | ... | inj₁ _ = refl _ 161 | ... | inj₂ ¬s = 162 | word⇒petal xs 163 | ≡⟨ sym $ trans-trans-symʳʳ (word⇒petal xs) (petal i) ⟩ 164 | trans (word⇒petal (gen i ∷ xs)) (sym $ petal i) 165 | ≡⟨ cong (λ x → trans (word⇒petal (x ∷ xs)) (sym $ petal i)) $ sym $ ¬stable⇒flip ¬s ⟩∎ 166 | trans (word⇒petal (x ∷ xs)) (sym $ petal i) 167 | ∎ 168 | 169 | -- Data to be fed into Flower-elim 170 | free-group⇒petal-petal : ∀ {n} i (f : FreeGroup n) → 171 | subst (λ x → C x → core n ≡ x) (petal i) free-group⇒petal f ≡ free-group⇒petal f 172 | free-group⇒petal-petal {n} i f = 173 | subst (λ x → C x → core n ≡ x) (petal i) free-group⇒petal f 174 | ≡⟨ subst-func C (λ x → core n ≡ x) (petal i) free-group⇒petal (subst-C-petal-cons-inv i f) ⟩ 175 | subst (λ x → core n ≡ x) (petal i) (free-group⇒petal $ cons-inv i f) 176 | ≡⟨ subst-path[idʳ] (const $ core n) (petal i) (free-group⇒petal $ cons-inv i f) ⟩ 177 | trans (sym $ cong (const $ core n) $ petal i) (trans (free-group⇒petal $ cons-inv i f) $ petal i) 178 | ≡⟨ cong (λ p → trans (sym p) (trans (free-group⇒petal $ cons-inv i f) $ petal i)) $ cong-const (core n) (petal i) ⟩ 179 | trans (free-group⇒petal $ cons-inv i f) (petal i) 180 | ≡⟨ cong (λ p → trans p (petal i)) $ free-group⇒petal-cons-inv i f ⟩ 181 | trans (trans (free-group⇒petal f) (sym $ petal i)) (petal i) 182 | ≡⟨ trans-trans-symʳ (free-group⇒petal f) (petal i) ⟩∎ 183 | free-group⇒petal f 184 | ∎ 185 | 186 | -- A generalized Cx⇒petal ... Cx⇒fractal ! 187 | Cx⇒fractal : ∀ {n} (x : Flower n) → C x → core n ≡ x 188 | Cx⇒fractal {n} = Flower-elim (λ x → C x → core n ≡ x) free-group⇒petal (ext ⊙ free-group⇒petal-petal) 189 | 190 | -- A generalized S¹loop⇒Cx ... S¹path⇒Cx ! 191 | fractal⇒Cx : ∀ {n} (x : Flower n) → core n ≡ x → C x 192 | fractal⇒Cx x p = subst C p unit 193 | 194 | -- The generalized left inversion property (now easy?) 195 | left-inverse-of′ : ∀ {n} (x : Flower n) → (p : core n ≡ x) → Cx⇒fractal x (fractal⇒Cx x p) ≡ p 196 | left-inverse-of′ x = elim′ (λ {x} p → Cx⇒fractal x (fractal⇒Cx x p) ≡ p) (refl _) 197 | 198 | -- The left inversion property we want 199 | left-inverse-of : ∀ {n} → (p : core n ≡ core n) → free-group⇒petal (petal⇒free-group p) ≡ p 200 | left-inverse-of {n} = left-inverse-of′ $ core n 201 | 202 | -- We have all the ingredients now! 203 | Ω₁-flower↔free-group : ∀ n → Ω 1 (core n) ↔ FreeGroup n 204 | Ω₁-flower↔free-group n = 205 | record 206 | { surjection = record 207 | { to = petal⇒free-group 208 | ; from = free-group⇒petal 209 | ; right-inverse-of = right-inverse-of 210 | } 211 | ; left-inverse-of = left-inverse-of 212 | } 213 | -------------------------------------------------------------------------------- /Space/FreeGroup.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Free groups 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | open import Univalence 10 | 11 | module Space.FreeGroup 12 | (univ : ∀ {ℓ} (A B : Set ℓ) → Univalence-axiom A B) where 13 | 14 | open import Prelude renaming (_∘_ to _⊙_) 15 | 16 | open import Path 17 | open import Path.Sum 18 | open import Path.Lemmas 19 | 20 | import Univalence.Lemmas; open Univalence.Lemmas univ 21 | 22 | open import Space.Fin.Lemmas as F 23 | open import Space.List.Lemmas as L 24 | 25 | -- ... was Alphabet n = Fin n ⊎ Fin n before 26 | data Alphabet (n : ℕ) : Set where 27 | gen : Fin n → Alphabet n 28 | inv : Fin n → Alphabet n 29 | 30 | Word : ℕ → Set 31 | Word n = List (Alphabet n) 32 | 33 | -- Positive definition (following Dan Licata's advice) 34 | data Stable {n} : Alphabet n → Alphabet n → Set where 35 | gen-gen : ∀ {x y} → Stable (gen x) (gen y) 36 | inv-inv : ∀ {x y} → Stable (inv x) (inv y) 37 | gen-inv : ∀ {x y} → ¬ x ≡ y → Stable (gen x) (inv y) 38 | inv-gen : ∀ {x y} → ¬ x ≡ y → Stable (inv x) (gen y) 39 | 40 | decideStable : ∀ {n} (x y : Alphabet n) → Dec (Stable x y) 41 | decideStable (gen x) (gen y) = inj₁ gen-gen 42 | decideStable (inv x) (inv y) = inj₁ inv-inv 43 | decideStable {n} (gen x) (inv y) with F.decide-≡ x y 44 | ... | inj₁ eq = inj₂ λ{(gen-inv neq) → neq eq} 45 | ... | inj₂ neq = inj₁ $ gen-inv neq 46 | decideStable {n} (inv x) (gen y) with F.decide-≡ x y 47 | ... | inj₁ eq = inj₂ λ{(inv-gen neq) → neq eq} 48 | ... | inj₂ neq = inj₁ $ inv-gen neq 49 | 50 | -- Positive definition (following Dan Licata's advice) 51 | data Normalized {n} : Word n → Set where 52 | ∅ : Normalized [] 53 | ✭ : ∀ {x} → Normalized [ x ] 54 | _▸_ : ∀ {x₁ x₂} {xs} → Stable x₁ x₂ → 55 | Normalized (x₂ ∷ xs) → 56 | Normalized (x₁ ∷ x₂ ∷ xs) 57 | 58 | FreeGroup : ℕ → Set 59 | FreeGroup n = Σ (Word n) Normalized 60 | 61 | ------------------------------------------------------------------------ 62 | -- misc 63 | 64 | stable-sym : ∀ {n} {x y : Alphabet n} → Stable x y → Stable y x 65 | stable-sym gen-gen = gen-gen 66 | stable-sym inv-inv = inv-inv 67 | stable-sym (gen-inv neq) = inv-gen $ neq ⊙ sym 68 | stable-sym (inv-gen neq) = gen-inv $ neq ⊙ sym 69 | 70 | private 71 | -- Negative definition 72 | data Unstable {n} : Alphabet n → Alphabet n → Set where 73 | gen-inv : ∀ {x} → Unstable (gen x) (inv x) 74 | inv-gen : ∀ {x} → Unstable (inv x) (gen x) 75 | 76 | ¬stable⇒unstable : ∀ {n} {x y : Alphabet n} → ¬ Stable x y → Unstable x y 77 | ¬stable⇒unstable {x = gen x} {y = inv y} ns with F.decide-≡ x y 78 | ... | inj₁ eq = subst (λ y → Unstable (gen x) (inv y)) eq gen-inv 79 | ... | inj₂ ne = ⊥-elim $ ns $ gen-inv ne 80 | ¬stable⇒unstable {x = inv x} {y = gen y} ns with F.decide-≡ x y 81 | ... | inj₁ eq = subst (λ y → Unstable (inv x) (gen y)) eq inv-gen 82 | ... | inj₂ ne = ⊥-elim $ ns $ inv-gen ne 83 | ¬stable⇒unstable {x = gen x} {y = gen y} ns = ⊥-elim $ ns gen-gen 84 | ¬stable⇒unstable {x = inv x} {y = inv y} ns = ⊥-elim $ ns inv-inv 85 | 86 | unstable-involutive : ∀ {n} {x y z : Alphabet n} → Unstable x y → Unstable y z → x ≡ z 87 | unstable-involutive gen-inv inv-gen = refl _ 88 | unstable-involutive inv-gen gen-inv = refl _ 89 | 90 | ¬stable-involutive : ∀ {n} {x y z : Alphabet n} → ¬ Stable x y → ¬ Stable y z → x ≡ z 91 | ¬stable-involutive ns₁ ns₂ = unstable-involutive (¬stable⇒unstable ns₁) (¬stable⇒unstable ns₂) 92 | 93 | ------------------------------------------------------------------------ 94 | -- head, tail, cons 95 | 96 | private 97 | data HeadStable {n} : Word n → Word n → Set where 98 | heads : ∀ {x y} {xs ys} → Stable x y → 99 | HeadStable (x ∷ xs) (y ∷ ys) 100 | left : ∀ {xs} → HeadStable xs [] 101 | right : ∀ {xs} → HeadStable [] xs 102 | 103 | norm-head : ∀ {n} {x} {xs : Word n} → 104 | Normalized (x ∷ xs) → ∀ ys → HeadStable xs (x ∷ ys) 105 | norm-head ✭ _ = right 106 | norm-head (s ▸ _) _ = heads $ stable-sym s 107 | 108 | norm-tail : ∀ {n} {x} {xs : Word n} → 109 | Normalized (x ∷ xs) → Normalized xs 110 | norm-tail ✭ = ∅ 111 | norm-tail (_ ▸ n) = n 112 | 113 | norm-cons : ∀ {n} {x} {xs ys : Word n} → HeadStable (x ∷ xs) ys → Normalized ys → 114 | Normalized (x ∷ ys) 115 | norm-cons left ∅ = ✭ 116 | norm-cons (heads s) n = s ▸ n 117 | 118 | ------------------------------------------------------------------------ 119 | -- head-reduce 120 | 121 | private 122 | word-head-reduce : ∀ {n} → Alphabet n → Word n → Word n 123 | word-head-reduce x [] = [ x ] 124 | word-head-reduce x (y ∷ ys) with decideStable x y 125 | ... | inj₁ _ = x ∷ y ∷ ys 126 | ... | inj₂ _ = ys 127 | 128 | norm-head-reduce : ∀ {n} x {xs : Word n} → Normalized xs → 129 | Normalized (word-head-reduce x xs) 130 | norm-head-reduce x ∅ = ✭ 131 | norm-head-reduce x {y ∷ _} n with decideStable x y 132 | ... | inj₁ s = s ▸ n 133 | ... | inj₂ _ = norm-tail n 134 | 135 | word-head-reduce-¬stable : ∀ {n} {x y : Alphabet n} → ¬ Stable x y → 136 | ∀ {ys} → word-head-reduce x (y ∷ ys) ≡ ys 137 | word-head-reduce-¬stable {x = x} {y} ¬s with decideStable x y 138 | ... | inj₁ s = ⊥-elim $ ¬s s 139 | ... | inj₂ _ = refl _ 140 | 141 | word-head-reduce₂-¬stable : ∀ {n} {x y : Alphabet n} → ¬ Stable x y → 142 | ∀ {zs} → Normalized zs → 143 | word-head-reduce x (word-head-reduce y zs) ≡ zs 144 | word-head-reduce₂-¬stable {x = x} {y} nsxy ∅ with decideStable x y 145 | ... | inj₁ s = ⊥-elim $ nsxy s 146 | ... | inj₂ _ = refl _ 147 | word-head-reduce₂-¬stable {x = x} {y} nsxy {z ∷ []} ✭ with decideStable y z 148 | ... | inj₂ nsyz = cong [_] $ ¬stable-involutive nsxy nsyz 149 | ... | inj₁ _ with decideStable x y 150 | ... | inj₁ s = ⊥-elim $ nsxy s 151 | ... | inj₂ _ = refl _ 152 | word-head-reduce₂-¬stable {x = x} {y} nsxy {z₁ ∷ z₂ ∷ zs} (s ▸ _) with decideStable y z₁ 153 | ... | inj₁ _ with decideStable x y 154 | ... | inj₁ sxy = ⊥-elim $ nsxy sxy 155 | ... | inj₂ _ = refl _ 156 | word-head-reduce₂-¬stable {x = x} {y} nsxy {z₁ ∷ z₂ ∷ zs} (s ▸ _) | inj₂ u with decideStable x z₂ 157 | ... | inj₁ _ = cong (λ x → x ∷ z₂ ∷ zs) $ ¬stable-involutive nsxy u 158 | ... | inj₂ u₂ = ⊥-elim $ subst (λ x → ¬ Stable x z₂) (¬stable-involutive nsxy u) u₂ s 159 | 160 | word-head-reduce₂-stable : ∀ {n} x₁ x₂ (xs : Word n) → Stable x₁ x₂ → 161 | word-head-reduce x₁ (x₂ ∷ xs) ≡ x₁ ∷ x₂ ∷ xs 162 | word-head-reduce₂-stable x₁ x₂ x s with decideStable x₁ x₂ 163 | ... | inj₁ _ = refl _ 164 | ... | inj₂ u = ⊥-elim $ u s 165 | 166 | ------------------------------------------------------------------------ 167 | -- append-reduce 168 | 169 | private 170 | word-append-reduce : ∀ {n} → Word n → Word n → Word n 171 | word-append-reduce xs ys = foldr word-head-reduce ys xs 172 | 173 | norm-append-reduce : ∀ {n} xs {ys : Word n} → Normalized ys → 174 | Normalized (word-append-reduce xs ys) 175 | norm-append-reduce [] n = n 176 | norm-append-reduce (x ∷ xs) n = norm-head-reduce x $ norm-append-reduce xs n 177 | 178 | word-snoc-append-reduce : ∀ {n} xs x (ys : Word n) → 179 | word-append-reduce (snoc xs x) ys ≡ 180 | word-append-reduce xs (word-head-reduce x ys) 181 | word-snoc-append-reduce [] x′ ys = refl _ 182 | word-snoc-append-reduce (x ∷ xs) x′ ys = 183 | cong (word-head-reduce x) $ word-snoc-append-reduce xs x′ ys 184 | 185 | word-append-reduce-head-reduce : ∀ {n} x (ys : Word n) {zs} → Normalized zs → 186 | word-append-reduce (word-head-reduce x ys) zs ≡ 187 | word-head-reduce x (word-append-reduce ys zs) 188 | word-append-reduce-head-reduce _ [] _ = refl _ 189 | word-append-reduce-head-reduce x (y ∷ ys) nz with decideStable x y 190 | ... | inj₁ _ = refl _ 191 | ... | inj₂ u = sym $ word-head-reduce₂-¬stable u (norm-append-reduce ys nz) 192 | 193 | word-append-reduce-assoc : ∀ {n} (xs ys : Word n) {zs} → Normalized zs → 194 | word-append-reduce (word-append-reduce xs ys) zs ≡ 195 | word-append-reduce xs (word-append-reduce ys zs) 196 | word-append-reduce-assoc [] _ nz = refl _ 197 | word-append-reduce-assoc (x ∷ xs) ys {zs} nz = 198 | word-append-reduce (word-head-reduce x $ word-append-reduce xs ys) zs 199 | ≡⟨ word-append-reduce-head-reduce x (word-append-reduce xs ys) nz ⟩ 200 | word-head-reduce x (word-append-reduce (word-append-reduce xs ys) zs) 201 | ≡⟨ cong (word-head-reduce x) $ word-append-reduce-assoc xs ys nz ⟩∎ 202 | word-head-reduce x (word-append-reduce xs $ word-append-reduce ys zs) 203 | ∎ 204 | 205 | word-append-reduce-[] : ∀ {n} {xs : Word n} → Normalized xs → 206 | word-append-reduce xs [] ≡ xs 207 | word-append-reduce-[] ∅ = refl _ 208 | word-append-reduce-[] ✭ = refl _ 209 | word-append-reduce-[] {xs = x₁ ∷ x₂ ∷ xs} (s ▸ n) = 210 | word-head-reduce x₁ (word-append-reduce (x₂ ∷ xs) []) 211 | ≡⟨ cong (word-head-reduce x₁) $ word-append-reduce-[] n ⟩ 212 | word-head-reduce x₁ (x₂ ∷ xs) 213 | ≡⟨ word-head-reduce₂-stable x₁ x₂ xs s ⟩∎ 214 | x₁ ∷ x₂ ∷ xs 215 | ∎ 216 | 217 | ------------------------------------------------------------------------ 218 | -- rev 219 | 220 | private 221 | private 222 | lemma₀ : ∀ {n} (xs ys : Word n) → Normalized xs → Normalized ys → HeadStable xs ys → 223 | Normalized (rev xs ++ ys) 224 | lemma₀ [] ys _ ny _ = ny 225 | lemma₀ (x ∷ xs) ys nx ny hs = 226 | subst Normalized (sym $ snoc-++ (rev xs) x ys) $ 227 | lemma₀ xs (x ∷ ys) (norm-tail nx) (norm-cons hs ny) (norm-head nx ys) 228 | 229 | norm-rev-++ : ∀ {n} {xs ys : Word n} → Normalized xs → Normalized ys → HeadStable xs ys → 230 | Normalized (rev xs ++ ys) 231 | norm-rev-++ nx ny hs = lemma₀ _ _ nx ny hs 232 | 233 | norm-rev : ∀ {n} {xs : Word n} → Normalized xs → Normalized (rev xs) 234 | norm-rev nx = subst Normalized (++-[] _) $ norm-rev-++ nx ∅ left 235 | 236 | ------------------------------------------------------------------------ 237 | -- flip 238 | 239 | alphabet-flip : ∀ {n} → Alphabet n → Alphabet n 240 | alphabet-flip (gen x) = inv x 241 | alphabet-flip (inv x) = gen x 242 | 243 | alphabet-flip-flip : ∀ {n} (x : Alphabet n) → alphabet-flip (alphabet-flip x) ≡ x 244 | alphabet-flip-flip (gen x) = refl _ 245 | alphabet-flip-flip (inv x) = refl _ 246 | 247 | stable-flip : ∀ {n} {x₁ x₂ : Alphabet n} → Stable x₁ x₂ → 248 | Stable (alphabet-flip x₁) (alphabet-flip x₂) 249 | stable-flip gen-gen = inv-inv 250 | stable-flip inv-inv = gen-gen 251 | stable-flip (gen-inv neq) = inv-gen neq 252 | stable-flip (inv-gen neq) = gen-inv neq 253 | 254 | stable-flip⁻¹ : ∀ {n} {x₁ x₂ : Alphabet n} → 255 | Stable (alphabet-flip x₁) (alphabet-flip x₂) → 256 | Stable x₁ x₂ 257 | stable-flip⁻¹ s = 258 | subst id (cong₂ Stable (alphabet-flip-flip _) (alphabet-flip-flip _)) $ 259 | stable-flip s 260 | 261 | word-flip : ∀ {n} → Word n → Word n 262 | word-flip xs = map alphabet-flip xs 263 | 264 | word-flip-flip : ∀ {n} (xs : Word n) → word-flip (word-flip xs) ≡ xs 265 | word-flip-flip [] = refl _ 266 | word-flip-flip (x ∷ xs) = cong₂ (λ x xs → x ∷ xs) (alphabet-flip-flip x) (word-flip-flip xs) 267 | 268 | norm-flip : ∀ {n} {xs : Word n} → Normalized xs → 269 | Normalized (word-flip xs) 270 | norm-flip ∅ = ∅ 271 | norm-flip ✭ = ✭ 272 | norm-flip (s ▸ n) = stable-flip s ▸ norm-flip n 273 | 274 | word-flip-snoc : ∀ {n} (xs : Word n) y → word-flip (snoc xs y) ≡ snoc (word-flip xs) (alphabet-flip y) 275 | word-flip-snoc [] _ = refl _ 276 | word-flip-snoc (x ∷ xs) y = cong (λ zs → alphabet-flip x ∷ zs) $ word-flip-snoc xs y 277 | 278 | word-flip-rev : ∀ {n} (xs : Word n) → word-flip (rev xs) ≡ rev (word-flip xs) 279 | word-flip-rev [] = refl _ 280 | word-flip-rev (x ∷ xs) = 281 | word-flip (snoc (rev xs) x) 282 | ≡⟨ word-flip-snoc (rev xs) x ⟩ 283 | snoc (word-flip $ rev xs) (alphabet-flip x) 284 | ≡⟨ cong (λ xs → snoc xs (alphabet-flip x)) $ word-flip-rev xs ⟩∎ 285 | snoc (rev $ word-flip xs) (alphabet-flip x) 286 | ∎ 287 | 288 | ¬stable-flipˡ : ∀ {n} (x : Alphabet n) → ¬ Stable (alphabet-flip x) x 289 | ¬stable-flipˡ (gen x) s = ¬stable-inv-gen (refl _) s 290 | where 291 | ¬stable-inv-gen : ∀ {n} {i j : Fin n} → i ≡ j → ¬ Stable (inv i) (gen j) 292 | ¬stable-inv-gen eq (inv-gen neq) = neq eq 293 | ¬stable-flipˡ (inv x) s = ¬stable-gen-inv (refl _) s 294 | where 295 | ¬stable-gen-inv : ∀ {n} {i j : Fin n} → i ≡ j → ¬ Stable (gen i) (inv j) 296 | ¬stable-gen-inv eq (gen-inv neq) = neq eq 297 | 298 | ¬stable-flipʳ : ∀ {n} (x : Alphabet n) → ¬ Stable x (alphabet-flip x) 299 | ¬stable-flipʳ x = ¬stable-flipˡ x ⊙ stable-sym 300 | 301 | ¬stable⇒flip : ∀ {n} {x y : Alphabet n} → ¬ Stable x y → y ≡ alphabet-flip x 302 | ¬stable⇒flip ¬s with ¬stable⇒unstable ¬s 303 | ... | gen-inv = refl _ 304 | ... | inv-gen = refl _ 305 | 306 | ------------------------------------------------------------------------ 307 | -- inverse 308 | 309 | private 310 | word-inverse : ∀ {n} → Word n → Word n 311 | word-inverse xs = rev $ word-flip xs 312 | 313 | word-inverse-inverse : ∀ {n} (xs : Word n) → word-inverse (word-inverse xs) ≡ xs 314 | word-inverse-inverse xs = 315 | rev (word-flip $ rev $ word-flip xs) 316 | ≡⟨ cong rev $ word-flip-rev $ word-flip xs ⟩ 317 | (rev $ rev $ word-flip $ word-flip xs) 318 | ≡⟨ rev-rev (word-flip $ word-flip xs) ⟩ 319 | (word-flip $ word-flip xs) 320 | ≡⟨ word-flip-flip xs ⟩∎ 321 | xs 322 | ∎ 323 | 324 | norm-inverse : ∀ {n} {xs : Word n} → Normalized xs → Normalized (word-inverse xs) 325 | norm-inverse n = norm-rev $ norm-flip n 326 | 327 | private 328 | lemma₁ : ∀ {n} x (xs : Word n) → word-head-reduce (alphabet-flip x) (x ∷ xs) ≡ xs 329 | lemma₁ (gen x) xs with F.decide-≡ x x 330 | ... | inj₁ _ = refl _ 331 | ... | inj₂ neq = ⊥-elim $ neq $ refl _ 332 | lemma₁ (inv x) xs with F.decide-≡ x x 333 | ... | inj₁ _ = refl _ 334 | ... | inj₂ neq = ⊥-elim $ neq $ refl _ 335 | 336 | word-inverse-append-reduce : ∀ {n} (xs : Word n) → word-append-reduce (word-inverse xs) xs ≡ [] 337 | word-inverse-append-reduce [] = refl _ 338 | word-inverse-append-reduce (x ∷ xs) = 339 | word-append-reduce (snoc (word-inverse xs) (alphabet-flip x)) (x ∷ xs) 340 | ≡⟨ word-snoc-append-reduce (word-inverse xs) (alphabet-flip x) (x ∷ xs) ⟩ 341 | word-append-reduce (word-inverse xs) (word-head-reduce (alphabet-flip x) (x ∷ xs)) 342 | ≡⟨ cong (word-append-reduce (word-inverse xs)) $ lemma₁ x xs ⟩ 343 | word-append-reduce (word-inverse xs) xs 344 | ≡⟨ word-inverse-append-reduce xs ⟩∎ 345 | [] 346 | ∎ 347 | 348 | word-append-reduce-inverse : ∀ {n} (xs : Word n) → word-append-reduce xs (word-inverse xs) ≡ [] 349 | word-append-reduce-inverse xs = 350 | word-append-reduce xs (word-inverse xs) 351 | ≡⟨ cong (λ xs₁ → word-append-reduce xs₁ (word-inverse xs)) $ sym $ word-inverse-inverse xs ⟩ 352 | word-append-reduce (word-inverse $ word-inverse xs) (word-inverse xs) 353 | ≡⟨ word-inverse-append-reduce (word-inverse xs) ⟩∎ 354 | [] 355 | ∎ 356 | 357 | ------------------------------------------------------------------------ 358 | -- the public interface for a group 359 | 360 | unit : ∀ {n} → FreeGroup n 361 | unit = [] , ∅ 362 | 363 | _∘_ : ∀ {n} → FreeGroup n → FreeGroup n → FreeGroup n 364 | (xs₁ , n₁) ∘ (xs₂ , n₂) = word-append-reduce xs₁ xs₂ , norm-append-reduce xs₁ n₂ 365 | 366 | private 367 | unique-stable-proof : ∀ {n} {x y : Alphabet n} (s₁ s₂ : Stable x y) → s₁ ≡ s₂ 368 | unique-stable-proof gen-gen gen-gen = refl _ 369 | unique-stable-proof inv-inv inv-inv = refl _ 370 | unique-stable-proof (gen-inv neq₁) (gen-inv neq₂) = cong gen-inv $ unique-neq-proof neq₁ neq₂ 371 | unique-stable-proof (inv-gen neq₁) (inv-gen neq₂) = cong inv-gen $ unique-neq-proof neq₁ neq₂ 372 | 373 | unique-normalized-proof : ∀ {n} {xs : Word n} (n₁ n₂ : Normalized xs) → n₁ ≡ n₂ 374 | unique-normalized-proof ∅ ∅ = refl _ 375 | unique-normalized-proof ✭ ✭ = refl _ 376 | unique-normalized-proof (s₁ ▸ n₁) (s₂ ▸ n₂) = 377 | cong₂ _▸_ (unique-stable-proof s₁ s₂) (unique-normalized-proof n₁ n₂) 378 | 379 | word-≡⇒free-group-≡ : ∀ {n} {f₁ f₂ : FreeGroup n} → proj₁ f₁ ≡ proj₁ f₂ → f₁ ≡ f₂ 380 | word-≡⇒free-group-≡ {f₁ = xs₁ , n₁} {f₂ = xs₂ , n₂} eq = 381 | Σ≡⇒≡Σ Normalized $ eq , unique-normalized-proof (subst Normalized eq n₁) n₂ 382 | 383 | ∘-unitˡ : ∀ {n} (f : FreeGroup n) → unit ∘ f ≡ f 384 | ∘-unitˡ _ = refl _ 385 | 386 | ∘-unitʳ : ∀ {n} (f : FreeGroup n) → f ∘ unit ≡ f 387 | ∘-unitʳ (_ , n) = word-≡⇒free-group-≡ $ word-append-reduce-[] n 388 | 389 | ∘-assoc : ∀ {n} (f₁ f₂ f₃ : FreeGroup n) → (f₁ ∘ f₂) ∘ f₃ ≡ f₁ ∘ (f₂ ∘ f₃) 390 | ∘-assoc (xs₁ , _) (xs₂ , _) (_ , n₃) = 391 | word-≡⇒free-group-≡ $ word-append-reduce-assoc xs₁ xs₂ n₃ 392 | 393 | _⁻¹ : ∀ {n} → FreeGroup n → FreeGroup n 394 | (xs , n) ⁻¹ = word-inverse xs , norm-inverse n 395 | 396 | ∘-invˡ : ∀ {n} (f : FreeGroup n) → (f ⁻¹) ∘ f ≡ unit 397 | ∘-invˡ (xs , _) = word-≡⇒free-group-≡ $ word-inverse-append-reduce xs 398 | 399 | ∘-invʳ : ∀ {n} (f : FreeGroup n) → f ∘ (f ⁻¹) ≡ unit 400 | ∘-invʳ (xs , _) = word-≡⇒free-group-≡ $ word-append-reduce-inverse xs 401 | 402 | cons : ∀ {n} → Alphabet n → FreeGroup n → FreeGroup n 403 | cons x (xs , n) = word-head-reduce x xs , norm-head-reduce x n 404 | 405 | private 406 | cons₂-stable′ : ∀ {n} {x y} (s : Stable x y) (ys : Word n) → 407 | word-head-reduce x (y ∷ ys) ≡ x ∷ y ∷ ys 408 | cons₂-stable′ {x = x} {y} s ys with decideStable x y 409 | ... | inj₁ _ = refl _ 410 | ... | inj₂ u = ⊥-elim $ u s 411 | 412 | cons-stable : ∀ {n} {x y} (s : Stable x y) {ys : Word n} n → 413 | cons x (y ∷ ys , n) ≡ (x ∷ y ∷ ys , s ▸ n) 414 | cons-stable s {ys} _ = word-≡⇒free-group-≡ $ cons₂-stable′ s ys 415 | 416 | cons-¬stable : ∀ {n} {x y} (us : ¬ Stable x y) {ys : Word n} {n} → 417 | cons x (y ∷ ys , n) ≡ (ys , norm-tail n) 418 | cons-¬stable ¬s = word-≡⇒free-group-≡ $ word-head-reduce-¬stable ¬s 419 | 420 | cons₂-flipʳ : ∀ {n} x (f : FreeGroup n) → cons x (cons (alphabet-flip x) f) ≡ f 421 | cons₂-flipʳ x (xs , n) = word-≡⇒free-group-≡ $ word-head-reduce₂-¬stable (¬stable-flipʳ x) n 422 | 423 | cons₂-flipˡ : ∀ {n} x (f : FreeGroup n) → cons (alphabet-flip x) (cons x f) ≡ f 424 | cons₂-flipˡ x (xs , n) = word-≡⇒free-group-≡ $ word-head-reduce₂-¬stable (¬stable-flipˡ x) n 425 | -------------------------------------------------------------------------------- /Space/Integer.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Integers and their friends 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Space.Integer where 10 | 11 | open import Prelude renaming (suc to ℕsuc; zero to ℕzero) 12 | 13 | ------------------------------------------------------------------------ 14 | -- Integers 15 | 16 | -- Note that ⟦ pos n ⟧ = n + 1 and ⟦ neg n ⟧ = - (n + 1) 17 | 18 | data ℤ : Set where 19 | pos : ℕ → ℤ 20 | neg : ℕ → ℤ 21 | zero : ℤ 22 | 23 | -- An experimental ℤ that might be suitable for induction (??) 24 | data ℤ″ : ℕ → Set where 25 | pos′ : (n : ℕ) → ℤ″ (ℕsuc n) 26 | neg′ : (n : ℕ) → ℤ″ (ℕsuc n) 27 | zero′ : ℤ″ ℕzero 28 | 29 | ℤ′ : Set 30 | ℤ′ = ∃ ℤ″ 31 | -------------------------------------------------------------------------------- /Space/Interval.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Spheres. 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- Currently we have S¹ and an experimental S² 8 | 9 | {-# OPTIONS --without-K #-} 10 | 11 | module Space.Interval where 12 | 13 | open import Prelude hiding (zero) 14 | open import Path 15 | open import Path.Lemmas 16 | 17 | ------------------------------------------------------------------------ 18 | -- Formation rules and introduction rules 19 | 20 | private 21 | data I′ : Set where 22 | zero′ : I′ 23 | one′ : I′ 24 | 25 | I : Set 26 | I = I′ 27 | 28 | zero : I 29 | zero = zero′ 30 | 31 | one : I 32 | one = one′ 33 | 34 | postulate 35 | seg : zero ≡ one 36 | 37 | I-rec : ∀ {ℓ} (P : I → Set ℓ) (pzero : P zero) (pone : P one) → subst P seg pzero ≡ pone → (x : I) → P x 38 | I-rec P pzero _ _ zero′ = pzero 39 | I-rec P _ pone _ one′ = pone 40 | 41 | postulate 42 | I-rec-seg : ∀ {ℓ} (P : I → Set ℓ) (pzero : P zero) (pone : P one) (pseg : subst P seg pzero ≡ pone) 43 | → cong[dep] P (I-rec P pzero pone pseg) seg ≡ pseg 44 | -------------------------------------------------------------------------------- /Space/List/Lemmas.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Lemmas for List 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Space.List.Lemmas where 10 | 11 | open import Prelude 12 | open import Path 13 | 14 | snoc : ∀ {ℓ} {A : Set ℓ} → List A → A → List A 15 | snoc [] x = x ∷ [] 16 | snoc (x ∷ xs) v = x ∷ snoc xs v 17 | 18 | rev : ∀ {ℓ} {A : Set ℓ} → List A → List A 19 | rev [] = [] 20 | rev (x ∷ xs) = snoc (rev xs) x 21 | 22 | rev-snoc : ∀ {ℓ} {A : Set ℓ} (l : List A) (x : A) → rev (snoc l x) ≡ x ∷ rev l 23 | rev-snoc [] y = refl _ 24 | rev-snoc (x ∷ xs) y = 25 | snoc (rev (snoc xs y)) x ≡⟨ cong (λ l → snoc l x) $ rev-snoc xs y ⟩∎ 26 | y ∷ snoc (rev xs) x ∎ 27 | 28 | rev-rev : ∀ {ℓ} {A : Set ℓ} (l : List A) → rev (rev l) ≡ l 29 | rev-rev [] = refl _ 30 | rev-rev (x ∷ xs) = 31 | rev (snoc (rev xs) x) ≡⟨ rev-snoc (rev xs) x ⟩ 32 | x ∷ rev (rev xs) ≡⟨ cong (λ l → x ∷ l) $ rev-rev xs ⟩∎ 33 | x ∷ xs ∎ 34 | 35 | snoc-++ : ∀ {ℓ} {A : Set ℓ} l₁ (x : A) l₂ → snoc l₁ x ++ l₂ ≡ l₁ ++ x ∷ l₂ 36 | snoc-++ [] _ _ = refl _ 37 | snoc-++ (x ∷ xs) y ys = cong (λ zs → x ∷ zs) $ snoc-++ xs y ys 38 | 39 | ++-[] : ∀ {ℓ} {A : Set ℓ} (l : List A) → l ++ [] ≡ l 40 | ++-[] [] = refl _ 41 | ++-[] (x ∷ xs) = cong (λ xs → x ∷ xs) $ ++-[] xs 42 | -------------------------------------------------------------------------------- /Space/Nat/Lemmas.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Basic facts about natural numbers. 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Space.Nat.Lemmas where 10 | 11 | open import Prelude 12 | open import Path 13 | 14 | n+0 : ∀ n → n + 0 ≡ n 15 | n+0 0 = refl _ 16 | n+0 (suc n) = cong suc $ n+0 n 17 | 18 | n+suc : ∀ n m → n + suc m ≡ suc n + m 19 | n+suc 0 m = refl _ 20 | n+suc (suc n) m = cong suc $ n+suc n m 21 | -------------------------------------------------------------------------------- /Space/Sphere.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Spheres (base + loop) 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- Currently we have S¹ and an experimental S² 8 | 9 | {-# OPTIONS --without-K #-} 10 | 11 | module Space.Sphere where 12 | 13 | open import Prelude 14 | open import Path 15 | open import Path.Lemmas 16 | open import Path.Sum 17 | open import Path.HigherOrder 18 | 19 | ------------------------------------------------------------------------ 20 | -- Safe area: S¹ only 21 | ------------------------------------------------------------------------ 22 | 23 | ------------------------------------------------------------------------ 24 | -- Formation rules and introduction rules 25 | 26 | private 27 | data S¹′ : Set where 28 | base¹′ : S¹′ 29 | 30 | S¹ : Set 31 | S¹ = S¹′ 32 | 33 | base¹ : S¹ 34 | base¹ = base¹′ 35 | 36 | base : S¹ 37 | base = base¹ 38 | 39 | -- Agda seems to think loop and refl are NOT definitionally equal 40 | -- outside of this abstract block 41 | postulate 42 | loop¹ : base¹ ≡ base¹ 43 | 44 | loop : base ≡ base 45 | loop = loop¹ 46 | 47 | ------------------------------------------------------------------------ 48 | -- Elimination rules and computation rules 49 | 50 | -- Dependent version 51 | 52 | S¹-elim : ∀ {ℓ} (P : S¹ → Set ℓ) (pbase : P base) → subst P loop¹ pbase ≡ pbase → (x : S¹) → P x 53 | S¹-elim _ pbase _ base¹′ = pbase 54 | 55 | -- This is actually definitionally equality... 56 | postulate 57 | S¹-elim-loop : ∀ {ℓ} (P : S¹ → Set ℓ) (pbase : P base) (ploop : subst P loop¹ pbase ≡ pbase) 58 | → cong[dep] P (S¹-elim P pbase ploop) loop¹ ≡ ploop 59 | 60 | -- Non-dependent version 61 | 62 | S¹-elim[simp] : ∀ {ℓ} {P : Set ℓ} (pbase : P) → pbase ≡ pbase → S¹ → P 63 | S¹-elim[simp] {P = P} pbase ploop = S¹-elim (const P) pbase (trans boring-loop ploop) 64 | where 65 | boring-loop = subst-const loop pbase 66 | 67 | -- This propositional equality is derivable from the dependent version. 68 | S¹-elim[simp]-loop : ∀ {ℓ} {P : Set ℓ} (pbase : P) (ploop : pbase ≡ pbase) 69 | → cong (S¹-elim[simp] pbase ploop) loop¹ ≡ ploop 70 | S¹-elim[simp]-loop {P = P} pbase ploop = 71 | cong dcircle loop¹ ≡⟨ sym $ trans-sym-trans boring-loop _ ⟩ 72 | trans (sym $ boring-loop) (trans boring-loop $ cong dcircle loop¹) ≡⟨ cong (trans $ sym $ boring-loop) $ sym $ cong[dep]-const dcircle loop¹ ⟩ 73 | trans (sym $ boring-loop) (cong[dep] (const P) dcircle loop¹) ≡⟨ cong (trans $ sym $ boring-loop) $ S¹-elim-loop (const P) _ _ ⟩ 74 | trans (sym $ boring-loop) (trans boring-loop ploop) ≡⟨ trans-sym-trans boring-loop _ ⟩∎ 75 | ploop ∎ 76 | where 77 | boring-loop = subst-const loop¹ pbase 78 | dcircle = S¹-elim[simp] pbase ploop 79 | 80 | ------------------------------------------------------------------------ 81 | -- Experimental area: S² and beyond! 82 | ------------------------------------------------------------------------ 83 | 84 | ------------------------------------------------------------------------ 85 | -- Lemmas for building the endpoint tower. They are put here to make 86 | -- sure that they are true without access to the secret structure of 87 | -- Sⁿ. lemma₁ ~ lemma₆ are to convert different kiends of towers/paths. 88 | 89 | private 90 | -- non-dependent tower 91 | S-endpoints⇑ : ∀ {ℓ} n {S : Set ℓ} → S → Endpoints⇑ n S 92 | S-loop⇑ : ∀ {ℓ} n {S : Set ℓ} (base : S) → Path⇑ n (S-endpoints⇑ n base) 93 | 94 | S-endpoints⇑ 0 base = lift tt 95 | S-endpoints⇑ (suc n) base = (S-endpoints⇑ n base , S-loop⇑ n base , S-loop⇑ n base) 96 | 97 | S-loop⇑ 0 base = base 98 | S-loop⇑ (suc n) base = refl (S-loop⇑ n base) 99 | 100 | -- dependent tower 101 | S-endpoints[dep]⇑ : ∀ {ℓ} n {S : Set} (P : S → Set ℓ) (base : S) (pbase : P base) → 102 | Endpoints[dep]⇑ n P (S-endpoints⇑ n base) 103 | S-loop[dep]⇑ : ∀ {ℓ} n {S : Set} (P : S → Set ℓ) (base : S) (pbase : P base) → 104 | Path[dep]⇑ n P (S-loop⇑ n base) (S-endpoints[dep]⇑ n P base pbase) 105 | 106 | S-endpoints[dep]⇑ 0 P base pbase = lift tt 107 | S-endpoints[dep]⇑ (suc n) P base pbase = (S-endpoints[dep]⇑ n P base pbase , 108 | S-loop[dep]⇑ n P base pbase , 109 | S-loop[dep]⇑ n P base pbase) 110 | 111 | S-loop[dep]⇑ 0 P base pbase = pbase 112 | S-loop[dep]⇑ (suc n) P base pbase = refl (S-loop[dep]⇑ n P base pbase) 113 | 114 | -- lemmas for paths in (ctx, path, path) 115 | cong-,, : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {x y : A} (p : x ≡ y) 116 | {x₁ : B x} {y₁ : B y} (p₁ : subst B p x₁ ≡ y₁) 117 | {x₂ : B x} {y₂ : B y} (p₂ : subst B p x₂ ≡ y₂) → 118 | (x , x₁ , x₂) ≡ (y , y₁ , y₂) 119 | cong-,, B = elim 120 | (λ {x y} (p : x ≡ y) → 121 | ∀ {x₁} {y₁} (p₁ : subst B p x₁ ≡ y₁) 122 | {x₂} {y₂} (p₂ : subst B p x₂ ≡ y₂) → 123 | (x , x₁ , x₂) ≡ (y , y₁ , y₂)) 124 | (λ x p₁ p₂ → cong₂ (λ x₁ x₂ → (x , x₁ , x₂)) p₁ p₂) 125 | 126 | subst-cong-,, : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) {x y : A} (p : x ≡ y) 127 | {x₁₂ : B x} {y₁₂ : B y} (p₁₂ : subst B p x₁₂ ≡ y₁₂) → 128 | subst (λ s → proj₁ (proj₂ s) ≡ proj₂ (proj₂ s)) (cong-,, B p p₁₂ p₁₂) (refl x₁₂) ≡ refl y₁₂ 129 | subst-cong-,, B = elim 130 | (λ {x y} p → 131 | {x₁₂ : B x} {y₁₂ : B y} (p₁₂ : subst B p x₁₂ ≡ y₁₂) → 132 | subst (λ s → proj₁ (proj₂ s) ≡ proj₂ (proj₂ s)) (cong-,, B p p₁₂ p₁₂) (refl x₁₂) ≡ refl y₁₂) 133 | (λ x → elim 134 | (λ {x₁₂ y₁₂} p₁₂ → 135 | subst (λ s → proj₁ (proj₂ s) ≡ proj₂ (proj₂ s)) (cong-,, B (refl x) p₁₂ p₁₂) (refl x₁₂) ≡ refl y₁₂) 136 | (λ x₁₂ → refl _)) 137 | 138 | -- lemma₁ and lemma₂ : cong[dep] <-> S-endpoints[dep] 139 | lemma₁ : ∀ {ℓ} m {S} (P : S → Set ℓ) (f : (x : S) → P x) (b : S) → 140 | cong-endpoints[dep]⇑ m P f (S-endpoints⇑ m b) ≡ S-endpoints[dep]⇑ m P b (f b) 141 | lemma₂ : ∀ {ℓ} m {S} (P : S → Set ℓ) (f : (x : S) → P x) (b : S) → 142 | subst (Path[dep]⇑ m P (S-loop⇑ m b)) (lemma₁ m P f b) 143 | (cong[dep]⇑ m P f (S-endpoints⇑ m b) (S-loop⇑ m b)) ≡ 144 | S-loop[dep]⇑ m P b (f b) 145 | 146 | lemma₁ 0 _ _ _ = refl _ 147 | lemma₁ (suc m) P f b = cong-,, (Path[dep]⇑ m P (S-loop⇑ m b)) (lemma₁ m P f b) (lemma₂ m P f b) (lemma₂ m P f b) 148 | 149 | lemma₂ 0 _ _ _ = refl _ 150 | lemma₂ (suc m) P f b = subst-cong-,, (Path[dep]⇑ m P (S-loop⇑ m b)) (lemma₁ m P f b) (lemma₂ m P f b) 151 | 152 | -- lemma₃ and lemma₄ : cong <-> S-endpoints 153 | lemma₃ : ∀ {ℓ} m {S : Set} {P : Set ℓ} (f : S → P) (b : S) → 154 | cong-endpoints⇑ m f (S-endpoints⇑ m b) ≡ S-endpoints⇑ m (f b) 155 | lemma₄ : ∀ {ℓ} m {S : Set} {P : Set ℓ} (f : S → P) (b : S) → 156 | subst (Path⇑ m) (lemma₃ m f b) 157 | (cong⇑ m f (S-endpoints⇑ m b) (S-loop⇑ m b)) ≡ 158 | S-loop⇑ m (f b) 159 | 160 | lemma₃ 0 _ _ = refl _ 161 | lemma₃ (suc m) f b = cong-,, (Path⇑ m) (lemma₃ m f b) (lemma₄ m f b) (lemma₄ m f b) 162 | 163 | lemma₄ 0 _ _ = refl _ 164 | lemma₄ (suc m) f b = subst-cong-,, (Path⇑ m) (lemma₃ m f b) (lemma₄ m f b) 165 | 166 | -- lemma₅ and lemma₆ : S-endpoints[dep] <-> S-endpoints 167 | -- when P is a constant function 168 | lemma₅ : ∀ {ℓ} m {S : Set} {P : Set ℓ} (b : S) (pb : P) (loop : Path⇑ m (S-endpoints⇑ m b)) → 169 | Path[dep]⇑ m (const P) loop (S-endpoints[dep]⇑ m (const P) b pb) ≡ Path⇑ m (S-endpoints⇑ m pb) 170 | lemma₆ : ∀ {ℓ} m {S : Set} {P : Set ℓ} (b : S) (pb : P) {free} (loop : S-loop⇑ m b ≡ free) → 171 | subst id (lemma₅ m b pb free) 172 | (subst (λ x → Path[dep]⇑ m (const P) x (S-endpoints[dep]⇑ m (const P) b pb)) 173 | loop 174 | (S-loop[dep]⇑ m (const P) b pb)) ≡ 175 | S-loop⇑ m pb 176 | 177 | lemma₆-refl : ∀ {ℓ} m {S : Set} {P : Set ℓ} (b : S) (pb : P) → 178 | subst id (lemma₅ m b pb (S-loop⇑ m b)) (S-loop[dep]⇑ m (const P) b pb) ≡ S-loop⇑ m pb 179 | 180 | lemma₅ 0 _ _ _ = refl _ 181 | lemma₅ (suc m) b pb loop = cong-≡ (lemma₅ m b pb (S-loop⇑ m b)) (lemma₆ m b pb loop) (lemma₆-refl m b pb) 182 | 183 | lemma₆ m {P = P} b pb loop = elim′ 184 | (λ {free} loop → 185 | subst id (lemma₅ m b pb free) 186 | (subst (λ x → Path[dep]⇑ m (const P) x (S-endpoints[dep]⇑ m (const P) b pb)) 187 | loop 188 | (S-loop[dep]⇑ m (const P) b pb)) ≡ 189 | S-loop⇑ m pb) 190 | (lemma₆-refl m b pb) 191 | loop 192 | 193 | lemma₆-refl 0 b pb = refl _ 194 | lemma₆-refl (suc m) b pb = subst-cong-≡ (lemma₅ m b pb (S-loop⇑ m b)) (lemma₆-refl m b pb) 195 | 196 | ------------------------------------------------------------------------ 197 | -- Formation rules and introduction rules 198 | 199 | private 200 | data Sⁿ′ (n : ℕ) : Set where 201 | baseⁿ′ : Sⁿ′ n 202 | 203 | Sⁿ : ℕ → Set 204 | Sⁿ = Sⁿ′ 205 | 206 | baseⁿ : ∀ n → Sⁿ n 207 | baseⁿ n = baseⁿ′ 208 | 209 | private 210 | Sⁿ-endpoints : ∀ n → Endpoints⇑ n (Sⁿ n) 211 | Sⁿ-endpoints n = S-endpoints⇑ n (baseⁿ n) 212 | 213 | -- loop! 214 | postulate 215 | loopⁿ : ∀ n → Path⇑ n (Sⁿ-endpoints n) 216 | 217 | ------------------------------------------------------------------------ 218 | -- Elimination rules and computation rules 219 | 220 | private 221 | Sⁿ-endpoints[dep] : ∀ {ℓ} n (P : Sⁿ n → Set ℓ) → P (baseⁿ n) → Endpoints[dep]⇑ n P (Sⁿ-endpoints n) 222 | Sⁿ-endpoints[dep] n P pbase = S-endpoints[dep]⇑ n P (baseⁿ n) pbase 223 | 224 | -- Dependent version 225 | Sⁿ-elim : ∀ {ℓ} n (P : Sⁿ n → Set ℓ) (pbase : P (baseⁿ n)) → Path[dep]⇑ n P (loopⁿ n) (Sⁿ-endpoints[dep] n P pbase) → ∀ x → P x 226 | Sⁿ-elim n P pbase _ baseⁿ′ = pbase 227 | 228 | postulate 229 | Sⁿ-elim-loop : ∀ {ℓ} n (P : Sⁿ n → Set ℓ) (pbase : P (baseⁿ n)) (ploop : Path[dep]⇑ n P (loopⁿ n) (Sⁿ-endpoints[dep] n P pbase)) → 230 | subst (Path[dep]⇑ n P (loopⁿ n)) (lemma₁ n P (Sⁿ-elim n P pbase ploop) (baseⁿ n)) 231 | (cong[dep]⇑ n P (Sⁿ-elim n P pbase ploop) _ (loopⁿ n)) ≡ 232 | ploop 233 | 234 | -- Non-dependent version 235 | Sⁿ-elim[simp] : ∀ {ℓ} n {P : Set ℓ} (pbase : P) → Path⇑ n (S-endpoints⇑ n pbase) → Sⁿ n → P 236 | Sⁿ-elim[simp] n {P} pbase ploop base′ = Sⁿ-elim n (const P) pbase (subst id (sym $ lemma₅ n (baseⁿ n) pbase (loopⁿ n)) ploop) base′ 237 | 238 | -- TODO This rule should be derived from the dependent elimination rule 239 | postulate 240 | Sⁿ-elim[simp]-loop : ∀ {ℓ} n {P : Set ℓ} (pbase : P) (ploop : Path⇑ n (S-endpoints⇑ n pbase)) 241 | → subst (Path⇑ n) (lemma₃ n (Sⁿ-elim[simp] n pbase ploop) (baseⁿ n)) 242 | (cong⇑ n (Sⁿ-elim[simp] n pbase ploop) _ (loopⁿ n)) ≡ ploop 243 | 244 | ------------------------------------------------------------------------ 245 | -- Test cases for S⁰, S¹, and S² 246 | 247 | private 248 | 249 | -- Test cases for "Path⇑ n (Sⁿ-endpoints n)" 250 | module Test₁ where 251 | test₀ : Path⇑ 0 (Sⁿ-endpoints 0) ≡ Sⁿ 0 252 | test₀ = refl _ 253 | 254 | test₁ : Path⇑ 1 (Sⁿ-endpoints 1) ≡ (baseⁿ 1 ≡ baseⁿ 1) 255 | test₁ = refl _ 256 | 257 | test₂ : Path⇑ 2 (Sⁿ-endpoints 2) ≡ (refl (baseⁿ 2) ≡ refl (baseⁿ 2)) 258 | test₂ = refl _ 259 | 260 | -- Test cases for "Path[dep]⇑ n P (loopⁿ n) (Sⁿ-endpoints[dep] n P pbase)" 261 | module Test₂ where 262 | test₀ : ∀ {ℓ} (P : Sⁿ 0 → Set ℓ) (pbase : P (baseⁿ 0)) → 263 | Path[dep]⇑ 0 P (loopⁿ 0) (Sⁿ-endpoints[dep] 0 P pbase) ≡ P (loopⁿ 0) 264 | test₀ _ _ = refl _ 265 | 266 | test₁ : ∀ {ℓ} (P : Sⁿ 1 → Set ℓ) (pbase : P (baseⁿ 1)) → 267 | Path[dep]⇑ 1 P (loopⁿ 1) (Sⁿ-endpoints[dep] 1 P pbase) ≡ (subst P (loopⁿ 1) pbase ≡ pbase) 268 | test₁ _ _ = refl _ 269 | 270 | test₂ : ∀ {ℓ} (P : Sⁿ 2 → Set ℓ) (pbase : P (baseⁿ 2)) → 271 | Path[dep]⇑ 2 P (loopⁿ 2) (Sⁿ-endpoints[dep] 2 P pbase) ≡ (subst (λ p → subst P p pbase ≡ pbase) (loopⁿ 2) (refl pbase) ≡ refl pbase) 272 | test₂ _ _ = refl _ 273 | 274 | -- Test cases for the type of the LHS of "Sⁿ-elim-loop" 275 | module Test₃ where 276 | test₀ : ∀ {ℓ} (P : Sⁿ 0 → Set ℓ) (pbase : P (baseⁿ 0)) (ploop : Path[dep]⇑ 0 P (loopⁿ 0) (Sⁿ-endpoints[dep] 0 P pbase)) → 277 | subst (Path[dep]⇑ 0 P (loopⁿ 0)) (lemma₁ 0 P (Sⁿ-elim 0 P pbase ploop) (baseⁿ 0)) 278 | (cong[dep]⇑ 0 P (Sⁿ-elim 0 P pbase ploop) (Sⁿ-endpoints 0) (loopⁿ 0)) 279 | ≡ (Sⁿ-elim 0 P pbase ploop) (loopⁿ 0) 280 | test₀ _ _ _ = refl _ 281 | 282 | test₁ : ∀ {ℓ} (P : Sⁿ 1 → Set ℓ) (pbase : P (baseⁿ 1)) (ploop : Path[dep]⇑ 1 P (loopⁿ 1) (Sⁿ-endpoints[dep] 1 P pbase)) → 283 | subst (Path[dep]⇑ 1 P (loopⁿ 1)) (lemma₁ 1 P (Sⁿ-elim 1 P pbase ploop) (baseⁿ 1)) 284 | (cong[dep]⇑ 1 P (Sⁿ-elim 1 P pbase ploop) (Sⁿ-endpoints 1) (loopⁿ 1)) 285 | ≡ cong[dep] P (Sⁿ-elim 1 P pbase ploop) (loopⁿ 1) 286 | test₁ _ _ _ = refl _ 287 | 288 | test₂ : ∀ {ℓ} (P : Sⁿ 2 → Set ℓ) (pbase : P (baseⁿ 2)) (ploop : Path[dep]⇑ 2 P (loopⁿ 2) (Sⁿ-endpoints[dep] 2 P pbase)) → 289 | subst (Path[dep]⇑ 2 P (loopⁿ 2)) (lemma₁ 2 P (Sⁿ-elim 2 P pbase ploop) (baseⁿ 2)) 290 | (cong[dep]⇑ 2 P (Sⁿ-elim 2 P pbase ploop) (Sⁿ-endpoints 2) (loopⁿ 2)) 291 | ≡ cong[dep] (λ p → subst P p pbase ≡ pbase) (cong[dep] P (Sⁿ-elim 2 P pbase ploop)) (loopⁿ 2) 292 | test₂ _ _ _ = refl _ 293 | 294 | -- Test cases for the type of the LHS of "Sⁿ-elim[simp]" 295 | module Test₄ where 296 | test₀ : ∀ {ℓ} (P : Set ℓ) (pbase : P) → 297 | Path⇑ 0 (S-endpoints⇑ 0 pbase) ≡ P 298 | test₀ _ _ = refl _ 299 | 300 | test₁ : ∀ {ℓ} (P : Set ℓ) (pbase : P) → 301 | Path⇑ 1 (S-endpoints⇑ 1 pbase) ≡ (pbase ≡ pbase) 302 | test₁ _ _ = refl _ 303 | 304 | test₂ : ∀ {ℓ} (P : Set ℓ) (pbase : P) → 305 | Path⇑ 2 (S-endpoints⇑ 2 pbase) ≡ (refl pbase ≡ refl pbase) 306 | test₂ _ _ = refl _ 307 | 308 | -- Test cases for the type of the LHS of "Sⁿ-elim-loop[simp]" 309 | module Test₅ where 310 | test₀ : ∀ {ℓ} (P : Set ℓ) (pbase : P) (ploop : Path⇑ 0 (S-endpoints⇑ 0 pbase)) → 311 | subst (Path⇑ 0) (lemma₃ 0 (Sⁿ-elim[simp] 0 pbase ploop) (baseⁿ 0)) 312 | (cong⇑ 0 (Sⁿ-elim[simp] 0 pbase ploop) (Sⁿ-endpoints 0) (loopⁿ 0)) 313 | ≡ (Sⁿ-elim[simp] 0 pbase ploop) (loopⁿ 0) 314 | test₀ _ _ _ = refl _ 315 | 316 | test₁ : ∀ {ℓ} (P : Set ℓ) (pbase : P) (ploop : Path⇑ 1 (S-endpoints⇑ 1 pbase)) → 317 | subst (Path⇑ 1) (lemma₃ 1 (Sⁿ-elim[simp] 1 pbase ploop) (baseⁿ 1)) 318 | (cong⇑ 1 (Sⁿ-elim[simp] 1 pbase ploop) (Sⁿ-endpoints 1) (loopⁿ 1)) 319 | ≡ cong (Sⁿ-elim[simp] 1 pbase ploop) (loopⁿ 1) 320 | test₁ _ _ _ = refl _ 321 | 322 | test₂ : ∀ {ℓ} (P : Set ℓ) (pbase : P) (ploop : Path⇑ 2 (S-endpoints⇑ 2 pbase)) → 323 | subst (Path⇑ 2) (lemma₃ 2 (Sⁿ-elim[simp] 2 pbase ploop) (baseⁿ 2)) 324 | (cong⇑ 2 (Sⁿ-elim[simp] 2 pbase ploop) (Sⁿ-endpoints 2) (loopⁿ 2)) 325 | ≡ cong (cong (Sⁿ-elim[simp] 2 pbase ploop)) (loopⁿ 2) 326 | test₂ _ _ _ = refl _ 327 | -------------------------------------------------------------------------------- /Space/Sphere/HopfJunior.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Construction of Hopf junior 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- {-# OPTIONS --without-K #-} 8 | 9 | -- A warmup for the Hopf fibration (the Hopf junior fibration) 10 | -- which is a fibration for S⁰ → S¹ → S¹. The goal of this file 11 | -- is to give the construction and show that the total space is 12 | -- really S¹. 13 | 14 | -- Credits: 15 | -- * Peter Lumsdaine gave the construction of the Hopf (junior) 16 | -- * Danial Licata gave the outline of almost all interesting parts 17 | 18 | open import Univalence 19 | 20 | module Space.Sphere.HopfJunior 21 | (univ : ∀ {ℓ} (A B : Set ℓ) → Univalence-axiom A B) where 22 | 23 | open import Prelude 24 | open import Path 25 | open import Path.Lemmas 26 | open import Path.Sum 27 | open import Map.H-equivalence hiding (_∘_; id) 28 | open import Map.WeakEquivalence as Weak hiding (_∘_; id) 29 | 30 | import Univalence.Lemmas; open Univalence.Lemmas univ 31 | 32 | open import Space.Sphere 33 | 34 | ------------------------------------------------------------------------ 35 | -- Construction: apply S¹-elim to a non-trivial homotopy equivalence in S⁰ 36 | 37 | private 38 | not-not : ∀ x → not (not x) ≡ x 39 | not-not true = refl true 40 | not-not false = refl false 41 | 42 | not-↔ : Bool ↔ Bool 43 | not-↔ = record 44 | { surjection = record 45 | { to = not 46 | ; from = not 47 | ; right-inverse-of = not-not 48 | } 49 | ; left-inverse-of = not-not 50 | } 51 | 52 | not-≈ : Bool ≈ Bool 53 | not-≈ = ↔⇒≈ not-↔ 54 | 55 | not-≡ : Bool ≡ Bool 56 | not-≡ = ≈⇒≡ not-≈ 57 | 58 | -- Here's the Hopf junior 59 | Hj : S¹ → Set 60 | Hj = S¹-elim[simp] Bool not-≡ 61 | 62 | ------------------------------------------------------------------------ 63 | -- A map from S¹ to the total space of Hopf junior fibration 64 | 65 | private 66 | -- For some magical reasons this will speed up the checking! 67 | abstract 68 | -- "subst Hj loop" is "not" 69 | subst-Hj-loop : ∀ (x : Bool) → subst Hj loop x ≡ not x 70 | subst-Hj-loop x = 71 | subst Hj loop x ≡⟨ sym $ subst-cong id Hj loop x ⟩ 72 | subst id (cong Hj loop) x ≡⟨ cong (λ p → subst id p x) $ S¹-elim[simp]-loop Bool not-≡ ⟩ 73 | subst id not-≡ x ≡⟨ subst-id-univ not-≡ x ⟩ 74 | _≈_.to (≡⇒≈ not-≡) x ≡⟨ cong (λ weq → _≈_.to weq x) $ right-inverse-of not-≈ ⟩∎ 75 | -- to not-≈ x ≡⟨ refl (not x) ⟩∎ 76 | not x ∎ 77 | where 78 | open _≈_ (≡≈≈ Bool Bool) 79 | 80 | -- (base , true) 81 | -- These are used to fill in implicit arguments in declarations 82 | double-base-true : Σ S¹ Hj 83 | double-base-true = (base , true) 84 | 85 | -- (base , false) 86 | -- These are used to fill in implicit arguments in declarations 87 | double-base-false : Σ S¹ Hj 88 | double-base-false = (base , false) 89 | 90 | -- (base , true) ~~> (base , false) 91 | double-path-true→false : double-base-true ≡ double-base-false 92 | double-path-true→false = Σ≡⇒≡Σ Hj (loop , subst-Hj-loop true ) 93 | 94 | -- (base , false) ~~> (base , true) 95 | double-path-false→true : double-base-false ≡ double-base-true 96 | double-path-false→true = Σ≡⇒≡Σ Hj (loop , subst-Hj-loop false ) 97 | 98 | -- The base 99 | double-base : Σ S¹ Hj 100 | double-base = double-base-true 101 | 102 | -- The loop 103 | double-loop : double-base ≡ double-base 104 | double-loop = trans double-path-true→false double-path-false→true 105 | 106 | -- The map 107 | double : S¹ → Σ S¹ Hj 108 | double = S¹-elim[simp] double-base double-loop 109 | 110 | ------------------------------------------------------------------------ 111 | -- A map from the total space to S¹ 112 | 113 | private 114 | -- The base 115 | halve′-base : Bool → S¹ 116 | halve′-base _ = base 117 | 118 | -- The interesting loop 119 | -- (Actually the most interesting part in the construction!) 120 | halve′-loop′ : ∀ x → halve′-base x ≡ halve′-base x 121 | halve′-loop′ true = loop 122 | halve′-loop′ false = refl base 123 | 124 | halve′-loop : halve′-base ≡ halve′-base 125 | halve′-loop = ext halve′-loop′ 126 | 127 | -- The boring loop to transport the base 128 | halve′-boring-loop′ : ∀ x → subst (λ (x : S¹) → Hj x → S¹) loop halve′-base x ≡ halve′-base x 129 | halve′-boring-loop′ true = 130 | subst (λ (x : S¹) → Hj x → S¹) loop halve′-base true ≡⟨ sym $ subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop false) ⟩ 131 | subst (const S¹) loop (halve′-base (subst Hj loop false)) ≡⟨ refl _ ⟩ 132 | subst (const S¹) loop base ≡⟨ subst-const loop base ⟩∎ 133 | base ∎ 134 | halve′-boring-loop′ false = 135 | subst (λ (x : S¹) → Hj x → S¹) loop halve′-base false ≡⟨ sym $ subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop true) ⟩ 136 | subst (const S¹) loop (halve′-base (subst Hj loop true)) ≡⟨ refl _ ⟩ 137 | subst (const S¹) loop base ≡⟨ subst-const loop base ⟩∎ 138 | base ∎ 139 | 140 | halve′-boring-loop : subst (λ (x : S¹) → Hj x → S¹) loop halve′-base ≡ halve′-base 141 | halve′-boring-loop = ext halve′-boring-loop′ 142 | 143 | -- The curried version of the map 144 | halve′ : (x : S¹) → Hj x → S¹ 145 | halve′ = S¹-elim (λ x → Hj x → S¹) halve′-base (trans halve′-boring-loop halve′-loop) 146 | 147 | -- The map 148 | halve : Σ S¹ Hj → S¹ 149 | halve = uncurry halve′ 150 | 151 | ------------------------------------------------------------------------ 152 | -- Homotopy equivalence 153 | 154 | private 155 | -- This lemma is the most interesting (difficult) one! 156 | -- (Sorry for the ugly proof.) 157 | -- 158 | -- It is possible to prove both cases (true & false) at once, 159 | -- but one need to apply "not not b ≡ b" at several places. 160 | -- However in any specific case they are definitionally equal! 161 | -- For anyone interested in this approach, you might want to 162 | -- redefine "halve′-boring-loop′" so that it does not contain 163 | -- a hidden Bool-elim. 164 | 165 | cong-halve-double-path : ∀ b → cong halve (Σ≡⇒≡Σ Hj (loop , subst-Hj-loop b)) ≡ halve′-loop′ (not b) 166 | cong-halve-double-path true = 167 | cong halve (Σ≡⇒≡Σ Hj (loop , subst-Hj-loop true)) 168 | -- This pulled out the proof term 169 | ≡⟨ cong-Σ≡⇒≡Σ Hj S¹ halve′ loop (subst-Hj-loop true) ⟩ 170 | (base ≡⟨ line₁ ⟩ 171 | subst (const S¹) loop base ≡⟨ line₂ ⟩ 172 | subst (λ x → Hj x → S¹) loop halve′-base false ≡⟨ line₃ ⟩∎ 173 | base ∎) 174 | ≡⟨ refl _ ⟩ 175 | trans line₁ (trans line₂ line₃) 176 | ≡⟨ cong (trans line₁ ∘ trans line₂) $ lemma₃ ⟩ 177 | trans line₁ (trans line₂ (trans line₃′|₁ line₃′|₂)) 178 | ≡⟨ cong (trans line₁) $ trans-transʳ-sym line₂ line₃′|₂ ⟩ 179 | trans line₁ line₃′|₂ 180 | ≡⟨ trans-symˡ line₃′|₂ ⟩∎ 181 | refl _ 182 | ∎ 183 | where 184 | line₁ : base ≡ subst (const S¹) loop base 185 | line₁ = sym $ subst-const loop base 186 | line₂ : subst (const S¹) loop base ≡ subst (λ x → Hj x → S¹) loop halve′-base false 187 | line₂ = subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop true) 188 | line₃ : subst (λ x → Hj x → S¹) loop halve′-base false ≡ base 189 | line₃ = cong (λ f′ → f′ false) $ cong[dep] (λ x → Hj x → S¹) halve′ loop 190 | 191 | line₃′|₁ : subst (λ (x : S¹) → Hj x → S¹) loop halve′-base false ≡ subst (const S¹) loop base 192 | line₃′|₁ = sym $ subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop true) 193 | line₃′|₂ : subst (const S¹) loop base ≡ base 194 | line₃′|₂ = subst-const loop base 195 | 196 | lemma₃ : line₃ ≡ trans line₃′|₁ line₃′|₂ 197 | lemma₃ = 198 | line₃ 199 | ≡⟨ cong (cong (λ f → f false)) $ S¹-elim-loop (λ x → Hj x → S¹) halve′-base (trans halve′-boring-loop halve′-loop) ⟩ 200 | cong (λ f → f false) (trans halve′-boring-loop halve′-loop) 201 | ≡⟨ cong-trans (λ f → f false) halve′-boring-loop halve′-loop ⟩ 202 | trans (cong (λ f → f false) halve′-boring-loop) (cong (λ f → f false) halve′-loop) 203 | ≡⟨ cong (trans $ cong (λ f → f false) halve′-boring-loop) $ ext-comp halve′-loop′ false ⟩ 204 | trans (cong (λ f → f false) halve′-boring-loop) (halve′-loop′ false) 205 | ≡⟨ trans-reflʳ $ cong (λ f → f false) halve′-boring-loop ⟩ 206 | cong (λ f → f false) halve′-boring-loop 207 | ≡⟨ ext-comp halve′-boring-loop′ false ⟩ 208 | halve′-boring-loop′ false 209 | ≡⟨ refl _ ⟩∎ 210 | trans line₃′|₁ line₃′|₂ 211 | ∎ 212 | 213 | cong-halve-double-path false = 214 | cong halve (Σ≡⇒≡Σ Hj (loop , subst-Hj-loop false)) 215 | -- This pulled out the proof term 216 | ≡⟨ cong-Σ≡⇒≡Σ Hj S¹ halve′ loop (subst-Hj-loop false) ⟩ 217 | (base ≡⟨ line₁ ⟩ 218 | subst (const S¹) loop base ≡⟨ line₂ ⟩ 219 | subst (λ x → Hj x → S¹) loop halve′-base true ≡⟨ line₃ ⟩∎ 220 | base ∎) 221 | ≡⟨ refl _ ⟩ 222 | trans line₁ (trans line₂ line₃) 223 | ≡⟨ cong (trans line₁ ∘ trans line₂) $ lemma₃ ⟩ 224 | trans line₁ (trans line₂ (trans line₃′|₁ (trans line₃′|₂ loop))) 225 | ≡⟨ cong (trans line₁) $ trans-transʳ-sym line₂ (trans line₃′|₂ loop) ⟩ 226 | trans line₁ (trans line₃′|₂ loop) 227 | ≡⟨ trans-sym-trans line₃′|₂ loop ⟩∎ 228 | loop 229 | ∎ 230 | where 231 | line₁ : base ≡ subst (const S¹) loop base 232 | line₁ = sym $ subst-const loop base 233 | line₂ : subst (const S¹) loop base ≡ subst (λ x → Hj x → S¹) loop halve′-base true 234 | line₂ = subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop false) 235 | line₃ : subst (λ x → Hj x → S¹) loop halve′-base true ≡ base 236 | line₃ = cong (λ f′ → f′ true) $ cong[dep] (λ x → Hj x → S¹) halve′ loop 237 | 238 | line₃′|₁ : subst (λ (x : S¹) → Hj x → S¹) loop halve′-base true ≡ subst (const S¹) loop base 239 | line₃′|₁ = sym $ subst-app Hj (const S¹) loop halve′-base (subst-Hj-loop false) 240 | line₃′|₂ : subst (const S¹) loop base ≡ base 241 | line₃′|₂ = subst-const loop base 242 | 243 | lemma₃ : line₃ ≡ trans line₃′|₁ (trans line₃′|₂ loop) 244 | lemma₃ = 245 | line₃ 246 | ≡⟨ cong (cong (λ f → f true)) $ S¹-elim-loop (λ x → Hj x → S¹) halve′-base (trans halve′-boring-loop halve′-loop) ⟩ 247 | cong (λ f → f true) (trans halve′-boring-loop halve′-loop) 248 | ≡⟨ cong-trans (λ f → f true) halve′-boring-loop halve′-loop ⟩ 249 | trans (cong (λ f → f true) halve′-boring-loop) (cong (λ f → f true) halve′-loop) 250 | ≡⟨ cong (trans $ cong (λ f → f true) halve′-boring-loop) $ ext-comp halve′-loop′ true ⟩ 251 | trans (cong (λ f → f true) halve′-boring-loop) (halve′-loop′ true) 252 | ≡⟨ refl _ ⟩ 253 | trans (cong (λ f → f true) halve′-boring-loop) loop 254 | ≡⟨ cong (λ p → trans p loop) $ ext-comp halve′-boring-loop′ true ⟩ 255 | trans (halve′-boring-loop′ true) loop 256 | ≡⟨ trans-assoc line₃′|₁ line₃′|₂ loop ⟩∎ 257 | trans line₃′|₁ (trans line₃′|₂ loop) 258 | ∎ 259 | 260 | private 261 | cong-halve-double-loop : cong halve double-loop ≡ loop 262 | cong-halve-double-loop = 263 | cong halve double-loop 264 | ≡⟨ cong-trans halve double-path-true→false double-path-false→true ⟩ 265 | trans (cong halve double-path-true→false) (cong halve double-path-false→true) 266 | ≡⟨ cong (trans (cong halve double-path-true→false)) $ cong-halve-double-path false ⟩ 267 | trans (cong halve double-path-true→false) loop 268 | ≡⟨ cong (λ p → trans p loop) $ cong-halve-double-path true ⟩∎ 269 | loop 270 | ∎ 271 | 272 | cong-double-loop : cong double loop ≡ double-loop 273 | cong-double-loop = S¹-elim[simp]-loop double-base double-loop 274 | 275 | S¹↔ΣS¹Hj : S¹ ↔ Σ S¹ Hj 276 | S¹↔ΣS¹Hj = 277 | record 278 | { surjection = record 279 | { to = double 280 | ; from = halve 281 | ; right-inverse-of = right-inverse-of 282 | } 283 | ; left-inverse-of = left-inverse-of 284 | } 285 | where 286 | left-inverse-of : ∀ x → halve (double x) ≡ x 287 | left-inverse-of = S¹-elim (λ x → halve (double x) ≡ x) (refl base) ≡-loop 288 | where 289 | ≡-loop : subst (λ x → halve (double x) ≡ x) loop (refl base) ≡ refl base 290 | ≡-loop = 291 | subst (λ x → halve (double x) ≡ x) loop (refl base) ≡⟨ subst-path[idʳ] (halve ∘ double) loop (refl base) ⟩ 292 | trans (sym (cong (halve ∘ double) loop)) loop ≡⟨ cong (λ p → trans (sym p) loop) $ sym $ cong-cong halve double loop ⟩ 293 | trans (sym (cong halve $ cong double loop)) loop ≡⟨ cong (λ p → trans (sym (cong halve p)) loop) $ cong-double-loop ⟩ 294 | trans (sym (cong halve double-loop)) loop ≡⟨ cong (λ p → trans (sym p) loop) $ cong-halve-double-loop ⟩ 295 | trans (sym loop) loop ≡⟨ trans-symˡ loop ⟩∎ 296 | refl base ∎ 297 | 298 | right-inverse-of : ∀ s → double (halve s) ≡ s 299 | right-inverse-of = uncurry right-inverse-of′ 300 | where 301 | right-inverse-of′ = S¹-elim (λ x → (y : Hj x) → double (halve (x , y)) ≡ (x , y)) ≡-base ≡-loop 302 | where 303 | ≡-base : ∀ b → double (halve (base , b)) ≡ (base , b) 304 | ≡-base true = refl _ -- (base , true) ~~> (base , true) 305 | ≡-base false = double-path-true→false -- (base , true) ~~> (base , false) 306 | 307 | ≡-loop′ : ∀ b → subst (λ x → (y : Hj x) → double (halve (x , y)) ≡ (x , y)) loop ≡-base b ≡ ≡-base b 308 | 309 | -- (base , false) ~~> (base , true) 310 | ≡-loop′ true = 311 | subst (λ x → (y : Hj x) → double (halve (x , y)) ≡ (x , y)) loop ≡-base true 312 | ≡⟨ subst-Σfunc Hj (λ s → double (halve s) ≡ s) loop ≡-base (subst-Hj-loop false) ⟩ 313 | subst (λ s → double (halve s) ≡ s) double-path-false→true double-path-true→false 314 | ≡⟨ subst-path[idʳ] (double ∘ halve) double-path-false→true double-path-true→false ⟩ 315 | trans (sym (cong (double ∘ halve) double-path-false→true)) double-loop 316 | ≡⟨ cong (λ p → trans (sym p) double-loop) $ sym $ cong-cong double halve double-path-false→true ⟩ 317 | trans (sym (cong double $ cong halve double-path-false→true)) double-loop 318 | ≡⟨ cong (λ p → trans (sym (cong double $ p)) double-loop) $ cong-halve-double-path false ⟩ 319 | trans (sym (cong double loop)) double-loop 320 | ≡⟨ cong (λ p → trans (sym p) double-loop) cong-double-loop ⟩ 321 | trans (sym double-loop) double-loop 322 | ≡⟨ trans-symˡ double-loop ⟩∎ 323 | refl _ 324 | ∎ 325 | 326 | -- (base , false) ~~> (base , false) 327 | ≡-loop′ false = 328 | subst (λ x → (y : Hj x) → double (halve (x , y)) ≡ (x , y)) loop ≡-base false 329 | ≡⟨ subst-Σfunc Hj (λ s → double (halve s) ≡ s) loop ≡-base (subst-Hj-loop true) ⟩ 330 | subst (λ s → double (halve s) ≡ s) double-path-true→false (refl _) 331 | ≡⟨ subst-path[idʳ] (double ∘ halve) double-path-true→false (refl _) ⟩ 332 | trans (sym (cong (double ∘ halve) double-path-true→false)) double-path-true→false 333 | ≡⟨ cong (λ p → trans (sym p) double-path-true→false) $ sym $ cong-cong double halve double-path-true→false ⟩ 334 | trans (sym (cong double $ cong halve double-path-true→false)) double-path-true→false 335 | ≡⟨ cong (λ p → trans (sym (cong double $ p)) double-path-true→false) $ cong-halve-double-path true ⟩∎ 336 | double-path-true→false 337 | ∎ 338 | 339 | ≡-loop : subst (λ x → (y : Hj x) → double (halve (x , y)) ≡ (x , y)) loop ≡-base ≡ ≡-base 340 | ≡-loop = ext[dep] ≡-loop′ 341 | 342 | -------------------------------------------------------------------------------- /Space/Sphere/Omega1.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Paths in circles 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- Credits: 8 | -- * As far as I know, Peter Lumsdaine gave the construction. 9 | -- * J rule is used, following Danial Licata's simplification. 10 | 11 | open import Univalence 12 | 13 | module Space.Sphere.Omega1 14 | (univ : ∀ {ℓ} (A B : Set ℓ) → Univalence-axiom A B) where 15 | 16 | open import Prelude renaming (zero to ℕzero; suc to ℕsuc) 17 | open import Path 18 | open import Path.Lemmas 19 | open import Path.Sum 20 | open import Path.HigherOrder 21 | open import Map.H-equivalence hiding (_∘_; id) 22 | open import Map.WeakEquivalence as Weak hiding (_∘_; id) 23 | 24 | import Univalence.Lemmas; open Univalence.Lemmas univ 25 | 26 | open import Space.Integer 27 | open import Space.Sphere 28 | 29 | ------------------------------------------------------------------------ 30 | -- From ℤ to paths 31 | 32 | private 33 | repeat : ℕ → base ≡ base → base ≡ base 34 | repeat ℕzero _ = refl base 35 | repeat (ℕsuc n) loop = trans (repeat n loop) loop 36 | 37 | ℤ⇒S¹loop : ℤ → base ≡ base 38 | ℤ⇒S¹loop zero = refl base 39 | ℤ⇒S¹loop (pos n) = repeat (ℕsuc n) loop 40 | ℤ⇒S¹loop (neg n) = repeat (ℕsuc n) (sym loop) 41 | 42 | ------------------------------------------------------------------------ 43 | -- From paths to ℤ 44 | 45 | private 46 | suc : ℤ → ℤ 47 | suc zero = pos ℕzero 48 | suc (pos i) = pos (ℕsuc i) 49 | suc (neg ℕzero) = zero 50 | suc (neg (ℕsuc i)) = neg i 51 | 52 | pred : ℤ → ℤ 53 | pred zero = neg ℕzero 54 | pred (neg i) = neg (ℕsuc i) 55 | pred (pos ℕzero) = zero 56 | pred (pos (ℕsuc i)) = pos i 57 | 58 | suc-pred : ∀ (i : ℤ) → suc (pred i) ≡ i 59 | suc-pred zero = refl _ 60 | suc-pred (neg i) = refl _ 61 | suc-pred (pos ℕzero) = refl _ 62 | suc-pred (pos (ℕsuc i)) = refl _ 63 | 64 | pred-suc : ∀ (i : ℤ) → pred (suc i) ≡ i 65 | pred-suc zero = refl _ 66 | pred-suc (pos i) = refl _ 67 | pred-suc (neg ℕzero) = refl _ 68 | pred-suc (neg (ℕsuc i)) = refl _ 69 | 70 | -- The non-trivial homotopy equivalence between ℤ 71 | suc-↔ : ℤ ↔ ℤ 72 | suc-↔ = record 73 | { surjection = record 74 | { to = suc 75 | ; from = pred 76 | ; right-inverse-of = suc-pred 77 | } 78 | ; left-inverse-of = pred-suc 79 | } 80 | 81 | suc-≈ : ℤ ≈ ℤ 82 | suc-≈ = ↔⇒≈ suc-↔ 83 | 84 | suc-≡ : ℤ ≡ ℤ 85 | suc-≡ = ≈⇒≡ suc-≈ 86 | 87 | -- The standard trick to build a universal covering 88 | C : S¹ → Set 89 | C = S¹-elim[simp] ℤ suc-≡ 90 | 91 | abstract 92 | subst-C-loop : ∀ x → subst C loop x ≡ suc x 93 | subst-C-loop x = 94 | subst C loop x ≡⟨ sym $ subst-cong id C loop x ⟩ 95 | subst id (cong C loop) x ≡⟨ cong (λ p → subst id p x) $ S¹-elim[simp]-loop ℤ suc-≡ ⟩ 96 | subst id suc-≡ x ≡⟨ subst-id-univ suc-≡ x ⟩ 97 | _≈_.to (≡⇒≈ suc-≡) x ≡⟨ cong (λ weq → _≈_.to weq x) $ right-inverse-of suc-≈ ⟩ 98 | _≈_.to suc-≈ x ≡⟨ refl (suc x) ⟩∎ 99 | suc x ∎ 100 | where 101 | open _≈_ (≡≈≈ ℤ ℤ) 102 | 103 | subst-C-loop-pred : ∀ x → subst C loop (pred x) ≡ x 104 | subst-C-loop-pred x = 105 | subst C loop (pred x) ≡⟨ subst-C-loop (pred x) ⟩ 106 | suc (pred x) ≡⟨ suc-pred x ⟩∎ 107 | x ∎ 108 | 109 | subst-C-sym-loop : ∀ x → subst C (sym loop) x ≡ pred x 110 | subst-C-sym-loop x = subst-sym C loop x (pred x) $ subst-C-loop-pred x 111 | 112 | -- The counting function 113 | S¹loop⇒ℤ : base ≡ base → ℤ 114 | S¹loop⇒ℤ p = subst C p zero 115 | 116 | ------------------------------------------------------------------------ 117 | -- Homotopy equivalence 118 | 119 | private 120 | -- The right inverse (easy) 121 | right-inverse-of : ∀ z → S¹loop⇒ℤ (ℤ⇒S¹loop z) ≡ z 122 | right-inverse-of zero = refl _ 123 | right-inverse-of (pos n) = right-inverse-of-pos n 124 | where 125 | right-inverse-of-pos : ∀ n → S¹loop⇒ℤ (ℤ⇒S¹loop (pos n)) ≡ pos n 126 | right-inverse-of-pos ℕzero = 127 | subst C (trans (refl base) loop) zero ≡⟨ refl _ ⟩ 128 | subst C loop zero ≡⟨ subst-C-loop zero ⟩∎ 129 | pos ℕzero ∎ 130 | right-inverse-of-pos (ℕsuc n) = 131 | subst C (trans (repeat (ℕsuc n) loop) loop) zero ≡⟨ subst-trans C (repeat (ℕsuc n) loop) loop zero ⟩ 132 | subst C loop (subst C (repeat (ℕsuc n) loop) zero) ≡⟨ cong (subst C loop) $ right-inverse-of-pos n ⟩ 133 | subst C loop (pos n) ≡⟨ subst-C-loop (pos n) ⟩∎ 134 | pos (ℕsuc n) ∎ 135 | right-inverse-of (neg n) = right-inverse-of-neg n 136 | where 137 | right-inverse-of-neg : ∀ n → S¹loop⇒ℤ (ℤ⇒S¹loop (neg n)) ≡ neg n 138 | right-inverse-of-neg ℕzero = 139 | subst C (trans (refl base) (sym loop)) zero ≡⟨ refl _ ⟩ 140 | subst C (sym loop) zero ≡⟨ subst-C-sym-loop zero ⟩∎ 141 | neg ℕzero ∎ 142 | right-inverse-of-neg (ℕsuc n) = 143 | subst C (trans (repeat (ℕsuc n) (sym loop)) (sym loop)) zero ≡⟨ subst-trans C (repeat (ℕsuc n) (sym loop)) (sym loop) zero ⟩ 144 | subst C (sym loop) (subst C (repeat (ℕsuc n) (sym loop)) zero) ≡⟨ cong (subst C (sym loop)) $ right-inverse-of-neg n ⟩ 145 | subst C (sym loop) (neg n) ≡⟨ subst-C-sym-loop (neg n) ⟩∎ 146 | neg (ℕsuc n) ∎ 147 | 148 | -- Intuition: 149 | -- 150 | -- J rule works on paths, not loops... 151 | -- Let's free one of the end points! 152 | 153 | -- A trasporting loop to be fed into S¹-elim 154 | ℤ⇒S¹path-loop : ∀ z → subst (λ x → C x → base ≡ x) loop ℤ⇒S¹loop z ≡ ℤ⇒S¹loop z 155 | ℤ⇒S¹path-loop z = 156 | subst (λ x → C x → base ≡ x) loop ℤ⇒S¹loop z 157 | ≡⟨ sym $ subst-app C (λ x → base ≡ x) loop ℤ⇒S¹loop (subst-C-loop-pred z) ⟩ 158 | subst (λ x → base ≡ x) loop (ℤ⇒S¹loop (pred z)) 159 | ≡⟨ subst-path[idʳ] (const base) loop (ℤ⇒S¹loop (pred z)) ⟩ 160 | trans (sym (cong (const base) loop)) (trans (ℤ⇒S¹loop (pred z)) loop) 161 | ≡⟨ cong (λ p → trans (sym p) (trans (ℤ⇒S¹loop (pred z)) loop)) $ cong-const base loop ⟩ 162 | trans (ℤ⇒S¹loop (pred z)) loop 163 | ≡⟨ lemma z ⟩∎ 164 | ℤ⇒S¹loop z 165 | ∎ 166 | where 167 | lemma : ∀ z → trans (ℤ⇒S¹loop (pred z)) loop ≡ ℤ⇒S¹loop z 168 | lemma zero = trans-symˡ loop 169 | lemma (pos ℕzero) = refl _ 170 | lemma (pos (ℕsuc n)) = refl _ 171 | lemma (neg _) = trans-trans-symʳ _ _ 172 | 173 | -- A generalized Cx⇒S¹loop ... Cx⇒S¹path ! 174 | Cx⇒S¹path : ∀ (x : S¹) → C x → base ≡ x 175 | Cx⇒S¹path = S¹-elim (λ x → C x → base ≡ x) ℤ⇒S¹loop (ext ℤ⇒S¹path-loop) 176 | 177 | -- A generalized S¹loop⇒Cx ... S¹path⇒Cx ! 178 | S¹path⇒Cx : ∀ (x : S¹) → base ≡ x → C x 179 | S¹path⇒Cx x p = subst C p zero 180 | 181 | -- The generalized left inversion property (now easy?) 182 | left-inverse-of′ : ∀ (x : S¹) → (p : base ≡ x) → Cx⇒S¹path x (S¹path⇒Cx x p) ≡ p 183 | left-inverse-of′ x = elim′ (λ {x} p → Cx⇒S¹path x (S¹path⇒Cx x p) ≡ p) (refl _) 184 | 185 | -- The left inversion property we want 186 | left-inverse-of : (p : base ≡ base) → ℤ⇒S¹loop (S¹loop⇒ℤ p) ≡ p 187 | left-inverse-of = left-inverse-of′ base 188 | 189 | -- We have all the ingredients now! 190 | Ω₁S¹↔ℤ : Ω 1 base ↔ ℤ 191 | Ω₁S¹↔ℤ = 192 | record 193 | { surjection = record 194 | { to = S¹loop⇒ℤ 195 | ; from = ℤ⇒S¹loop 196 | ; right-inverse-of = right-inverse-of 197 | } 198 | ; left-inverse-of = left-inverse-of 199 | } 200 | -------------------------------------------------------------------------------- /Space/Sphere/TwoPoints.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Spheres (two-point version) 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | -- Currently we have S¹ and an experimental S² 8 | 9 | {-# OPTIONS --without-K #-} 10 | 11 | module Space.Sphere.TwoPoints where 12 | 13 | open import Prelude 14 | open import Path 15 | open import Path.Lemmas 16 | 17 | ------------------------------------------------------------------------ 18 | -- Formation rules and introduction rules 19 | 20 | private 21 | data S′¹′ : Set where 22 | north⁰′ : S′¹′ 23 | south⁰′ : S′¹′ 24 | 25 | S′¹ : Set 26 | S′¹ = S′¹′ 27 | 28 | north⁰ : S′¹ 29 | north⁰ = north⁰′ 30 | 31 | south⁰ : S′¹ 32 | south⁰ = south⁰′ 33 | 34 | postulate 35 | north¹ : south⁰ ≡ north⁰ 36 | south¹ : north⁰ ≡ south⁰ 37 | 38 | ------------------------------------------------------------------------ 39 | -- Elimination rules and computation rules 40 | 41 | -- Dependent version 42 | 43 | S′¹-elim : ∀ {ℓ} (P : S′¹ → Set ℓ) (pnorth⁰ : P north⁰) (psouth⁰ : P south⁰) → 44 | subst P north¹ psouth⁰ ≡ pnorth⁰ → subst P south¹ pnorth⁰ ≡ psouth⁰ → 45 | (x : S′¹) → P x 46 | S′¹-elim _ pnorth⁰ _ _ _ north⁰′ = pnorth⁰ 47 | S′¹-elim _ _ psouth⁰ _ _ south⁰′ = psouth⁰ 48 | 49 | -- They are actually definitionally equal... 50 | postulate 51 | S′¹-elim-north¹ : ∀ {ℓ} (P : S′¹ → Set ℓ) (pnorth⁰ : P north⁰) (psouth⁰ : P south⁰) 52 | (pnorth¹ : subst P north¹ psouth⁰ ≡ pnorth⁰) 53 | (psouth¹ : subst P south¹ pnorth⁰ ≡ psouth⁰) → 54 | cong[dep] P (S′¹-elim P pnorth⁰ psouth⁰ pnorth¹ psouth¹) north¹ ≡ pnorth¹ 55 | S′¹-elim-south¹ : ∀ {ℓ} (P : S′¹ → Set ℓ) (pnorth⁰ : P north⁰) (psouth⁰ : P south⁰) → 56 | (pnorth¹ : subst P north¹ psouth⁰ ≡ pnorth⁰) 57 | (psouth¹ : subst P south¹ pnorth⁰ ≡ psouth⁰) → 58 | cong[dep] P (S′¹-elim P pnorth⁰ psouth⁰ pnorth¹ psouth¹) south¹ ≡ psouth¹ 59 | 60 | -- Non-dependent version 61 | 62 | S′¹-elim[simp] : ∀ {ℓ} {P : Set ℓ} (pnorth psouth : P) → psouth ≡ pnorth → pnorth ≡ psouth → S′¹ → P 63 | S′¹-elim[simp] {P = P} pn ps pn¹ ps¹ = S′¹-elim (const P) pn ps (trans boring-n¹ pn¹) (trans boring-s¹ ps¹) 64 | where 65 | boring-n¹ = subst-const north¹ ps 66 | boring-s¹ = subst-const south¹ pn 67 | 68 | -- This propositional equality is derivable from the dependent version. 69 | postulate 70 | S′¹-elim[simp]-north¹ : ∀ {ℓ} {P : Set ℓ} (pnorth⁰ psouth⁰ : P) 71 | (pnorth¹ : psouth⁰ ≡ pnorth⁰) 72 | (psouth¹ : pnorth⁰ ≡ psouth⁰) → 73 | cong (S′¹-elim[simp] pnorth⁰ psouth⁰ pnorth¹ psouth¹) north¹ ≡ pnorth¹ 74 | S′¹-elim[simp]-south¹ : ∀ {ℓ} {P : Set ℓ} (pnorth⁰ psouth⁰ : P) → 75 | (pnorth¹ : psouth⁰ ≡ pnorth⁰) 76 | (psouth¹ : pnorth⁰ ≡ psouth⁰) → 77 | cong (S′¹-elim[simp] pnorth⁰ psouth⁰ pnorth¹ psouth¹) south¹ ≡ psouth¹ 78 | -------------------------------------------------------------------------------- /Space/Torus.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- Torus 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | 7 | {-# OPTIONS --without-K #-} 8 | 9 | module Space.Torus where 10 | 11 | open import Prelude 12 | 13 | open import Path 14 | open import Path.Lemmas 15 | open import Path.Symbol 16 | 17 | import Space.Sphere as S 18 | 19 | private 20 | data T²′ : Set where 21 | base′ : T²′ 22 | 23 | T² : Set 24 | T² = T²′ 25 | 26 | base : T² 27 | base = base′ 28 | 29 | postulate 30 | loopα : base ≡ base 31 | loopβ : base ≡ base 32 | cell : loopα ◇ loopβ ≡ loopβ ◇ loopα 33 | 34 | T²-elim : ∀ {ℓ} (P : T² → Set ℓ) (pbase : P base) 35 | (ploopα : subst P loopα pbase ≡ pbase) 36 | (ploopβ : subst P loopβ pbase ≡ pbase) 37 | (pcell : subst (λ p → subst P p pbase ≡ pbase) cell 38 | (subst-trans P loopα loopβ pbase ◇ cong (subst P loopβ) ploopα ◇ ploopβ) ≡ 39 | (subst-trans P loopβ loopα pbase ◇ cong (subst P loopα) ploopβ ◇ ploopα)) → 40 | ∀ x → P x 41 | T²-elim _ pbase′ _ _ _ base′ = pbase′ 42 | 43 | cong[dep]-trans : ∀ {ℓ₁} {ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) (f : (x : A) → B x) 44 | {x y} (loop₁ : x ≡ y) {z} (loop₂ : y ≡ z) → 45 | cong[dep] B f (loop₁ ◇ loop₂) ≡ 46 | subst-trans B loop₁ loop₂ (f x) ◇ cong (subst B loop₂) (cong[dep] B f loop₁) ◇ cong[dep] B f loop₂ 47 | cong[dep]-trans B f {x} = 48 | elim′ 49 | (λ {y} p₁ → ∀ {z} (p₂ : y ≡ z) → 50 | cong[dep] B f (p₁ ◇ p₂) ≡ 51 | subst-trans B p₁ p₂ (f x) ◇ cong (subst B p₂) (cong[dep] B f p₁) ◇ cong[dep] B f p₂) 52 | (elim′ 53 | (λ {z} p₂ → 54 | cong[dep] B f p₂ ≡ 55 | subst-trans B (refl _) p₂ (f x) ◇ cong[dep] B f p₂) 56 | (refl _)) 57 | 58 | cong[dep]-trans′ : ∀ {ℓ₁} {ℓ₂} {A : Set ℓ₁} (B : A → Set ℓ₂) (f : (x : A) → B x) 59 | {x y} (loop₁ : x ≡ y) {z} (loop₂ : y ≡ z) 60 | {bloop₁} (bloop₁-eq : cong[dep] B f loop₁ ≡ bloop₁) 61 | {bloop₂} (bloop₂-eq : cong[dep] B f loop₂ ≡ bloop₂) → 62 | cong[dep] B f (loop₁ ◇ loop₂) ≡ 63 | subst-trans B loop₁ loop₂ (f x) ◇ cong (subst B loop₂) bloop₁ ◇ bloop₂ 64 | cong[dep]-trans′ B f {x} = 65 | elim′ 66 | (λ {y} p₁ → ∀ {z} (p₂ : y ≡ z) {bp₁} (bp₁-eq : cong[dep] B f p₁ ≡ bp₁) {bp₂} (bp₂-eq : cong[dep] B f p₂ ≡ bp₂) → 67 | cong[dep] B f (p₁ ◇ p₂) ≡ 68 | subst-trans B p₁ p₂ (f x) ◇ cong (subst B p₂) bp₁ ◇ bp₂) 69 | (elim′ 70 | (λ {z} p₂ → ∀ {bp₁} (bp₁-eq : refl (f x) ≡ bp₁) {bp₂} (bp₂-eq : cong[dep] B f p₂ ≡ bp₂) → 71 | cong[dep] B f p₂ ≡ 72 | cong (subst B p₂) bp₁ ◇ bp₂) 73 | (λ {bp₁} bp₁-eq {bp₂} bp₂-eq → 74 | bp₂-eq ◇ cong (λ p → p ◇ bp₂) (bp₁-eq ◇ sym (cong-id bp₁)))) 75 | 76 | postulate 77 | T²-elim-loopα : ∀ {ℓ} (P : T² → Set ℓ) (pbase : P base) 78 | (ploopα : subst P loopα pbase ≡ pbase) 79 | (ploopβ : subst P loopβ pbase ≡ pbase) 80 | (pcell : subst (λ p → subst P p pbase ≡ pbase) cell 81 | (subst-trans P loopα loopβ pbase ◇ cong (subst P loopβ) ploopα ◇ ploopβ) ≡ 82 | (subst-trans P loopβ loopα pbase ◇ cong (subst P loopα) ploopβ ◇ ploopα)) → 83 | cong[dep] P (T²-elim P pbase ploopα ploopβ pcell) loopα ≡ ploopα 84 | 85 | T²-elim-loopβ : ∀ {ℓ} (P : T² → Set ℓ) (pbase : P base) 86 | (ploopα : subst P loopα pbase ≡ pbase) 87 | (ploopβ : subst P loopβ pbase ≡ pbase) 88 | (pcell : subst (λ p → subst P p pbase ≡ pbase) cell 89 | (subst-trans P loopα loopβ pbase ◇ cong (subst P loopβ) ploopα ◇ ploopβ) ≡ 90 | (subst-trans P loopβ loopα pbase ◇ cong (subst P loopα) ploopβ ◇ ploopα)) → 91 | cong[dep] P (T²-elim P pbase ploopα ploopβ pcell) loopβ ≡ ploopβ 92 | 93 | T²-elim-cell : ∀ {ℓ} (P : T² → Set ℓ) (pbase : P base) 94 | (ploopα : subst P loopα pbase ≡ pbase) 95 | (ploopβ : subst P loopβ pbase ≡ pbase) 96 | (pcell : subst (λ p → subst P p pbase ≡ pbase) cell 97 | (subst-trans P loopα loopβ pbase ◇ cong (subst P loopβ) ploopα ◇ ploopβ) ≡ 98 | (subst-trans P loopβ loopα pbase ◇ cong (subst P loopα) ploopβ ◇ ploopα)) → 99 | let 100 | eliminator = T²-elim P pbase ploopα ploopβ pcell 101 | elim-loopα = T²-elim-loopα P pbase ploopα ploopβ pcell 102 | elim-loopβ = T²-elim-loopβ P pbase ploopα ploopβ pcell 103 | convertαβ = cong[dep]-trans′ P eliminator loopα loopβ elim-loopα elim-loopβ 104 | convertβα = cong[dep]-trans′ P eliminator loopβ loopα elim-loopβ elim-loopα 105 | in 106 | cong (subst (λ p → subst P p pbase ≡ pbase) cell) (sym convertαβ) ◇ 107 | cong[dep] (λ p → subst P p pbase ≡ pbase) (cong[dep] P eliminator) cell ◇ 108 | convertβα ≡ 109 | pcell 110 | -------------------------------------------------------------------------------- /Univalence.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- The univalence axiom 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | -- Partly based on Voevodsky's work on univalent foundations. 11 | 12 | module Univalence where 13 | 14 | open import Prelude 15 | open import Path 16 | open import Path.Lemmas 17 | open import Map.WeakEquivalence as Weak hiding (_∘_; id) 18 | 19 | ------------------------------------------------------------------------ 20 | -- The univalence axiom 21 | 22 | -- If two sets are equal, then they are weakly equivalent. 23 | 24 | ≡⇒≈ : ∀ {ℓ} {A B : Set ℓ} → A ≡ B → A ≈ B 25 | ≡⇒≈ = elim (λ {A B} _ → A ≈ B) (λ _ → Weak.id) 26 | 27 | -- The univalence axiom states that ≡⇒≈ is a weak equivalence. 28 | 29 | Univalence-axiom : ∀ {ℓ} → Set ℓ → Set ℓ → Set (lsuc ℓ) 30 | Univalence-axiom A B = Is-weak-equivalence (≡⇒≈ {A = A} {B = B}) 31 | 32 | ------------------------------------------------------------------------ 33 | -- Experimental Area 34 | -- 35 | -- Another insight given by Robert Harper: 36 | -- The Univalence axiom is equivalence induction! 37 | -- 38 | -- Note: Here we show that "weak-equivalent induction" 39 | -- and the Univalence axiom are mutually derivable. 40 | 41 | private 42 | 43 | -- The definition of "weak-equivalence induction" 44 | ≈-Elim : ∀ (ℓ₁ ℓ₂ : Level) → Set (lsuc (ℓ₁ ⊔ ℓ₂)) 45 | ≈-Elim ℓ₁ ℓ₂ = ∀ (P : {A B : Set ℓ₁} → A ≈ B → Set ℓ₂) → 46 | ((A : Set ℓ₁) → P (Weak.id {A = A})) → 47 | ∀ {A B : Set ℓ₁} (A≈B : A ≈ B) → P A≈B 48 | 49 | ≈-Elim-id : ∀ (ℓ₁ ℓ₂ : Level) → ≈-Elim ℓ₁ ℓ₂ → Set (lsuc (ℓ₁ ⊔ ℓ₂)) 50 | ≈-Elim-id ℓ₁ ℓ₂ ≈-elim = ∀ (P : {A B : Set ℓ₁} → A ≈ B → Set ℓ₂) → 51 | (pid : (A : Set ℓ₁) → P (Weak.id {A = A})) → 52 | ∀ {A} → ≈-elim P pid Weak.id ≡ pid A 53 | 54 | -- The elimination and computation rules imply Univalence 55 | weq-elim⇒univ : ∀ {ℓ} (≈-elim : ≈-Elim ℓ (lsuc ℓ)) → ≈-Elim-id ℓ (lsuc ℓ) ≈-elim → 56 | ∀ (A B : Set ℓ) → Univalence-axiom A B 57 | weq-elim⇒univ {ℓ} ≈-elim ≈-elim-id A B = 58 | _≈_.is-weak-equivalence $ ↔⇒≈ $ record 59 | { surjection = record 60 | { to = ≡⇒≈ 61 | ; from = ≈⇒≡ 62 | ; right-inverse-of = right-inverse-of 63 | } 64 | ; left-inverse-of = left-inverse-of 65 | } 66 | where 67 | ≈⇒≡ : ∀ {A B : Set ℓ} → A ≈ B → A ≡ B 68 | ≈⇒≡ = ≈-elim (λ {A} {B} A≈B → A ≡ B) (λ A → refl A) 69 | 70 | ≈⇒≡-id : ∀ (A : Set ℓ) → ≈⇒≡ Weak.id ≡ refl A 71 | ≈⇒≡-id A = ≈-elim-id (λ {A B} A≈B → A ≡ B) (λ A → refl A) 72 | 73 | left-inverse-of : ∀ (A≡B : A ≡ B) → ≈⇒≡ (≡⇒≈ A≡B) ≡ A≡B 74 | left-inverse-of A≡B = elim 75 | (λ {A B} A≡B → ≈⇒≡ (≡⇒≈ A≡B) ≡ A≡B) 76 | (λ A → ≈⇒≡-id A) 77 | A≡B 78 | 79 | -- Here we use ↑ to handle the universe mismatch. 80 | right-inverse-of : ∀ (A≈B : A ≈ B) → ≡⇒≈ (≈⇒≡ A≈B) ≡ A≈B 81 | right-inverse-of A≈B = ↑.lower $ ≈-elim 82 | (λ {A B : Set ℓ} A≈B → ↑ (lsuc ℓ) (≡⇒≈ (≈⇒≡ A≈B) ≡ A≈B)) 83 | (λ A → lift (cong ≡⇒≈ (≈⇒≡-id A))) 84 | A≈B 85 | 86 | -- The other way around. The Univalence axiom implies both. 87 | -- 88 | -- Maybe we should have one big lemma outputting a pair containing 89 | -- the elimination rule and the computation rule directly? 90 | univ⇒weq-elim : ∀ {ℓ ℓ′} → (∀ (A B : Set ℓ) → Univalence-axiom A B) → ≈-Elim ℓ ℓ′ 91 | univ⇒weq-elim {ℓ} {ℓ′} univ P pid {A} {B} A≈B = 92 | subst P right-inverse (elim (λ {A B : Set ℓ} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ A≈B)) 93 | where 94 | A≡B≈A≈B : (A ≡ B) ≈ (A ≈ B) 95 | A≡B≈A≈B = weq ≡⇒≈ (univ A B) 96 | 97 | ≈⇒≡ : A ≈ B → A ≡ B 98 | ≈⇒≡ = _≈_.from A≡B≈A≈B 99 | 100 | right-inverse : ≡⇒≈ (≈⇒≡ A≈B) ≡ A≈B 101 | right-inverse = _≈_.right-inverse-of A≡B≈A≈B A≈B 102 | 103 | univ⇒weq-elim-id : ∀ {ℓ ℓ′} (univ : ∀ (A B : Set ℓ) → Univalence-axiom A B) → 104 | ≈-Elim-id ℓ ℓ′ (univ⇒weq-elim univ) 105 | univ⇒weq-elim-id {ℓ} {ℓ′} univ P pid {A} = 106 | subst P right-inverse (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ Weak.id)) 107 | ≡⟨ cong (λ p → subst P p (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ Weak.id))) $ sym cong-left ⟩ 108 | subst P (cong ≡⇒≈ left-inverse) (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ Weak.id)) 109 | ≡⟨ subst-cong P ≡⇒≈ left-inverse (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ Weak.id)) ⟩ 110 | subst (P ∘ ≡⇒≈) left-inverse (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (≈⇒≡ Weak.id)) 111 | ≡⟨ cong[dep] (P ∘ ≡⇒≈) (elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid) left-inverse ⟩∎ 112 | elim (λ {_ _} A≡B → P (≡⇒≈ A≡B)) pid (refl A) 113 | ∎ 114 | where 115 | A≡A≈A≈A : (A ≡ A) ≈ (A ≈ A) 116 | A≡A≈A≈A = weq ≡⇒≈ (univ A A) 117 | 118 | ≈⇒≡ : A ≈ A → A ≡ A 119 | ≈⇒≡ = _≈_.from A≡A≈A≈A 120 | 121 | left-inverse : ≈⇒≡ Weak.id ≡ refl A 122 | left-inverse = _≈_.left-inverse-of A≡A≈A≈A (refl A) 123 | 124 | right-inverse : ≡⇒≈ (≈⇒≡ Weak.id) ≡ Weak.id 125 | right-inverse = _≈_.right-inverse-of A≡A≈A≈A Weak.id 126 | 127 | cong-left : cong ≡⇒≈ left-inverse ≡ right-inverse 128 | cong-left = _≈_.left-right-lemma A≡A≈A≈A (refl A) 129 | -------------------------------------------------------------------------------- /Univalence/Lemmas.agda: -------------------------------------------------------------------------------- 1 | ------------------------------------------------------------------------ 2 | -- More lemmas 3 | ------------------------------------------------------------------------ 4 | 5 | -- Copyright (c) 2012 Favonia 6 | -- Copyright (c) 2011-2012 Nils Anders Danielsson 7 | 8 | {-# OPTIONS --without-K #-} 9 | 10 | open import Univalence 11 | 12 | module Univalence.Lemmas 13 | (univ : ∀ {ℓ} (A B : Set ℓ) → Univalence-axiom A B) where 14 | 15 | open import Prelude 16 | open import Path 17 | open import Path.Lemmas 18 | open import Map.H-equivalence hiding (_∘_; id) 19 | open import Map.WeakEquivalence as Weak hiding (_∘_; id) 20 | 21 | ------------------------------------------------------------------------ 22 | -- Conversions between homotopy equivalences, weak equivalences, and identities 23 | 24 | ≡≈≈ : ∀ {ℓ} (A B : Set ℓ) → (A ≡ B) ≈ (A ≈ B) 25 | ≡≈≈ A B = weq ≡⇒≈ (univ A B) 26 | 27 | ≈⇒≡ : ∀ {ℓ} {A B : Set ℓ} → A ≈ B → A ≡ B 28 | ≈⇒≡ = _≈_.from (≡≈≈ _ _) 29 | 30 | subst-id-univ : ∀ {ℓ} {A B : Set ℓ} (A≡B : A ≡ B) (x : A) → subst id A≡B x ≡ _≈_.to (≡⇒≈ A≡B) x 31 | subst-id-univ {ℓ} = 32 | elim 33 | (λ {A B : Set ℓ} (A≡B : A ≡ B) → (x : A) → subst id A≡B x ≡ _≈_.to (≡⇒≈ A≡B) x) 34 | (λ A x → 35 | subst id (refl A) x ≡⟨ refl _ ⟩ 36 | x ≡⟨ refl x ⟩ 37 | (_≈_.to Weak.id) x ≡⟨ refl _ ⟩∎ 38 | (_≈_.to (≡⇒≈ (refl A))) x ∎) 39 | 40 | postulate 41 | ext : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} → Extensionality A B 42 | 43 | postulate 44 | ext-comp : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : Set ℓ₂} {f g : A → B} 45 | (reason : ∀ x → f x ≡ g x) → 46 | ∀ x → cong (λ f → f x) (ext reason) ≡ reason x 47 | 48 | postulate 49 | ext[dep] : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : A → Set ℓ₂} → Extensionality[dep] A B 50 | 51 | postulate 52 | ext-comp[dep] : ∀ {ℓ₁ ℓ₂} {A : Set ℓ₁} {B : A → Set ℓ₂} {f g : (x : A) → B x} 53 | (reason : ∀ x → f x ≡ g x) → 54 | ∀ x → cong (λ f → f x) (ext[dep] reason) ≡ reason x 55 | 56 | unique-neq-proof : ∀ {ℓ} {A : Set ℓ} {a₁ a₂ : A} (neq₁ neq₂ : ¬ a₁ ≡ a₂) → neq₁ ≡ neq₂ 57 | unique-neq-proof neq₁ neq₂ = ext $ ⊥-elim ∘ neq₁ 58 | --------------------------------------------------------------------------------