diff --git a/src/CubicalHott/Chapter2.lagda.md b/src/CubicalHott/Chapter2.lagda.md index 1ecb81b..1fca2d5 100644 --- a/src/CubicalHott/Chapter2.lagda.md +++ b/src/CubicalHott/Chapter2.lagda.md @@ -186,6 +186,34 @@ module lemma2∙4∙12 where -- in g ∘ f , qinv-to-isequiv (mkQinv (f-inv ∘ g-inv) forward backward) ``` +### Lemma 2.9.2 + +``` +happly : {A B : Type} + → {f g : A → B} + → (p : f ≡ g) + → (x : A) + → f x ≡ g x +happly {A} {B} {f} {g} p x = ap (λ h → h x) p + +happlyd : {A : Type} {B : A → Type} + → {f g : (a : A) → B a} + → (p : f ≡ g) + → (x : A) + → f x ≡ g x +happlyd {A} {B} {f} {g} p x = ap (λ h → h x) p +``` + +### Theorem 2.9.3 (Function extensionality) + +``` +funext : {l : Level} {A B : Type l} + → {f g : A → B} + → ((x : A) → f x ≡ g x) + → f ≡ g +funext h i x = h x i +``` + ## 2.10 Universes and the univalence axiom ### Lemma 2.10.1 @@ -222,6 +250,31 @@ module axiom2∙10∙3 where open axiom2∙10∙3 ``` +### Theorem 2.11.2 + +``` +module lemma2∙11∙2 where + open ≡-Reasoning + + i : {l : Level} {A : Type l} {a x1 x2 : A} + → (p : x1 ≡ x2) + → (q : a ≡ x1) + → transport (λ y → a ≡ y) p q ≡ q ∙ p + i {l} {A} {a} {x1} {x2} p q j = {! !} + + ii : {l : Level} {A : Type l} {a x1 x2 : A} + → (p : x1 ≡ x2) + → (q : x1 ≡ a) + → transport (λ y → y ≡ a) p q ≡ sym p ∙ q + ii {l} {A} {a} {x1} {x2} p q = {! !} + + iii : {l : Level} {A : Type l} {a x1 x2 : A} + → (p : x1 ≡ x2) + → (q : x1 ≡ x1) + → transport (λ y → y ≡ y) p q ≡ sym p ∙ q ∙ p + iii {l} {A} {a} {x1} {x2} p q = {! !} +``` + ### Remark 2.12.6 ``` diff --git a/src/CubicalHott/Chapter6.lagda.md b/src/CubicalHott/Chapter6.lagda.md index c62650a..3025b3b 100644 --- a/src/CubicalHott/Chapter6.lagda.md +++ b/src/CubicalHott/Chapter6.lagda.md @@ -32,4 +32,22 @@ lemma6∙4∙1 p = p ≡⟨ {! !} ⟩ p ≡⟨ {! !} ⟩ refl ∎ +``` + +### Lemma 6.4.2 + +``` +lemma6∙4∙2 : Σ ((x : S¹) → x ≡ x) (λ y → y ≢ (λ x → refl)) +lemma6∙4∙2 = f , g + where + f : (x : S¹) → x ≡ x + f base = loop + f (loop i) i₁ = loop {! !} + + g : f ≢ (λ x → refl) + g p = let + z = happlyd p + z2 = z base + z3 = lemma6∙4∙1 z2 + in z3 ``` \ No newline at end of file diff --git a/src/HoTTEST/Agda/Agda-installation.md b/src/HoTTEST/Agda/Agda-installation.md new file mode 100644 index 0000000..4ed3b9e --- /dev/null +++ b/src/HoTTEST/Agda/Agda-installation.md @@ -0,0 +1,78 @@ +# Instructions on how to get Agda 2.6.2.2 + +## Online + +Thanks to Amélia Liao, there's an [online version](https://hcomp.io/) available, +which means you won't have to install anything. But if you would like to +install Agda locally, please see below. + +## Mac + +If you're a Mac user who already uses homebrew, the process is painless: just +run `brew install` agda and get ready to wait a while. + +If you're a Mac user who doesn't already use homebrew, follow the [instructions +here](https://brew.sh/) to get it working. + +## Linux + +If you're an Arch or Fedora user, then `pacman -S agda` or `yum install agda` +should work. See also [theses Arch instructions](https://pastebin.com/jj2c2dqR). + +If you're Debian/Ubuntu, then you should install Agda through `cabal`, because +the Agda in the repositories of the operating system is outdated. See [these +instructions](https://agda.readthedocs.io/en/v2.6.2.2/getting-started/installation.html). + +## Windows + +For Windows, see [these instructions](https://pastebin.com/Zguv4743). (Thanks +Todd Waugh Ambridge!) + +## Editing Agda code + +Most people use emacs with agda-mode to edit, but you can also use neovim with +cornelis or vscode with the agda-mode port. + +To get emacs with agda-mode set up, there's a [tutorial +here](https://agda.readthedocs.io/en/v2.6.2.2/getting-started/installation.html#running-the-agda-mode-program) + +To get neovim with corenlis setup, there's installation instructions on [their +GitHub](https://github.com/isovector/cornelis) + +To get the vscode plugin setup, see +[here](https://marketplace.visualstudio.com/items?itemName=banacorn.agda-mode) + +## Agda in Markdown files + +If you're using emacs, then add the following at the very end of your `.emacs` +file (in your home directory) + +``` +(add-to-list 'auto-mode-alist '("\\.lagda.md\\'" . agda2-mode)) +``` + +to enable agda-mode for Markdown files with Agda code, like the lecture notes +and exercises. + +## Checking that everything works + +After installing Agda, you should check that it's installed correctly by +type-checking the [Hello, Agda! +program](https://agda.readthedocs.io/en/v2.6.2.2/getting-started/hello-world.html#hello-agda) + +## Working outside the repository + +If you would like to work on the exercises and have the lecture notes available +outside (for `open import`) the repository, then you should register the +repository as a library: + +In `~/.agda/libraries` add the line +``` +~/HoTTEST-Summer-School/Agda/agda-lectures.agda-lib +``` +and then add the line `agda-lectures` to `~/.agda/defaults`. + +## Getting help + +If you're having trouble installing Agda, please ask for help in the +`#agda-installation` channel on the summer school's Discord server. diff --git a/src/HoTTEST/Agda/Auxiliary/Hurkens.lagda.md b/src/HoTTEST/Agda/Auxiliary/Hurkens.lagda.md new file mode 100644 index 0000000..830f436 --- /dev/null +++ b/src/HoTTEST/Agda/Auxiliary/Hurkens.lagda.md @@ -0,0 +1,69 @@ +# Hurkens' paradox +(File provided by Ulrik Buchholtz) + +Hurkens' simplification of Girard's paradox. +This derives a contradiction using only Π-types +and “type in type”, a universe that contains itself. +For a description of how it works, and the +similarity to Burali-Forti's paradox, see: + +https://www.cs.cmu.edu/~kw/scans/hurkens95tlca.pdf + +```agda +{-# OPTIONS --without-K --type-in-type #-} +module Hurkens where + +Type = Set + +P : Type → Type +P X = X → Type + +PP : Type → Type +PP X = P (P X) + +⊥ : Type +⊥ = (X : Type) → X + +¬ : Type → Type +¬ X = X → ⊥ + +U : Type +U = (X : Type) → (PP X → X) → PP X + +τ : PP U → U +τ t X f p = t λ x → p (f (x X f)) + +σ : U → PP U +σ s = s U λ t → τ t + +Δ : P U +Δ y = ¬ ((p : P U) → σ y p → p (τ (σ y))) + +Ω : U +Ω = τ λ p → (x : U) → σ x p → p x + +D : Type +D = (p : P U) → σ Ω p → p (τ (σ Ω)) + +lemma : (p : P U) → ((x : U) → σ x p → p x) → p Ω +lemma p H = H Ω λ x → H (τ (σ x)) + +nd : ¬ D +nd = lemma Δ λ x H K → K Δ H λ p → K λ y → p (τ (σ y)) + +d : D +d p = lemma λ y → p (τ (σ y)) + +boom : ⊥ +boom = nd d +``` + +With this encoding of false we can of course inhabit all types, +inlucing the inductively defined empty type. + +```agda +data ∅ : Type where + +boom' : ∅ +boom' = boom ∅ +``` diff --git a/src/HoTTEST/Agda/Auxiliary/README.md b/src/HoTTEST/Agda/Auxiliary/README.md new file mode 100644 index 0000000..89358b5 --- /dev/null +++ b/src/HoTTEST/Agda/Auxiliary/README.md @@ -0,0 +1,10 @@ +# Auxiliary Agda files + +This folder contains some auxiliary files +for the Agda track of the HoTTEST Summer School +2022. +Everything here is optional material. + +1. Paradoxes derived from “type in type”: + 1. [Hurkens' paradox](Hurkens.lagda.md) + 1. [Russell's paradox](Russell.lagda.md) diff --git a/src/HoTTEST/Agda/Auxiliary/Russell.lagda.md b/src/HoTTEST/Agda/Auxiliary/Russell.lagda.md new file mode 100644 index 0000000..d1b4db8 --- /dev/null +++ b/src/HoTTEST/Agda/Auxiliary/Russell.lagda.md @@ -0,0 +1,83 @@ +# Russell's paradox +(File provided by Ulrik Buchholtz) + +Here we derive a contradiction from “type in type”, +a universe that contains itself, assuming that +this is additionally closed under generalized +inductive types. +Then we can directly encode Russell's paradox. + +This development is inspired by Aczel's +[“sets as trees” interpretation of constructive set +theory](https://doi.org/10.1016/S0049-237X(08)71989-X), +from which this very simple paradox was +[deduced by Thierry Coquand](https://doi.org/10.1007/BF01995104). +N.G. de Bruijn also noticed a similar paradox. + +```agda +{-# OPTIONS --without-K --type-in-type #-} +module Russell where + +Type = Set +``` + +We need some standard inductive types: + Σ, ∅, and identity types. + +```agda +data Σ (A : Type) (B : A → Type) : Type where + _,_ : (a : A) → B a → Σ A B + +pr₁ : {A : Type} {B : A → Type} → Σ A B → A +pr₁ (a , b) = a + +data ∅ : Type where + +data _≡_ {A : Type} (a : A) : A → Type where + refl : a ≡ a +``` + +Then we use the following generalized inductive type. +(It's called generalized because it has a recursive +argument that is a function type.) +It represents trees that can branch according to +any type in the universe. +Without “type in type”, these would be well-founded +trees, like the membership trees in the cumulative +hierarchy model of set theory. + +```agda +data V : Type where + sup : (X : Type) → (X → V) → V +``` + +We can similarly define membership and non-membership. + +```agda +_∈_ : V → V → Type +x ∈ sup X f = Σ X λ x' → x ≡ f x' + +_∉_ : V → V → Type +x ∉ y = x ∈ y → ∅ +``` + +Now it's straight-forward to define Russell's +paradoxical “set” of all sets that don't contain +themselves, and derive a contradiction. + +```agda +R : V +R = sup (Σ V λ x → x ∉ x) pr₁ + +x∈R→x∉x : (x : V) → x ∈ R → x ∉ x +x∈R→x∉x x ((.x , x∉x) , refl) = x∉x + +x∉x→x∈R : (x : V) → x ∉ x → x ∈ R +x∉x→x∈R x x∉x = (x , x∉x) , refl + +R∉R : R ∉ R +R∉R R∈R = x∈R→x∉x R R∈R R∈R + +boom : ∅ +boom = R∉R (x∉x→x∈R R R∉R) +``` diff --git a/src/HoTTEST/Agda/Cubical/Exercises7.lagda.md b/src/HoTTEST/Agda/Cubical/Exercises7.lagda.md new file mode 100644 index 0000000..6dbe8a4 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Exercises7.lagda.md @@ -0,0 +1,224 @@ +# Week 7 - Cubical Agda Exercises + +## Standard disclaimer: + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +In case you haven't done the other Agda exercises: +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**When solving the problems, +please make a copy of this file to work in, so that it doesn't get overwritten +(in case we update the exercises through `git`)!** + + +```agda +{-# OPTIONS --cubical --allow-unsolved-metas #-} + +module Exercises7 where + +open import cubical-prelude +open import Lecture7-notes +``` + +```agda +private + variable + A : Type ℓ + B : A → Type ℓ +``` + +## Part I: Generalizing to dependent functions + +### Exercise 1 (★) + +State and prove funExt for dependent functions `f g : (x : A) → B x` + +### Exercise 2 (★) + +Generalize the type of ap to dependent function `f : (x : A) → B x` +(hint: the result should be a `PathP`) + + +## Part II: Some facts about (homotopy) propositions and sets + +The first three homotopy levels `isContr`, `isProp` and `isSet` +are defined in `cubical-prelude` in the usual way +(only using path types of course). + +### Exercise 3 (★) + +State and prove that inhabited propositions are contractible + + +### Exercise 4 (★) + +Prove + +```agda +isPropΠ : (h : (x : A) → isProp (B x)) → isProp ((x : A) → B x) +isPropΠ = {!!} +``` + +### Exercise 5 (★) + +Prove the inverse of `funExt` (sometimes called `happly`): + +```agda +funExt⁻ : {f g : (x : A) → B x} → f ≡ g → ((x : A) → f x ≡ g x) +funExt⁻ = {!!} +``` + +### Exercise 6 (★★) + +Use funExt⁻ to prove isSetΠ: + +```agda +isSetΠ : (h : (x : A) → isSet (B x)) → isSet ((x : A) → B x) +isSetΠ = {!!} +``` + +### Exercise 7 (★★★): alternative contractibility of singletons + +We could have defined the type of singletons as follows + +```agda +singl' : {A : Type ℓ} (a : A) → Type ℓ +singl' {A = A} a = Σ x ꞉ A , x ≡ a +``` + +Prove the corresponding version of contractibility of singetons for +`singl'` (hint: use a suitable combinations of connections and `~_`) + +```agda +isContrSingl' : (x : A) → isContr (singl' x) +isContrSingl' x = {!!} +``` + +## Part III: Equality in Σ-types +### Exercise 8 (★★) + +Having the primitive notion of equality be heterogeneous is an +easily overlooked, but very important, strength of cubical type +theory. This allows us to work directly with equality in Σ-types +without using transport. + +Fill the following holes (some of them are easier than you might think): + +```agda +module _ {A : Type ℓ} {B : A → Type ℓ'} {x y : Σ A B} where + + ΣPathP : Σ p ꞉ pr₁ x ≡ pr₁ y , PathP (λ i → B (p i)) (pr₂ x) (pr₂ y) + → x ≡ y + ΣPathP = {!!} + + PathPΣ : x ≡ y + → Σ p ꞉ pr₁ x ≡ pr₁ y , PathP (λ i → B (p i)) (pr₂ x) (pr₂ y) + PathPΣ = {!!} + + ΣPathP-PathPΣ : ∀ p → PathPΣ (ΣPathP p) ≡ p + ΣPathP-PathPΣ = {!!} + + PathPΣ-ΣPathP : ∀ p → ΣPathP (PathPΣ p) ≡ p + PathPΣ-ΣPathP = {!!} +``` + +If one looks carefully the proof of prf in Lecture 7 uses ΣPathP +inlined, in fact this is used all over the place when working +cubically and simplifies many proofs which in Book HoTT requires long +complicated reasoning about transport. + + +## Part IV: Some HITs + +Now we want prove some identities of HITs analogous to `Torus ≡ S¹ × S¹` +Hint: you can just use `isoToPath` in all of them. + + +### Exercise 9 (★★) + +We saw two definitions of the torus: +`Torus` having a dependent path constructor `square` +and `Torus'` with a path constructor `square` that involves composition. + +Using these two ideas, define the *Klein bottle* in two different ways. + +### Exercise 10 (★★) + +Prove + +```agda +suspUnitChar : Susp Unit ≡ Interval +suspUnitChar = {!!} +``` + + +### Exercise 11 (★★★) + +Define suspension using the Pushout HIT and prove that it's equal to Susp. + + +### Exercise 12 (🌶) + +The goal of this exercise is to prove + +```agda +suspBoolChar : Susp Bool ≡ S¹ +suspBoolChar = {!!} +``` + +For the map `Susp Bool → S¹`, we have to specify the behavior of two +path constructors. The idea is to map one to `loop` and one to `refl`. + +For the other direction, we have to fix one base point in `Susp Bool` +and give a non-trivial loop on it, i.e. one that shouldn't cancel out +to `refl`. + +Proving that the two maps cancel each other requires some primitives +of `Cubical Agda` that we won't really discuss in the lectures, +namely `hcomp` and `hfill`. These are used (among other things) +to define path composition and ensure that it behaves correctly. + +Path composition corresponds to the top of the following square: + +```text + p∙q + x --------> z + ^ ^ + ¦ ¦ + refl ¦ sq ¦ q + ¦ ¦ + ¦ ¦ + x --------> y + p +``` + +The type of `sq` is a `PathP` sending `p` to `p ∙ q` +over `q` and the following lemma lets us construct such a *square*: + +```agda +comp-filler : {x y z : A} (p : x ≡ y) (q : y ≡ z) + → PathP (λ j → refl {x = x} j ≡ q j) p (p ∙ q) +comp-filler {x = x} p q j i = hfill (λ k → λ { (i = i0) → x + ; (i = i1) → q k }) + (inS (p i)) j +``` + +You can use this `comp-filler` as a black-box for proving the round-trips +in this exercise. + +**Hint:** For one of the round-trips you only need the following familiar +result, that is a direct consequence of `comp-filler` in `Cubical Agda` + +```agda +rUnit : {x y : A} (p : x ≡ y) → p ∙ refl ≡ p +rUnit p = sym (comp-filler p refl) +``` diff --git a/src/HoTTEST/Agda/Cubical/Exercises8.lagda.md b/src/HoTTEST/Agda/Cubical/Exercises8.lagda.md new file mode 100644 index 0000000..dcce9e2 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Exercises8.lagda.md @@ -0,0 +1,126 @@ +# Week 8 - Cubical Agda Exercises + +## Standard disclaimer: + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +In case you haven't done the other Agda exercises: +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**When solving the problems, +please make a copy of this file to work in, so that it doesn't get overwritten +(in case we update the exercises through `git`)!** + + +```agda +{-# OPTIONS --cubical --allow-unsolved-metas #-} + +module Exercises8 where + +open import cubical-prelude +open import Lecture7-notes +open import Lecture8-notes +open import Solutions7 +``` + +## Part I: `transp` and `hcomp` + +### Exercise 1 (★) + +Prove the propositional computation law for `J`: + +```agda +JRefl : {A : Type ℓ} {x : A} (P : (z : A) → x ≡ z → Type ℓ'') + (d : P x refl) → J P d refl ≡ d +JRefl P d = {!!} +``` + +### Exercise 2 (★★) + +Using `transp`, construct a method for turning dependent paths into paths. + +**Hint**: +Transport the point `p i` to the fibre `A i1`, and set `φ = i` such that the +transport computes away at `i = i1`. +```text + x ----(p i)----> y + A i0 A i A i1 +``` + +```agda +fromPathP : {A : I → Type ℓ} {x : A i0} {y : A i1} → + PathP A x y → transport (λ i → A i) x ≡ y +fromPathP {A = A} p i = {!!} +``` + +### Exercise 3 (★★★) + +Using `hcomp`, cunstruct a method for turning paths into dependent paths. + +**Hint**: +At each point `i`, the verical fibre of the following diagram should lie in +`A i`. The key is to parametrise the bottom line with a dependent path from `x` +to `transport A x`. This requires writing a `transp` that computes away at +`i = i0`. + +```text + x - - - -> y + ^ ^ + ¦ ¦ + refl ¦ ¦ p i + ¦ ¦ + ¦ ¦ + x ---> transport A x +``` + + +```agda +toPathP : {A : I → Type ℓ} {x : A i0} {y : A i1} → + transport (λ i → A i) x ≡ y → PathP A x y +toPathP {A = A} {x = x} p i = + hcomp + (λ {j (i = i0) → {!!} ; + j (i = i1) → {!!} }) + {!!} +``` + +### Exercise 4 (★) + +Using `toPathP`, prove the following lemma, which lets you fill in dependent +lines in hProps, provided their boundary. + +```agda +isProp→PathP : {A : I → Type ℓ} (p : (i : I) → isProp (A i)) + (a₀ : A i0) (a₁ : A i1) → PathP A a₀ a₁ +isProp→PathP p a₀ a₁ = {!!} +``` + +### Exercise 5 (★★) + +Prove the following lemma charictarising equality in subtypes: + +```agda +Σ≡Prop : {A : Type ℓ} {B : A → Type ℓ'} {u v : Σ A B} (h : (x : A) → isProp (B x)) + → (p : pr₁ u ≡ pr₁ v) → u ≡ v +Σ≡Prop {B = B} {u = u} {v = v} h p = {!!} +``` + +### Exercise 6 (★★★) + +Prove that `isContr` is a proposition: + +**Hint**: +This requires drawing a cube (yes, an actual 3D one)! + +```agda +isPropIsContr : {A : Type ℓ} → isProp (isContr A) +isPropIsContr (c0 , h0) (c1 , h1) j = {!!} +``` diff --git a/src/HoTTEST/Agda/Cubical/Exercises9.lagda.md b/src/HoTTEST/Agda/Cubical/Exercises9.lagda.md new file mode 100644 index 0000000..17696c7 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Exercises9.lagda.md @@ -0,0 +1,205 @@ +# Week 9 - Cubical Agda Exercises + +## Standard disclaimer: + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +In case you haven't done the other Agda exercises: +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**When solving the problems, +please make a copy of this file to work in, so that it doesn't get overwritten +(in case we update the exercises through `git`)!** + + +```agda +{-# OPTIONS --cubical --allow-unsolved-metas #-} + +module Exercises9 where + +open import cubical-prelude +open import Lecture7-notes +open import Lecture8-notes +open import Lecture9-notes +open import Solutions7 hiding (rUnit) +open import Solutions8 +open import Lecture9-live using (SemiGroupℕ) +``` + +## Part I: More hcomps + +### Exercise 1 (★★) +### (a) +Show the left cancellation law for path composition using an hcomp. +Hint: one hcomp should suffice. Use `comp-filler` and connections + +```agda +lUnit : {A : Type ℓ} {x y : A} (p : x ≡ y) → refl ∙ p ≡ p +lUnit = {!!} + +``` +### (b) +Try to mimic the construction of lUnit for rUnit (i.e. redefine it) +in such a way that `rUnit refl ≡ lUnit refl` holds by `refl`. +Hint: use (almost) the exact same hcomp. + +```agda +rUnit : {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ refl ≡ p +rUnit = {!!} + +-- uncomment to see if it type-checks + +{- +rUnit≡lUnit : ∀ {ℓ} {A : Type ℓ} {x : A} → rUnit (refl {x = x}) ≡ lUnit refl +rUnit≡lUnit = refl +-} + +``` + + +### Exercise 2 (★★) +Show the associativity law for path composition +Hint: one hcomp should suffice. This one can be done without connections + (but you might need comp-filler in more than one place) + +```agda +assoc : {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙ (q ∙ r) ≡ (p ∙ q) ∙ r +assoc = {!!} + +``` + +### Exercise 3 (Master class in connections) (🌶) +The goal of this exercise is to give a cubical proof of the Eckmann-Hilton argument, +which says that path composition for higher loops is commutative + +(a) While we cannot get `p ∙ q ≡ q ∙ p` as a one-liner, we can get a +one-liner showing that the identiy holds up to some annoying +coherences. Try to understand the following statement (and why it's +well-typed). After that, fill the holes + +Hint: each hole will need a `∨` or a `∧` + +```agda +pre-EH : {A : Type ℓ} {x : A} (p q : refl {x = x} ≡ refl) + → ap (λ x → x ∙ refl) p ∙ ap (λ x → refl ∙ x) q + ≡ ap (λ x → refl ∙ x) q ∙ ap (λ x → x ∙ refl) p +pre-EH {x = x} p q i = (λ j → p {!!} ∙ q {!!}) + ∙ (λ j → p {!!} ∙ q {!!}) + +``` +(b) If we manage to cancel out all of the annoying aps, we get Eckmann-Hilton: +For paths (p q : refl ≡ refl), we have p ∙ q ≡ q ∙ p. Try to prove this, using the above lemma. + +Hint: Use the pre-EH as the bottom of an hcomp (one should be enough). +For the sides, use lUnit and rUnit wherever they're needed. Note that this will only work out smoothly if +you've solved Exercise 1 (b). + +```agda +Eckmann-Hilton : {A : Type ℓ} {x : A} (p q : refl {x = x} ≡ refl) → p ∙ q ≡ q ∙ p +Eckmann-Hilton = {!!} + +``` +# Part 2: Binary numbers as a HIT +Here is another HIT describing binary numbers. The idea is that a binary number is a list of booleans, modulo trailing zeros. + +For instance, `true ∷ true ∷ true ∷ []` is the binary number 110 ... +... and so is `true ∷ true ∷ false ∷ false ∷ false ∷ []` + +(!) Note that we're interpreting 110 as 1·2⁰ + 1·2¹ + 0·2² here. + +```agda +0B = false +1B = true + +data ListBin : Type where + [] : ListBin + _∷_ : (x : Bool) (xs : ListBin) → ListBin + drop0 : 0B ∷ [] ≡ [] + +-- 1 as a binary number +1LB : ListBin +1LB = 1B ∷ [] +``` +### Exercise 4 (★) +Define the successor function on ListBin +```agda + +sucListBin : ListBin → ListBin +sucListBin = {!!} + +``` +### Exercise 5 (★★) +Define an addition `+LB` on ListBin and prove that `x +LB [] ≡ x` +Do this by mutual induction! Make sure the three cases for the right unit law hold by refl. +```agda + +_+LB_ : ListBin → ListBin → ListBin +rUnit+LB : (x : ListBin) → x +LB [] ≡ x +x +LB y = {!!} +rUnit+LB = {!!} + +``` +(c) Prove that sucListBin is left distributive over `+LB` +Hint: If you pattern match deep enough, there should be a lot of refls... +```agda + +sucListBinDistrL : (x y : ListBin) → sucListBin (x +LB y) ≡ (sucListBin x +LB y) +sucListBinDistrL = {!!} +``` + +### Exercise 6 (★) +Define a map `LB→ℕ : ListBin → ℕ` and show that it preserves addition + +```agda +ℕ→ListBin : ℕ → ListBin +ℕ→ListBin = {!!} + +ℕ→ListBin-pres+ : (x y : ℕ) → ℕ→ListBin (x + y) ≡ (ℕ→ListBin x +LB ℕ→ListBin y) +ℕ→ListBin-pres+ = {!!} + +``` + +### Exercise 7 (★★★) +Show that `ℕ ≃ ListBin`. + +```agda + +ListBin→ℕ : ListBin → ℕ +ListBin→ℕ = {!!} + +ListBin→ℕ→ListBin : (x : ListBin) → ℕ→ListBin (ListBin→ℕ x) ≡ x +ListBin→ℕ→ListBin = {!!} + +ℕ→ListBin→ℕ : (x : ℕ) → ListBin→ℕ (ℕ→ListBin x) ≡ x +ℕ→ListBin→ℕ = {!!} + +ℕ≃ListBin : ℕ ≃ ListBin +ℕ≃ListBin = {!!} + +``` +# Part 3: The SIP +### Exericise 8 (★★) +Show that, using an SIP inspired argument, if `(A , _+A_)` is a semigroup and `(B , _+B_)` is some other type with a composition satisfying: + +(i) `e : A ≃ B` + +(ii) `((x y : A) → e (x +A y) ≡ e x +B e y` + +then `(B , _+B_)` defines a semigroup. + +Conclude that `(ListBin , _+LB_)` is a semigroup + +For inspiration, see Lecture9-notes +```agda + +SemiGroupListBin : SemiGroup ListBin +SemiGroupListBin = {!!} diff --git a/src/HoTTEST/Agda/Cubical/Lecture7-live.agda b/src/HoTTEST/Agda/Cubical/Lecture7-live.agda new file mode 100644 index 0000000..2b9eca5 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture7-live.agda @@ -0,0 +1,196 @@ +-- Lecture 7: Cubical Agda + +{-# OPTIONS --cubical #-} + +module Lecture7-live where + +-- Today: +-- - Path types and the cubical interval +-- - Cubical higher inductive types + +open import cubical-prelude + +-- Libraries: +-- - agda/cubical library +-- - 1lab + +-- The interval in Cubical Agda is written I +-- Endpoints are called i0 and i1 + +apply0 : (A : Type) → (p : I → A) → A +apply0 A p = p i0 + +-- Path types: +-- p : x ≡ y consists of p : I → A s.t. p i0 ≐ x and p i1 ≐ y + +refl : {A : Type} {x : A} → x ≡ x +refl {x = x} = λ i → x + +-- Cannot pattern-match on r: +-- oops : {A : Type} → I → A +-- oops r = {!r!} + +variable + A B : Type ℓ -- ℓ is \ell + +-- ap is called cong in the agda/cubical library +ap : (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f p i = f (p i) + +-- Satisfies definitional/judgmental/strict computation rules (see exercises) + +-- Funext can be seen as turning +-- A → (I → B) +-- into +-- I → (A → B) + +funExt : {f g : A → B} → (p : (x : A) → f x ≡ g x) → f ≡ g +funExt p = λ i → λ x → p x i +-- To convince oneself that this makes sense: plug in i0 and i1 for i + +-- In Cubical Agda we also have three operations on I + +-- Minimum: _∧_ : I → I → I +-- Maximum: _∨_ : I → I → I +-- Symmetry: ~_ : I → I + +-- ! or ⁻¹ +sym : {x y : A} → x ≡ y → y ≡ x +sym p i = p (~ i) + +-- For examples of _∧_ and _∨_ see the notes for now + +-- Path overs in Book HoTT are called PathP in Cubical Agda +-- In fact, x ≡ y is short for PathP (λ i → A) x y + +reflP : {x : A} → PathP (λ i → A) x x +reflP {x = x} = λ i → x + + + +-- Cubical higher inductive types + +-- HITs: S1, Circle2, Torus, ... + +-- Dan introduced these using postulates + +-- In Cubical Agda you can just write them as data types, +-- in particular you don't have postulate the recursor/eliminator +-- and can instead use pattern matching + +data S¹ : Type where + base : S¹ + loop : base ≡ base -- I → S¹ with endpoints base and base + +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i + +helix : S¹ → Type +helix base = ℤ +helix (loop i) = sucPath i + +winding : base ≡ base → ℤ +winding p = transp (λ i → helix (p i)) i0 (pos 0) + +_ : winding (loop ∙ loop) ≡ pos 2 +_ = refl + +-- Following Dan's lecture one can prove that winding is an equivalence +-- using the encode-decode, see Cubical.HITs.S1.Base + + +-- Torus + + -- line1 + -- p ----------> p + -- ^ ^ + -- ¦ ¦ + -- line2 ¦ square ¦ line2 + -- ¦ ¦ + -- p ----------> p + -- line1 + +data Torus : Type where + point : Torus + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 i ≡ line1 i) line2 line2 + +-- Exercise: define the Klein bottle + +t2c : Torus → S¹ × S¹ +t2c point = base , base +t2c (line1 i) = (loop i) , base +t2c (line2 i) = base , (loop i) +t2c (square i j) = (loop i) , (loop j) + +c2t : S¹ × S¹ → Torus +c2t (base , base) = point +c2t (base , loop i) = line2 i +c2t (loop i , base) = line1 i +c2t (loop i , loop j) = square i j + +c2t-t2c : (t : Torus) → c2t (t2c t) ≡ t +c2t-t2c point = refl +c2t-t2c (line1 i) = refl +c2t-t2c (line2 i) = refl +c2t-t2c (square i j) = refl + +t2c-c2t : (t : S¹ × S¹) → t2c (c2t t) ≡ t +t2c-c2t (base , base) = refl +t2c-c2t (base , loop i) = refl +t2c-c2t (loop i , base) = refl +t2c-c2t (loop i , loop j) = refl + +Torus≡S¹×S¹ : Torus ≡ S¹ × S¹ +Torus≡S¹×S¹ = isoToPath (iso t2c c2t t2c-c2t c2t-t2c) + + +-- Alternative definition: +data Torus' : Type where + point : Torus' + line1 : point ≡ point + line2 : point ≡ point + square : line1 ∙ line2 ≡ line2 ∙ line1 + +-- (Probably very hard) Exercise: prove that Torus ≃ Torus' + + +-- More HITs: + +-- Suspension + +data Susp (A : Type) : Type where + north : Susp A + south : Susp A + merid : (a : A) → north ≡ south + +flip : Susp S¹ → Susp S¹ +flip north = south +flip south = north +flip (merid a i) = merid a (~ i) + +-- Pushout + +-- f +-- A ---> B +-- | | +-- g| | +-- V V +-- C ---> Pushout + + +data Pushout {A B C : Type} (f : A → B) (g : A → C) : Type where + inl : B → Pushout f g + inr : C → Pushout f g + push : (a : A) → inl (f a) ≡ inr (g a) + + +-- Today: +-- - Interval and path types +-- - Cubical HITs + +-- Next time (Friday) +-- - Cubical transport and path induction +-- - Set truncated HITs and how to work with them diff --git a/src/HoTTEST/Agda/Cubical/Lecture7-notes.lagda.md b/src/HoTTEST/Agda/Cubical/Lecture7-notes.lagda.md new file mode 100644 index 0000000..64b299b --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture7-notes.lagda.md @@ -0,0 +1,529 @@ +# Lecture 7: Cubical Agda + +Contents: + +- The interval and Path/PathP types +- Cubical higher inductive types + +## Some introductory pointers for further reading + +Documentation of the Cubical Agda mode can be found +[here](https://agda.readthedocs.io/en/v2.6.2.2/language/cubical.html). + +These lectures about Cubical Agda will be inspired by my material from +the [2020 EPIT school on HoTT](https://github.com/HoTT/EPIT-2020/tree/main/04-cubical-type-theory). + +For students interested in a more in depth introduction to cubical +type theory see [my lecture notes for the 2019 HoTT school](https://www.cambridge.org/core/journals/mathematical-structures-in-computer-science/article/cubical-methods-in-homotopy-type-theory-and-univalent-foundations/ECB3FE6B4A0B19AED2D3A2D785C38AF9). +Those notes contain a lot more background and motivation to cubical +methods in HoTT and an extensive list of references for those that +want to read more background material. + +If one wants a library to work with when doing Cubical Agda there is +the [agda/cubical](https://github.com/agda/cubical/) library that I +started developing with Andrea Vezzosi (the implementor of Cubical +Agda) in 2018 and which has now over 70 contributors. It contains a +variety of things, including data structures, algebra, synthetic +homotopy and cohomology theory, etc.. + +For a slower-paced introduction to mathematics in Cubical Agda, there +is the more recent [1lab](https://1lab.dev/). Although meant to be +read on the web, it also doubles as a community-maintained library of +formalized mathematics, most notably category theory. + +# Cubical Agda + +To make Agda cubical simply add the following option: + +```agda +{-# OPTIONS --cubical #-} + +module Lecture7-notes where +``` + +We also have a small cubical prelude which sets things up to work +nicely and which provides whatever we might need for the lectures. + +```agda +open import cubical-prelude +``` + +The key idea in cubical type theories like Cubical Agda is to not have +equality be inductively defined as in Book HoTT, but rather we assume +that there is a primitive interval and define equality literally as +paths, i.e. as functions out of the interval. By iterating these paths +we get squares, cubes, hypercubes, ..., making the type theory +inherently cubical. + +# The interval and path types + +The interval is a primitive concept in Cubical Agda. It's written `I`. +It has two endpoints: + +```text + i0 : I + i1 : I +``` + +These stand for "interval 0" and "interval 1". + +We can apply a function out of the interval to an endpoint just +like we would with any Agda function: + +```agda +apply0 : (A : Type ℓ) (p : I → A) → A +apply0 A p = p i0 +``` + +The equality type `_≡_` is not inductively defined in Cubical Agda, +instead it's a builtin primitive notion defined using the interval. An +element of `x ≡ y` consists of a function `p : I → A` such that `p i0` +is definitionally `x` and `p i1` is definitionally `y`. The check that +the endpoints are correct when we provide a `p : I → A` is +automatically performed by Agda during typechecking, so introducing an +element of `x ≡ y` is done just like how we introduce elements of +`I → A` but Agda will check the side conditions. + +We can hence write paths using λ-abstraction: + +```agda +mypath : {A : Type ℓ} (x : A) → x ≡ x +mypath x = λ i → x +``` + +As explained above Agda checks that whatever we written as definition +matches the path that we have provided (so the endpoints have to be +correct). In this case everything is fine and mypath can be thought of +as a proof reflexivity. Let's give it a nicer name and more implicit +arguments: + +```agda +refl : {A : Type ℓ} {x : A} → x ≡ x +refl {x = x} = λ i → x +``` + +The notation `{x = x}` lets us access the implicit argument `x` (the +`x` in the LHS of `x = x`) and rename it to `x` (the `x` in the RHS +`x = x`) in the body of `refl`. We could just as well have written: + +```text +refl : {A : Type ℓ} {x : A} → x ≡ x +refl {x = y} = λ i → y +``` + +Note that we cannot pattern-match on interval variables as `I` is not +inductively defined. Try uncommenting and typing `C-c C-c` in the hole: + +```text +oops : {A : Type} → I → A +oops r = {!r!} +``` + +It quickly gets tiring to write `{A : Type ℓ}` everywhere, so let's +assume that we have some types (in fact, we've already assumed that +`ℓ` is a `Level` in the cubical-prelude): + +```agda +private + variable + A B : Type ℓ +``` + +This will make `A` and `B` elements of different universes (all +arguments is maximally generalized) and all definitions that use them +will have them as implicit arguments. + +We can now implement some basic operations on `_≡_`. Let's start with +`ap`: + +```agda +ap : (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f p i = f (p i) +``` + +Note that the definition differs from the Book HoTT definition in that +it is not defined by path induction or pattern-matching on `p`, but +rather it's just a direct definition as a composition of functions. +Agda treats `p : x ≡ y` like any function, so we can apply it to `i` +to get an element of `A` which at `i0` is `x` and at `i1` is `y`. By +applying `f` to this element we hence get an element of `B` which at +`i0` is `f x` and at `i1` is `f y`. + +As this is just function composition it satisfies lots of nice +definitional equalities, see the exercises. Some of these are not +satisfied by the Book HoTT definition of `ap`. + +In Book HoTT function extensionality is proved as a consequence of +univalence using a rather ingenious proof due to Voevodsky, but in +cubical systems it has a much more direct proof. As paths are just +functions we can get it by swapping the arguments to `p`: + +```agda +funExt : {f g : A → B} (p : (x : A) → f x ≡ g x) → f ≡ g +funExt p i x = p x i +``` + +To see that this has the correct type, note that when `i` is `i0` we +have `p x i0 = f x` and when `i` is `i1` we have `p x i1 = g x`, so by +η for function types we have a path `f ≡ g` as desired. + +The interval has additional operations: + +```text +Minimum: _∧_ : I → I → I (corresponds to min(i,j)) +Maximum: _∨_ : I → I → I (corresponds to max(i,j)) +Symmetry: ~_ : I → I (corresponds to 1 - i) +``` + +These satisfy the equations of a De Morgan algebra (i.e. a +distributive lattice (_∧_ , _∨_ , i0 , i1) with an "De Morgan" +involution ~). This just means that we have the following kinds of +equations definitionally: + +```text +i0 ∨ i ≐ i +i ∨ i1 ≐ i1 +i ∨ j ≐ j ∨ i +i0 ∧ i ≐ i0 +i1 ∧ i ≐ i +i ∧ j ≐ j ∧ i +~ (~ i) ≐ i +i0 ≐ ~ i1 +~ (i ∨ j) ≐ ~ i ∧ ~ j +~ (i ∧ j) ≐ ~ i ∨ ~ j +``` + +However, we do not have `i ∨ ~ i = i1` and `i ∧ ~ i = i0`. The reason +is that I represents an abstract interval, so we if we think of it as +the real interval [0,1] ⊂ ℝ we clearly don't always have +"max(i,1-i) = 1" or "min(i,1-i) = 0)" for all i ∈ [0,1]. + +These operations on `I` are very useful as they let us define even +more things directly. For example symmetry of paths is easily defined +using `~_`. + +```agda +sym : {x y : A} → x ≡ y → y ≡ x +sym p i = p (~ i) +``` + +Remark: this has been called `⁻¹` and `!` in the previous +lectures. Here we stick to `sym` for the cubical version following the +agda/cubical notation. + +The operations `_∧_` and `_∨_` are called *connections* and let us +build higher dimensional cubes from lower dimensional ones, for +example if we have a path `p : x ≡ y` then + +```text + sq i j = p (i ∧ j) +``` + +is a square (as we've parametrized by `i` and `j`) with the following +boundary: + +```text + sq i0 j = p (i0 ∧ j) = p i0 = x + sq i1 j = p (i1 ∧ j) = p j + sq i i0 = p (i ∧ i0) = p i0 = x + sq i i1 = p (i ∧ i1) = p i +``` + +If we draw this we get: + +```text + p + x --------> y + ^ ^ + ¦ ¦ + refl ¦ sq ¦ p + ¦ ¦ + ¦ ¦ + x --------> x + refl +``` + +Being able to make this square directly is very useful. It for +example let's prove that singletons are contractible (a.k.a. based +path induction). + +We define the type of singletons as follows + +```agda +singl : {A : Type ℓ} (a : A) → Type ℓ +singl {A = A} a = Σ x ꞉ A , a ≡ x +``` + +To show that this type is contractible we need to provide a center +of contraction and the fact that any element of it is path-equal to +the center + +```agda +isContrSingl : (x : A) → isContr (singl x) +isContrSingl x = ctr , prf + where + -- The center is just a pair with x and refl + ctr : singl x + ctr = x , refl + + -- We then need to prove that ctr is equal to any element s : singl x. + -- This is an equality in a Σ-type, so the first component is a path + -- and the second is a path over the path we pick as first argument, + -- i.e. the second component is a square. In fact, we need a square + -- relating refl and pax over a path between refl and pax, so we can + -- use an _∧_ connection. + prf : (s : singl x) → ctr ≡ s + prf (y , pax) i = (pax i) , λ j → pax (i ∧ j) +``` + +As we saw in the second component of prf we often need squares when +proving things. In fact, `pax (i ∧ j)` is a path relating `refl` to +`pax` *over* another path `λ j → x ≡ pax j`. This notion of path over +a path is very useful when working in Book HoTT as we've seen in the +previous lectures, this is also the case when working cubically. In +Cubical Agda path-overs are a primitive notion called `PathP` ("Path +over a Path"). In general `PathP A x y` has + +```text + A : I → Type ℓ + x : A i0 + y : A i1 +``` + +So PathP lets us natively define heteorgeneous paths, i.e. paths +where the endpoints are in different types. This allows us to +specify the type of the second component of `prf`: + +```agda +prf' : (x : A) (s : singl x) → (x , refl) ≡ s +prf' x (y , pax) i = (pax i) , λ j → goal i j + where + goal : PathP (λ j → x ≡ pax j) refl pax + goal i j = pax (i ∧ j) +``` + +Just like `_×_` is a special case of Σ-types we have that `_≡_` is a +special case of PathP. In fact, `x ≡ y` is just short for +`PathP (λ _ → A) x y`: + +```agda +reflP : {x : A} → PathP (λ _ → A) x x +reflP = refl +``` + +Working directly with paths and equalities makes many proofs from Book +HoTT very short: + +```agda +isContrΠ : {B : A → Type ℓ} (h : (x : A) → isContr (B x)) + → isContr ((x : A) → B x) +isContrΠ h = (λ x → pr₁ (h x)) , (λ f i x → pr₂ (h x) (f x) i) +``` + +# Cubical higher inductive types + +We have seen various HITs earlier in the course. These were added +axiomatically to Agda by postulating their existence together with +suitable elimination/induction principles. In Cubical Agda they are +instead added just like any inductive data type, but with path +constructors. This is made possible by the fact that paths in Cubical +Agda are just fancy functions. + +## The circle + +We can define the circle as the following simple data declaration: + +```agda +data S¹ : Type₀ where + base : S¹ + loop : base ≡ base +``` + +We can write functions on `S¹` using pattern-matching: + +```agda +double : S¹ → S¹ +double base = base +double (loop i) = (loop ∙ loop) i +``` + +Note that loop takes an `i : I` argument. This is not very surprising +as it's a path of type `base ≡ base`, but it's an important difference +to Book HoTT where we instead would have to state the equation using +`ap`. Having the native notion of equality be heterogeneous makes it +possible to quite directly define a general schema for a large class +of HITs and use it in the implementation of a system like Cubical Agda. + +Let's use univalence to compute some winding numbers on the +circle. We first define a family of types over the circle with +fibers being the integers. + +```agda +helix : S¹ → Type₀ +helix base = ℤ +helix (loop i) = sucPath i +``` + +Here univalence is baked into `sucPath : ℤ ≡ ℤ`. The loopspace of the +circle is then defined as + +```agda +ΩS¹ : Type₀ +ΩS¹ = base ≡ base +``` + +and we can then define a function computing how many times we've +looped around the circle by: + +```agda +winding : ΩS¹ → ℤ +winding p = transp (λ i → helix (p i)) i0 (pos 0) +``` + +Here `transp` is a cubical transport function. We'll talk about it in +more detail in the next lecture, but for now we can observe that it +reduces as expected: + +```agda +_ : winding (λ i → double ((loop ∙ loop) i)) ≡ pos 4 +_ = refl +``` + +This would not reduce definitionally in Book HoTT as univalence is an +axiom. Having things compute definitionally makes it possible to +substantially simplify many proofs from Book HoTT in Cubical Agda. + +We can in fact prove that `winding` is an equivalence, this is very +similar to the Book HoTT proof and uses the encode-decode method. For +details about how this proof looks in Cubical Agda see the +Cubical.HITs.S1.Base file in the agda/cubical library. + +## The torus + +We can define the torus as: + +```agda +data Torus : Type₀ where + point : Torus + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 i ≡ line1 i) line2 line2 +``` + +The square corresponds to the usual folding diagram from topology +(where `p` is short for `point`): + +```text + line1 + p ----------> p + ^ ^ + ¦ ¦ + line2 ¦ ¦ line2 + ¦ ¦ + p ----------> p + line1 +``` + +Proving that it is equivalent to two circles is pretty much trivial as +we have definitional computation rules for all constructors, including +higher ones: + +```agda +t2c : Torus → S¹ × S¹ +t2c point = (base , base) +t2c (line1 i) = (loop i , base) +t2c (line2 j) = (base , loop j) +t2c (square i j) = (loop i , loop j) + +c2t : S¹ × S¹ → Torus +c2t (base , base) = point +c2t (loop i , base) = line1 i +c2t (base , loop j) = line2 j +c2t (loop i , loop j) = square i j + +c2t-t2c : (t : Torus) → c2t (t2c t) ≡ t +c2t-t2c point = refl +c2t-t2c (line1 _) = refl +c2t-t2c (line2 _) = refl +c2t-t2c (square _ _) = refl + +t2c-c2t : (p : S¹ × S¹) → t2c (c2t p) ≡ p +t2c-c2t (base , base) = refl +t2c-c2t (base , loop _) = refl +t2c-c2t (loop _ , base) = refl +t2c-c2t (loop _ , loop _) = refl +``` + +Using univalence we get the following equality: + +```agda +Torus≡S¹×S¹ : Torus ≡ S¹ × S¹ +Torus≡S¹×S¹ = isoToPath (iso t2c c2t t2c-c2t c2t-t2c) +``` + +We can also directly compute winding numbers on the torus + +```agda +windingTorus : point ≡ point → ℤ × ℤ +windingTorus l = ( winding (λ i → pr₁ (t2c (l i))) + , winding (λ i → pr₂ (t2c (l i)))) + +_ : windingTorus (line1 ∙ sym line2) ≡ (pos 1 , negsuc 0) +_ = refl +``` + +# Bonus content if there is time + +## Interval + +```agda +data Interval : Type₀ where + zero : Interval + one : Interval + seg : zero ≡ one +``` + +## Suspension + +```agda +data Susp (A : Type ℓ) : Type ℓ where + north : Susp A + south : Susp A + merid : (a : A) → north ≡ south +``` + +We can define Dan's Circle2 as the suspension of Bool, or we can do it directly as: + +```agda +data Circle2 : Type₀ where + north : Circle2 + south : Circle2 + west : north ≡ south + east : north ≡ south +``` + +## Pushouts + +```agda +data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) : Type (ℓ ⊔ ℓ' ⊔ ℓ'') where + inl : B → Pushout f g + inr : C → Pushout f g + push : (a : A) → inl (f a) ≡ inr (g a) +``` + +## Relation quotient + +```agda +data _/ₜ_ {ℓ ℓ'} (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ ⊔ ℓ') where + [_] : (a : A) → A /ₜ R + eq/ : (a b : A) → (r : R a b) → [ a ] ≡ [ b ] +``` + +## Propositional truncation + +```agda +data ∥_∥₋₁ {ℓ} (A : Type ℓ) : Type ℓ where + ∣_∣₋₁ : A → ∥ A ∥₋₁ + squash₁ : (x y : ∥ A ∥₋₁) → x ≡ y +``` diff --git a/src/HoTTEST/Agda/Cubical/Lecture8-live.agda b/src/HoTTEST/Agda/Cubical/Lecture8-live.agda new file mode 100644 index 0000000..afc2822 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture8-live.agda @@ -0,0 +1,232 @@ +-- Lecture 8: Cubical Agda + +{-# OPTIONS --cubical #-} + +module Lecture8-live where + +{- + +Last time: +- Path types and the cubical interval (I with endpoints i0, i1. _≡_ is now a PathP) +- Cubical higher inductive types + +Today: +- Set quotients +- Cubical transport (`transp`) and path induction +- Homogeneous composition (`hcomp`) + +-} + +open import cubical-prelude +open import Lecture7-notes + +private + variable + A B : Type ℓ + + +-- Set quotients + +-- A Type +-- R : A → A → Type (hProp) +-- Want a quotient type A / R + +data _/_ (A : Type) (R : A → A → Type) : Type where + [_] : A → A / R + eq/ : (a b : A) → R a b → [ a ] ≡ [ b ] + trunc : isSet (A / R) -- (a b : A / R) (p q : a ≡ b) → p ≡ q + + + +-- (a , b) ~ (c , d) iff a + d ≡ c + b +ℤ' : Type +ℤ' = (ℕ × ℕ) / rel + where + rel : ℕ × ℕ → ℕ × ℕ → Type -- Press C-c C-s and get the type automagically + rel (a , b) (c , d) = a + d ≡ b + c + +-- Exercise: prove ℤ ≃ ℤ' +-- Exercise: define 0, 1, +, -, * and prove standard properties + +-- Can also do ℤ / nℤ + +-- Similarly we can define ℚ as a quotient pairs of a number with a nonzero number + +-- And so on... + +-- All of these can be defined without quotients in type theory, but +-- the quotient definition is sometimes more efficient. + +-- ℚ can be defined not as a quotient by taking pairs of coprime numbers (i.e. as normalized fractions) + +-- Problem: need to normalize constantly + +-- With the quotient definition we don't have to normalize unless we want to + +-- Practical benefit with quotients: can get more efficient because you don't have to normalize + + + +-- Q: What happens if A is a set and R is prop valued, do you need trunc? + +-- A = Unit +-- R is the total + +-- Unit / total ≃ S¹ (if we don't have trunc constructor) + + +-- Q: What happens if A is not a set, but R is not prop valued? + +-- It will be a set by def. + +-- Nothing too bad... But A / R won't have all the properties you want. For example: + +-- effectivity : (a b : A) → [ a ] ≡ [ b ] → R a b + +-- (Also need R equivalence relation, and proof of effectivity uses univalence for hProp (prop. ext.)) + + + + + + + +-- Another example: finite multisets + +infixr 5 _∷_ + +data FMSet (A : Type) : Type where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : (x y : A) (xs : FMSet A) → x ∷ y ∷ xs ≡ y ∷ x ∷ xs + trunc : isSet (FMSet A) + + +-- Didn't use _/_ because it gets a bit longer and annoying +-- Also nice to have a simple path constructor comm to pattern-match on + +infixr 30 _++_ + +_++_ : (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ (xs ++ ys) +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = + trunc (xs ++ ys) (zs ++ ys) (λ k → p k ++ ys) (λ k → q k ++ ys) i j + +-- unitr-++ : (xs : FMSet A) → xs ++ [] ≡ xs +-- unitr-++ [] = refl +-- unitr-++ (x ∷ xs) = ap (x ∷_) (unitr-++ xs) +-- unitr-++ (comm x y xs i) j = comm x y (unitr-++ xs j) i +-- unitr-++ (trunc xs xs₁ x y i i₁) = {!!} -- Ugh... + +-- unitr-++ : (xs : FMSet A) → xs ++ [] ≡ xs +-- unitr-++ = ElimProp.f (trunc _ _) refl (λ x p → ap (x ∷_) p) + +-- Can do set quotients this way, but need good lemmas and combinators! + + +-- In the notes there's a more efficient representation of FMSet, but don't have time to cover it now... + + +-- Set quotients: +-- Either use _/_ -- get good lemmas, but maybe not the constructors you want +-- Or write your own -- need to prove good lemmas (but they are boilerplate), but get good constructors to pattern-match on + + + + +-- Cubical transport and path induction + +-- Last time we saw that we can prove many things for path types using +-- the operations on the interval I (like ~_ , _∧_ , _∨_) + +-- But we cannot yet compose paths or transport along paths or prove things by path induction... + + +-- Basic/primitive operation in Cubical Agda for transport is called `transp` + +-- Special case is cubical transport: + +transport : A ≡ B → A → B +transport p a = transp (λ i → p i) i0 a + +-- In general transp has this (slightly simplified type): + +-- transp : (L : I → Type) → (r : I) → L i0 → L i1 + +-- The r lets us specify where the transp is the identity function +-- +-- transp L i1 a ≐ a +-- +-- For this to make sense L has to be "constant" whenever r = i1 + + +-- What we cubists call subst, Book HoTT calls transport + +subst : (B : A → Type) {x y : A} → x ≡ y → B x → B y +subst B p bx = transport (ap B p) bx + + +-- Because we are *not* defining things by path-induction transport along refl +-- isn't the identity function + +transportRefl : (x : A) → transport refl x ≡ x +transportRefl {A = A} x i = transp (λ _ → A) i x + + +-- Path-induction is not a primitive concept in Cubical Agda, rather +-- we have to prove it + +-- General fact: subst + contractibility of singletons ⇒ path induction + +-- Book HoTT version: transport + singleton contractibility ⇒ based path induction ⇒ path induction + +-- Path-induction has traditionally been called J in the type theory community + +J : {x : A} + (P : (y : A) → x ≡ y → Type) + (d : P x refl) + {y : A} + (p : x ≡ y) + → + P y p +J {A = A} {x = x} P d p = subst (λ X → P (pr₁ X) (pr₂ X)) (lem .pr₂ (_ , p)) d + where + lem : isContr (Σ y ꞉ A , x ≡ y) + lem = isContrSingl x + + +-- A little bit annoying to use because it doesn't satisfy the computation rule at refl + +-- Also can lead to very big terms when Agda unfolds things (bad for readability and efficiency) + +-- Often better to avoid J and stick to cubical primitives directly + + + + +-- Homogenous composition (`hcomp`) + +-- Special case is path composition + +compPath : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +compPath {x = x} p q i = hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → q j }) + (p i) + +{- + + hcomp + x - - - - -> z + ^ ^ + ¦ ¦ + x ¦ ¦ q j + ¦ ¦ + x ----------> y + p i + + +-} + + diff --git a/src/HoTTEST/Agda/Cubical/Lecture8-notes.lagda.md b/src/HoTTEST/Agda/Cubical/Lecture8-notes.lagda.md new file mode 100644 index 0000000..91f99fe --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture8-notes.lagda.md @@ -0,0 +1,381 @@ +# Lecture 8 + +Contents: + +- Set quotients +- Cubical transport and path induction +- Homogeneous composition (`hcomp`) + +```agda +{-# OPTIONS --cubical #-} + +module Lecture8-notes where + +open import cubical-prelude +open import Lecture7-notes + +private + variable + A B : Type ℓ +``` + +## Set quotients + +Last time we saw some HITs that are useful for topology and homotopy +theory. Now we'll look at some HITs that are not very interesting from +a homotopical perspective, but still very useful for other purposes. +In particular we will look at how we can construct quotient types and +in order for the result to be a set we will also set truncate the +type. This kind of types has many applications in both computer +science and mathematics. + +In general these are written as: + +```text +data _/_ (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ ⊔ ℓ') where + [_] : A → A / R + eq/ : (a b : A) → R a b → [ a ] ≡ [ b ] + trunc : (a b : A / R) (p q : a ≡ b) → p ≡ q +``` + +The type of the `trunc` constructor can simply be written as: + +```agda +data _/_ (A : Type ℓ) (R : A → A → Type ℓ') : Type (ℓ ⊔ ℓ') where + [_] : A → A / R + eq/ : (a b : A) → R a b → [ a ] ≡ [ b ] + trunc : isSet (A / R) +``` + +The idea is that `[_]` injects elements of `A` into the quotient while +`eq/` ensures that elements that are related by `R` are sent to equal +elements. Finally we have to truncate the type to be a set in order +for this to be a set quotient. If we leave out the `trunc` constructor +quite weird things can happen, for example if one quotients the unit +type by the total relation then one obtains a type which is equivalent +to the circle! The `trunc` constructor is hence crucial to ensure that +the result is a set **even** if `A` is one already (so quotienting can +raise the truncation level). + +Various sources require that `R` is proposition-valued, i.e. `R : A → +A → hProp`. However this is not necessary to define `_/_` as seen +above. There are however various important properties that are only +satisfied when one quotients by a prop-valued relation. The key +example being effectivity, i.e. that + +```text +(a b : A) → [ a ] ≡ [ b ] → R a b +``` + +To prove this one first of all needs that `R` is prop-valued and the +proof then uses univalence for propositions (i.e., logically +equivalent propositions are equal, a.k.a. "propositional +extensionality"). + +Having the ability to define set quotients using `_/_` lets us do many +examples from mathematics and computer science. For example we can +define the integers as: + +```agda +ℤ' : Type +ℤ' = (ℕ × ℕ) / rel + where + rel : ℕ × ℕ → ℕ × ℕ → Type + rel (x₀ , y₀) (x₁ , y₁) = x₀ + y₁ ≡ x₁ + y₀ +``` + +If you haven't seen this construction before see +https://en.wikipedia.org/wiki/Integer#Construction + +Exercises: define `0`, `1`, `+`, `-`, `*` + +Similarly we can define the rational numbers as a quotient of pairs of +a number and a nonzero number (and more generally we can define the +field of fractions of a commutative ring R using `_/_`). Both the +integers and rationals can of course be defined without using +quotients in type theory, but in the case of the rationals this has +some efficiency problems. Indeed, when defining the rationals not as a +quotient we need to do it as normalized fractions, that is, pairs of +coprime numbers. All operations, like addition and multiplication, +then have to ensure that the resulting fractions are normalized which +can become quite inefficient because of repeated GCD computations. A +more efficient way is to work with the equivalence classes of +not-necessarily-normalized fractions and then normalize whenever one +wants to. This can be done when defining the rationals using `_/_`. + +Let us now look at a more computer science inspired example: finite +multisets. These can be represented as lists modulo permutations. +Defining these using `_/_` is a bit annoying, luckily there is an +easier way: + +```agda +infixr 5 _∷_ + +data FMSet (A : Type ℓ) : Type ℓ where + [] : FMSet A + _∷_ : (x : A) → (xs : FMSet A) → FMSet A + comm : (x y : A) (xs : FMSet A) → x ∷ y ∷ xs ≡ y ∷ x ∷ xs + trunc : (xs ys : FMSet A) (p q : xs ≡ ys) → p ≡ q +``` + +Programming and proving with this is quite straightforward: + +```agda +infixr 30 _++_ + +_++_ : (xs ys : FMSet A) → FMSet A +[] ++ ys = ys +(x ∷ xs) ++ ys = x ∷ xs ++ ys +comm x y xs i ++ ys = comm x y (xs ++ ys) i +trunc xs zs p q i j ++ ys = + trunc (xs ++ ys) (zs ++ ys) (λ k → p k ++ ys) (λ k → q k ++ ys) i j + +unitl-++ : (xs : FMSet A) → [] ++ xs ≡ xs +unitl-++ xs = refl + +unitr-++ : (xs : FMSet A) → xs ++ [] ≡ xs +unitr-++ [] = refl +unitr-++ (x ∷ xs) = ap (x ∷_) (unitr-++ xs) +unitr-++ (comm x y xs i) j = comm x y (unitr-++ xs j) i +unitr-++ (trunc xs ys p q i j) k = + trunc (unitr-++ xs k) (unitr-++ ys k) + (λ l → unitr-++ (p l) k) (λ l → unitr-++ (q l) k) i j +``` + +Filling the goals for `comm` and `trunc` quickly gets tiresome and +when working with HITs like this it's very strongly recommended to +prove special lemmas for eliminating into propositions (which is the +case above as `FMSet` is a set, so its `≡`-type is a proposition). If +we do this the proof of `unitr-++` can be simplified to a one-liner: + +```text +unitr-++ : (xs : FMSet A) → xs ++ [] ≡ xs +unitr-++ = ElimProp.f (trunc _ _) refl (λ x p → cong (_∷_ x) p) +``` + +We recommend the interested reader to look at the code in +Cubical.HITs.FiniteMultiset.Base and +Cubical.HITs.FiniteMultiset.Properties in agda/cubical to see how +these lemmas are stated and proved. This is a very common pattern when +working with set truncated HITs: first define the HIT, then prove +special purpose recursors and eliminators for eliminating into types +of different truncation levels. All definitions are then written using +these recursors and eliminators and one get very short proofs. + +A more efficient version of finite multisets based on association +lists can be found in Cubical.HITs.AssocList.Base. It looks like this: + +```agda +data AssocList (A : Type ℓ) : Type ℓ where + ⟨⟩ : AssocList A + ⟨_,_⟩∷_ : (a : A) (n : ℕ) (xs : AssocList A) → AssocList A + per : (a b : A) (m n : ℕ) (xs : AssocList A) + → ⟨ a , m ⟩∷ ⟨ b , n ⟩∷ xs ≡ ⟨ b , n ⟩∷ ⟨ a , m ⟩∷ xs + agg : (a : A) (m n : ℕ) (xs : AssocList A) + → ⟨ a , m ⟩∷ ⟨ a , n ⟩∷ xs ≡ ⟨ a , m + n ⟩∷ xs + del : (a : A) (xs : AssocList A) → ⟨ a , 0 ⟩∷ xs ≡ xs + trunc : (xs ys : AssocList A) (p q : xs ≡ ys) → p ≡ q +``` + +Programming and proving is more complicated with `AssocList` compared +to `FMSet`. This kind of example occurs a lot in programming and +mathematics: one representation is easier to work with, but not +efficient, while another is efficient but difficult to work with. +Next time we will see how we can use univalence and the structure +identity principle (SIP) to get the best of both worlds (if someone +can't wait they can look at https://dl.acm.org/doi/10.1145/3434293 for +more details). + +It's sometimes easier to work directly with `_/_` instead of defining +special HITs as one can reuse lemmas for `_/_` instead of reproving +things. For example, general lemmas about eliminating into +propositions have already been proved for `_/_` so we do not have to +reprove them for our special purpose HIT. However on the other hand it +can sometimes be much easier to write the HIT directly (imagine +implementing `AssocList` using `_/_`...) and one also gets the benefit +that one can pattern-match directly with nice names for the +constructors (it might still be better to use the handcrafted +recursor/eliminator for the HIT to avoid having to give cases for path +construtors though). + +## Cubical transport and path induction + +While path types are great for reasoning about equality they don't let +us transport along paths between types or even compose paths. +Furthermore, as paths are not inductively defined we don't +automatically get an induction principle for them. In order to remedy +this Cubical Agda also has a built-in cubical transport operation and +homogeneous composition operation from which the induction principle +is derivable (and much more!). + +The basic operation is called `transp` and we will soon explain it, +but let's first focus on the special case of cubical transport which +we used to define `winding` last time: + +```agda +transport : A ≡ B → A → B +transport p a = transp (λ i → p i) i0 a +``` + +This is a more primitive operation than "transport" in Book HoTT as it +only lets us turn a path into a function. However, the transport of +HoTT can easily be proved from cubical transport and in order to avoid +a name clash we call it "subst": + +```agda +subst : (B : A → Type ℓ') {x y : A} (p : x ≡ y) → B x → B y +subst B p pa = transport (λ i → B (p i)) pa +``` + +The `transp` operation is a generalized transport in the sense that it +lets us specify where the transport is the identity function (this is +why there is an `i0` in the definition of `transport` above). The +general type of `transp` is: + +```text +transp : (A : I → Type ℓ) (r : I) (a : A i0) → A i1 +``` + +There is an additional side condition which has to be satisfied for +Cubical Agda to typecheck `transp A r a`. This is that `A` has to be +"constant" on `r`. This means that `A` should be a constant function +whenever `r = i1` is satisfied. This side condition is vacuously true +when `r = i0`, so there is nothing to check when writing transport as +above. However, when `r` is equal to `i1` the `transp` function will +compute as the identity function. + +```text +transp A i1 a ≐ a +``` + +(Here `≐` is definitional/judgmental equality) + +Having this extra generality is useful for quite technical reasons, +for instance we can easily relate a term `a` with its transport over a +path `p`: + +```agda +transportFill : (p : A ≡ B) (a : A) → PathP (λ i → p i) a (transport p a) +transportFill p a i = transp (λ j → p (i ∧ j)) (~ i) a +``` + +Another result that follows easily from `transp` is that transporting +in a constant family is the identity function (up to a path). Note +that this is *not* proved by `refl`. This is maybe not so surprising +as `transport` is not defined by pattern-matching on `p`. + +```agda +transportRefl : (x : A) → transport refl x ≡ x +transportRefl {A = A} x i = transp (λ _ → A) i x +``` + +Having `transp` lets us prove many more useful lemmas like this. For +details see Cubical.Foundations.Transport in the agda/cubical library. + +Using contractibility of singletons we can also define the `J` +eliminator for paths (a.k.a. path induction): + +```agda +J : {x : A} (P : (y : A) → x ≡ y → Type ℓ'') + (d : P x refl) {y : A} (p : x ≡ y) → P y p +J {x = x} P d p = subst (λ X → P (pr₁ X) (pr₂ X)) (isContrSingl x .pr₂ (_ , p)) d +``` + +Unfolded version: + +```text +transport (λ i → P (p i) (λ j → p (i ∧ j))) d +``` + +So `J` is provable, but it doesn't satisfy the computation rule of +`refl` definitionally as `_≡_` is not inductively defined. See +exercises for how to prove it. Not having this hold definitionally is +almost never a problem in practice as the cubical primitives satisfy +many new definitional equalities (c.f. `ap`). + +## Homogeneous composition (`hcomp`) + +As we now have `J` we can define path concatenation and many more +things, however this is not the way to do things in Cubical Agda. One +of the key features of cubical type theory is that the `transp` +primitive reduces differently for different types formers (see CCHM +[1] or the Cubical Agda paper [2] for details). For paths it reduces +to another primitive operation called `hcomp`. This primitive is much +better suited for concatenating paths than `J` as it is much more +general. In particular, it lets us compose multiple higher dimensional +cubes directly. We will explain it by example. + +In order to compose two paths we write: + +```agda +compPath : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +compPath {x = x} p q i = hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → q j }) + (p i) +``` + +This is best understood with the following drawing: + +```text + x z + ^ ^ + ¦ ¦ + x ¦ ¦ q j + ¦ ¦ + x ----------> y + p i +``` + +In the drawing the direction `i` goes left-to-right and `j` goes +bottom-to-top. As we are constructing a path from `x` to `z` along `i` +we have `i : I` in the context already and we put `p i` as bottom. The +direction `j` that we are doing the composition in is abstracted in +the first argument to `hcomp`. + +Having `hcomp` as a primitive operation lets us prove many things very +directly. For instance, we can prove that any proposition is also a +set using a higher dimensional `hcomp`. + +```agda +isProp→isSet : isProp A → isSet A +isProp→isSet h a b p q j i = + hcomp (λ k → λ { (i = i0) → h a a k + ; (i = i1) → h a b k + ; (j = i0) → h a (p i) k + ; (j = i1) → h a (q i) k }) a +``` + +Geometric picture: start with a square with `a` everywhere as base, +then change its sides so that they connect `p` with `q` over `refl a` +and `refl b`. + +This has some useful consequences: + +```agda +isPropIsProp : isProp (isProp A) +isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i + +isPropIsSet : isProp (isSet A) +isPropIsSet h1 h2 i x y = isPropIsProp (h1 x y) (h2 x y) i +``` + +In order to really understand what the second argument to `hcomp` is +and how to use `hfill` we will need to explain partial elements and +cubical subtypes. This is documented in the Cubical Agda documentation +(https://agda.readthedocs.io/en/v2.6.2.2/language/cubical.html#partial-elements) +and in lecture 9 these will be discussed in more detail. + + +However, beginners often doesn't have to write `hcomp` to prove things +as the library provides many basic lemmas. This is especially true +when reasoning about sets. But when reasoning about types that have +higher truncation level it's very convenient to be able to construct +squares and cubes directly and being able to use `hcomp` is quite +necessary (see exercise 12 on +[Exercises7](https://github.com/martinescardo/HoTTEST-Summer-School/blob/main/Agda/Cubical/Exercises7.lagda.md)). + +References +========== + +[1] https://arxiv.org/abs/1611.02108 +[2] https://staff.math.su.se/anders.mortberg/papers/cubicalagda2.pdf diff --git a/src/HoTTEST/Agda/Cubical/Lecture9-live.agda b/src/HoTTEST/Agda/Cubical/Lecture9-live.agda new file mode 100644 index 0000000..13fe671 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture9-live.agda @@ -0,0 +1,309 @@ +-- Lecture 9: Cubical Agda + +{-# OPTIONS --cubical #-} + +module Lecture9-live where + +open import cubical-prelude hiding (_∙_) +open import Lecture7-notes +open import Lecture8-notes hiding (compPath) + +private + variable + A B : Type ℓ + +-- Last time: +-- - Set quotients +-- - Cubical transport and path induction +-- - Homogeneous composition (`hcomp`) + + +-- Today: +-- - More about homogeneous composition (`hcomp`) +-- - Cubical univalence (Glue types) +-- - The structure identity principle (SIP) + + + + +-- More about hcomp + + +compPath : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +compPath {x = x} p q i = + hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → q j }) -- System of sides + (p i) -- The base of the hcomp + +{- + x z + ^ ^ + ¦ ¦ + x ¦ ¦ q j + ¦ ¦ + x ----------> y + p i +-} + +_∙∙_∙∙_ : {x y z w : A} → x ≡ y → y ≡ z → z ≡ w → x ≡ w +_∙∙_∙∙_ p q r i = + hcomp (λ j → λ { (i = i0) → p (~ j) + ; (i = i1) → r j }) + (q i) + +{- + x w + ^ ^ + ¦ ¦ + p (~ j) ¦ ¦ r j + ¦ ¦ + y ----------> z + q i +-} + +_∙_ : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ p q = refl ∙∙ p ∙∙ q + + +-- The hcomp operation is a cubical version of lifting conditions for Kan complexes + + +-- What is the type of the weird first argument to hcomp? + +-- Answer: I → .(IsOne φ) → A + +-- IsOne φ represents (φ = i1) + +-- Partial φ A the type of partial elements of A, i.e. the type of cubes in A that are defined when IsOne φ holds + +-- Question: is Partial φ A the same as (IsOne φ) → A ? +-- Answer: not exactly, Partial φ A is "more extensional" (elements equal up to permutation, duplication...) + +-- Partial (i ∨ ~ i) Bool ≐ Partial (~ i ∨ i) Bool + +partialBool : (i : I) → Partial (i ∨ ~ i) Bool +partialBool i = λ { (i = i0) → true + ; (i = i1) → false } + +partialBool' : (i j : I) → Partial (~ i ∨ i ∨ (i ∧ j)) Bool +partialBool' i j (i = i0) = true +partialBool' i j (i = i1) = true +partialBool' i j (i = i1) (j = i1) = true + + +-- partialBoolBad : (i : I) → Bool +-- partialBoolBad i0 = true +-- partialBoolBad i1 = false + + + +-- Cubical subtypes + +{- + +_[_↦_] : (A : Type ℓ) (φ : I) (u : Partial φ A) → SSet ℓ +A [ φ ↦ u ] = Sub A φ u + +inS : {A : Type ℓ} {φ : I} (u : A) → A [ φ ↦ (λ _ → u) ] + +outS : {A : Type ℓ} {φ : I} {u : Partial φ A} → A [ φ ↦ u ] → A + +They satisfy the following equalities: + +outS (inS a) ≐ a + +inS {φ = φ} (outS {φ = φ} a) ≐ a + +outS {φ = i1} {u} _ ≐ u 1=1 + +-} + + +hcomp' : {A : Type} {φ : I} (u : I → Partial φ A) (u0 : A [ φ ↦ u i0 ]) → A [ φ ↦ u i1 ] +hcomp' u u0 = inS (hcomp u (outS u0)) + + + + + +-- Cubical univalence (Glue types) + + +ua : {A B : Type} → A ≃ B → A ≡ B +ua {A = A} {B = B} e i = + Glue B λ { (i = i0) → A , e + ; (i = i1) → B , idEquiv B } + + +uaβ : (e : A ≃ B) → (a : A) → transport (ua e) a ≡ pr₁ e a +uaβ e a = transportRefl (pr₁ e a) + + +uaβℕ : (e : ℕ ≃ ℕ) → (a : ℕ) → transport (ua e) a ≡ pr₁ e a +uaβℕ e a = refl + + +-- Fact: ua + uaβ ⇒ univalence axiom + + +-- Examples + +not : Bool → Bool +not true = false +not false = true + +notPath : Bool ≡ Bool +notPath = ua (isoToEquiv (iso not not rem rem)) + where + rem : (b : Bool) → not (not b) ≡ b + rem true = refl + rem false = refl + + + +-- The structure identity principle (SIP) + + +substEquiv : (S : Type → Type) (e : A ≃ B) → S A → S B +substEquiv S e = subst S (ua e) + + +-- Easy exercise (in the notes): substEquiv induces and equivalence + +-- Example: S could be IsMonoid + + +-- This is useful for many things! + +-- One application: Can program with one type and prove using another equivalent type + +{- + +data Pos : Type where + pos1 : Pos + x0 : Pos → Pos + x1 : Pos → Pos + +-- x1 (x0 (x1 pos1)) = 1101 = 13 + +data Bin : Type where + bin0 : Bin + binPos : Pos → Bin + +-} + + +ℕ≃Bin : ℕ ≃ Bin +ℕ≃Bin = isoToEquiv (iso ℕ→Bin Bin→ℕ Bin→ℕ→Bin ℕ→Bin→ℕ) + + +SemiGroup : Type → Type +SemiGroup A = Σ _·_ ꞉ (A → A → A) , ((x y z : A) → x · (y · z) ≡ (x · y) · z) + +SemiGroupℕ : SemiGroup ℕ +SemiGroupℕ = _+_ , +-assoc + +SemiGroupBin : SemiGroup Bin +SemiGroupBin = substEquiv SemiGroup ℕ≃Bin SemiGroupℕ + +_+Bin_ : Bin → Bin → Bin +_+Bin_ = pr₁ SemiGroupBin + ++Bin-assoc : (x y z : Bin) → x +Bin (y +Bin z) ≡ (x +Bin y) +Bin z ++Bin-assoc = pr₂ SemiGroupBin + + + +mutual + _+P_ : Pos → Pos → Pos + pos1 +P y = sucPos y + x0 x +P pos1 = x1 x + x0 x +P x0 y = x0 (x +P y) + x0 x +P x1 y = x1 (x +P y) + x1 x +P pos1 = x0 (sucPos x) + x1 x +P x0 y = x1 (x +P y) + x1 x +P x1 y = x0 (x +PC y) + + _+B_ : Bin → Bin → Bin + bin0 +B y = y + x +B bin0 = x + binPos x +B binPos y = binPos (x +P y) + + -- Add with carry + _+PC_ : Pos → Pos → Pos + pos1 +PC pos1 = x1 pos1 + pos1 +PC x0 y = x0 (sucPos y) + pos1 +PC x1 y = x1 (sucPos y) + x0 x +PC pos1 = x0 (sucPos x) + x0 x +PC x0 y = x1 (x +P y) + x0 x +PC x1 y = x0 (x +PC y) + x1 x +PC pos1 = x1 (sucPos x) + x1 x +PC x0 y = x0 (x +PC y) + x1 x +PC x1 y = x1 (x +PC y) + + + + + +-- How to prove that +B is associative? (By hand = total pain) + + + + + + +-- Correctness: ++PC-suc : (x y : Pos) → x +PC y ≡ sucPos (x +P y) ++PC-suc pos1 pos1 = refl ++PC-suc pos1 (x0 y) = refl ++PC-suc pos1 (x1 y) = refl ++PC-suc (x0 x) pos1 = refl ++PC-suc (x0 x) (x0 y) = refl ++PC-suc (x0 x) (x1 y) = ap x0 (+PC-suc x y) ++PC-suc (x1 x) pos1 = refl ++PC-suc (x1 x) (x0 y) = ap x0 (+PC-suc x y) ++PC-suc (x1 x) (x1 y) = refl + +sucPos-+P : (x y : Pos) → sucPos (x +P y) ≡ sucPos x +P y +sucPos-+P pos1 pos1 = refl +sucPos-+P pos1 (x0 y) = refl +sucPos-+P pos1 (x1 y) = refl +sucPos-+P (x0 x) pos1 = refl +sucPos-+P (x0 x) (x0 y) = refl +sucPos-+P (x0 x) (x1 y) = ap x0 (sym (+PC-suc x y)) +sucPos-+P (x1 x) pos1 = refl +sucPos-+P (x1 x) (x0 y) = ap x0 (sucPos-+P x y) +sucPos-+P (x1 x) (x1 y) = ap x1 (+PC-suc x y ∙ sucPos-+P x y) + +ℕ→Pos-+P : (x y : ℕ) → ℕ→Pos (suc x + suc y) ≡ ℕ→Pos (suc x) +P ℕ→Pos (suc y) +ℕ→Pos-+P zero _ = refl +ℕ→Pos-+P (suc x) y = ap sucPos (ℕ→Pos-+P x y) ∙ sucPos-+P (ℕ→Pos (suc x)) (ℕ→Pos (suc y)) + +ℕ→Bin-+B : (x y : ℕ) → ℕ→Bin (x + y) ≡ ℕ→Bin x +B ℕ→Bin y +ℕ→Bin-+B zero y = refl +ℕ→Bin-+B (suc x) zero = ap (λ x → binPos (ℕ→Pos (suc x))) (+-zero x) +ℕ→Bin-+B (suc x) (suc y) = ap binPos (ℕ→Pos-+P x y) + + + + ++B≡+Bin : _+B_ ≡ _+Bin_ ++B≡+Bin i x y = goal x y i + where + goal : (x y : Bin) → x +B y ≡ ℕ→Bin (Bin→ℕ x + Bin→ℕ y) + goal x y = (λ i → Bin→ℕ→Bin x (~ i) +B Bin→ℕ→Bin y (~ i)) + ∙ sym (ℕ→Bin-+B (Bin→ℕ x) (Bin→ℕ y)) + + + + ++B-assoc : (m n o : Bin) → m +B (n +B o) ≡ (m +B n) +B o ++B-assoc m n o = + (λ i → +B≡+Bin i m (+B≡+Bin i n o)) + ∙∙ +Bin-assoc m n o + ∙∙ (λ i → +B≡+Bin (~ i) (+B≡+Bin (~ i) m n) o) + + + +-- The agda/cubical library has convenient automation for making a lot +-- of this easier diff --git a/src/HoTTEST/Agda/Cubical/Lecture9-notes.lagda.md b/src/HoTTEST/Agda/Cubical/Lecture9-notes.lagda.md new file mode 100644 index 0000000..46748be --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Lecture9-notes.lagda.md @@ -0,0 +1,608 @@ +# Lecture 9 + +Contents: + +- More about homogeneous composition (`hcomp`) +- Cubical univalence (Glue types) +- The structure identity principle (SIP) + +```agda +{-# OPTIONS --cubical #-} + +module Lecture9-notes where + +open import cubical-prelude +open import Lecture7-notes +open import Lecture8-notes hiding (compPath) + +private + variable + A B : Type ℓ +``` + +## More about homogeneous composition (`hcomp`) + +Recall from Lecture 8: in order to compose two paths we write: + +```agda +compPath : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +compPath {x = x} p q i = hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → q j }) + (p i) +``` + +This is best understood with the following drawing: + +```text + x z + ^ ^ + ¦ ¦ + x ¦ ¦ q j + ¦ ¦ + x ----------> y + p i +``` + +In the drawing the direction `i` goes left-to-right and `j` goes +bottom-to-top. As we are constructing a path from `x` to `z` along `i` +we have `i : I` in the context already and we put `p i` as bottom. The +direction `j` that we are doing the composition in is abstracted in +the first argument to `hcomp`. + +As we can see here the `hcomp` operation has a very natural geometric +motivation. The following YouTube video might be very helpful to +clarify this: https://www.youtube.com/watch?v=MVtlD22Y8SQ + +Side remark: this operation is related to lifting conditions for Kan +cubical sets, i.e. it's a form of open box filling analogous to horn +filling in Kan complexes. + +A more natural form of composition of paths in Cubical Agda is +ternary composition: + +```text + x w + ^ ^ + ¦ ¦ + p (~ j) ¦ ¦ r j + ¦ ¦ + y ----------> z + q i +``` + +This is written `p ∙∙ q ∙∙ r` in the agda/cubical library and +implemented by: + +```agda +_∙∙_∙∙_ : {x y z w : A} → x ≡ y → y ≡ z → z ≡ w → x ≡ w +(p ∙∙ q ∙∙ r) i = hcomp (λ j → λ { (i = i0) → p (~ j) + ; (i = i1) → r j }) + (q i) +``` + +Using this we can define `compPath` much slicker: + +```text +_∙_ : {x y z : A} → x ≡ y → y ≡ z → x ≡ z +p ∙ q = refl ∙∙ p ∙∙ q +``` + +In order to understand the second argument to `hcomp` we need to talk +about partial elements and cubical subtypes. These allow us to write +partially specified n-dimensional cubes (i.e. cubes where some faces +are missing as in the input to `hcomp`). + +### Partial elements + +Given an element of the interval `r : I` there is a predicate `IsOne` +which represents the constraint `r = i1`. This comes with a proof that +`i1` is in fact equal to `i1` called `1=1 : IsOne i1`. We use Greek +letters like `φ` or `ψ` when such an `r` should be thought of as being +in the domain of `IsOne`. + +Using this we introduce a type of partial elements called `Partial φ +A`, this is a special version of `IsOne φ → A` with a more extensional +judgmental equality (two elements of `Partial φ A` are considered +equal if they represent the same subcube, so the faces of the cubes +can for example be given in different order and the two elements will +still be considered the same). The idea is that `Partial φ A` is the +type of cubes in `A` that are only defined when `IsOne φ`. There is also a +dependent version of this called `PartialP φ A` which allows `A` to be +defined only when `IsOne φ`. + +```text +Partial : ∀ {ℓ} → I → Set ℓ → SSet ℓ + +PartialP : ∀ {ℓ} → (φ : I) → Partial φ (Set ℓ) → SSet ℓ +``` + +Here `SSet` is the universe of *strict* sets. + +Cubical Agda has a new form of pattern matching that can be used to +introduce partial elements: + +```agda +partialBool : (i : I) → Partial (i ∨ ~ i) Bool +partialBool i (i = i0) = true +partialBool i (i = i1) = false +``` + +The term `partialBool i` should be thought of a boolean with different +values when `(i = i0)` and `(i = i1)`. Note that this is different +from pattern matching on the interval which is not allowed, so we +couldn't have written: + +```text +partialBool : (i : I) → Bool +partialBool i0 = true +partialBool i1 = false +``` + +Terms of type `Partial φ A` can also be introduced using a pattern +matching lambda and this is what we have been using when we wrote +`hcomp`'s above. + +```agda +partialBool' : (i : I) → Partial (i ∨ ~ i) Bool +partialBool' i = λ { (i = i0) → true + ; (i = i1) → false } +``` + +When the cases overlap they must agree (note that the order of the +cases doesn’t have to match the interval formula exactly): + +```agda +partialBool'' : (i j : I) → Partial (~ i ∨ i ∨ (i ∧ j)) Bool +partialBool'' i j = λ { (i = i1) → true + ; (i = i1) (j = i1) → true + ; (i = i0) → false } +``` + +Furthermore `IsOne i0` is actually absurd. + +```agda +_ : {A : Type} → Partial i0 A +_ = λ { () } +``` + +With this cubical infrastructure we can now describe the type of the +`hcomp` operation. + +```text +hcomp : {A : Type ℓ} {φ : I} (u : I → Partial φ A) (u0 : A) → A +``` + +When calling `hcomp {φ = φ} u u0` Agda makes sure that `u0` agrees +with `u i0` on `φ`. The result is then an element of `A` which agrees +with `u i1` on `φ`. The idea being that `u0` is the base and `u` +specifies the sides of an open box while `hcomp u u0` is the lid of +the box. In fact, we can use yet another cubical construction to +specify these side conditions in the type of `hcomp`. For this we need +to talk about cubical subtypes. + +### Cubical subtypes and filling + +Cubical Agda also has cubical subtypes as in the CCHM type theory: + +```text +_[_↦_] : (A : Type ℓ) (φ : I) (u : Partial φ A) → SSet ℓ +A [ φ ↦ u ] = Sub A φ u +``` + +A term `v : A [ φ ↦ u ]` should be thought of as a term of type `A` +which is definitionally equal to `u : A` when `IsOne φ` is +satisfied. Any term `u : A` can be seen as a term of `A [ φ ↦ u ]` +which agrees with itself on `φ`: + +```text +inS : {A : Type ℓ} {φ : I} (u : A) → A [ φ ↦ (λ _ → u) ] +``` + +One can also forget that a partial element agrees with `u` on `φ`: + +```text +outS : {A : Type ℓ} {φ : I} {u : Partial φ A} → A [ φ ↦ u ] → A +``` + +They satisfy the following equalities: + +```text +outS (inS a) ≐ a + +inS {φ = φ} (outS {φ = φ} a) ≐ a + +outS {φ = i1} {u} _ ≐ u 1=1 +``` + +Note that given `a : A [ φ ↦ u ]` and `α : IsOne φ`, it is not the case +that `outS a ≐ u α`; however, underneath the pattern binding `(φ = i1)`, +one has `outS a ≐ u 1=1`. + +With this we can now give `hcomp` the following more specific type: + +```agda +hcomp' : {A : Type ℓ} {φ : I} (u : I → Partial φ A) (u0 : A [ φ ↦ u i0 ]) → A [ φ ↦ u i1 ] +hcomp' u u0 = inS (hcomp u (outS u0)) +``` + +This more specific type is of course more informative, but it quickly +gets quite annoying to write `inS`/`outS` everywhere. So the builtin +`hcomp` operation comes with the slightly less informative type and +the side conditions are then implicit and checked internally. + +Another very useful operation is open box *filling*. This produces an +element corresponding to the interior of an open box: + +```agda +hfill' : {A : Type ℓ} + {φ : I} + (u : ∀ i → Partial φ A) + (u0 : A [ φ ↦ u i0 ]) + (i : I) → + A [ φ ∨ ~ i ∨ i ↦ + (λ { (φ = i1) → u i 1=1 + ; (i = i0) → outS u0 + ; (i = i1) → hcomp u (outS u0) }) ] +hfill' {φ = φ} u u0 i = inS (hcomp (λ j → λ { (φ = i1) → u (i ∧ j) 1=1 + ; (i = i0) → outS u0 }) + (outS u0)) +``` + +This has a slightly less informative type in the cubical-prelude: + +```text +hfill : {A : Type ℓ} + {φ : I} + (u : ∀ i → Partial φ A) + (u0 : A [ φ ↦ u i0 ]) + (i : I) → + A +hfill {φ = φ} u u0 i = + hcomp (λ j → λ { (φ = i1) → u (i ∧ j) 1=1 + ; (i = i0) → outS u0 }) + (outS u0) +``` + +Having defined this operation we can prove various groupoid laws for +`_≡_` very cubically as `_∙_` is defined using `hcomp`. + +```agda +compPathRefl : {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ refl ≡ p +compPathRefl {x = x} {y = y} p j i = hfill (λ _ → λ { (i = i0) → x + ; (i = i1) → y }) + (inS (p i)) + (~ j) +``` + +For more examples see Cubical.Foundations.GroupoidLaws in agda/cubical. + + +## Cubical univalence (Glue types) + SIP + +A key concept in HoTT/UF is univalence. As we have seen earlier in the +course this is assumed as an axiom in Book HoTT. In Cubical Agda it is +instead provable and has computational content. This means that +transporting with paths constructed using univalence reduces as +opposed to Book HoTT where they would be stuck. This simplifies many +proofs and make it possible to actually do concrete computations using +univalence. + +The part of univalence which is most useful for our applications is to +be able to turn equivalences (written `_≃_` and defined as a Σ-type of +a function and a proof that it has contractible fibers) into paths: + +```agda +ua : {A B : Type ℓ} → A ≃ B → A ≡ B +ua {A = A} {B = B} e i = Glue B (λ { (i = i0) → (A , e) + ; (i = i1) → (B , idEquiv B) }) +``` + +This term is defined using "Glue types" which were introduced in the +CCHM paper. We won't have time to go into too much details about them +today, but for practical applications one can usually forget about +them and use `ua` as a black box. The key idea though is that they are +similar to `hcomp`'s, but instead of attaching (higher dimensional) +paths to a base we attach equivalences to a type. One of the original +applications of Glue types was to give computational meaning to +transporting in a line of types constructed by `hcomp` in the universe +`Type`. + +For us today the important point is that transporting along the path +constructed using `ua` applies the function underlying the +equivalence. This is easily proved using `transportRefl`: + +```agda +uaβ : (e : A ≃ B) (x : A) → transport (ua e) x ≡ pr₁ e x +uaβ e x = transportRefl (equivFun e x) +``` + +Note that for concrete types this typically holds definitionally, but +for an arbitrary equivalence `e` between abstract types `A` and `B` we +have to prove it. + +```agda +uaβℕ : (e : ℕ ≃ ℕ) (x : ℕ) → transport (ua e) x ≡ pr₁ e x +uaβℕ e x = refl +``` + +The fact that we have both `ua` and `uaβ` suffices to be able to prove +the standard formulation of univalence. For details see +Cubical.Foundations.Univalence in the agda/cubical library. + +The standard way of constructing equivalences is to start with an +isomorphism and then improve it into an equivalence. The lemma in the +agda/cubical library that does this is + +```text +isoToEquiv : {A B : Type ℓ} → Iso A B → A ≃ B +``` + +Composing this with `ua` lets us directly turn isomorphisms into paths: + +```text +isoToPath : {A B : Type ℓ} → Iso A B → A ≡ B +isoToPath e = ua (isoToEquiv e) +``` + +Time for an example! + +```agda +not : Bool → Bool +not false = true +not true = false + +notPath : Bool ≡ Bool +notPath = isoToPath (iso not not rem rem) + where + rem : (b : Bool) → not (not b) ≡ b + rem false = refl + rem true = refl + +_ : transport notPath true ≡ false +_ = refl +``` + +Another example, integers: + +```agda +_ : transport sucPath (pos 0) ≡ pos 1 +_ = refl + +_ : transport (sucPath ∙ sucPath) (pos 0) ≡ pos 2 +_ = refl + +_ : transport (sym sucPath) (pos 0) ≡ negsuc 0 +_ = refl +``` + +Note that we have already used this in the `winding` function in +Lecture 7. + +One can do more fun things with `sucPath`. For example by composing it +with itself `n` times and then transporting we get a new addition +function `_+ℤ'_`. It is direct to prove `isEquiv (λ n → n +ℤ' m)` as +`_+ℤ'_` is defined by `transport`. + +```agda +addPath : ℕ → ℤ ≡ ℤ +addPath zero = refl +addPath (suc n) = addPath n ∙ sucPath + +predPath : ℤ ≡ ℤ +predPath = isoToPath (iso predℤ sucℤ predSuc sucPred) + +subPath : ℕ → ℤ ≡ ℤ +subPath zero = refl +subPath (suc n) = subPath n ∙ predPath + +_+ℤ'_ : ℤ → ℤ → ℤ +m +ℤ' pos n = transport (addPath n) m +m +ℤ' negsuc n = transport (subPath (suc n)) m + +-- This agrees with regular addition defined by pattern-matching ++ℤ'≡+ℤ : _+ℤ'_ ≡ _+ℤ_ ++ℤ'≡+ℤ i m (pos zero) = m ++ℤ'≡+ℤ i m (pos (suc n)) = sucℤ (+ℤ'≡+ℤ i m (pos n)) ++ℤ'≡+ℤ i m (negsuc zero) = predℤ m ++ℤ'≡+ℤ i m (negsuc (suc n)) = predℤ (+ℤ'≡+ℤ i m (negsuc n)) + +-- We can prove that transport is an equivalence easily +isEquivTransport : {A B : Type ℓ} (p : A ≡ B) → isEquiv (transport p) +isEquivTransport p = + transport (λ i → isEquiv (transp (λ j → p (i ∧ j)) (~ i))) (idEquiv _ .pr₂) + +-- So +ℤ' with a fixed element is an equivalence +isEquivAddℤ' : (m : ℤ) → isEquiv (λ n → n +ℤ' m) +isEquivAddℤ' (pos n) = isEquivTransport (addPath n) +isEquivAddℤ' (negsuc n) = isEquivTransport (subPath (suc n)) + +isEquivAddℤ : (m : ℤ) → isEquiv (λ n → n +ℤ m) +isEquivAddℤ = subst (λ add → (m : ℤ) → isEquiv (λ n → add n m)) +ℤ'≡+ℤ isEquivAddℤ' +``` + +### The structure identity principle + +Combining `subst` and `ua` lets us transport any structure on `A` to +get a structure on an equivalent type `B`: + +```agda +substEquiv : (S : Type → Type) (e : A ≃ B) → S A → S B +substEquiv S e = subst S (ua e) +``` + +In fact this induces an equivalence: + +```agda +substEquiv≃ : (S : Type → Type) (e : A ≃ B) → S A ≃ S B +substEquiv≃ S e = (substEquiv S e) , (isEquivTransport (ap S (ua e))) +``` + +What this says is that any structure on types must be invariant under +equivalence. We can for example take `IsMonoid` for `S` and get that +any two monoid structures on equivalent types `A` and `B` are +themselves equivalent. This is a simple version of the *structure +identity principle* (SIP). There are various more high powered and +generalized versions which also work for *structured* types +(e.g. monoids, groups, etc.) together with their corresponding notions +of *structured* equivalences (e.g. monoid and group isomorphism, +etc.). We will look more at this later, but for now let's look at a +nice example where we can use `substEquiv` to transport functions and +their properties between equivalent types. + +When programming and proving properties of programs there are usually +various tradeoffs between different data representations. Very often +one version is suitable for proofs while the other is more suitable +for efficient programming. The standard example of them is unary and +binary numbers. One way to define binary numbers in Agda is as: + +```text +data Pos : Type where + pos1 : Pos + x0 : Pos → Pos + x1 : Pos → Pos + +data Bin : Type where + bin0 : Bin + binPos : Pos → Bin +``` + +With some work one can prove that this is equivalent to unary numbers +(see the cubical-prelude for details): + +```agda +ℕ≃Bin : ℕ ≃ Bin +ℕ≃Bin = isoToEquiv (iso ℕ→Bin Bin→ℕ Bin→ℕ→Bin ℕ→Bin→ℕ) +``` + +We can now use `substEquiv` to transport addition on `ℕ` together with +the fact that it's associative over to `Bin`: + +```agda +SemiGroup : Type → Type +SemiGroup X = Σ _+_ ꞉ (X → X → X) , ((x y z : X) → x + (y + z) ≡ (x + y) + z) + +SemiGroupBin : SemiGroup Bin +SemiGroupBin = substEquiv SemiGroup ℕ≃Bin (_+_ , +-assoc) + +_+Bin_ : Bin → Bin → Bin +_+Bin_ = pr₁ SemiGroupBin + ++Bin-assoc : (m n o : Bin) → m +Bin (n +Bin o) ≡ (m +Bin n) +Bin o ++Bin-assoc = pr₂ SemiGroupBin +``` + +This is nice as it helps us avoid having to repeat work on `Bin` that +we have already done on `ℕ`. This is however not always what we want +as `_+Bin_` is not very efficient as an addition function on binary +numbers. In fact, it will translate its input to unary numbers, add +using unary addition, and then translate back. This is of course very +naive and what we instead want to do is to use efficient addition on +binary numbers, but get the associativity proof for free. + +This can be achieved by first defining our fast addition `_+B_ : Bin → +Bin → Bin` and then prove that the map `ℕ→Bin : ℕ → Bin` maps `x + y : +ℕ` to `ℕ→Bin x +B ℕ→Bin y : Bin`. + +``` +mutual + _+P_ : Pos → Pos → Pos + pos1 +P y = sucPos y + x0 x +P pos1 = x1 x + x0 x +P x0 y = x0 (x +P y) + x0 x +P x1 y = x1 (x +P y) + x1 x +P pos1 = x0 (sucPos x) + x1 x +P x0 y = x1 (x +P y) + x1 x +P x1 y = x0 (x +PC y) + + _+B_ : Bin → Bin → Bin + bin0 +B y = y + x +B bin0 = x + binPos x +B binPos y = binPos (x +P y) + + -- Add with carry + _+PC_ : Pos → Pos → Pos + pos1 +PC pos1 = x1 pos1 + pos1 +PC x0 y = x0 (sucPos y) + pos1 +PC x1 y = x1 (sucPos y) + x0 x +PC pos1 = x0 (sucPos x) + x0 x +PC x0 y = x1 (x +P y) + x0 x +PC x1 y = x0 (x +PC y) + x1 x +PC pos1 = x1 (sucPos x) + x1 x +PC x0 y = x0 (x +PC y) + x1 x +PC x1 y = x1 (x +PC y) + +-- Correctness: ++PC-suc : (x y : Pos) → x +PC y ≡ sucPos (x +P y) ++PC-suc pos1 pos1 = refl ++PC-suc pos1 (x0 y) = refl ++PC-suc pos1 (x1 y) = refl ++PC-suc (x0 x) pos1 = refl ++PC-suc (x0 x) (x0 y) = refl ++PC-suc (x0 x) (x1 y) = ap x0 (+PC-suc x y) ++PC-suc (x1 x) pos1 = refl ++PC-suc (x1 x) (x0 y) = ap x0 (+PC-suc x y) ++PC-suc (x1 x) (x1 y) = refl + +sucPos-+P : (x y : Pos) → sucPos (x +P y) ≡ sucPos x +P y +sucPos-+P pos1 pos1 = refl +sucPos-+P pos1 (x0 y) = refl +sucPos-+P pos1 (x1 y) = refl +sucPos-+P (x0 x) pos1 = refl +sucPos-+P (x0 x) (x0 y) = refl +sucPos-+P (x0 x) (x1 y) = ap x0 (sym (+PC-suc x y)) +sucPos-+P (x1 x) pos1 = refl +sucPos-+P (x1 x) (x0 y) = ap x0 (sucPos-+P x y) +sucPos-+P (x1 x) (x1 y) = ap x1 (+PC-suc x y ∙ sucPos-+P x y) + +ℕ→Pos-+P : (x y : ℕ) → ℕ→Pos (suc x + suc y) ≡ ℕ→Pos (suc x) +P ℕ→Pos (suc y) +ℕ→Pos-+P zero _ = refl +ℕ→Pos-+P (suc x) y = ap sucPos (ℕ→Pos-+P x y) ∙ sucPos-+P (ℕ→Pos (suc x)) (ℕ→Pos (suc y)) + +ℕ→Bin-+B : (x y : ℕ) → ℕ→Bin (x + y) ≡ ℕ→Bin x +B ℕ→Bin y +ℕ→Bin-+B zero y = refl +ℕ→Bin-+B (suc x) zero = ap (λ x → binPos (ℕ→Pos (suc x))) (+-zero x) +ℕ→Bin-+B (suc x) (suc y) = ap binPos (ℕ→Pos-+P x y) +``` + +Having done this it's now straightforward to prove that `_+B_` is +associative using the fact that `_+Bin_` is: + +```agda ++B≡+Bin : _+B_ ≡ _+Bin_ ++B≡+Bin i x y = goal x y i + where + goal : (x y : Bin) → x +B y ≡ ℕ→Bin (Bin→ℕ x + Bin→ℕ y) + goal x y = (λ i → Bin→ℕ→Bin x (~ i) +B Bin→ℕ→Bin y (~ i)) + ∙ sym (ℕ→Bin-+B (Bin→ℕ x) (Bin→ℕ y)) + ++B-assoc : (m n o : Bin) → m +B (n +B o) ≡ (m +B n) +B o ++B-assoc m n o = + (λ i → +B≡+Bin i m (+B≡+Bin i n o)) + ∙∙ +Bin-assoc m n o + ∙∙ (λ i → +B≡+Bin (~ i) (+B≡+Bin (~ i) m n) o) +``` + +This kind of proofs quickly get tedious and it turns out we can +streamline them using a more high prowered version of the SIP. +Indeed, what we really want to do is to characterize the equality of +T-structured types, i.e. we want a proof that two types equipped with +T-structures are equal if there is a T-structure preserving +equivalence between them. In the semigroup example above `T` would +just be the structure of a magma (i.e. a type with a binary operation) +together with magma preserving equivalence so that we can transport +for example the fact that the magma is a semigroup over to the other +magma. + +We formalize this and make it more convenient to use using a lot of +automation in the agda/cubical library. This is documented with +various more examples in: + +Internalizing Representation Independence with Univalence +Carlo Angiuli, Evan Cavallo, Anders Mörtberg, Max Zeuner +https://dl.acm.org/doi/10.1145/3434293 + +The binary numbers example with efficient addition is spelled out in +detail in Section 2.1.1 of: + +https://www.doi.org/10.1017/S0956796821000034 +(Can be downloaded from: https://staff.math.su.se/anders.mortberg/papers/cubicalagda2.pdf) diff --git a/src/HoTTEST/Agda/Cubical/Solutions7.lagda.md b/src/HoTTEST/Agda/Cubical/Solutions7.lagda.md new file mode 100644 index 0000000..3cd7517 --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Solutions7.lagda.md @@ -0,0 +1,249 @@ +# Week 7 - Cubical Agda Exercises - Solutions + +```agda +{-# OPTIONS --cubical #-} + +module Solutions7 where + +open import cubical-prelude +open import Lecture7-notes +``` + +```agda +private + variable + A : Type ℓ + B : A → Type ℓ +``` + +## Part I: + +### Exercise 1 +```agda +depFunExt : {f g : (x : A) → B x} + → (∀ x → f x ≡ g x) + → f ≡ g +depFunExt h i x = h x i +``` + +### Exercise 2 +```agda +apd : (f : (x : A) → B x) {x y : A} (p : x ≡ y) → PathP (λ i → B (p i)) (f x) (f y) +apd f p i = f (p i) + +``` + + +## Part II + +### Exercise 3 +```agda +inhPropIsContr : isProp A → A → isContr A +inhPropIsContr h a = a , h a +``` + +### Exercise 4 +```agda +isPropΠ : (h : (x : A) → isProp (B x)) → isProp ((x : A) → B x) +isPropΠ h f g i x = h x (f x) (g x) i +``` + +### Exercise 5 +```agda +funExt⁻ : {f g : (x : A) → B x} → f ≡ g → ((x : A) → f x ≡ g x) +funExt⁻ p x i = p i x +``` + +### Exercise 6 +```agda +isSetΠ : (h : (x : A) → isSet (B x)) → isSet ((x : A) → B x) +isSetΠ h f g p q i j x = h x (f x) (g x) (funExt⁻ p x) (funExt⁻ q x) i j +``` + +### Exercise 7 +```agda +singl' : {A : Type ℓ} (a : A) → Type ℓ +singl' {A = A} a = Σ x ꞉ A , x ≡ a + +isContrSingl' : (x : A) → isContr (singl' x) +isContrSingl' x = ctr , prf + where + ctr : singl' x + ctr = (x , refl) + + prf : (s : singl' x) → ctr ≡ s + prf (y , p) i = p (~ i) , λ j → p ((~ i) ∨ j) +``` + + +## Part III: Equality in Σ-types +### Exercise 8 +```agda +module _ {A : Type ℓ} {B : A → Type ℓ'} {x y : Σ A B} where + + ΣPathP : Σ p ꞉ pr₁ x ≡ pr₁ y , PathP (λ i → B (p i)) (pr₂ x) (pr₂ y) + → x ≡ y + ΣPathP (p , q) i = p i , q i + + PathPΣ : x ≡ y + → Σ p ꞉ pr₁ x ≡ pr₁ y , PathP (λ i → B (p i)) (pr₂ x) (pr₂ y) + PathPΣ p = (ap pr₁ p) , (apd pr₂ p) + + ΣPathP-PathPΣ : ∀ p → PathPΣ (ΣPathP p) ≡ p + ΣPathP-PathPΣ _ = refl + + PathPΣ-ΣPathP : ∀ p → ΣPathP (PathPΣ p) ≡ p + PathPΣ-ΣPathP _ = refl +``` + + +## Part III: Some HITs +### Exercise 9 +```agda +{- + + +The torus: + + ∙ ->->->-> ∙ + ∧ ∧ + | | + ∧ ∧ + | | + ∙ ->->->-> ∙ + + + + +The Klein bottle: + + ∙ <-<-<-<- ∙ + ∧ ∧ + | | + ∧ ∧ + | | + ∙ ->->->-> ∙ + + +-} +data KleinBottle : Type where + point : KleinBottle + line1 : point ≡ point + line2 : point ≡ point + square : PathP (λ i → line1 (~ i) ≡ line1 i) line2 line2 + +data KleinBottle' : Type where + point : KleinBottle' + line1 : point ≡ point + line2 : point ≡ point + square : (sym line1) ∙ line2 ∙ line1 ≡ line2 +``` + + + +### Exercise 10 +```agda +suspUnitChar : Susp Unit ≡ Interval +suspUnitChar = isoToPath (iso to fro toFro froTo) + where + to : Susp Unit → Interval + to north = zero + to south = one + to (merid ⋆ i) = seg i + + fro : Interval → Susp Unit + fro zero = north + fro one = south + fro (seg i) = merid ⋆ i + + toFro : (b : Interval) → to (fro b) ≡ b + toFro zero = refl + toFro one = refl + toFro (seg i) = refl + + froTo : (s : Susp Unit) → fro (to s) ≡ s + froTo north = refl + froTo south = refl + froTo (merid ⋆ i) = refl +``` + + + +### Exercise 11 +```agda +SuspPushout : Type ℓ → Type ℓ +SuspPushout A = Pushout {A = A} {B = Unit} {C = Unit} (λ a → ⋆) (λ a → ⋆) + +Susp≡SuspPushout : (A : Type ℓ) → Susp A ≡ SuspPushout A +Susp≡SuspPushout A = isoToPath (iso to fro toFro froTo) + where + to : Susp A → SuspPushout A + to north = inl ⋆ + to south = inr ⋆ + to (merid a i) = push a i + + fro : SuspPushout A → Susp A + fro (inl ⋆) = north + fro (inr ⋆) = south + fro (push a i) = merid a i + + toFro : (b : SuspPushout A) → to (fro b) ≡ b + toFro (inl ⋆) = refl + toFro (inr ⋆) = refl + toFro (push a i) = refl + + froTo : (s : Susp A) → fro (to s) ≡ s + froTo north = refl + froTo south = refl + froTo (merid a i) = refl +``` + +### Exercise 12 + +```agda +comp-filler : {x y z : A} (p : x ≡ y) (q : y ≡ z) + → PathP (λ j → refl {x = x} j ≡ q j) p (p ∙ q) +comp-filler {x = x} p q j i = hfill (λ k → λ { (i = i0) → x + ; (i = i1) → q k }) + (inS (p i)) j + +rUnit : {x y : A} (p : x ≡ y) → p ∙ refl ≡ p +rUnit p = sym (comp-filler p refl) + + + +suspBoolChar : Susp Bool ≡ S¹ +suspBoolChar = isoToPath (iso to fro toFro froTo) + where + to : Susp Bool → S¹ + to north = base + to south = base + to (merid true i) = loop i + to (merid false i) = base + + fro : S¹ → Susp Bool + fro base = north + fro (loop i) = (merid true ∙ sym (merid false)) i + + toFro : (b : S¹) → to (fro b) ≡ b + toFro base = refl + toFro (loop i) j = mySquare j i + where + mySquare : ap to (merid true ∙ sym (merid false)) ≡ loop + mySquare = ap to (merid true ∙ sym (merid false)) ≡⟨ refl ⟩ + ap to (merid true) ∙ ap to (sym (merid false)) ≡⟨ refl ⟩ + ap to (merid true) ∙ refl ≡⟨ refl ⟩ + loop ∙ refl ≡⟨ rUnit loop ⟩ + loop ∎ + + froTo : (s : Susp Bool) → fro (to s) ≡ s + froTo north = refl + froTo south = merid false + froTo (merid true i) j = mySquare (~ j) i + where + mySquare : PathP (λ j → north ≡ sym (merid false) j) + (merid true) + (merid true ∙ sym (merid false)) + mySquare = comp-filler (merid true) (sym (merid false)) + froTo (merid false i) j = merid false (i ∧ j) +``` diff --git a/src/HoTTEST/Agda/Cubical/Solutions8.lagda.md b/src/HoTTEST/Agda/Cubical/Solutions8.lagda.md new file mode 100644 index 0000000..589d0ce --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Solutions8.lagda.md @@ -0,0 +1,114 @@ +# Week 8 - Cubical Agda Exercises - Solutions + +```agda +{-# OPTIONS --cubical #-} + +module Solutions8 where + +open import cubical-prelude +open import Lecture7-notes +open import Lecture8-notes +``` + +## Part I: `transp` and `hcomp` + +### Exercise 1 (★) + +Prove the propositional computation law for `J`: + +```agda +-- J = λ {x} P d p → transp (λ i → P (p i) (λ j → p (i ∧ j))) i0 d + +JRefl : {A : Type ℓ} {x : A} (P : (z : A) → x ≡ z → Type ℓ'') + (d : P x refl) → J P d refl ≡ d +JRefl {x = x} P d i = transp (λ _ → P x (λ _ → x)) i d +``` + +### Exercise 2 (★★) + +Using `transp`, construct a method for turning dependent paths into paths. + +**Hint**: +Transport the point `p i` to the fibre `A i1`, and set `φ = i` such that the +transport computes away at `i = i1`. +```text + x ----(p i)----> y + A i0 A i A i1 +``` + +```agda +fromPathP : {A : I → Type ℓ} {x : A i0} {y : A i1} → + PathP A x y → transport (λ i → A i) x ≡ y +fromPathP {A = A} p i = transp (λ j → A (i ∨ j)) i (p i) +``` + +### Exercise 3 (★★★) + +Using `hcomp`, construct a method for turning paths into dependent paths. + +**Hint**: +At each point `i`, the verical fibre of the following diagram should lie in +`A i`. The key is to parametrise the bottom line with a dependent path from `x` +to `transport A x`. This requires writing a `transp` that computes away at +`i = i0`. + +```text + x - - - -> y + ^ ^ + ¦ ¦ + refl ¦ ¦ p i + ¦ ¦ + ¦ ¦ + x ---> transport A x +``` + + +```agda +toPathP : {A : I → Type ℓ} {x : A i0} {y : A i1} → + transport (λ i → A i) x ≡ y → PathP A x y +toPathP {A = A} {x = x} p i = + hcomp + (λ {j (i = i0) → x ; + j (i = i1) → p j }) + (transpFill i) + where + transpFill : PathP A x (transport (λ i → A i) x) + transpFill i = transp (λ j → A (i ∧ j)) (~ i) x +``` + +### Exercise 4 (★) + +Using `toPathP`, prove the following lemma, which lets you fill in dependent +lines in hProps, provided their boundary. + +```agda +isProp→PathP : {A : I → Type ℓ} (p : (i : I) → isProp (A i)) + (a₀ : A i0) (a₁ : A i1) → PathP A a₀ a₁ +isProp→PathP p a₀ a₁ = toPathP (p i1 _ a₁) +``` + +### Exercise 5 (★★) + +Prove the following lemma characterizing equality in subtypes: + +```agda +Σ≡Prop : {A : Type ℓ} {B : A → Type ℓ'} {u v : Σ A B} (h : (x : A) → isProp (B x)) + → (p : pr₁ u ≡ pr₁ v) → u ≡ v +Σ≡Prop {B = B} {u = u} {v = v} h p i = p i , isProp→PathP (λ i → h (p i)) (pr₂ u) (pr₂ v) i +``` + +### Exercise 6 (★★★) + +Prove that `isContr` is a proposition: + +**Hint**: +This requires drawing a cube (yes, an actual 3D one)! + +```agda +isPropIsContr : {A : Type ℓ} → isProp (isContr A) +isPropIsContr (c0 , h0) (c1 , h1) j = h0 c1 j , λ y i → + hcomp (λ {k (i = i0) → h0 c1 j ; + k (i = i1) → h1 y k ; + k (j = i0) → h0 (h1 y k) i ; + k (j = i1) → h1 y (i ∧ k) }) (h0 c1 (i ∨ j)) +``` diff --git a/src/HoTTEST/Agda/Cubical/Solutions9.lagda.md b/src/HoTTEST/Agda/Cubical/Solutions9.lagda.md new file mode 100644 index 0000000..d69809c --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/Solutions9.lagda.md @@ -0,0 +1,359 @@ +# Week 9 - Cubical Agda Exercises + +## Standard disclaimer: + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +In case you haven't done the other Agda exercises: +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**When solving the problems, +please make a copy of this file to work in, so that it doesn't get overwritten +(in case we update the exercises through `git`)!** + + +```agda +{-# OPTIONS --cubical --allow-unsolved-metas #-} + +module Solutions9 where + +open import cubical-prelude +open import Lecture7-notes +open import Lecture8-notes +open import Lecture9-notes +open import Solutions7 hiding (rUnit) +open import Solutions8 +open import Lecture9-live using (SemiGroupℕ) +``` + +## Part I: More hcomps + +### Exercise 1 (★★) +### (a) +Show the left cancellation law for path composition using an hcomp. +Hint: one hcomp should suffice. Use `comp-filler` and connections + +```agda +lUnit : {A : Type ℓ} {x y : A} (p : x ≡ y) → refl ∙ p ≡ p +lUnit {x = x} {y = y} p i j = + hcomp (λ k → λ {(i = i0) → comp-filler refl p k j + ; (i = i1) → p (k ∧ j) + ; (j = i0) → x + ; (j = i1) → p k}) + x + + +``` +### (b) +Try to mimic the construction of lUnit for rUnit (i.e. redefine it) +in such a way that `rUnit refl ≡ lUnit refl` holds by `refl`. +Hint: use (almost) the exact same hcomp. + +```agda +rUnit : {A : Type ℓ} {x y : A} (p : x ≡ y) → p ∙ refl ≡ p +rUnit {x = x} {y = y} p i j = + hcomp (λ k → λ {(i = i0) → comp-filler p refl k j + ; (i = i1) → p j + ; (j = i0) → x + ; (j = i1) → y}) + (p j) + +-- uncomment to see if it type-checks + +{- +rUnit≡lUnit : ∀ {ℓ} {A : Type ℓ} {x : A} → rUnit (refl {x = x}) ≡ lUnit refl +rUnit≡lUnit = refl +-} + +``` + + +### Exercise 2 (★★) +Show the associativity law for path composition +Hint: one hcomp should suffice. This one can be done without connections + (but you might need comp-filler in more than one place) + +```agda +assoc : {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙ (q ∙ r) ≡ (p ∙ q) ∙ r +assoc {x = x} {y = y} p q r i j = + hcomp (λ k → λ {(i = i0) → (p ∙ comp-filler q r k) j + ; (i = i1) → comp-filler (p ∙ q) r k j + ; (j = i0) → x + ; (j = i1) → r k}) + ((p ∙ q) j) + +``` + +### Exercise 3 (Master class in connections) (🌶) +The goal of this exercise is to give a cubical proof of the Eckmann-Hilton argument, +which says that path composition for higher loops is commutative + +(a) While we cannot get `p ∙ q ≡ q ∙ p` as a one-liner, we can get a +one-liner showing that the identiy holds up to some annoying +coherences. Try to understand the following statement (and why it's +well-typed). After that, fill the holes + +Hint: each hole will need a `∨` or a `∧` + +```agda +pre-EH : {A : Type ℓ} {x : A} (p q : refl {x = x} ≡ refl) + → ap (λ x → x ∙ refl) p ∙ ap (λ x → refl ∙ x) q + ≡ ap (λ x → refl ∙ x) q ∙ ap (λ x → x ∙ refl) p +pre-EH {x = x} p q i = (λ j → p (~ i ∧ j) ∙ q (i ∧ j)) + ∙ (λ j → p (~ i ∨ j) ∙ q (i ∨ j)) + +``` +(b) If we manage to cancel out all of the annoying aps, we get Eckmann-Hilton: +For paths (p q : refl ≡ refl), we have p ∙ q ≡ q ∙ p. Try to prove this, using the above lemma. + +Hint: Use the pre-EH as the bottom of an hcomp (one should be enough). +For the sides, use lUnit and rUnit wherever they're needed. Note that this will only work out smoothly if +you've solved Exercise 1 (b). + +```agda +Eckmann-Hilton : {A : Type ℓ} {x : A} (p q : refl {x = x} ≡ refl) → p ∙ q ≡ q ∙ p +Eckmann-Hilton {x = x} p q k i = + hcomp (λ r → λ {(i = i0) → rUnit (refl {x = x}) r + ; (i = i1) → rUnit (refl {x = x}) r + ; (k = i0) → ((λ j → rUnit (p j) r) ∙ (λ j → lUnit (q j) r)) i + ; (k = i1) → ((λ j → lUnit (q j) r) ∙ (λ j → rUnit (p j) r)) i}) + (pre-EH p q k i) + +``` +# # Part 2: Binary numbers as a HIT +Here is another HIT describing binary numbers. The idea is that a binary number is a list of booleans, modulo trailing zeros. + +For instance, `true ∷ true ∷ true ∷ []` is the binary number 110 ... +... and so is `true ∷ true ∷ false ∷ false ∷ false ∷ []` + +(!) Note that we're interpreting 110 as 1·2⁰ + 1·2¹ + 0·2² here. + +```agda +0B = false +1B = true + +data ListBin : Type where + [] : ListBin + _∷_ : (x : Bool) (xs : ListBin) → ListBin + drop0 : 0B ∷ [] ≡ [] + +-- 1 as a binary number +1LB : ListBin +1LB = 1B ∷ [] +``` +### Exercise 4 (★) +Define the successor function on ListBin +```agda + + + +sucListBin : ListBin → ListBin +sucListBin [] = 1LB +sucListBin (true ∷ xs) = false ∷ sucListBin xs +sucListBin (false ∷ xs) = 1B ∷ xs +sucListBin (drop0 i) = 1LB + +``` +### Exercise 5 (★★) +Define an addition `+LB` on ListBin and prove that `x +LB [] ≡ x` +Do this by mutual induction! Make sure the three cases for the right unit law hold by refl. +```agda + +_+LB_ : ListBin → ListBin → ListBin +rUnit+LB : (x : ListBin) → x +LB [] ≡ x +[] +LB y = y +(x ∷ xs) +LB [] = x ∷ xs +(true ∷ xs) +LB (true ∷ ys) = false ∷ sucListBin (xs +LB ys) +(true ∷ xs) +LB (false ∷ ys) = true ∷ (xs +LB ys) +(false ∷ xs) +LB (y ∷ ys) = y ∷ (xs +LB ys) +(true ∷ xs) +LB drop0 i = true ∷ (rUnit+LB xs i) +(false ∷ xs) +LB drop0 i = false ∷ (rUnit+LB xs i) +drop0 i +LB [] = drop0 i +drop0 i +LB (y ∷ ys) = y ∷ ys +drop0 i +LB drop0 j = drop0 (j ∧ i) +rUnit+LB [] = refl +rUnit+LB (x ∷ x₁) = refl +rUnit+LB (drop0 i) = refl + +``` +(c) Prove that sucListBin is left distributive over `+LB` +Hint: If you pattern match deep enough, there should be a lot of refls... +```agda + +sucListBinDistrL : (x y : ListBin) → sucListBin (x +LB y) ≡ (sucListBin x +LB y) +sucListBinDistrL [] [] = refl +sucListBinDistrL [] (true ∷ y) = refl +sucListBinDistrL [] (false ∷ y) = refl +sucListBinDistrL [] (drop0 i) = refl +sucListBinDistrL (true ∷ xs) [] = refl +sucListBinDistrL (false ∷ xs) [] = refl +sucListBinDistrL (true ∷ xs) (true ∷ y) = ap (1B ∷_) (sucListBinDistrL xs y) +sucListBinDistrL (true ∷ xs) (false ∷ y) = ap (0B ∷_) (sucListBinDistrL xs y) +sucListBinDistrL (false ∷ xs) (true ∷ y) = refl +sucListBinDistrL (false ∷ xs) (false ∷ y) = refl +sucListBinDistrL (true ∷ []) (drop0 i) = refl +sucListBinDistrL (true ∷ (true ∷ xs)) (drop0 i) = refl +sucListBinDistrL (true ∷ (false ∷ xs)) (drop0 i) = refl +sucListBinDistrL (true ∷ drop0 i) (drop0 j) = refl +sucListBinDistrL (false ∷ xs) (drop0 i) = refl +sucListBinDistrL (drop0 i) [] = refl +sucListBinDistrL (drop0 i) (true ∷ y) = refl +sucListBinDistrL (drop0 i) (false ∷ y) = refl +sucListBinDistrL (drop0 i) (drop0 j) = refl +``` + +### Exercise 6 (★) +Define a map `LB→ℕ : ListBin → ℕ` and show that it preserves addition + +```agda +ℕ→ListBin : ℕ → ListBin +ℕ→ListBin zero = [] +ℕ→ListBin (suc x) = sucListBin (ℕ→ListBin x) + +ℕ→ListBin-pres+ : (x y : ℕ) → ℕ→ListBin (x + y) ≡ (ℕ→ListBin x +LB ℕ→ListBin y) +ℕ→ListBin-pres+ zero y = refl +ℕ→ListBin-pres+ (suc x) zero = + ap sucListBin (ap ℕ→ListBin (+-comm x zero)) + ∙ sym (rUnit+LB (sucListBin (ℕ→ListBin x))) +ℕ→ListBin-pres+ (suc x) (suc y) = + ap sucListBin (ℕ→ListBin-pres+ x (suc y)) + ∙ sucListBinDistrL (ℕ→ListBin x) (sucListBin (ℕ→ListBin y)) + +``` + +### Exercise 7 (★★★) +Show that `ℕ ≃ ListBin`. + +```agda + +-- useful lemma +move4 : (x y z w : ℕ) → (x + y) + (z + w) ≡ (x + z) + (y + w) +move4 x y z w = + (x + y) + (z + w) ≡⟨ +-assoc (x + y) z w ⟩ + ((x + y) + z) + w ≡⟨ ap (_+ w) (sym (+-assoc x y z) + ∙ ap (x +_) (+-comm y z) + ∙ (+-assoc x z y)) ⟩ + ((x + z) + y) + w ≡⟨ sym (+-assoc (x + z) y w) ⟩ + ((x + z) + (y + w)) ∎ + +ListBin→ℕ : ListBin → ℕ +ListBin→ℕ [] = 0 +ListBin→ℕ (true ∷ xs) = suc (2 · ListBin→ℕ xs) +ListBin→ℕ (false ∷ xs) = 2 · ListBin→ℕ xs +ListBin→ℕ (drop0 i) = 0 + +ListBin→ℕ-suc : (x : ListBin) → ListBin→ℕ (sucListBin x) ≡ suc (ListBin→ℕ x) +ListBin→ℕ-suc [] = refl +ListBin→ℕ-suc (true ∷ xs) = + ap (λ x → x + x) (ListBin→ℕ-suc xs) + ∙ ap suc (sym (+-suc (ListBin→ℕ xs) (ListBin→ℕ xs))) +ListBin→ℕ-suc (false ∷ xs) = refl +ListBin→ℕ-suc (drop0 i) = refl + +x+x-charac : (xs : ListBin) → (xs +LB xs) ≡ (false ∷ xs) +x+x-charac [] = sym drop0 +x+x-charac (true ∷ xs) = ap (false ∷_) (ap sucListBin (x+x-charac xs)) +x+x-charac (false ∷ xs) = ap (false ∷_) (x+x-charac xs) +x+x-charac (drop0 i) j = + hcomp (λ k → λ {(i = i0) → false ∷ drop0 (~ j ∨ ~ k) + ; (i = i1) → drop0 (~ j) + ; (j = i0) → drop0 i + ; (j = i1) → false ∷ drop0 (i ∨ ~ k)}) + (drop0 (~ j ∧ i)) + +suc-x+x-charac : (xs : ListBin) → sucListBin (xs +LB xs) ≡ (true ∷ xs) +suc-x+x-charac xs = ap sucListBin (x+x-charac xs) + +ListBin→ℕ→ListBin : (x : ListBin) → ℕ→ListBin (ListBin→ℕ x) ≡ x +ListBin→ℕ→ListBin [] = refl +ListBin→ℕ→ListBin (true ∷ xs) = + ap sucListBin (ℕ→ListBin-pres+ (ListBin→ℕ xs) (ListBin→ℕ xs) + ∙ ap (λ x → x +LB x) (ListBin→ℕ→ListBin xs)) + ∙ suc-x+x-charac xs +ListBin→ℕ→ListBin (false ∷ xs) = + (ℕ→ListBin-pres+ (ListBin→ℕ xs) (ListBin→ℕ xs) + ∙ (ap (λ x → x +LB x) (ListBin→ℕ→ListBin xs))) + ∙ x+x-charac xs +ListBin→ℕ→ListBin (drop0 i) = help i + where + lem : (refl ∙ refl) ∙ sym drop0 ≡ sym drop0 + lem = ap (_∙ sym drop0) (rUnit refl) ∙ lUnit (sym drop0) + + help : PathP (λ i → [] ≡ drop0 i) ((refl ∙ refl) ∙ sym drop0) refl + help i j = hcomp (λ k → λ {(i = i0) → lem (~ k) j + ; (i = i1) → [] + ; (j = i0) → [] + ; (j = i1) → drop0 i}) + (drop0 (i ∨ ~ j)) + +ℕ→ListBin→ℕ : (x : ℕ) → ListBin→ℕ (ℕ→ListBin x) ≡ x +ℕ→ListBin→ℕ zero = refl +ℕ→ListBin→ℕ (suc x) = + ListBin→ℕ-suc (ℕ→ListBin x) + ∙ ap suc (ℕ→ListBin→ℕ x) + +ℕ≃ListBin : ℕ ≃ ListBin +ℕ≃ListBin = isoToEquiv (iso ℕ→ListBin ListBin→ℕ ListBin→ℕ→ListBin ℕ→ListBin→ℕ) + +``` +# Part 3: The SIP +### Exericise 8 (★★) +Show that, using an SIP inspired argument, if `(A , _+A_)` is a semigroup and `(B , _+B_)` is some other type with a composition satisfying: + +(i) `e : A ≃ B` + +(ii) `((x y : A) → e (x +A y) ≡ e x +B e y` + +then `(B , _+B_)` defines a semigroup. + +Conclude that `(ListBin , _+LB_)` is a semigroup + +For inspiration, see Lecture9-notes +```agda + + +module _ {A B : Type} (SA : SemiGroup A) (e : A ≃ B) + (_+B_ : B → B → B) + (hom : (x y : A) → pr₁ e (pr₁ SA x y) ≡ ((pr₁ e x) +B pr₁ e y)) + where + f = pr₁ e + f⁻ = invEq e + + _+A_ = pr₁ SA + + SemiGroupSIPAux : SemiGroup B + SemiGroupSIPAux = substEquiv SemiGroup e SA + + _+B'_ = pr₁ SemiGroupSIPAux + + +B'≡+B : (x y : B) → x +B' y ≡ (x +B y) + +B'≡+B x y = + (x +B' y) ≡⟨ transportRefl (f (f⁻ (transport refl x) + +A (f⁻ (transport refl y)))) ⟩ + f (f⁻ (transport refl x) +A f⁻ (transport refl y)) ≡⟨ (λ i → f (f⁻ (transportRefl x i) + +A f⁻ (transportRefl y i))) ⟩ + f (f⁻ x +A f⁻ y) ≡⟨ hom (f⁻ x) (f⁻ y) ⟩ + (f (f⁻ x) +B f (f⁻ y)) ≡⟨ (λ i → secEq e x i +B secEq e y i) ⟩ + x +B y ∎ + + assoc+B : ∀ m n o → m +B (n +B o) ≡ (m +B n) +B o + assoc+B m n o = + (m +B (n +B o)) ≡⟨ (λ i → +B'≡+B m (+B'≡+B n o (~ i)) (~ i)) ⟩ + (m +B' (n +B' o)) ≡⟨ pr₂ SemiGroupSIPAux m n o ⟩ + ((m +B' n) +B' o) ≡⟨ (λ i → +B'≡+B (+B'≡+B m n i) o i) ⟩ + ((m +B n) +B o) ∎ + + inducedSemiGroup : SemiGroup B + inducedSemiGroup = _+B_ , assoc+B + +SemiGroupListBin : SemiGroup ListBin +SemiGroupListBin = inducedSemiGroup SemiGroupℕ ℕ≃ListBin _+LB_ ℕ→ListBin-pres+ diff --git a/src/HoTTEST/Agda/Cubical/cubical-prelude.lagda.md b/src/HoTTEST/Agda/Cubical/cubical-prelude.lagda.md new file mode 100644 index 0000000..deb0c6e --- /dev/null +++ b/src/HoTTEST/Agda/Cubical/cubical-prelude.lagda.md @@ -0,0 +1,502 @@ +This file contains whatever is needed from the agda/cubical library to +get the lectures to typecheck. Warning: it is not very organized or +documented. + +```agda +{-# OPTIONS --cubical #-} + +module cubical-prelude where + +open import Agda.Builtin.Cubical.Path public +open import Agda.Builtin.Cubical.Sub public + renaming ( primSubOut to outS + ) +open import Agda.Primitive.Cubical public + renaming ( primIMin to _∧_ -- I → I → I + ; primIMax to _∨_ -- I → I → I + ; primINeg to ~_ -- I → I + ; isOneEmpty to empty + ; primComp to comp + ; primHComp to hcomp + ; primTransp to transp + ; itIsOne to 1=1 ) + + +open import Agda.Builtin.Cubical.Glue public + using ( isEquiv -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ ⊔ ℓ') + + ; equiv-proof -- ∀ (y : B) → isContr (fiber f y) + + ; _≃_ -- ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Type (ℓ ⊔ ℓ') + + ; equivFun -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → A → B + + ; equivProof -- ∀ {ℓ ℓ'} (T : Type ℓ) (A : Type ℓ') (w : T ≃ A) (a : A) φ → + -- Partial φ (fiber (equivFun w) a) → fiber (equivFun w) a + + ; primGlue -- ∀ {ℓ ℓ'} (A : Type ℓ) {φ : I} (T : Partial φ (Type ℓ')) + -- → (e : PartialP φ (λ o → T o ≃ A)) → Type ℓ' + + ; prim^unglue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} → primGlue A T e → A + + -- The ∀ operation on I. This is commented out as it is not currently used for anything + -- ; primFaceForall -- (I → I) → I + ) + renaming ( prim^glue to glue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} + -- → PartialP φ T → A → primGlue A T e + ) + + +open import Agda.Primitive public + using ( Level + ; SSet + ; lzero + ; lsuc + ; _⊔_ ) + renaming ( Set to Type ) + +-- We parametrize everything by some universe levels +variable + ℓ ℓ' ℓ'' : Level + +-- Non dependent path types + +Path : ∀ {ℓ} (A : Type ℓ) → A → A → Type ℓ +Path A a b = PathP (λ _ → A) a b + +-- PathP (λ i → A) x y gets printed as x ≡ y when A does not mention i. +-- _≡_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ +-- _≡_ {A = A} = PathP (λ _ → A) + +-- Path composition +_∙_ : {A : Type ℓ} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ {x = x} p q i = hcomp (λ j → λ { (i = i0) → x + ; (i = i1) → q j }) + (p i) + +infixr 30 _∙_ + + +-- Equality reasoning +infix 3 _∎ +infixr 2 _≡⟨_⟩_ _≡⟨⟩_ + +_≡⟨_⟩_ : {A : Type ℓ} (x : A) {y z : A} → x ≡ y → y ≡ z → x ≡ z +_ ≡⟨ x≡y ⟩ y≡z = x≡y ∙ y≡z + +≡⟨⟩-syntax : {A : Type ℓ} (x : A) {y z : A} → x ≡ y → y ≡ z → x ≡ z +≡⟨⟩-syntax = _≡⟨_⟩_ +infixr 2 ≡⟨⟩-syntax +syntax ≡⟨⟩-syntax x (λ i → B) y = x ≡[ i ]⟨ B ⟩ y + +_≡⟨⟩_ : {A : Type ℓ} (x : A) {y : A} → x ≡ y → x ≡ y +_ ≡⟨⟩ x≡y = x≡y + +_∎ : {A : Type ℓ} (x : A) → x ≡ x +x ∎ = λ _ → x + + +_[_↦_] : ∀ {ℓ} (A : Type ℓ) (φ : I) (u : Partial φ A) → SSet ℓ +A [ φ ↦ u ] = Sub A φ u + +infix 4 _[_↦_] + +--- Homogeneous filling +hfill : {A : Type ℓ} {φ : I} (u : ∀ i → Partial φ A) (u0 : A [ φ ↦ u i0 ]) (i : I) → A +hfill {φ = φ} u u0 i = + hcomp (λ j → λ { (φ = i1) → u (i ∧ j) 1=1 + ; (i = i0) → outS u0 }) + (outS u0) + +-- Use built in Σ types to avoid problems with the imported Cubical +-- Agda operations that use Σ types +open import Agda.Builtin.Sigma public renaming (fst to pr₁ ; snd to pr₂) + +Sigma : {l1 l2 : Level} (A : Type l1) (B : A → Type l2) → Type (l1 ⊔ l2) +Sigma {l1} {l2} A B = Σ {l1} {l2} A B + +syntax Sigma A (λ x → b) = Σ x ꞉ A , b + +infix -1 Sigma + +_×_ : ∀ {l1 l2} → Type l1 → Type l2 → Type (l1 ⊔ l2) +A × B = Sigma A (\ _ → B) + +infixr 5 _×_ + +-- Contractible types, propositions and sets: + +isContr : Type ℓ → Type ℓ +isContr A = Σ x ꞉ A , ((y : A) → x ≡ y) + +isProp : Type ℓ → Type ℓ +isProp A = (x y : A) → x ≡ y + +isSet : Type ℓ → Type ℓ +isSet A = (x y : A) → isProp (x ≡ y) + +-- Fibers +fiber : {A : Type ℓ} {B : Type ℓ'} (f : A → B) (y : B) → Set (ℓ ⊔ ℓ') +fiber {A = A} f y = Σ x ꞉ A , f x ≡ y + +-- In the agda/cubical library we call these h-levels following +-- Voevodsky instead of n-types and index by natural numbers instead +-- of ℕ₋₂. So isContr is isOfHLevel 0, isProp is isOfHLevel 1, isSet +-- is isOfHLevel 2, etc. For details see Cubical/Foundations/HLevels.agda + + +-- Sections and retractions +module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} where + section : (f : A → B) → (g : B → A) → Type ℓ' + section f g = ∀ b → f (g b) ≡ b + + -- NB: `g` is the retraction! + retract : (f : A → B) → (g : B → A) → Type ℓ + retract f g = ∀ a → g (f a) ≡ a + +module _ {A B : Type} {f : A → B} (equivF : isEquiv f) where + funIsEq : A → B + funIsEq = f + + invIsEq : B → A + invIsEq y = equivF .equiv-proof y .pr₁ .pr₁ + + secIsEq : section f invIsEq + secIsEq y = equivF .equiv-proof y .pr₁ .pr₂ + + retIsEq : retract f invIsEq + retIsEq a i = equivF .equiv-proof (f a) .pr₂ (a , λ _ → f a) i .pr₁ + +module _ {A B : Type} (w : A ≃ B) where + invEq : B → A + invEq = invIsEq (pr₂ w) + + retEq : retract (w .pr₁) invEq + retEq = retIsEq (pr₂ w) + + secEq : section (w .pr₁) invEq + secEq = secIsEq (pr₂ w) + +-- Isomorphisms +record Iso {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ ⊔ ℓ') where + no-eta-equality + constructor iso + field + fun : A → B + inv : B → A + rightInv : section fun inv + leftInv : retract fun inv + +-- Any iso is an equivalence (called "improve" by Dan, here we use the +-- contractible fibers version of equivalences) +module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (i : Iso A B) where + open Iso i renaming ( fun to f + ; inv to g + ; rightInv to s + ; leftInv to t) + + private + module _ (y : B) (x0 x1 : A) (p0 : f x0 ≡ y) (p1 : f x1 ≡ y) where + fill0 : I → I → A + fill0 i = hfill (λ k → λ { (i = i1) → t x0 k + ; (i = i0) → g y }) + (inS (g (p0 (~ i)))) + + fill1 : I → I → A + fill1 i = hfill (λ k → λ { (i = i1) → t x1 k + ; (i = i0) → g y }) + (inS (g (p1 (~ i)))) + + fill2 : I → I → A + fill2 i = hfill (λ k → λ { (i = i1) → fill1 k i1 + ; (i = i0) → fill0 k i1 }) + (inS (g y)) + + p : x0 ≡ x1 + p i = fill2 i i1 + + sq : I → I → A + sq i j = hcomp (λ k → λ { (i = i1) → fill1 j (~ k) + ; (i = i0) → fill0 j (~ k) + ; (j = i1) → t (fill2 i i1) (~ k) + ; (j = i0) → g y }) + (fill2 i j) + + sq1 : I → I → B + sq1 i j = hcomp (λ k → λ { (i = i1) → s (p1 (~ j)) k + ; (i = i0) → s (p0 (~ j)) k + ; (j = i1) → s (f (p i)) k + ; (j = i0) → s y k }) + (f (sq i j)) + + lemIso : (x0 , p0) ≡ (x1 , p1) + lemIso i .pr₁ = p i + lemIso i .pr₂ = λ j → sq1 i (~ j) + + isoToIsEquiv : isEquiv f + isoToIsEquiv .equiv-proof y .pr₁ .pr₁ = g y + isoToIsEquiv .equiv-proof y .pr₁ .pr₂ = s y + isoToIsEquiv .equiv-proof y .pr₂ z = lemIso y (g y) (pr₁ z) (s y) (pr₂ z) + + +isoToEquiv : {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} → Iso A B → A ≃ B +isoToEquiv i .pr₁ = i .Iso.fun +isoToEquiv i .pr₂ = isoToIsEquiv i + +Glue : (A : Type ℓ) {φ : I} + → (Te : Partial φ (Σ T ꞉ Type ℓ' , T ≃ A)) + → Type ℓ' +Glue A Te = primGlue A (λ x → Te x .pr₁) (λ x → Te x .pr₂) + +idEquiv : ∀ {ℓ} (A : Type ℓ) → A ≃ A +idEquiv A .pr₁ = λ x → x +equiv-proof (idEquiv A .pr₂) = λ y → (y , (λ i → y)) , (λ h i → (h .pr₂ (~ i)) , λ j → h .pr₂ ((j ∨ ~ i))) + +isoToPath : {A B : Type ℓ} → Iso A B → A ≡ B +isoToPath {A = A} {B = B} f i = + Glue B (λ { (i = i0) → (A , isoToEquiv f) + ; (i = i1) → (B , idEquiv B) }) + +-- Natural numbers. We use the built in ones to be able to use +-- numerals. +open import introduction public + using (ℕ; zero; suc; _+_) + renaming (_*_ to _·_) + +_-_ : ℕ → ℕ → ℕ +n - zero = n +zero - suc m = zero +suc n - suc m = n - m + +{-# BUILTIN NATMINUS _-_ #-} + + +-- Integers (slightly different from how Dan did them in order to have +-- one less constructor to pattern match on) +data ℤ : Type₀ where + pos : (n : ℕ) → ℤ + negsuc : (n : ℕ) → ℤ + +sucℤ : ℤ → ℤ +sucℤ (pos n) = pos (suc n) +sucℤ (negsuc zero) = pos zero +sucℤ (negsuc (suc n)) = negsuc n + +predℤ : ℤ → ℤ +predℤ (pos zero) = negsuc zero +predℤ (pos (suc n)) = pos n +predℤ (negsuc n) = negsuc (suc n) + +sucPred : ∀ i → sucℤ (predℤ i) ≡ i +sucPred (pos zero) = λ i → pos zero +sucPred (pos (suc n)) = λ i → pos (suc n) +sucPred (negsuc n) = λ i → negsuc n + +predSuc : ∀ i → predℤ (sucℤ i) ≡ i +predSuc (pos n) = λ i → pos n +predSuc (negsuc zero) = λ i → negsuc zero +predSuc (negsuc (suc n)) = λ i → negsuc (suc n) + +sucPath : ℤ ≡ ℤ +sucPath = isoToPath (iso sucℤ predℤ sucPred predSuc) + +_+ℤ_ : ℤ → ℤ → ℤ +m +ℤ pos n = m +pos n + where + _+pos_ : ℤ → ℕ → ℤ + z +pos 0 = z + z +pos (suc n) = sucℤ (z +pos n) +m +ℤ negsuc n = m +negsuc n + where + _+negsuc_ : ℤ → ℕ → ℤ + z +negsuc 0 = predℤ z + z +negsuc (suc n) = predℤ (z +negsuc n) + + + +-- 'Data' types from Martín's prelude +record Unit : Type where + constructor + ⋆ + +open Unit public + +data Bool : Type where + true false : Bool + +if_then_else_ : {A : Type ℓ} → Bool → A → A → A +if true then x else y = x +if false then x else y = y +``` + + +```agda + + +funExt₂ : {A : Type ℓ} {B : A → Type} {C : (x : A) → B x → I → Type} + {f : (x : A) → (y : B x) → C x y i0} + {g : (x : A) → (y : B x) → C x y i1} + → ((x : A) (y : B x) → PathP (C x y) (f x y) (g x y)) + → PathP (λ i → ∀ x y → C x y i) f g +funExt₂ p i x y = p x y i + +doubleℕ : ℕ → ℕ +doubleℕ zero = zero +doubleℕ (suc x) = suc (suc (doubleℕ x)) + ++-suc : ∀ m n → suc (m + n) ≡ m + suc n ++-suc zero n i = suc n ++-suc (suc m) n i = suc (+-suc m n i) + ++-comm : ∀ m n → m + n ≡ n + m ++-comm zero zero i = 0 ++-comm zero (suc n) i = suc (+-comm zero n i) ++-comm (suc m) zero i = suc (+-comm m zero i) ++-comm (suc m) (suc n) i = + suc (((λ i → (+-suc m n) (~ i)) + ∙ (λ j → suc (+-comm m n j)) + ∙ +-suc n m) i) + ++-zero : ∀ m → m + 0 ≡ m ++-zero zero i = zero ++-zero (suc m) i = suc (+-zero m i) + ++-assoc : ∀ m n o → m + (n + o) ≡ (m + n) + o ++-assoc zero n o i = n + o ++-assoc (suc m) n o i = suc (+-assoc m n o i) + + +data ⊥ : Type₀ where + + +⊥-elim : {A : ⊥ → Type ℓ} → (x : ⊥) → A x +⊥-elim () + +⊥-rec : {A : Type ℓ} → ⊥ → A +⊥-rec () + +¬_ : Type ℓ → Type ℓ +¬ A = A → ⊥ + +``` + + +Binary numbers code: + +```agda + +localrefl : {A : Type} {x : A} → x ≡ x +localrefl {x = x} = λ i → x + +localap : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +localap f p i = f (p i) + +localfunExt : {A B : Type} {f g : A → B} (p : (x : A) → f x ≡ g x) → f ≡ g +localfunExt p i x = p x i + + +data Pos : Type where + pos1 : Pos + x0 : Pos → Pos + x1 : Pos → Pos + +sucPos : Pos → Pos +sucPos pos1 = x0 pos1 +sucPos (x0 ps) = x1 ps +sucPos (x1 ps) = x0 (sucPos ps) + +Pos→ℕ : Pos → ℕ +Pos→ℕ pos1 = suc zero +Pos→ℕ (x0 ps) = doubleℕ (Pos→ℕ ps) +Pos→ℕ (x1 ps) = suc (doubleℕ (Pos→ℕ ps)) + +posInd : {P : Pos → Type} → P pos1 → ((p : Pos) → P p → P (sucPos p)) → (p : Pos) → P p +posInd {P} h1 hs = f + where + H : (p : Pos) → P (x0 p) → P (x0 (sucPos p)) + H p hx0p = hs (x1 p) (hs (x0 p) hx0p) + + f : (ps : Pos) → P ps + f pos1 = h1 + f (x0 ps) = posInd (hs pos1 h1) H ps + f (x1 ps) = hs (x0 ps) (posInd (hs pos1 h1) H ps) + +Pos→ℕsucPos : (p : Pos) → Pos→ℕ (sucPos p) ≡ suc (Pos→ℕ p) +Pos→ℕsucPos pos1 = localrefl +Pos→ℕsucPos (x0 p) = localrefl +Pos→ℕsucPos (x1 p) = λ i → doubleℕ (Pos→ℕsucPos p i) + +caseNat : ∀ {ℓ} → {A : Type ℓ} → (a0 aS : A) → ℕ → A +caseNat a0 aS zero = a0 +caseNat a0 aS (suc n) = aS + +-- zero is not in the image of Pos→ℕ. +znots : {n : ℕ} → ¬ (0 ≡ suc n) +znots eq = transp (λ i → caseNat ℕ ⊥ (eq i)) i0 0 + +zero≠Pos→ℕ : (p : Pos) → ¬ (zero ≡ Pos→ℕ p) +zero≠Pos→ℕ p = posInd (λ prf → znots prf) hs p + where + hs : (p : Pos) → ¬ (zero ≡ Pos→ℕ p) → zero ≡ Pos→ℕ (sucPos p) → ⊥ + hs p neq ieq = ⊥-rec (znots (ieq ∙ Pos→ℕsucPos p)) + +ℕ→Pos : ℕ → Pos +ℕ→Pos zero = pos1 +ℕ→Pos (suc zero) = pos1 +ℕ→Pos (suc (suc n)) = sucPos (ℕ→Pos (suc n)) + +ℕ→PosSuc : ∀ n → ¬ (zero ≡ n) → ℕ→Pos (suc n) ≡ sucPos (ℕ→Pos n) +ℕ→PosSuc zero neq = ⊥-elim (neq localrefl) +ℕ→PosSuc (suc n) neq = localrefl + +Pos→ℕ→Pos : (p : Pos) → ℕ→Pos (Pos→ℕ p) ≡ p +Pos→ℕ→Pos p = posInd localrefl hs p + where + hs : (p : Pos) → ℕ→Pos (Pos→ℕ p) ≡ p → ℕ→Pos (Pos→ℕ (sucPos p)) ≡ sucPos p + hs p hp = + ℕ→Pos (Pos→ℕ (sucPos p)) ≡⟨ localap ℕ→Pos (Pos→ℕsucPos p) ⟩ + ℕ→Pos (suc (Pos→ℕ p)) ≡⟨ ℕ→PosSuc (Pos→ℕ p) (zero≠Pos→ℕ p) ⟩ + sucPos (ℕ→Pos (Pos→ℕ p)) ≡⟨ localap sucPos hp ⟩ + sucPos p ∎ + +ℕ→Pos→ℕ : (n : ℕ) → Pos→ℕ (ℕ→Pos (suc n)) ≡ suc n +ℕ→Pos→ℕ zero = localrefl +ℕ→Pos→ℕ (suc n) = + Pos→ℕ (sucPos (ℕ→Pos (suc n))) ≡⟨ Pos→ℕsucPos (ℕ→Pos (suc n)) ⟩ + suc (Pos→ℕ (ℕ→Pos (suc n))) ≡⟨ localap suc (ℕ→Pos→ℕ n) ⟩ + suc (suc n) ∎ + +-- Binary numbers +data Bin : Type where + bin0 : Bin + binPos : Pos → Bin + +ℕ→Bin : ℕ → Bin +ℕ→Bin zero = bin0 +ℕ→Bin (suc n) = binPos (ℕ→Pos (suc n)) + +Bin→ℕ : Bin → ℕ +Bin→ℕ bin0 = zero +Bin→ℕ (binPos x) = Pos→ℕ x + +ℕ→Bin→ℕ : (n : ℕ) → Bin→ℕ (ℕ→Bin n) ≡ n +ℕ→Bin→ℕ zero = localrefl +ℕ→Bin→ℕ (suc zero) = localrefl +ℕ→Bin→ℕ (suc (suc n)) = + Pos→ℕ (sucPos (ℕ→Pos (suc n))) ≡⟨ Pos→ℕsucPos (ℕ→Pos (suc n)) ⟩ + suc (Pos→ℕ (ℕ→Pos (suc n))) ≡⟨ localap suc (ℕ→Bin→ℕ (suc n)) ⟩ + suc (suc n) ∎ + +Bin→ℕ→Bin : (n : Bin) → ℕ→Bin (Bin→ℕ n) ≡ n +Bin→ℕ→Bin bin0 = localrefl +Bin→ℕ→Bin (binPos p) = posInd localrefl (λ p _ → rem p) p + where + rem : (p : Pos) → ℕ→Bin (Pos→ℕ (sucPos p)) ≡ binPos (sucPos p) + rem p = + ℕ→Bin (Pos→ℕ (sucPos p)) ≡⟨ localap ℕ→Bin (Pos→ℕsucPos p) ⟩ + binPos (ℕ→Pos (suc (Pos→ℕ p))) ≡⟨ localap binPos (ℕ→PosSuc (Pos→ℕ p) (zero≠Pos→ℕ p) ∙ + (localap sucPos (Pos→ℕ→Pos p))) ⟩ + binPos (sucPos p) ∎ +``` diff --git a/src/HoTTEST/Agda/Exercises/01-Exercises.lagda.md b/src/HoTTEST/Agda/Exercises/01-Exercises.lagda.md new file mode 100644 index 0000000..c5fc2a7 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/01-Exercises.lagda.md @@ -0,0 +1,279 @@ +# Week 01 - Agda Exercises + +## Please read before starting the exercises + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**Please make a copy of this file to work in, so that it doesn't get overwritten + (in case we update the exercises through `git`)!** + +```agda +{-# OPTIONS --without-K --allow-unsolved-metas #-} + +module 01-Exercises where + +open import prelude hiding (not-is-involution) +``` + +## Part I: Writing functions on Booleans, natural numbers and lists (★/★★) + +### Exercise 1 (★) + +In the lectures we defined `&&` (logical and) on `Bool` by pattern matching on +the leftmost argument only. + +*NB*: We didn't get round to doing this in the lecture, but see `_&&_` in + [introduction.lagda.md](../Lecture-Notes/files/introduction.lagda.md). + +**Define** the same operation but this time by pattern matching (case splitting) + on both arguments. + +```agda +_&&'_ : Bool → Bool → Bool +a &&' b = {!!} +``` + +One advantage of this definition is that it reads just like a Boolean truth +table. Later on in this exercise sheet, we will see a disadvantange of this more +verbose definition. + +### Exercise 2 (★) + +**Define** `xor` (excluse or) on `Bool`. Exclusive or is true if and only if +*exactly one* of its arguments is true. + +```agda +_xor_ : Bool → Bool → Bool +a xor b = {!!} +``` + +### Exercise 3 (★) + +**Define** the exponential and factorial functions on natural numbers. + +If you do things correctly, then the examples should compute correctly, i.e. the +proof that 3 ^ 4 ≡ 81 should simply be given by `refl _` which says that the +left hand side and the right hand side compute to the same value. + +```agda +_^_ : ℕ → ℕ → ℕ +n ^ m = {!!} + +^-example : 3 ^ 4 ≡ 81 +^-example = {!!} -- refl 81 should fill the hole here + +_! : ℕ → ℕ +n ! = {!!} + +!-example : 4 ! ≡ 24 +!-example = {!!} -- refl 24 should fill the hole here +``` + +### Exercise 4 (★) + +We can recursively compute the maximum of two natural numbers as follows. +```agda +max : ℕ → ℕ → ℕ +max zero m = m +max (suc n) zero = suc n +max (suc n) (suc m) = suc (max n m) +``` + +**Define** the minimum of two natural numbers analogously. + +```agda +min : ℕ → ℕ → ℕ +min = {!!} + +min-example : min 5 3 ≡ 3 +min-example = {!!} -- refl 3 should fill the hole here +``` + +### Exercise 5 (★) + +Use pattern matching on lists to **define** `map`. + +This function should behave as follows: +`map f [x₁ , x₂ , ... , xₙ] = [f x₁ , f x₂ , ... , f xₙ]`. +That is, `map f xs` applies the given function `f` to every +element of the list `xs` and returns the resulting list. + +```agda +map : {X Y : Type} → (X → Y) → List X → List Y +map f xs = {!!} + +map-example : map (_+ 3) (1 :: 2 :: 3 :: []) ≡ 4 :: 5 :: 6 :: [] +map-example = {!!} -- refl _ should fill the hole here + + -- We write the underscore, because we don't wish to repeat + -- the relatively long "4 :: 5 :: 6 :: []" and Agda can + -- figure out what is supposed to go there. +``` + +### Exercise 6 (★★) + +**Define** a function `filter` that takes predicate `p : X → Bool` and a list + `xs` that returns the list of elements of `xs` for which `p` is true. + +For example, filtering the non-zero elements of the list [4 , 3 , 0 , 1 , 0] +should return [4 , 3 , 1], see the code below. + +```agda +filter : {X : Type} (p : X → Bool) → List X → List X +filter = {!!} + +is-non-zero : ℕ → Bool +is-non-zero zero = false +is-non-zero (suc _) = true + +filter-example : filter is-non-zero (4 :: 3 :: 0 :: 1 :: 0 :: []) ≡ 4 :: 3 :: 1 :: [] +filter-example = {!!} -- refl _ should fill the hole here +``` + +## Part II: The identity type of the Booleans (★/★★) + +In the lectures we saw a definition of `≣` on natural numbers where the idea was +that `x ≣ y` is a type which either has precisely one element, if `x` and `y` +are the same natural number, or else is empty, if `x` and `y` are different. + +### Exercise 1 (★) + +**Define** `≣` for Booleans this time. + +```agda +_≣_ : Bool → Bool → Type +a ≣ b = {!!} +``` + +### Exercise 2 (★) + +**Show** that for every Boolean `b` we can find an element of the type `b ≣ b`. + +```agda +Bool-refl : (b : Bool) → b ≣ b +Bool-refl = {!!} +``` + +### Exercise 3 (★★) + +Just like we did in the lectures for natural numbers, **show** that we can go +back and forth between `a ≣ b` and `a ≡ b`. + +*NB*: Again, we didn't have time to do this in the lectures, but see + [introduction.lagda.md](../Lecture-Notes/files/introduction.lagda.md), + specifically the functions `back` and `forth` there. + +```agda +≡-to-≣ : (a b : Bool) → a ≡ b → a ≣ b +≡-to-≣ = {!!} + +≣-to-≡ : (a b : Bool) → a ≣ b → a ≡ b +≣-to-≡ = {!!} +``` + +## Part III: Proving in Agda (★★/★★★) + +We now turn to *proving* things in Agda: one of its key features. + +For example, here is a proof that `not (not b) ≡ b` for every Boolean `b`. + +```agda +not-is-involution : (b : Bool) → not (not b) ≡ b +not-is-involution true = refl true +not-is-involution false = refl false +``` + +### Exercise 1 (★★) + +Use pattern matching to **prove** that `||` is commutative. + +```agda +||-is-commutative : (a b : Bool) → a || b ≡ b || a +||-is-commutative a b = {!!} +``` + +### Exercise 2 (★★) + +Taking inspiration from the above, try to **state** and **prove** that `&&` is +commutative. + +```agda +&&-is-commutative : {!!} +&&-is-commutative = {!!} +``` + +### Exercise 3 (★★) + +**Prove** that `&&` and `&&'` are both associative. + +```agda +&&-is-associative : (a b c : Bool) → a && (b && c) ≡ (a && b) && c +&&-is-associative = {!!} + +&&'-is-associative : (a b c : Bool) → a &&' (b &&' c) ≡ (a &&' b) &&' c +&&'-is-associative = {!!} +``` + +**Question**: Can you spot a downside of the more verbose definition of `&&'` + now? + +### Exercise 4 (★★★) + +Another key feature of Agda is its ability to carry out inductive proofs. For +example, here is a commented inductive proof that `max` is commutative. + +```agda +max-is-commutative : (n m : ℕ) → max n m ≡ max m n +max-is-commutative zero zero = refl zero +max-is-commutative zero (suc m) = refl (suc m) +max-is-commutative (suc n) zero = refl (suc n) +max-is-commutative (suc n) (suc m) = to-show + where + IH : max n m ≡ max m n -- induction hypothesis + IH = max-is-commutative n m -- recursive call on smaller arguments + to-show : suc (max n m) ≡ suc (max m n) + to-show = ap suc IH -- ap(ply) suc on both sides of the equation +``` + +**Prove** analogously that `min` is commutative. + +```agda +min-is-commutative : (n m : ℕ) → min n m ≡ min m n +min-is-commutative = {!!} +``` + +### Exercise 5 (★★★) + +Using another inductive proof, **show** that `n ≡ n + 0` for every natural +number `n`. + +```agda +0-right-neutral : (n : ℕ) → n ≡ n + 0 +0-right-neutral = {!!} +``` + +### Exercise 6 (★★★) + +The function `map` on lists is a so-called *functor*, which means that it +respects the identity and composition, as formally expressed below. + +Try to **prove** these equations using pattern matching and inductive proofs. + +```agda +map-id : {X : Type} (xs : List X) → map id xs ≡ xs +map-id xs = {!!} + +map-comp : {X Y Z : Type} (f : X → Y) (g : Y → Z) + (xs : List X) → map (g ∘ f) xs ≡ map g (map f xs) +map-comp f g xs = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/01-Solutions.lagda.md b/src/HoTTEST/Agda/Exercises/01-Solutions.lagda.md new file mode 100644 index 0000000..fac047d --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/01-Solutions.lagda.md @@ -0,0 +1,182 @@ +# Week 01 - Agda Exercises - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module 01-Solutions where + +open import prelude hiding (not-is-involution) +``` + +## Part I + +```agda +_&&'_ : Bool → Bool → Bool +true &&' true = true +true &&' false = false +false &&' true = false +false &&' false = false + +_xor_ : Bool → Bool → Bool +true xor true = false +true xor false = true +false xor true = true +false xor false = false + +_^_ : ℕ → ℕ → ℕ +n ^ zero = 1 +n ^ suc m = n * (n ^ m) + +^-example : 3 ^ 4 ≡ 81 +^-example = refl 81 + +_! : ℕ → ℕ +zero ! = 1 +suc n ! = suc n * (n !) + +!-example : 4 ! ≡ 24 +!-example = refl 24 + +max : ℕ → ℕ → ℕ +max zero m = m +max (suc n) zero = suc n +max (suc n) (suc m) = suc (max n m) + +min : ℕ → ℕ → ℕ +min zero m = zero +min (suc n) zero = zero +min (suc n) (suc m) = suc (min n m) + +min-example : min 5 3 ≡ 3 +min-example = refl 3 + +map : {X Y : Type} → (X → Y) → List X → List Y +map f [] = [] +map f (x :: xs) = f x :: map f xs + +map-example : map (_+ 3) (1 :: 2 :: 3 :: []) ≡ 4 :: 5 :: 6 :: [] +map-example = refl _ + +filter : {X : Type} (p : X → Bool) → List X → List X +filter {X} p [] = [] +filter {X} p (x :: xs) = if (p x) then (x :: ys) else ys + where + ys : List X + ys = filter p xs + +is-non-zero : ℕ → Bool +is-non-zero zero = false +is-non-zero (suc _) = true + +filter-example : filter is-non-zero (4 :: 3 :: 0 :: 1 :: 0 :: []) ≡ 4 :: 3 :: 1 :: [] +filter-example = refl _ +``` + +## Part II + +```agda +_≣_ : Bool → Bool → Type +true ≣ true = 𝟙 +true ≣ false = 𝟘 +false ≣ true = 𝟘 +false ≣ false = 𝟙 + +Bool-refl : (b : Bool) → b ≣ b +Bool-refl true = ⋆ +Bool-refl false = ⋆ + +≡-to-≣ : (a b : Bool) → a ≡ b → a ≣ b +≡-to-≣ b b (refl b) = Bool-refl b + +≣-to-≡ : (a b : Bool) → a ≣ b → a ≡ b +≣-to-≡ true true ⋆ = refl true +≣-to-≡ false false ⋆ = refl false +``` + +## Part III + +```agda +not-is-involution : (b : Bool) → not (not b) ≡ b +not-is-involution true = refl true +not-is-involution false = refl false + +||-is-commutative : (a b : Bool) → a || b ≡ b || a +||-is-commutative true true = refl true +||-is-commutative true false = refl true +||-is-commutative false true = refl true +||-is-commutative false false = refl false + +&&-is-commutative : (a b : Bool) → a && b ≡ b && a +&&-is-commutative true true = refl true +&&-is-commutative true false = refl false +&&-is-commutative false true = refl false +&&-is-commutative false false = refl false + +-- Notice how we choose to pattern match on the leftmost argument, because +-- that's also how we defined &&. +&&-is-associative : (a b c : Bool) → a && (b && c) ≡ (a && b) && c +&&-is-associative true b c = refl (b && c) +&&-is-associative false b c = refl false + +-- Because &&' is defined by pattern matching on all its arguments, we now need +-- to consider 2^3 = 8 cases. +&&'-is-associative : (a b c : Bool) → a &&' (b &&' c) ≡ (a &&' b) &&' c +&&'-is-associative true true true = refl true +&&'-is-associative true true false = refl false +&&'-is-associative true false true = refl false +&&'-is-associative true false false = refl false +&&'-is-associative false true true = refl false +&&'-is-associative false true false = refl false +&&'-is-associative false false true = refl false +&&'-is-associative false false false = refl false + +max-is-commutative : (n m : ℕ) → max n m ≡ max m n +max-is-commutative zero zero = refl zero +max-is-commutative zero (suc m) = refl (suc m) +max-is-commutative (suc n) zero = refl (suc n) +max-is-commutative (suc n) (suc m) = to-show + where + IH : max n m ≡ max m n -- induction hypothesis + IH = max-is-commutative n m -- recursive call on smaller arguments + to-show : suc (max n m) ≡ suc (max m n) + to-show = ap suc IH -- ap(ply) suc on both sides of the equation + +min-is-commutative : (n m : ℕ) → min n m ≡ min m n +min-is-commutative zero zero = refl zero +min-is-commutative zero (suc m) = refl zero +min-is-commutative (suc n) zero = refl zero +min-is-commutative (suc n) (suc m) = to-show + where + IH : min n m ≡ min m n -- induction hypothesis + IH = min-is-commutative n m -- recursive call on smaller arguments + to-show : suc (min n m) ≡ suc (min m n) + to-show = ap suc IH -- ap(ply) suc on both sides of the equation + +0-right-neutral : (n : ℕ) → n ≡ n + 0 +0-right-neutral zero = refl zero +0-right-neutral (suc n) = to-show + where + IH : n ≡ n + 0 -- induction hypothesis + IH = 0-right-neutral n -- recursive call on smaller argument + to-show : suc n ≡ suc (n + 0) + to-show = ap suc IH -- ap(ply) suc on both sides of the equation + +map-id : {X : Type} (xs : List X) → map id xs ≡ xs +map-id [] = refl [] +map-id (x :: xs) = to-show + where + IH : map id xs ≡ xs + IH = map-id xs + to-show : x :: map id xs ≡ x :: xs + to-show = ap (x ::_) IH + +map-comp : {X Y Z : Type} (f : X → Y) (g : Y → Z) + (xs : List X) → map (g ∘ f) xs ≡ map g (map f xs) +map-comp f g [] = refl [] +map-comp f g (x :: xs) = to-show + where + IH : map (g ∘ f) xs ≡ map g (map f xs) + IH = map-comp f g xs + to-show : g (f x) :: map (g ∘ f) xs ≡ g (f x) :: map g (map f xs) + to-show = ap (g (f x) ::_) IH +``` diff --git a/src/HoTTEST/Agda/Exercises/02-Exercises.lagda.md b/src/HoTTEST/Agda/Exercises/02-Exercises.lagda.md new file mode 100644 index 0000000..1303126 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/02-Exercises.lagda.md @@ -0,0 +1,182 @@ +# Week 02 - Agda Exercises + +## Please read before starting the exercises + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**Please make a copy of this file to work in, so that it doesn't get overwritten + (in case we update the exercises through `git`)!** + +```agda +{-# OPTIONS --without-K --allow-unsolved-metas #-} + +module 02-Exercises where + +open import prelude +open import decidability +open import sums +``` + +## Part I: Propositions as types + + +### Exercise 1 (★) + +Prove +```agda +uncurry : {A B X : Type} → (A → B → X) → (A × B → X) +uncurry = {!!} + +curry : {A B X : Type} → (A × B → X) → (A → B → X) +curry = {!!} +``` +You might know these functions from programming e.g. in Haskell. +But what do they say under the propositions-as-types interpretation? + + +### Exercise 2 (★) + +Consider the following goals: +```agda +[i] : {A B C : Type} → (A × B) ∔ C → (A ∔ C) × (B ∔ C) +[i] = {!!} + +[ii] : {A B C : Type} → (A ∔ B) × C → (A × C) ∔ (B × C) +[ii] = {!!} + +[iii] : {A B : Type} → ¬ (A ∔ B) → ¬ A × ¬ B +[iii] = {!!} + +[iv] : {A B : Type} → ¬ (A × B) → ¬ A ∔ ¬ B +[iv] = {!!} + +[v] : {A B : Type} → (A → B) → ¬ B → ¬ A +[v] = {!!} + +[vi] : {A B : Type} → (¬ A → ¬ B) → B → A +[vi] = {!!} + +[vii] : {A B : Type} → ((A → B) → A) → A +[vii] = {!!} + +[viii] : {A : Type} {B : A → Type} + → ¬ (Σ a ꞉ A , B a) → (a : A) → ¬ B a +[viii] = {!!} + +[ix] : {A : Type} {B : A → Type} + → ¬ ((a : A) → B a) → (Σ a ꞉ A , ¬ B a) +[ix] = {!!} + +[x] : {A B : Type} {C : A → B → Type} + → ((a : A) → (Σ b ꞉ B , C a b)) + → Σ f ꞉ (A → B) , ((a : A) → C a (f a)) +[x] = {!!} +``` +For each goal determine whether it is provable or not. +If it is, fill it. If not, explain why it shouldn't be possible. +Propositions-as-types might help. + + +### Exercise 3 (★★) + +```agda +¬¬_ : Type → Type +¬¬ A = ¬ (¬ A) + +¬¬¬ : Type → Type +¬¬¬ A = ¬ (¬¬ A) +``` +In the lecture we have discussed that we can't prove `∀ {A : Type} → ¬¬ A → A`. +What you can prove however, is +```agda +tne : ∀ {A : Type} → ¬¬¬ A → ¬ A +tne = {!!} +``` + + +### Exercise 4 (★★★) +Prove +```agda +¬¬-functor : {A B : Type} → (A → B) → ¬¬ A → ¬¬ B +¬¬-functor = {!!} + +¬¬-kleisli : {A B : Type} → (A → ¬¬ B) → ¬¬ A → ¬¬ B +¬¬-kleisli = {!!} +``` +Hint: For the second goal use `tne` from the previous exercise + + + + + +## Part II: `_≡_` for `Bool` + +**In this exercise we want to investigate what equality of booleans looks like. +In particular we want to show that for `true false : Bool` we have `true ≢ false`.** + +### Exercise 1 (★) + +Under the propositions-as-types paradigm, an inhabited type corresponds +to a true proposition while an uninhabited type corresponds to a false proposition. +With this in mind construct a family +```agda +bool-as-type : Bool → Type +bool-as-type = {!!} +``` +such that `bool-as-type true` corresponds to "true" and +`bool-as-type false` corresponds to "false". (Hint: +we have seen canonical types corresponding true and false in the lectures) + + +### Exercise 2 (★★) + +Prove +```agda +bool-≡-char₁ : ∀ (b b' : Bool) → b ≡ b' → (bool-as-type b ⇔ bool-as-type b') +bool-≡-char₁ = {!!} +``` + + +### Exercise 3 (★★) + +Using ex. 2, conclude that +```agda +true≢false : ¬ (true ≡ false) +true≢false () +``` +You can actually prove this much easier! How? + + +### Exercise 4 (★★★) + +Finish our characterisation of `_≡_` by proving +```agda +bool-≡-char₂ : ∀ (b b' : Bool) → (bool-as-type b ⇔ bool-as-type b') → b ≡ b' +bool-≡-char₂ = {!!} +``` + + +## Part III (🌶) +A type `A` is called *discrete* if it has decidable equality. +Consider the following predicate on types: +```agda +has-bool-dec-fct : Type → Type +has-bool-dec-fct A = Σ f ꞉ (A → A → Bool) , (∀ x y → x ≡ y ⇔ (f x y) ≡ true) +``` + +Prove that + +```agda +decidable-equality-char : (A : Type) → has-decidable-equality A ⇔ has-bool-dec-fct A +decidable-equality-char = ? +``` diff --git a/src/HoTTEST/Agda/Exercises/02-Solutions.lagda.md b/src/HoTTEST/Agda/Exercises/02-Solutions.lagda.md new file mode 100644 index 0000000..694b74d --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/02-Solutions.lagda.md @@ -0,0 +1,166 @@ +# Week 02 - Agda Exercises - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module 02-Solutions where + +open import prelude +open import decidability +open import sums + +``` + +## Part I + +### Exercise 1 +```agda +uncurry : {A B X : Type} → (A → B → X) → (A × B → X) +uncurry f (a , b) = f a b + +curry : {A B X : Type} → (A × B → X) → (A → B → X) +curry f a b = f (a , b) +``` +Under the propositions-as-types interpretation curry and uncurry +tell us that "if A then if B then X" is (logically) equivalent +to "if A and B then X" + +### Exercise 2 +```agda +[i] : {A B C : Type} → (A × B) ∔ C → (A ∔ C) × (B ∔ C) +[i] (inl (a , b)) = inl a , inl b +[i] (inr c) = inr c , inr c + +[ii] : {A B C : Type} → (A ∔ B) × C → (A × C) ∔ (B × C) +[ii] (inl a , c) = inl (a , c) +[ii] (inr b , c) = inr (b , c) + +[iii] : {A B : Type} → ¬ (A ∔ B) → ¬ A × ¬ B +pr₁ ([iii] f) a = f (inl a) +pr₂ ([iii] f) b = f (inr b) +``` + +The next goal `[iv] : {A B : Type} → ¬ (A × B) → ¬ A ∔ ¬ B` +is not provable. Under propositions-as-types it says that +"not (A and B) implies not A or not B", which is not true +in constructive logic. At the end we have to give a term of +the form `inl ...` or `inr ...` but for abstract `A B` we +can not say of which form it should be. +```agda +[v] : {A B : Type} → (A → B) → ¬ B → ¬ A -- also known as contraposition +[v] f g a = g (f a) +``` + +Neither of `[vi] : {A B : Type} → (¬ A → ¬ B) → B → A` +and `[vii] : {A B : Type} → ((A → B) → A) → A` are provable +Under propositions-as-types `[vi]` is known as *inverse contraposition* +and `[vii]` is known as *Peirce's law*. At the end we have to construct +something of type `A` but this is not possible with all the assumptions +being functions. +```agda +[viii] : {A : Type} {B : A → Type} + → ¬ (Σ a ꞉ A , B a) → (a : A) → ¬ B a +[viii] f a bₐ = f (a , bₐ) +``` +The next goal +`[ix] : {A : Type} {B : A → Type} → ¬ ((a : A) → B a) → (Σ a ꞉ A , ¬ B a)` +reads as: "If not for all a, B(a), then there is an a such that not B(a)" +This is not true in constructive logic. Again, we have to construct +an `a : A` as the first projection of the Sigma-type in the conclusion, +which is not possible from our assumptions. + +```agda +[x] : {A B : Type} {C : A → B → Type} + → ((a : A) → (Σ b ꞉ B , C a b)) + → Σ f ꞉ (A → B) , ((a : A) → C a (f a)) +pr₁ ([x] g) a = g a .pr₁ +pr₂ ([x] g) a = g a .pr₂ +``` +Note that under propositions-as-types `[x]` reads somewhat like the +*axiom of choice*. Yet it is still provable. This result is often +referred to as the *distributivity of Π over Σ* and shows that +propositions-as-types should be taken with a grain of salt sometimes. + +### Exercise 3 +```agda +¬¬_ : Type → Type +¬¬ A = ¬ (¬ A) + +¬¬¬ : Type → Type +¬¬¬ A = ¬ (¬¬ A) + +tne : ∀ {A : Type} → ¬¬¬ A → ¬ A +tne f a = f (λ g → g a) +``` + +### Exercise 4 +```agda +¬¬-functor : {A B : Type} → (A → B) → ¬¬ A → ¬¬ B +¬¬-functor f = [v] ([v] f) + +¬¬-kleisli : {A B : Type} → (A → ¬¬ B) → ¬¬ A → ¬¬ B +¬¬-kleisli f g = tne (¬¬-functor f g) +``` + + + + +## Part II + +```agda +bool-as-type : Bool → Type +bool-as-type true = 𝟙 +bool-as-type false = 𝟘 + +bool-≡-char₁ : ∀ (b b' : Bool) → b ≡ b' → (bool-as-type b ⇔ bool-as-type b') +bool-≡-char₁ b .b (refl .b) = ⇔-refl + where + ⇔-refl : {X : Type} → X ⇔ X + pr₁ ⇔-refl x = x + pr₂ ⇔-refl x = x + +true≢false : ¬ (true ≡ false) +true≢false p = bool-≡-char₁ true false p .pr₁ ⋆ +-- also true≢false () + +bool-≡-char₂ : ∀ (b b' : Bool) → (bool-as-type b ⇔ bool-as-type b') → b ≡ b' +bool-≡-char₂ true true (b→b' , b'→b) = refl true +bool-≡-char₂ true false (b→b' , b'→b) = 𝟘-elim (b→b' ⋆) +bool-≡-char₂ false true (b→b' , b'→b) = 𝟘-elim (b'→b ⋆) +bool-≡-char₂ false false (b→b' , b'→b) = refl false +``` + + + + +## Part III + +```agda +has-bool-dec-fct : Type → Type +has-bool-dec-fct A = Σ f ꞉ (A → A → Bool) , (∀ x y → x ≡ y ⇔ (f x y) ≡ true) + +decidable-equality-char : (A : Type) → has-decidable-equality A ⇔ has-bool-dec-fct A +pr₁ (decidable-equality-char A) discA = f , f-dec -- left to right implication + where + f' : (a b : A) → is-decidable (a ≡ b) → Bool + f' a b (inl _) = true + f' a b (inr _) = false + + f'-refl : (x : A) (d : is-decidable (x ≡ x)) → f' x x d ≡ true + f'-refl x (inl _) = refl true + f'-refl x (inr x≢x) = 𝟘-nondep-elim (x≢x (refl x)) + + f : A → A → Bool + f a b = f' a b (discA a b) + + f-dec : ∀ x y → x ≡ y ⇔ (f x y) ≡ true + pr₁ (f-dec x .x) (refl .x) = f'-refl x (discA x x) + pr₂ (f-dec x y) with discA x y + ... | (inl p) = λ _ → p + ... | (inr _) = λ q → 𝟘-nondep-elim (true≢false (q ⁻¹)) + +pr₂ (decidable-equality-char A) (f , f-dec) x y -- right to left implication + with Bool-has-decidable-equality (f x y) true +... | (inl p) = inl (f-dec x y .pr₂ p) +... | (inr g) = inr λ p → g (f-dec x y .pr₁ p) +``` diff --git a/src/HoTTEST/Agda/Exercises/03-Exercises.lagda.md b/src/HoTTEST/Agda/Exercises/03-Exercises.lagda.md new file mode 100644 index 0000000..61f9055 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/03-Exercises.lagda.md @@ -0,0 +1,349 @@ +# Week 03 - Agda Exercises + +## Please read before starting the exercises + +**The exercises are designed to increase in difficulty so that we can cater to +our large and diverse audience. This also means that it is *perfectly fine* if +you don't manage to do all exercises: some of them are definitely a bit hard for +beginners and there are likely too many exercises! You *may* wish to come back +to them later when you have learned more.** + +Having said that, here we go! + +This is a markdown file with Agda code, which means that it displays nicely on +GitHub, but at the same time you can load this file in Agda and fill the holes +to solve exercises. + +**Please make a copy of this file to work in, so that it doesn't get overwritten + (in case we update the exercises through `git`)!** + +```agda +{-# OPTIONS --without-K --allow-unsolved-metas #-} + +module 03-Exercises where + +open import prelude hiding (_∼_) +``` + +## Part I -- Homotopies + +It is often convenient to work with *pointwise equality* of functions, defined as follows. +```agda +module _ {A : Type} {B : A → Type} where + _∼_ : ((x : A) → B x) → ((x : A) → B x) → Type + f ∼ g = ∀ x → f x ≡ g x +``` +An element of `f ∼ g` is usually called a homotopy. + +### Exercise 1 (⋆⋆) + +Unsurprisingly, many properties of this type of pointwise equalities +can be inferred directly from the same operations on paths. + +Try to prove reflexivity, symmetry and transitivity of `_∼_` by filling these holes. +```agda + ∼-refl : (f : (x : A) → B x) → f ∼ f + ∼-refl f = {!!} + + ∼-inv : (f g : (x : A) → B x) → (f ∼ g) → (g ∼ f) + ∼-inv f g H x = {!!} + + ∼-concat : (f g h : (x : A) → B x) → f ∼ g → g ∼ h → f ∼ h + ∼-concat f g h H K x = {!!} + + infix 0 _∼_ +``` + +## Part II -- Isomorphisms + +A function `f : A → B` is called a *bijection* if there is a function `g : B → A` in the opposite direction such that `g ∘ f ∼ id` and `f ∘ g ∼ id`. Recall that `_∼_` is [pointwise equality](../Lecture-Notes/files/identity-type.lagda.md) and that `id` is the [identity function](../Lecture-Notes/files/products.lagda.md). This means that we can convert back and forth between the types `A` and `B` landing at the same element we started with, either from `A` or from `B`. In this case, we say that the types `A` and `B` are *isomorphic*, and we write `A ≅ B`. Bijections are also called type *isomorphisms*. We can define these concepts in Agda using [sum types](../Lecture-Notes/files/sums.lagda.md) or [records](https://agda.readthedocs.io/en/latest/language/record-types.html). We will adopt the latter, but we include both definitions for the sake of illustration. +Recall that we [defined](../Lecture-Notes/files/general-notation.lagda.md) the domain of a function `f : A → B` to be `A` and its codomain to be `B`. + +We adopt this definition of isomorphisms using records. +```agda +record is-bijection {A B : Type} (f : A → B) : Type where + constructor + Inverse + field + inverse : B → A + η : inverse ∘ f ∼ id + ε : f ∘ inverse ∼ id + +record _≅_ (A B : Type) : Type where + constructor + Isomorphism + field + bijection : A → B + bijectivity : is-bijection bijection + +infix 0 _≅_ +``` + +### Exercise 2 (⋆) + +Reformulate the same definition using Sigma-types. +```agda +is-bijection' : {A B : Type} → (A → B) → Type +is-bijection' f = {!!} + +_≅'_ : Type → Type → Type +A ≅' B = {!!} +``` +The definition with `Σ` is probably more intuitive, but, as discussed above, +the definition with a record is often easier to work with, +because we can easily extract the components of the definitions using the names of the fields. +It also often allows Agda to infer more types, and to give us more sensible goals in the +interactive development of Agda programs and proofs. + +Notice that `inverse` plays the role of `g`. +The reason we use `inverse` is that then we can use the word +`inverse` to extract the inverse of a bijection. +Similarly we use `bijection` for `f`, as we can use the word +`bijection` to extract the bijection from a record. + +This type can be defined to be `𝟙 ∔ 𝟙` using the coproduct, +but we give a direct definition which will allow us to discuss some relationships between the various type formers of Basic MLTT. + +```agda +data 𝟚 : Type where + 𝟎 𝟏 : 𝟚 +``` + +### Exercise 3 (⋆⋆) + +Prove that 𝟚 and Bool are isomorphic + +```agda +Bool-𝟚-isomorphism : Bool ≅ 𝟚 +Bool-𝟚-isomorphism = record { bijection = {!!} ; bijectivity = {!!} } + where + f : Bool → 𝟚 + f false = {!!} + f true = {!!} + + g : 𝟚 → Bool + g 𝟎 = {!!} + g 𝟏 = {!!} + + gf : g ∘ f ∼ id + gf true = {!!} + gf false = {!!} + + fg : f ∘ g ∼ id + fg 𝟎 = {!!} + fg 𝟏 = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = {!!} ; η = {!!} ; ε = {!!} } +``` + + +## Part III - Finite Types + +In the last HoTT Exercises we encountered two definitions of the finite types. +Here we prove that they are isomorphic. +Note that `zero` was called `pt` and suc `i` on the HoTT exercise sheet. + +```agda +data Fin : ℕ → Type where + zero : {n : ℕ} → Fin (suc n) + suc : {n : ℕ} → Fin n → Fin (suc n) +``` + +### Exercise 4 (⋆) + +Prove the elimination principle of `Fin`. +```agda +Fin-elim : (A : {n : ℕ} → Fin n → Type) + → ({n : ℕ} → A {suc n} zero) + → ({n : ℕ} (k : Fin n) → A k → A (suc k)) + → {n : ℕ} (k : Fin n) → A k +Fin-elim A a f = h + where + h : {n : ℕ} (k : Fin n) → A k + h zero = {!!} + h (suc k) = {!!} +``` + +We give the other definition of the finite types and introduce some notation. + +```agda +Fin' : ℕ → Type +Fin' 0 = 𝟘 +Fin' (suc n) = 𝟙 ∔ Fin' n + +zero' : {n : ℕ} → Fin' (suc n) +zero' = inl ⋆ + +suc' : {n : ℕ} → Fin' n → Fin' (suc n) +suc' = inr +``` + +### Exercise 5 (⋆⋆⋆) + +Prove that `Fin n` and `Fin' n` are isomorphic for every `n`. + +```agda +Fin-isomorphism : (n : ℕ) → Fin n ≅ Fin' n +Fin-isomorphism n = record { bijection = f n ; bijectivity = f-is-bijection n } + where + f : (n : ℕ) → Fin n → Fin' n + f (suc n) zero = {!!} + f (suc n) (suc k) = {!!} + + g : (n : ℕ) → Fin' n → Fin n + g (suc n) (inl ⋆) = {!!} + g (suc n) (inr k) = {!!} + + gf : (n : ℕ) → g n ∘ f n ∼ id + gf (suc n) zero = {!!} + gf (suc n) (suc k) = γ + where + IH : g n (f n k) ≡ k + IH = gf n k + + γ = g (suc n) (f (suc n) (suc k)) ≡⟨ {!!} ⟩ + g (suc n) (suc' (f n k)) ≡⟨ {!!} ⟩ + suc (g n (f n k)) ≡⟨ {!!} ⟩ + suc k ∎ + + fg : (n : ℕ) → f n ∘ g n ∼ id + fg (suc n) (inl ⋆) = {!!} + fg (suc n) (inr k) = γ + where + IH : f n (g n k) ≡ k + IH = fg n k + + γ = f (suc n) (g (suc n) (suc' k)) ≡⟨ {!!} ⟩ + f (suc n) (suc (g n k)) ≡⟨ {!!} ⟩ + suc' (f n (g n k)) ≡⟨ {!!} ⟩ + suc' k ∎ + + f-is-bijection : (n : ℕ) → is-bijection (f n) + f-is-bijection n = record { inverse = g n ; η = gf n ; ε = fg n} +``` + +## Part IV -- minimal elements in the natural numbers + +In this section we establish some definitions which are needed to state and prove the well-ordering +principle of the natural numbers. + +### Exercise 6 (⋆) + +Give the recursive definition of the less than or equals relation on the natural numbers. + +```agda +_≤₁_ : ℕ → ℕ → Type +0 ≤₁ y = {!!} +suc x ≤₁ 0 = {!!} +suc x ≤₁ suc y = {!!} +``` + +### Exercise 7 (⋆) + +Given a type family `P` over the naturals, a lower bound `n` is a natural number such that +for all other naturals `m`, we have that if `P(m)` holds, then`n ≤₁ m`. +Translate this definition into HoTT. + +```agda +is-lower-bound : (P : ℕ → Type) (n : ℕ) → Type +is-lower-bound P n = {!!} +``` + +We define the type of minimal elements of a type family over the naturals. +```agda +minimal-element : (P : ℕ → Type) → Type +minimal-element P = Σ n ꞉ ℕ , (P n) × (is-lower-bound P n) +``` + +### Exercise 8 (⋆) + +Prove that all numbers are at least as large as zero. +```agda +leq-zero : (n : ℕ) → 0 ≤₁ n +leq-zero n = {!!} +``` + + +## Part V -- The well-ordering principle of ℕ + +Classically, the well-ordering principle states that every non-empty set +of natural numbers has a least element. + +In HoTT, such subsets can be translated as decidable type family. +Recall the definition of decidability: +```agda +open import decidability + using (is-decidable; is-decidable-predicate) +``` + +The well-ordering principle reads +```agda +Well-ordering-principle = (P : ℕ → Type) → (d : is-decidable-predicate P) → (n : ℕ) → P n → minimal-element P +``` + +We shall prove this statement via induction on `n`. +In order to make the proof more readable, we first prove two lemmas. + +### Exercise 9 (🌶) + +What is the statement of `is-minimal-element-suc` +under the Curry-Howard interpretation? +Prove this lemma. + +```agda +is-minimal-element-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (m : ℕ) (pm : P (suc m)) + (is-lower-bound-m : is-lower-bound (λ x → P (suc x)) m) → + ¬ (P 0) → is-lower-bound P (suc m) +is-minimal-element-suc P d m pm is-lower-bound neg-p0 = {! !} +``` + +### Exercise 10 (🌶) + +What is the statement of `well-ordering-principle-suc` +under the Curry-Howard interpretation? +Prove this lemma. + +```agda +well-ordering-principle-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (n : ℕ) (p : P (suc n)) → + is-decidable (P 0) → + minimal-element (λ m → P (suc m)) → minimal-element P +well-ordering-principle-suc P d n p (inl p0) _ = {!!} +well-ordering-principle-suc P d n p (inr neg-p0) (m , (pm , is-min-m)) = {!!} +``` + +### Exercise 11 (🌶) + +Use the previous two lemmas to prove the well-ordering principle +```agda +well-ordering-principle : (P : ℕ → Type) → (d : is-decidable-predicate P) → (n : ℕ) → P n → minimal-element P +well-ordering-principle P d 0 p = {!!} +well-ordering-principle P d (suc n) p = well-ordering-principle-suc P d n p (d 0) {!!} +``` + +### Exercise 12 (🌶) + +Prove that the well-ordering principle returns 0 if `P 0` holds. + +```agda +is-zero-well-ordering-principle-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (n : ℕ) (p : P (suc n)) (d0 : is-decidable (P 0)) → + (x : minimal-element (λ m → P (suc m))) (p0 : P 0) → + (pr₁ (well-ordering-principle-suc P d n p d0 x)) ≡ 0 +is-zero-well-ordering-principle-suc P d n p (inl p0) x q0 = {!!} +is-zero-well-ordering-principle-suc P d n p (inr np0) x q0 = {!!} + +is-zero-well-ordering-principle : + (P : ℕ → Type) (d : is-decidable-predicate P) → + (n : ℕ) → (pn : P n) → + P 0 → + pr₁ (well-ordering-principle P d n pn) ≡ 0 +is-zero-well-ordering-principle P d zero p p0 = {! !} +is-zero-well-ordering-principle P d (suc m) pm = + is-zero-well-ordering-principle-suc P d m pm (d 0) {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/03-Solutions.lagda.md b/src/HoTTEST/Agda/Exercises/03-Solutions.lagda.md new file mode 100644 index 0000000..12da7fc --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/03-Solutions.lagda.md @@ -0,0 +1,359 @@ +# Week 03 - Agda Exercises - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module 03-Solutions where + +open import prelude hiding (_∼_) +``` + +## Part I -- Homotopies + +It is often convenient to work with *pointwise equality* of functions, defined as follows. +```agda +module _ {A : Type} {B : A → Type} where + _∼_ : ((x : A) → B x) → ((x : A) → B x) → Type + f ∼ g = ∀ x → f x ≡ g x +``` +An element of `f ∼ g` is usually called a homotopy. + +### Exercise 1 (⋆⋆) + +Unsurprisingly, many properties of this type of pointwise equalities +can be inferred directly from the same operations on paths. + +Try to prove reflexivity, symmetry and transitivity of `_∼_` by filling these holes. +```agda + ∼-refl : (f : (x : A) → B x) → f ∼ f + ∼-refl f = λ x → refl (f x) + + ∼-inv : (f g : (x : A) → B x) → (f ∼ g) → (g ∼ f) + ∼-inv f g H x = sym (H x) + + ∼-concat : (f g h : (x : A) → B x) → f ∼ g → g ∼ h → f ∼ h + ∼-concat f g h H K x = trans (H x) (K x) + + infix 0 _∼_ +``` + +## Part II -- Isomorphisms + +A function `f : A → B` is called a *bijection* if there is a function `g : B → A` in the opposite direction such that `g ∘ f ∼ id` and `f ∘ g ∼ id`. Recall that `_∼_` is [pointwise equality](../Lecture-Notes/files/identity-type.lagda.md) and that `id` is the [identity function](../Lecture-Notes/files/products.lagda.md). This means that we can convert back and forth between the types `A` and `B` landing at the same element we started with, either from `A` or from `B`. In this case, we say that the types `A` and `B` are *isomorphic*, and we write `A ≅ B`. Bijections are also called type *isomorphisms*. We can define these concepts in Agda using [sum types](../Lecture-Notes/files/sums.lagda.md) or [records](https://agda.readthedocs.io/en/latest/language/record-types.html). We will adopt the latter, but we include both definitions for the sake of illustration. +Recall that we [defined](../Lecture-Notes/files/general-notation.lagda.md) the domain of a function `f : A → B` to be `A` and its codomain to be `B`. + +We adopt this definition of isomorphisms using records. +```agda +record is-bijection {A B : Type} (f : A → B) : Type where + constructor + Inverse + field + inverse : B → A + η : inverse ∘ f ∼ id + ε : f ∘ inverse ∼ id + +record _≅_ (A B : Type) : Type where + constructor + Isomorphism + field + bijection : A → B + bijectivity : is-bijection bijection + +infix 0 _≅_ +``` + +### Exercise 2 (⋆) +Reformulate the same definition using Sigma-types. +```agda +is-bijection' : {A B : Type} → (A → B) → Type +is-bijection' f = Σ g ꞉ (codomain f → domain f) , ((g ∘ f ∼ id) × (f ∘ g ∼ id)) + +_≅'_ : Type → Type → Type +A ≅' B = Σ f ꞉ (A → B) , is-bijection' f +``` +The definition with `Σ` is probably more intuitive, but, as discussed above, +the definition with a record is often easier to work with, +because we can easily extract the components of the definitions using the names of the fields. +It also often allows Agda to infer more types, and to give us more sensible goals in the +interactive development of Agda programs and proofs. + +Notice that `inverse` plays the role of `g`. +The reason we use `inverse` is that then we can use the word +`inverse` to extract the inverse of a bijection. +Similarly we use `bijection` for `f`, as we can use the word +`bijection` to extract the bijection from a record. + +This type can be defined to be `𝟙 ∔ 𝟙` using the coproduct, +but we give a direct definition which will allow us to discuss some relationships between the various type formers of Basic MLTT. +```agda +data 𝟚 : Type where + 𝟎 𝟏 : 𝟚 +``` + +### Exercise 3 (⋆⋆) +Prove that 𝟚 and Bool are isomorphic + +```agda +Bool-𝟚-isomorphism : Bool ≅ 𝟚 +Bool-𝟚-isomorphism = record { bijection = f ; bijectivity = f-is-bijection } + where + f : Bool → 𝟚 + f false = 𝟎 + f true = 𝟏 + + g : 𝟚 → Bool + g 𝟎 = false + g 𝟏 = true + + gf : g ∘ f ∼ id + gf true = refl true + gf false = refl false + + fg : f ∘ g ∼ id + fg 𝟎 = refl 𝟎 + fg 𝟏 = refl 𝟏 + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + + +## Part III - Finite Types + +In the last HoTT Exercises we encountered two definitions of the finite types. +Here we prove that they are isomorphic. +Note that `zero` was called `pt` and suc `i` on the HoTT exercise sheet. + +```agda +data Fin : ℕ → Type where + zero : {n : ℕ} → Fin (suc n) + suc : {n : ℕ} → Fin n → Fin (suc n) +``` + +### Exercise 4 (⋆) + +Prove the elimination principle of `Fin`. +```agda +Fin-elim : (A : {n : ℕ} → Fin n → Type) + → ({n : ℕ} → A {suc n} zero) + → ({n : ℕ} (k : Fin n) → A k → A (suc k)) + → {n : ℕ} (k : Fin n) → A k +Fin-elim A a f = h + where + h : {n : ℕ} (k : Fin n) → A k + h zero = a + h (suc k) = f k (h k) +``` + +We give the other definition of the finite types and introduce some notation. + +```agda +Fin' : ℕ → Type +Fin' 0 = 𝟘 +Fin' (suc n) = 𝟙 ∔ Fin' n + +zero' : {n : ℕ} → Fin' (suc n) +zero' = inl ⋆ + +suc' : {n : ℕ} → Fin' n → Fin' (suc n) +suc' = inr +``` + +### Exercise 5 (⋆⋆⋆) + +Prove that `Fin n` and `Fin' n` are isomorphic for every `n`. + +```agda +Fin-isomorphism : (n : ℕ) → Fin n ≅ Fin' n +Fin-isomorphism n = record { bijection = f n ; bijectivity = f-is-bijection n } + where + f : (n : ℕ) → Fin n → Fin' n + f (suc n) zero = zero' + f (suc n) (suc k) = suc' (f n k) + + g : (n : ℕ) → Fin' n → Fin n + g (suc n) (inl ⋆) = zero + g (suc n) (inr k) = suc (g n k) + + gf : (n : ℕ) → g n ∘ f n ∼ id + gf (suc n) zero = refl zero + gf (suc n) (suc k) = γ + where + IH : g n (f n k) ≡ k + IH = gf n k + + γ = g (suc n) (f (suc n) (suc k)) ≡⟨ refl _ ⟩ + g (suc n) (suc' (f n k)) ≡⟨ refl _ ⟩ + suc (g n (f n k)) ≡⟨ ap suc IH ⟩ + suc k ∎ + + fg : (n : ℕ) → f n ∘ g n ∼ id + fg (suc n) (inl ⋆) = refl (inl ⋆) + fg (suc n) (inr k) = γ + where + IH : f n (g n k) ≡ k + IH = fg n k + + γ = f (suc n) (g (suc n) (suc' k)) ≡⟨ refl _ ⟩ + f (suc n) (suc (g n k)) ≡⟨ refl _ ⟩ + suc' (f n (g n k)) ≡⟨ ap suc' IH ⟩ + suc' k ∎ + + f-is-bijection : (n : ℕ) → is-bijection (f n) + f-is-bijection n = record { inverse = g n ; η = gf n ; ε = fg n} +``` + +## Part IV -- minimal elements in the natural numbers + +In this section we establish some definitions which are needed to state and prove the well-ordering +principle of the natural numbers. + +### Exercise 6 (⋆) + +Give the recursive definition of the less than or equals relation on the natural numbers. + +```agda +_≤₁_ : ℕ → ℕ → Type +0 ≤₁ y = 𝟙 +suc x ≤₁ 0 = 𝟘 +suc x ≤₁ suc y = x ≤₁ y +``` + +### Exercise 7 (⋆) + +Given a type family `P` over the naturals, a lower bound `n` is a natural number such that +for all other naturals `m`, we have that if `P(m)` holds, then`n ≤₁ m`. +Translate this definition into HoTT. + +```agda +is-lower-bound : (P : ℕ → Type) (n : ℕ) → Type +is-lower-bound P n = (m : ℕ) → P m → n ≤₁ m +``` + +We define the type of minimal elements of a type family over the naturals. +```agda +minimal-element : (P : ℕ → Type) → Type +minimal-element P = Σ n ꞉ ℕ , (P n) × (is-lower-bound P n) +``` + +### Exercise 8 (⋆) + +Prove that all numbers are at least as large as zero. +```agda +leq-zero : (n : ℕ) → 0 ≤₁ n +leq-zero n = ⋆ +``` + + +## Part V -- The well-ordering principle of ℕ + +Classically, the well-ordering principle states that every non-empty set +of natural numbers has a least element. + +In HoTT, such subsets can be translated as decidable type family. +Recall the definition of decidability: +```agda +open import decidability + using (is-decidable; is-decidable-predicate) +``` + +The well-ordering principle reads +```agda +Well-ordering-principle = (P : ℕ → Type) → (d : is-decidable-predicate P) → (n : ℕ) → P n → minimal-element P +``` + +We shall prove this statement via induction on `n`. +In order to make the proof more readable, we first prove two lemmas. + +### Exercise 9 (🌶) + +What is the statement of `is-minimal-element-suc` +under the Curry-Howard interpretation? +Prove this lemma. + +```agda +is-minimal-element-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (m : ℕ) (pm : P (suc m)) + (is-lower-bound-m : is-lower-bound (λ x → P (suc x)) m) → + ¬ (P 0) → is-lower-bound P (suc m) +is-minimal-element-suc P d m pm is-lower-bound-m neg-p0 0 p0 = 𝟘-nondep-elim (neg-p0 p0) +-- In the previous clause, 𝟘-nondep-elim is superfluous, because neg-p0 p0 : ∅ already. +is-minimal-element-suc P d 0 pm is-lower-bound-m neg-p0 (suc n) psuccn = ⋆ +is-minimal-element-suc P d (suc m) pm is-lower-bound-m neg-p0 (suc n) psuccn = h + where + h : suc m ≤₁ n + h = is-minimal-element-suc (λ x → P (suc x)) + (λ x → d (suc x)) m pm + (λ m → is-lower-bound-m (suc m)) + (is-lower-bound-m 0) + (n) + (psuccn) + -- alternative solution + h' : suc m ≤₁ n + h' = is-lower-bound-m n psuccn +``` +The lemma states that for a decidable type family `P`, if `m` is a lower bound +for `P ∘ suc`, and `P 0` is false, then `m + 1` is a lower bound for `P`. +Note that the assumptions `d` and `pm` are not used. + +### Exercise 10 (🌶) + +What is the statement of `well-ordering-principle-suc` +under the Curry-Howard interpretation? +Prove this lemma. + +```agda +well-ordering-principle-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (n : ℕ) (p : P (suc n)) → + is-decidable (P 0) → + minimal-element (λ m → P (suc m)) → minimal-element P +well-ordering-principle-suc P d n p (inl p0) _ = 0 , (p0 , (λ m q → leq-zero m)) +well-ordering-principle-suc P d n p (inr neg-p0) (m , (pm , is-min-m)) = (suc m) , (pm , h) + where + h : is-lower-bound P (suc m) + h = is-minimal-element-suc P d m pm is-min-m neg-p0 + + -- alternative solution + h' : is-lower-bound P (suc m) + h' zero q = 𝟘-nondep-elim (neg-p0 q) + h' (suc k) q = is-min-m k q +``` +This lemma states that for a decidable type family `P`, if `P ∘ suc` is true for some `n`, +and `P 0` is decidable, then minimal elements of `P ∘ suc` yield minimal elements of `P`. +Note that `d` and `p` are not used. + +### Exercise 11 (🌶) + +Use the previous two lemmas to prove the well-ordering principle +```agda +well-ordering-principle : (P : ℕ → Type) → (d : is-decidable-predicate P) → (n : ℕ) → P n → minimal-element P +well-ordering-principle P d 0 p = 0 , (p , (λ m q → leq-zero m)) +well-ordering-principle P d (suc n) p = well-ordering-principle-suc P d n p (d 0) (well-ordering-principle (λ m → P (suc m)) (λ m → d (suc m)) n p) +``` + +### Exercise 12 (⋆⋆⋆) + +Prove that the well-ordering principle returns 0 if `P 0` holds. + +```agda +is-zero-well-ordering-principle-suc : + (P : ℕ → Type) (d : is-decidable-predicate P) + (n : ℕ) (p : P (suc n)) (d0 : is-decidable (P 0)) → + (x : minimal-element (λ m → P (suc m))) (p0 : P 0) → + (pr₁ (well-ordering-principle-suc P d n p d0 x)) ≡ 0 +is-zero-well-ordering-principle-suc P d n p (inl p0) x q0 = refl 0 +is-zero-well-ordering-principle-suc P d n p (inr np0) x q0 = 𝟘-nondep-elim (np0 q0) + +is-zero-well-ordering-principle : + (P : ℕ → Type) (d : is-decidable-predicate P) → + (n : ℕ) → (pn : P n) → P 0 → pr₁ (well-ordering-principle P d n pn) ≡ 0 +is-zero-well-ordering-principle P d 0 p p0 = refl 0 +is-zero-well-ordering-principle P d (suc m) pm = + is-zero-well-ordering-principle-suc P d m pm (d 0) + (well-ordering-principle + (λ z → P (suc z)) + (λ x → d (suc x)) + m pm) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2-solutions.lagda.md new file mode 100644 index 0000000..17c9d1a --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2-solutions.lagda.md @@ -0,0 +1,69 @@ +Week 2 - Homework Sheet - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework2-solutions where + +open import prelude hiding (id ; _∘_) + +ap² : {A : Type} (f : A → A) {x y : A} → x ≡ y → f (f x) ≡ f (f y) +ap² f e = ap f (ap f e) + +double-not-eq : {a b : Bool} → a ≡ b → not (not a) ≡ not (not b) +double-not-eq e = ap² not e + +transport-vec-A : {A : Type} {n m : ℕ} → n ≡ m + → Vector A n → Vector A m +transport-vec-A (refl n) v = v + + +map : {A B : Type} → (A → B) → List A → List B +map f [] = [] +map f (x :: xs) = f x :: map f xs + +id : {A : Type} → A → A +id = λ x → x + +_∘_ : {A B C : Type} → (g : B → C) → (f : A → B)→ A → C +g ∘ f = λ x → g (f x) + +map-law1 : {A : Type} (xs : List A) → map id xs ≡ xs +map-law1 [] = refl [] +map-law1 (x :: xs) = ap (x ::_) IH + where + IH : map id xs ≡ xs -- IH is short for induction hypothesis + IH = map-law1 xs + +map-law2 : {A B C : Type} (g : B → C) (f : A → B) (xs : List A) + → map (g ∘ f) xs ≡ map g (map f xs) +map-law2 g f [] = refl [] +map-law2 g f (x :: xs) = ap (g (f x) ::_) IH + where + IH : map (g ∘ f) xs ≡ map g (map f xs) + IH = map-law2 g f xs + + +and : List Bool → Bool +and [] = true +and (true :: bs) = and bs +and (false :: bs) = false + +and-example1 : and (true :: false :: []) ≡ false +and-example1 = refl false + +and-example2 : and [] ≡ true +and-example2 = refl true + +for-all : {A : Type} → (A → Bool) → List A → Bool +for-all f xs = and (map f xs) + +filter : {A : Type} → (A → Bool) → List A → List A +filter p [] = [] +filter p (x :: xs) = if (p x) then x :: ys else ys + where + ys = filter p xs + +filter-soundness : {A : Type} → Type +filter-soundness {A} = (p : A → Bool) (xs : List A) + → for-all p (filter p xs) ≡ true diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2.lagda.md new file mode 100644 index 0000000..a0beb10 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework2.lagda.md @@ -0,0 +1,182 @@ +# Week 2 - Homework Sheet + +Please complete Part II of this week's Lab Sheet before moving on to these exercises. + +## Section 1: Identity Type + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework2 where + +open import prelude hiding (id ; _∘_) +``` + +### Exercise 1.1 + +Recall the definitions in `trans`, `sym` and especially `ap` from the +[introduction](../introduction.lagda.md). + +We could define "double `ap`" like so: + +```agda +ap² : {A : Type} (f : A → A) {x y : A} → x ≡ y → f (f x) ≡ f (f y) +ap² f e = {!!} +``` + +For example, this would help us prove the following without pattern matching: + +```agda +double-not-eq : {a b : Bool} → a ≡ b → not (not a) ≡ not (not b) +double-not-eq e = ap² not e +``` + +**Define** "double `ap`" *without* using pattern matching. + +### Exercise 1.2 + +We have seen that `ap` allows us to "transport" an equality along a function. + +We will now see how we can "transport" equalities along types; in our example, +we use the type for Vectors. + +```agda +transport-vec-A : {A : Type} {n m : ℕ} → n ≡ m + → Vector A n → Vector A m +transport-vec-A e v = {!!} +``` + +**Complete** the above function. + +## Section 2: Laws for `map` + +Recall that we defined `map` as follows: +```agda +map : {A B : Type} → (A → B) → List A → List B +map f [] = [] +map f (x :: xs) = f x :: map f xs +``` + +In Haskell, `map` is supposed to satisfy, for every list `xs`, two laws: + + 1. `map id xs = xs` (where `id` the identity function); + 1. `map (g . f) xs = map g (map f xs)` (where `.` is function composition). + +But checking these laws is left to a pen-and-paper check by the programmer. + +In Agda, we can both *state* and *prove* these laws. + +We first define function composition and the identity function. + +```agda +id : {A : Type} → A → A +id = λ x → x + +_∘_ : {A B C : Type} → (g : B → C) → (f : A → B)→ A → C +g ∘ f = λ x → g (f x) +``` + +Tip: Use `\comp` to type `∘`. + +### Exercise 2.1 + +The first `map` law can now be written in Agda as +```agda +map-law1 : {A : Type} (xs : List A) → map id xs ≡ xs +map-law1 xs = {!!} +``` + +**Define** this function. (Hint: An induction hypothesis comes in helpful). + +### Exercise 2.2 + +A partial statement of the second `map` law is the following: +```agda +map-law2 : {A B C : Type} (g : {!!}) (f : {!!}) (xs : List A) + → {!!} ≡ {!!} +map-law2 g f xs = {!!} +``` + +**Complete** the holes in `map-law2 : ...`, i.e. write down the types of `g` and +`f`, and translate the second `map` law `map (g ∘ f) xs = map g (map f xs)` to +Agda. + + +### Exercise 2.3 + +Now **complete** the proof of `map-law2`. + +## Section 3: Filtering lists + +**Define** a function `and` that takes a list of Booleans and computes the +conjunction of all the Booleans in it: + +### Exercise 3.1 + +```agda +and : List Bool → Bool +and bs = {!!} +``` + +For example, `and true false` should equal to `false` + +```agda +and-example1 : and (true :: false :: []) ≡ false +and-example1 = {! refl false !} + -- make sure that your code is correct by checking + -- that the equality in this type is definitional. + -- The proof should be just `refl false`. +``` + +```agda +and-example2 : and [] ≡ true +and-example2 = {! refl true!} + -- make sure that your code is correct by checking + -- that the equality in this type is definitional. + -- The proof should be just `refl true`. +``` + +### Exercise 3.2 + +**Define** a function `for-all` that takes a list of element of some type `A` +together with a predicate `p : A → Bool`, and computes the Boolean truth value +denoting if *everything* in the list satisfy the predicate `p`. You should be +able to this just by composing `map` and your implementation of `and`. + +```agda +for-all : {A : Type} → (A → Bool) → List A → Bool +for-all = {!!} +``` + +Next, we will define a function `filter` that filters out a subset of a list of +elements that satisfy a certain predicate. You have probably worked with a +function like this in Haskell before. + +### Exercise 3.3 + +```agda +filter : {A : Type} → (A → Bool) → List A → List A +filter = {!!} +``` + +`filter p xs` should give you the sublist of `xs` consisting of elements that +satisfy `p`. + +### Exercise 3.4 + +You have probably worked with similar functions in Haskell before. The nice +thing about working in Agda is that, we can actually write down *specifications* +of such functions. + +Using your `for-all`, **define** a function that expresses: *for every predicate +`p` and every list `xs`, everything in `filter p xs` satisfies the predicate +`p`*. Call this type `filter-soundness`. + +```agda +filter-soundness : {A : Type} → Type +filter-soundness {A} = {!!} +``` + +Note that the type of the function here is `Type`. This means that you are not +writing an inhabitant of a type but you are *writing a type*, i.e. you are not +asked to *prove* filter soundness, but rather to *state* it. diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3-solutions.lagda.md new file mode 100644 index 0000000..5e9603c --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3-solutions.lagda.md @@ -0,0 +1,79 @@ +# Week 3 - Homework Sheet - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework3-solutions where + +open import prelude hiding (𝟘-nondep-elim) +open import Pool.Lab.lab3-solutions + +∔-assoc : {A B C : Type} → A ∔ (B ∔ C) → (A ∔ B) ∔ C +∔-assoc (inl a) = inl (inl a) +∔-assoc (inr (inl b)) = inl (inr b) +∔-assoc (inr (inr c)) = inr c + +×-assoc : {A B C : Type} → A × (B × C) → (A × B) × C +×-assoc (a , (b , c)) = (a , b) , c + +∔-comm : {A B : Type} → A ∔ B → B ∔ A +∔-comm (inl a) = inr a +∔-comm (inr b) = inl b + +×-comm : {A B : Type} → A × B → B × A +×-comm (a , b) = b , a + + +not-A-and-not-A : {A : Type} → ¬ (A × ¬ A) +not-A-and-not-A (a , p) = p a + +A-and-not-A-implies-B : {A B : Type} → A × ¬ A → B +A-and-not-A-implies-B p = 𝟘-nondep-elim (not-A-and-not-A p) + +LEM : {A : Type} → A ∔ ¬ A +LEM = {!!} + +not-not-LEM : {A : Type} → ¬¬ (A ∔ ¬ A) +not-not-LEM p = p (inr (λ a → p (inl a))) + +DNE : {A : Type} → ¬¬ A → A +DNE = {!!} + +LEM' : {A : Type} → A ∔ ¬ A +LEM' = DNE not-not-LEM + +DNE' : {A : Type} → ¬¬ A → A +DNE' {A} p = γ LEM + where + γ : A ∔ ¬ A → A + γ (inl a) = a + γ (inr q) = 𝟘-nondep-elim (p q) + +not-not-DNE : {A : Type} → ¬¬ (¬¬ A → A) +not-not-DNE {A} p = p dne + where + r : ¬ A + r a = p (λ _ → a) + dne : ¬¬ A → A + dne q = 𝟘-nondep-elim (q r) + + +Σ-∔-distributivity : {A : Type} {B C : A → Type} + → (Σ a ꞉ A , (B a ∔ C a)) + → (Σ a ꞉ A , B a) ∔ (Σ a ꞉ A , C a) +Σ-∔-distributivity (a , inl b) = inl (a , b) +Σ-∔-distributivity (a , inr c) = inr (a , c) + +¬Σ-if-forall-not : {A : Type} {B : A → Type} + → ((a : A) → ¬ B a) → ¬ (Σ a ꞉ A , B a) +¬Σ-if-forall-not p (a , b) = p a b + +forall-not-if-¬Σ : {A : Type} {B : A → Type} + → ¬ (Σ a ꞉ A , B a) → (a : A) → ¬ B a +forall-not-if-¬Σ p a b = p (a , b) + +Π-Σ-distributivity : {A B : Type} {C : A → B → Type} + → ((a : A) → (Σ b ꞉ B , C a b)) + → Σ f ꞉ (A → B) , ((a : A) → C a (f a)) +Π-Σ-distributivity p = (λ a → fst (p a)) , (λ a → snd (p a)) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3.lagda.md new file mode 100644 index 0000000..7cb1b35 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework3.lagda.md @@ -0,0 +1,199 @@ +# Week 3 - Homework Sheet + +**Please finish the lab sheet before moving on to these exercises.** + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework3 where + +open import prelude +open import Pool.Lab.lab3-solutions +``` + +## Part I: Associativity and Commutativity of ∔ and × + +We have already seen that the Boolean operations `_||_` and `_&&_` are +associative and commutative. + +The type formers that represent these logical connectives are also associative +and commutative. + +### Exercise 1.1 + +**Prove** that `_∔_` is associative. + +```agda +∔-assoc : {A B C : Type} → A ∔ (B ∔ C) → (A ∔ B) ∔ C +∔-assoc = {!!} +``` + +### Exercise 1.2 + +**Prove** that `_×_` is associative. + +```agda +×-assoc : {A B C : Type} → A × (B × C) → (A × B) × C +×-assoc = {!!} +``` + +### Exercise 1.3 + +**Prove** that `_∔_` is commutative. + +```agda +∔-comm : {A B : Type} → A ∔ B → B ∔ A +∔-comm = {!!} +``` +### Exercise 1.4 + +**Prove** that `_×_` is commutative. + +```agda +×-comm : {A B : Type} → A × B → B × A +×-comm = {!!} +``` + +## Part II: Law of excluded middle and double-negation elimination + +In classical logic, we have the law of excluded middle (LEM): `A + ¬ A`, for any +logical formula `A`. + +### Exercise 2.1 + +This seems intuitive to us, as having both `A` and `¬ A` gives us a +contradiction. + +```agda +not-A-and-not-A : {A : Type} → ¬ (A × ¬ A) +not-A-and-not-A = {!!} +``` + +**Complete** the proof that `¬ (A x ¬ A)`. + +### Exercise 2.2 + +Furthermore, if we had both `A` and `¬ A`, we could prove anything. + +```agda +A-and-not-A-implies-B : {A B : Type} → A × ¬ A → B +A-and-not-A-implies-B p = {!!} +``` + +**Complete** the above statement *without* using pattern matching. + +Hint: Use `𝟘-nondep-elim`. + +### Exercise 2.3 + +However, it turns out that LEM is not provable (or disprovable) in Agda. + +You can try this out yourself (but you won't succeed): + +```agda +LEM : {A : Type} → A ∔ ¬ A +LEM = {!!} -- You are not expected to complete this hole. + -- In fact, it's impossible. +``` + +However, we *can* prove the *double-negation* of `LEM`. + +```agda +not-not-LEM : {A : Type} → ¬¬ (A ∔ ¬ A) +not-not-LEM = {!!} +``` + +**Prove** the double-negation of the law of excluded middle. + +### Exercise 2.4 + +If we had access to double-negation elimination (DNE), as in classical logic, we +could give `LEM`. + +Note: Do not try to prove DNE (see Exercise 2.5). + +**Complete** `LEM'` using `DNE`. + +```agda +DNE : {A : Type} → ¬¬ A → A +DNE = {!!} -- You are not expected to complete this hole. + -- In fact, it's impossible. + +LEM' : {A : Type} → A ∔ ¬ A +LEM' = {!!} +``` + +### Exercise 2.5 + +However, `DNE` is *also* not provable or disprovable in Agda. + +It is the case, however, that if we had access to `LEM`, we could prove `DNE`. + +```agda +DNE' : {A : Type} → ¬¬ A → A +DNE' {A} = {!!} +``` + +**Complete** `DNE'` using `LEM`. + +### Exercise 2.6 + +So we have seen that `LEM` and `DNE` are both not provable in Agda, yet are +equivalent in the sense that having one gives us the other. + +Finally, we can also show that the *double-negation* of DNE *is* provable in +Agda. + +```agda +not-not-DNE : {A : Type} → ¬¬ (¬¬ A → A) +not-not-DNE {A} = {!!} +``` + +**Prove** the double-negation of the law of excluded middle. + +## Part III: Sums and products + +### Exercise 3.1 + +**Complete** the following proof of distributivity of `Σ` over `_∔_`. + +```agda +Σ-∔-distributivity : {A : Type} {B C : A → Type} + → (Σ a ꞉ A , (B a ∔ C a)) + → (Σ a ꞉ A , B a) ∔ (Σ a ꞉ A , C a) +Σ-∔-distributivity = {!!} +``` + +### Exercise 3.2 + +If, for every `a : A` we have `¬ B a` (the type `B a` is empty), then there +does not exist any `a : A` satisfying `B a` (the type `Σ B` is empty). + +```agda +¬Σ-if-forall-not : {A : Type} {B : A → Type} + → ((a : A) → ¬ B a) → ¬ (Σ a ꞉ A , B a) +¬Σ-if-forall-not = {!!} +``` + +**Complete** the proof of the above statement. + +### Exercise 3.3 + +**Prove** that the converse of the above also holds. + +```agda +forall-not-if-¬Σ : {A : Type} {B : A → Type} + → ¬ (Σ a ꞉ A , B a) → (a : A) → ¬ B a +forall-not-if-¬Σ = {!!} +``` + +### Exercise 3.4 + +Finally, **prove** that `Σ` distributes over "for all". + +```agda +Π-Σ-distributivity : {A B : Type} {C : A → B → Type} + → ((a : A) → (Σ b ꞉ B , C a b)) + → Σ f ꞉ (A → B) , ((a : A) → C a (f a)) +Π-Σ-distributivity = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4-solutions.lagda.md new file mode 100644 index 0000000..556b119 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4-solutions.lagda.md @@ -0,0 +1,79 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework4-solutions where + +open import prelude +open import List-functions hiding (++-assoc) + ++-comm : (x y : ℕ) → x + y ≡ y + x ++-comm 0 0 = refl zero ++-comm 0 (suc y) = ap suc (+-comm zero y) ++-comm (suc x) 0 = ap suc (+-comm x zero) ++-comm (suc x) (suc y) + = suc (x + suc y) ≡⟨ ap suc (+-comm x (suc y)) ⟩ + suc (suc (y + x)) ≡⟨ ap (suc ∘ suc) (+-comm y x) ⟩ + suc (suc x + y) ≡⟨ ap suc (+-comm (suc x) y) ⟩ + suc (y + suc x) ∎ + +data _≼_ {X : Type} : List X → List X → Type where + []-is-prefix : (xs : List X) → [] ≼ xs + ::-is-prefix : (x : X) (xs ys : List X) + → xs ≼ ys → (x :: xs) ≼ (x :: ys) + +_≼'_ : {X : Type} → List X → List X → Type +_≼'_ {X} xs ys = Σ zs ꞉ List X , xs ++ zs ≡ ys + + +++-assoc : {X : Type} (xs ys zs : List X) + → (xs ++ ys) ++ zs ≡ xs ++ (ys ++ zs) +++-assoc [] ys zs = refl (ys ++ zs) +++-assoc (x :: xs) ys zs = ap (λ - → x :: -) IH + where + IH : (xs ++ ys) ++ zs ≡ xs ++ (ys ++ zs) + IH = ++-assoc xs ys zs + +≼'-is-transitive : {X : Type} (xs ys zs : List X) + → xs ≼' ys → ys ≼' zs → xs ≼' zs +≼'-is-transitive xs ys zs (l , e) (l' , e') = ((l ++ l') , e'') + where + e'' : xs ++ l ++ l' ≡ zs + e'' = xs ++ l ++ l' ≡⟨ sym (++-assoc xs l l') ⟩ + (xs ++ l) ++ l' ≡⟨ ap (λ - → - ++ l') e ⟩ + ys ++ l' ≡⟨ e' ⟩ + zs ∎ + + +≼-++ : {X : Type} (xs ys : List X) → xs ≼ (xs ++ ys) +≼-++ [] ys = []-is-prefix ys +≼-++ (x :: xs) ys = ::-is-prefix x xs (xs ++ ys) IH + where + IH : xs ≼ (xs ++ ys) + IH = ≼-++ xs ys + +≼-unprime : {X : Type} (xs ys : List X) → xs ≼' ys → xs ≼ ys +≼-unprime xs _ (zs , refl _) = ≼-++ xs zs + + +≼'-[] : {X : Type} (xs : List X) → [] ≼' xs +≼'-[] xs = (xs , refl xs) + +≼'-cons : {X : Type} (x : X) (xs ys : List X) + → xs ≼' ys + → (x :: xs) ≼' (x :: ys) +≼'-cons x xs ys (zs , e) = (zs , ap (λ - → x :: -) e) + +≼-prime : {X : Type} (xs ys : List X) → xs ≼ ys → xs ≼' ys +≼-prime [] ys ([]-is-prefix ys) = ≼'-[] ys +≼-prime (x :: xs) (x :: ys) (::-is-prefix x xs ys h) = ≼'-cons x xs ys IH + where + IH : xs ≼' ys + IH = ≼-prime xs ys h + + +≼-is-transitive : {X : Type} (xs ys zs : List X) + → xs ≼ ys → ys ≼ zs → xs ≼ zs +≼-is-transitive xs ys zs p q = ≼-unprime xs zs + (≼'-is-transitive xs ys zs (≼-prime xs ys p) + (≼-prime ys zs q)) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4.lagda.md new file mode 100644 index 0000000..665e4ad --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework4.lagda.md @@ -0,0 +1,191 @@ +# Week 4 - Homework Sheet + +**Please finish the lab sheet before moving on to these exercises.** + +```agda +{-# OPTIONS --without-K --safe #-} +module Pool.Homework.homework4 where + +open import prelude +open import List-functions hiding (++-assoc) +``` + +## Part I: Commutativity of addition of natural numbers + +We wish to prove the commutativity of `_+_` on the natural numbers. + +The following proof skeleton has been provided for you, using +[notation for equational reasoning](https://git.cs.bham.ac.uk/mhe/afp-learning/-/blob/main/files/LectureNotes/files/identity-type.lagda.md#notation-for-equality-reasoning). + +```agda ++-comm : (x y : ℕ) → x + y ≡ y + x ++-comm 0 0 = {!!} ++-comm 0 (suc y) = {!!} ++-comm (suc x) 0 = {!!} ++-comm (suc x) (suc y) + = suc (x + suc y) ≡⟨ {!!} ⟩ + suc (suc (y + x)) ≡⟨ {!!} ⟩ + suc (suc x + y) ≡⟨ {!!} ⟩ + suc (y + suc x) ∎ +``` + +**Fill** the holes of our proof that `_+_` is commutative. + +## Part II: Prefixes of lists + +In this part we will study two ways of expressing that a list is prefix of +another list. + +This will be similar to how we had two ways of expressing less-than-or-equal +`≤` on natural numbers, as seen in the Lab Sheet for Week 4. In particular, +you will notice how to the structure of the proofs below mirror the structure +of the proofs in this week's Lab Sheet. + +The first definition is an inductive one. + +```agda +data _≼_ {X : Type} : List X → List X → Type where + []-is-prefix : (xs : List X) → [] ≼ xs + ::-is-prefix : (x : X) (xs ys : List X) + → xs ≼ ys → (x :: xs) ≼ (x :: ys) +``` + +It says: +1. that the empty list is a prefix of any list; +1. if `xs` is a prefix of `ys`, then `x :: xs` is a prefix of `x :: ys`. + +The second item says that you can extend prefixes by adding the same element at +the start. + +The second definition uses a `Σ`-type. + +```agda +_≼'_ : {X : Type} → List X → List X → Type +_≼'_ {X} xs ys = Σ zs ꞉ List X , xs ++ zs ≡ ys +``` + +It says that `xs` is a prefix of `ys` if we have another list `zs` such that +`xs ++ zs ≡ ys`. In other words, `xs` can be extended to `ys.` + +### Examples + +Here are some examples of prefixes of lists. + +```agda +≼'-example₀ : [] ≼' (1 :: 2 :: [ 3 ]) +≼'-example₀ = ((1 :: 2 :: [ 3 ]) , refl (1 :: 2 :: [ 3 ])) + +≼'-example₁ : [ 1 ] ≼' (1 :: [ 2 ]) +≼'-example₁ = ([ 2 ] , refl (1 :: [ 2 ])) + +≼'-example₂ : (7 :: [ 3 ]) ≼' (7 :: 3 :: 4 :: [ 9 ]) +≼'-example₂ = ((4 :: [ 9 ]) , refl (7 :: 3 :: 4 :: [ 9 ])) +``` + +For comparison, here are some examples using `≼` instead of `≼'`. + +```agda +≼-example₀ : [] ≼ (1 :: 2 :: [ 3 ]) +≼-example₀ = []-is-prefix (1 :: 2 :: [ 3 ]) + +≼-example₁ : [ 1 ] ≼ (1 :: [ 2 ]) +≼-example₁ = ::-is-prefix 1 [] [ 2 ] + ([]-is-prefix [ 2 ]) + +≼-example₂ : (7 :: [ 3 ]) ≼ (7 :: 3 :: 4 :: [ 9 ]) +≼-example₂ = ::-is-prefix 7 [ 3 ] (3 :: 4 :: [ 9 ]) + (::-is-prefix 3 [] (4 :: [ 9 ]) + ([]-is-prefix (4 :: [ 9 ]))) +``` + +Notice how we build up the proofs by consecutive uses of `::-is-prefix`, +finishing with `[]-is-prefix`. This reflects the inductive definition of `≼`. + +We will prove that `x ≼ y` and `x ≼' y` are logically equivalent, but first we +include a useful exercise on associativity. + +### Exercise 2.1 + +**Prove** that concatenation of lists, `++`, is associative. + +```agda +++-assoc : {X : Type} (xs ys zs : List X) + → (xs ++ ys) ++ zs ≡ xs ++ (ys ++ zs) +++-assoc = {!!} +``` + +### Exercise 2.2 + +**Complete** the holes in the equational reasoning below to prove that `≼'` is +transitive. + +```agda +≼'-is-transitive : {X : Type} (xs ys zs : List X) + → xs ≼' ys → ys ≼' zs → xs ≼' zs +≼'-is-transitive xs ys zs (l , e) (l' , e') = ((l ++ l') , e'') + where + e'' : xs ++ l ++ l' ≡ zs + e'' = xs ++ (l ++ l') ≡⟨ {!!} ⟩ + (xs ++ l) ++ l' ≡⟨ {!!} ⟩ + ys ++ l' ≡⟨ {!!} ⟩ + zs ∎ +``` + +### Exercise 2.3 + +**Prove** the following about `≼`. + +```agda +≼-++ : {X : Type} (xs ys : List X) → xs ≼ (xs ++ ys) +≼-++ [] ys = {!!} +≼-++ (x :: xs) ys = {!!} +``` + +### Exercise 2.4 + +**Complete** the function below, showing that we can go from `≼'` to `≼`. + +*Hint*: Use `≼-++`. + +```agda +≼-unprime : {X : Type} (xs ys : List X) → xs ≼' ys → xs ≼ ys +≼-unprime = {!!} +``` + +### Exercise 2.5 + +**Prove** the following facts about `≼'`. + +```agda +≼'-[] : {X : Type} (xs : List X) → [] ≼' xs +≼'-[] = {!!} + +≼'-cons : {X : Type} (x : X) (xs ys : List X) + → xs ≼' ys + → (x :: xs) ≼' (x :: ys) +≼'-cons x xs ys (zs , e) = {!!} +``` + +Note that these amount to the constructors of `≼`. + +### Exercise 2.6 + +**Complete** the function below, showing that we can go from `≼` to `≼'`. + +*Hint*: Use `≼'-[]` and `≼'-cons`. + +```agda +≼-prime : {X : Type} (xs ys : List X) → xs ≼ ys → xs ≼' ys +≼-prime = {!!} +``` + +### Exercise 2.7 + +Using the functions `≼-unprime` and `≼-prime`, and the fact that `≼'` is +transitive, **prove** that `≼` is transitive too. + +```agda +≼-is-transitive : {X : Type} (xs ys zs : List X) + → xs ≼ ys → ys ≼ zs → xs ≼ zs +≼-is-transitive = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5-solutions.lagda.md new file mode 100644 index 0000000..2fa3810 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5-solutions.lagda.md @@ -0,0 +1,240 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework5-solutions where + +open import prelude hiding (Bool-elim) +open import natural-numbers-functions hiding (_≤_ ; is-even ; +-assoc ; +-comm) +open import List-functions +open import isomorphisms +open import negation + +open import Pool.Homework.homework4-solutions +open import Pool.Lab.lab4-solutions +open import Pool.Lab.lab5-solutions + +length-of-reverse : {A : Type} (xs : List A) + → length (reverse xs) ≡ length xs +length-of-reverse [] = refl 0 +length-of-reverse (x :: xs) = + length (reverse xs ++ [ x ]) ≡⟨ length-of-++ (reverse xs) [ x ] ⟩ + length (reverse xs) + length [ x ] ≡⟨ refl _ ⟩ + length (reverse xs) + 1 ≡⟨ ap (_+ 1) IH ⟩ + length xs + 1 ≡⟨ +-comm (length xs) 1 ⟩ + 1 + length xs ∎ + where + IH : length (reverse xs) ≡ length xs + IH = length-of-reverse xs + +ℕ-[⋆]-iso : ℕ ≅ List 𝟙 +ℕ-[⋆]-iso = record { bijection = f ; bijectivity = f-is-bijection } + where + f : ℕ → List 𝟙 + f 0 = [] + f (suc n) = ⋆ :: f n + + g : List 𝟙 → ℕ + g [] = 0 + g (⋆ :: ⋆s) = suc (g ⋆s) + + gf : g ∘ f ∼ id + gf 0 = refl 0 + gf (suc n) = ap suc (gf n) + + fg : f ∘ g ∼ id + fg [] = refl [] + fg (⋆ :: ⋆s) = ap (⋆ ::_) (fg ⋆s) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } + +open _≅_ + +ℕ→[⋆]-preserves-length : (n : ℕ) → length (bijection ℕ-[⋆]-iso n) ≡ n +ℕ→[⋆]-preserves-length zero = refl 0 +ℕ→[⋆]-preserves-length (suc n) = ap suc (ℕ→[⋆]-preserves-length n) + +principle-of-bivalence : (b : Bool) → (b ≡ true) ∔ (b ≡ false) +principle-of-bivalence true = inl (refl true) +principle-of-bivalence false = inr (refl false) + +is-even-is-decidable : (n : ℕ) → is-decidable (is-even n) +is-even-is-decidable n = + ∔-nondep-elim goal₁ goal₂ (principle-of-bivalence (check-even n)) + where + goal₁ : check-even n ≡ true → is-decidable (is-even n) + goal₁ p = inl (check-even⇒even n p) + + goal₂ : check-even n ≡ false → is-decidable (is-even n) + goal₂ p = inr subgoal + where + subgoal : ¬ is-even n + subgoal q = true-is-not-false (true ≡⟨ sym (even⇒check-even n q) ⟩ + check-even n ≡⟨ p ⟩ + false ∎) + +filter : {A : Type} → (A → Bool) → List A → List A +filter P [] = [] +filter P (x :: xs) = if P x then (x :: ys) else ys + where + ys = filter P xs + +≤-suc-lemma : (n : ℕ) → n ≤ (1 + n) +≤-suc-lemma 0 = ≤-zero (1 + 0) +≤-suc-lemma (suc n) = goal + where + IH : n ≤ (1 + n) + IH = ≤-suc-lemma n + goal : suc n ≤ suc (suc n) + goal = ≤-suc n (suc n) IH + +Bool-elim : (A : Bool → Type) + → A false + → A true + → (x : Bool) → A x +Bool-elim A x₀ x₁ false = x₀ +Bool-elim A x₀ x₁ true = x₁ + +length-of-filter : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) ≤ length xs +length-of-filter P [] = ≤-zero 0 +length-of-filter P (x :: xs) = Bool-elim goal-statement false-case + true-case + (P x) + where + ys = filter P xs + + {- Note that by definition + filter P (x :: xs) = if P x then (x :: ys) else ys + and so goal-statement is almost + length (filter P (x :: xs)) ≤ length (x :: xs) + except that we replace "P x" by the Boolean "b". -} + goal-statement : Bool → Type + goal-statement b = length (if b then (x :: ys) else ys) ≤ length (x :: xs) + + IH : length ys ≤ length xs + IH = length-of-filter P xs + + {- The type of "false-case" is equal to "goal-statement false". -} + false-case : length ys ≤ length (x :: xs) + false-case = ≤-trans (length ys) (length xs) (length (x :: xs)) + IH (≤-suc-lemma (length xs)) + + {- The type of "true-case" is equal to "goal-statement true". -} + true-case : length (x :: ys) ≤ length (x :: xs) + true-case = ≤-suc (length ys) (length xs) IH + +{- Here is a version that uses Agda's built-in with-keyword (as shown by Eric in + the lab of 28 Feb) instead of Bool-elim. (Under the hood, they amount to the + same thing.) -} +length-of-filter' : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) ≤ length xs +length-of-filter' P [] = ≤-zero 0 + +length-of-filter' P (x :: xs) with P x +length-of-filter' P (x :: xs) | true = true-case + where + ys = filter P xs + + IH : length ys ≤ length xs + IH = length-of-filter' P xs + + true-case : length (x :: ys) ≤ length (x :: xs) + true-case = ≤-suc (length ys) (length xs) IH + +length-of-filter' P (x :: xs) | false = false-case + where + ys = filter P xs + + IH : length ys ≤ length xs + IH = length-of-filter' P xs + + false-case : length ys ≤ length (x :: xs) + false-case = ≤-trans (length ys) (length xs) (length (x :: xs)) + IH (≤-suc-lemma (length xs)) + +length-of-filters : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) + length (filter (not ∘ P) xs) + ≡ length xs +length-of-filters P [] = refl _ +length-of-filters P (x :: xs) = Bool-elim goal-statement false-case + true-case + (P x) + where + ys = filter P xs + ys' = filter (not ∘ P) xs + + IH : length ys + length ys' ≡ length xs + IH = length-of-filters P xs + + {- Note that by definition + filter P (x :: xs) = if P x then (x :: ys) else ys + and so goal-statement is almost + length (filter P (x :: xs)) + + length (filter (not ∘ P) (x :: xs)) + ≡ length (x :: xs) + except that we replace "P x" by the Boolean "b". -} + goal-statement : Bool → Type + goal-statement b = length (if b then (x :: ys ) else ys ) + + length (if not b then (x :: ys') else ys') + ≡ 1 + length xs + + {- The type of "false-case" is equal to "goal-statement false. -} + false-case : length ys + length (x :: ys') ≡ length (x :: xs) + false-case = + length ys + length (x :: ys') ≡⟨ refl _ ⟩ + length ys + (1 + length ys') ≡⟨ +-assoc (length ys) 1 (length ys') ⟩ + (length ys + 1) + length ys' ≡⟨ ap (_+ length ys') (+-comm (length ys) 1) ⟩ + (1 + length ys) + length ys' ≡⟨ sym (+-assoc 1 (length ys) (length ys')) ⟩ + 1 + (length ys + length ys') ≡⟨ ap (1 +_) IH ⟩ + 1 + length xs ∎ + + {- The type of "true-case" is equal to "goal-statement true". -} + true-case : length (x :: ys) + length ys' ≡ length (x :: xs) + true-case = + length (x :: ys) + length ys' ≡⟨ refl _ ⟩ + (1 + length ys) + length ys' ≡⟨ +-assoc 1 (length ys) (length ys') ⟩ + 1 + (length ys + length ys') ≡⟨ ap (1 +_) IH ⟩ + 1 + length xs ∎ + +{- Here is a version that uses Agda's built-in with-keyword (as shown by Eric in + the lab of 28 Feb) instead of Bool-elim. (Under the hood, they amount to the + same thing.) -} +length-of-filters' : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) + length (filter (not ∘ P) xs) + ≡ length xs +length-of-filters' P [] = refl _ + +length-of-filters' P (x :: xs) with P x +length-of-filters' P (x :: xs) | true = true-case + where + ys = filter P xs + ys' = filter (not ∘ P) xs + + IH : length ys + length ys' ≡ length xs + IH = length-of-filters P xs + + true-case : length (x :: ys) + length ys' ≡ length (x :: xs) + true-case = + length (x :: ys) + length ys' ≡⟨ refl _ ⟩ + (1 + length ys) + length ys' ≡⟨ +-assoc 1 (length ys) (length ys') ⟩ + 1 + (length ys + length ys') ≡⟨ ap (1 +_) IH ⟩ + 1 + length xs ∎ + +length-of-filters' P (x :: xs) | false = false-case + where + ys = filter P xs + ys' = filter (not ∘ P) xs + + IH : length ys + length ys' ≡ length xs + IH = length-of-filters P xs + + false-case : length ys + length (x :: ys') ≡ length (x :: xs) + false-case = + length ys + length (x :: ys') ≡⟨ refl _ ⟩ + length ys + (1 + length ys') ≡⟨ +-assoc (length ys) 1 (length ys') ⟩ + (length ys + 1) + length ys' ≡⟨ ap (_+ length ys') (+-comm (length ys) 1) ⟩ + (1 + length ys) + length ys' ≡⟨ sym (+-assoc 1 (length ys) (length ys')) ⟩ + 1 + (length ys + length ys') ≡⟨ ap (1 +_) IH ⟩ + 1 + length xs ∎ +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5.lagda.md new file mode 100644 index 0000000..53b6786 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework5.lagda.md @@ -0,0 +1,196 @@ +# Week 5 - Homework Sheet + +**Please finish the lab sheet before moving on to these exercises.** + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework5 where + +open import prelude hiding (Bool-elim) +open import natural-numbers-functions hiding (_≤_ ; is-even ; +-assoc ; +-comm) +open import List-functions +open import isomorphisms +``` +We will also want to use some things from the Lab and Homework sheet of Week 4: + +```agda +open import Pool.Homework.homework4-solutions +open import Pool.Lab.lab4-solutions +``` + +## Part I: More on length + +Besides `map`, the function `reverse` is another example of a length-preserving +operation. + +```agda +length-of-reverse : {A : Type} (xs : List A) + → length (reverse xs) ≡ length xs +length-of-reverse = {!!} +``` + +**Prove** the above. + +## Part II: More on isomorphisms + +### Exercise 2a + +```agda +ℕ-[⋆]-iso : ℕ ≅ List 𝟙 +ℕ-[⋆]-iso = record { bijection = f ; bijectivity = f-is-bijection } + where + f : ℕ → List 𝟙 + f = {!!} + + g : List 𝟙 → ℕ + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +**Show** that the type of natural numbers `ℕ` is isomorphic to the type of lists +over the unit type `𝟙`. + +Hint: The statement of Exercise 2b may help you. + +### Exercise 2b + +```agda +open _≅_ + +ℕ→[⋆]-preserves-length : (n : ℕ) → length (bijection ℕ-[⋆]-iso n) ≡ n +ℕ→[⋆]-preserves-length = {!!} +``` + +Notice how `bijection` extracts the function `f` you defined in `ℕ-[⋆]-iso`. + +**Prove** that for any `n : ℕ`, the length of the list `f n : List 𝟙` +(where `f : ℕ → List 𝟙` is as you defined in Exercise 3a) is `n`. + +## Part III: More on evenness + +In this exercise, we will continue where we left off in the lab exercises on +evenness. Recall the predicates `is-even` and `check-even`: + +```agda +is-even : ℕ → Type +is-even x = Σ y ꞉ ℕ , x ≡ 2 * y +``` + +```agda +check-even : ℕ → Bool +check-even zero = true +check-even (suc zero) = false +check-even (suc (suc n)) = check-even n +``` + +Now recall the discussion about decidable predicates that we had in the previous +weeks. When you proved that `check-even` and `is-even` are logically equivalent +in the lab, you have in fact implicitly proved that `is-even` is a decidable +predicate! In this exercise, we will make this implicit proof _explicit_. + +**Complete** the remaining holes in the following proof outline; starting with +proving a lemma stating that a Boolean is either `true` or `false`. + +```agda +principle-of-bivalence : (b : Bool) → (b ≡ true) ∔ (b ≡ false) +principle-of-bivalence = {!!} + +is-even-is-decidable : (n : ℕ) → is-decidable (is-even n) +is-even-is-decidable n = + ∔-nondep-elim goal₁ goal₂ (principle-of-bivalence (check-even n)) + where + goal₁ : check-even n ≡ true → is-decidable (is-even n) + goal₁ p = {!!} + + goal₂ : check-even n ≡ false → is-decidable (is-even n) + goal₂ p = inr subgoal + where + subgoal : ¬ is-even n + subgoal q = {!!} +``` + +## Part IV: Stretcher exercises on length + +*The following two exercises are rather hard and are should be interesting to +students that like a challenge.* + +Recall that we can define `filter` as +```agda +filter : {A : Type} → (A → Bool) → List A → List A +filter P [] = [] +filter P (x :: xs) = if P x then (x :: ys) else ys + where + ys = filter P xs +``` + +We also remind you of the inductively defined less-than-or-equal relation `≤` +on the natural numbers. + +```agdacode +data _≤_ : ℕ → ℕ → Type where + ≤-zero : ( y : ℕ) → 0 ≤ y + ≤-suc : (x y : ℕ) → x ≤ y → suc x ≤ suc y +``` + +Finally, the following lemmas are provided to you for your convenience. + +```agda +≤-suc-lemma : (n : ℕ) → n ≤ (1 + n) +≤-suc-lemma 0 = ≤-zero (1 + 0) +≤-suc-lemma (suc n) = goal + where + IH : n ≤ (1 + n) + IH = ≤-suc-lemma n + goal : suc n ≤ suc (suc n) + goal = ≤-suc n (suc n) IH + +Bool-elim : (A : Bool → Type) + → A false + → A true + → (x : Bool) → A x +Bool-elim A x₀ x₁ false = x₀ +Bool-elim A x₀ x₁ true = x₁ +``` + +### Exercise 4a (stretcher 🌶) + +**Prove** that filtering a list decreases its length. + +```agda +length-of-filter : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) ≤ length xs +length-of-filter = {!!} +``` + +*Hints*: +- You can use `≤-trans` from the [sample solutions to + Lab 4](lab4-solutions.lagda.md) if you need that `≤` is transitive. + (The sample solutions are already imported for you.) +- Think about how to use `Bool-elim`. + +### Exercise 4b (stretcher 🌶🌶) + +Given a predicate `P : A → Bool` and a list `xs : List A`, we could filter `xs` +by `P` and by `not ∘ P`. If we do this and compute the lengths of the resulting +lists, then we expect their sum to be equal to the length of the unfiltered list +`xs`. **Prove** this fact. + +```agda +length-of-filters : {A : Type} (P : A → Bool) (xs : List A) + → length (filter P xs) + length (filter (not ∘ P) xs) + ≡ length xs +length-of-filters = {!!} +``` + +*Hint*: You can use associativity (`+-assoc`) and commutativity (`+-comm`) from +the sample solutions to last week's exercises. (The necessary files are already +imported for you.) diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6-solutions.lagda.md new file mode 100644 index 0000000..992e63a --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6-solutions.lagda.md @@ -0,0 +1,96 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework6-solutions where + +open import prelude +open import isomorphisms +open import natural-numbers-functions +open import Fin +open import Vector + + +×-distributivity-+ : (X Y Z : Type) → (X ∔ Y) × Z ≅ (X × Z) ∔ (Y × Z) +×-distributivity-+ X Y Z = + record { bijection = f + ; bijectivity = record { inverse = g + ; η = section + ; ε = retraction } } + where + f : (X ∔ Y) × Z → (X × Z) ∔ (Y × Z) + f (inl x , z) = inl (x , z) + f (inr y , z) = inr (y , z) + + g : (X × Z) ∔ (Y × Z) → (X ∔ Y) × Z + g (inl (x , z)) = inl x , z + g (inr (y , z)) = inr y , z + + section : (g ∘ f) ∼ id + section (inl x , z) = refl (inl x , z) + section (inr y , z) = refl (inr y , z) + + retraction : (f ∘ g) ∼ id + retraction (inl (x , z)) = refl (inl (x , z)) + retraction (inr (y , z)) = refl (inr (y , z)) + +alternate : Type → Type → Bool → Type +alternate X _ true = X +alternate _ Y false = Y + +∔-isomorphic-to-Σ-bool : (X Y : Type) → (X ∔ Y) ≅ (Σ b ꞉ Bool , alternate X Y b) +∔-isomorphic-to-Σ-bool X Y = + record { bijection = f ; bijectivity = record { inverse = g + ; η = section + ; ε = retraction } } + where + f : X ∔ Y → Σ b ꞉ Bool , alternate X Y b + f (inl x) = true , x + f (inr y) = false , y + + g : (Σ b ꞉ Bool , alternate X Y b) → X ∔ Y + g (true , X) = inl X + g (false , Y) = inr Y + + section : (g ∘ f) ∼ id + section (inl x) = refl (inl x) + section (inr y) = refl (inr y) + + retraction : (f ∘ g) ∼ id + retraction (true , X) = refl (true , X) + retraction (false , Y) = refl (false , Y) + + +fib : ℕ → ℕ +fib 0 = 0 +fib 1 = 1 +fib (suc (suc n)) = fib n + fib (suc n) + +fib-aux : ℕ → ℕ → ℕ → ℕ +fib-aux a b 0 = b +fib-aux a b (suc n) = fib-aux (b + a) a n + +fib-fast : ℕ → ℕ +fib-fast = fib-aux 1 0 + +fib-aux-fib-lemma : (k n : ℕ) → fib-aux (fib (suc n)) (fib n) k ≡ fib (k + n) +fib-aux-fib-lemma zero n = refl (fib n) +fib-aux-fib-lemma (suc k) n = + fib-aux (fib n + fib (1 + n)) (fib (1 + n)) k ≡⟨ refl _ ⟩ + fib-aux (fib (2 + n)) (fib (1 + n)) k ≡⟨ IH ⟩ + fib (k + suc n) ≡⟨ ap fib (+-step k n) ⟩ + fib (suc (k + n)) ∎ + where + IH : fib-aux (fib (2 + n)) (fib (1 + n)) k ≡ fib (k + suc n) + IH = fib-aux-fib-lemma k (suc n) + +fib-fast-is-correct : (n : ℕ) → fib-fast n ≡ fib n +fib-fast-is-correct n = fib-fast n ≡⟨ refl _ ⟩ + fib-aux 1 0 n ≡⟨ claim ⟩ + fib (n + 0) ≡⟨ ap fib (+-base n) ⟩ + fib n ∎ + where + lemma : fib-aux (fib 1) (fib 0) n ≡ fib (n + 0) + lemma = fib-aux-fib-lemma n 0 + claim : fib-aux 1 0 n ≡ fib (n + 0) + claim = lemma +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6.lagda.md new file mode 100644 index 0000000..5caf9f6 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Homework/homework6.lagda.md @@ -0,0 +1,109 @@ +# Week 6 - Homework Sheet + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Homework.homework6 where + +open import prelude +open import isomorphisms +open import natural-numbers-functions +open import Fin +open import Vector +``` + +## Part I: Isomorphisms + +### Exercise 1.1 + +**Prove** the following isomorphism: + +```agda +×-distributivity-+ : (X Y Z : Type) → (X ∔ Y) × Z ≅ (X × Z) ∔ (Y × Z) +×-distributivity-+ = {!!} +``` + +### Exercise 1.2 + +We now define a function `alternate` that takes two types `X` and `Y` and +evaluates to `X` on `true` and evaluates to `Y` on `false`. + +```agda +alternate : Type → Type → Bool → Type +alternate X _ true = X +alternate _ Y false = Y +``` + +It can be proved that `Σ b ꞉ Bool , alternate X Y b` is the same thing as `X ∔ +Y`. **Prove** this by constructing the following isomorphism: + +```agda +∔-isomorphic-to-Σ-bool : (X Y : Type) → (X ∔ Y) ≅ (Σ b ꞉ Bool , alternate X Y b) +∔-isomorphic-to-Σ-bool = {!!} +``` + +## Part II: Proving correctness of efficient Fibonacci + +In Functional Programming you saw two ways of defining the Fibonacci function. +The first one, `fib` was fairly simple, but inefficient + +```agda +fib : ℕ → ℕ +fib 0 = 0 +fib 1 = 1 +fib (suc (suc n)) = fib n + fib (suc n) +``` + +Therefore, we introduce a second version, `fib-fast`, which uses an accumulating +parameter to make it more efficient. + +```agda +fib-aux : ℕ → ℕ → ℕ → ℕ +fib-aux a b 0 = b +fib-aux a b (suc n) = fib-aux (b + a) a n + +fib-fast : ℕ → ℕ +fib-fast = fib-aux 1 0 +``` + +It is not obvious, however, that `fib-fast` implements the same behaviour as +`fib`. In Agda, we can *prove* this, showing that `fib-fast` is correct. + +The following lemma relates `fib-aux` and `fib` and is fundamental in proving +the correctness of `fib-fast`. You will be asked to prove it later. + +```agda +fib-aux-fib-lemma : (k n : ℕ) → fib-aux (fib (suc n)) (fib n) k ≡ fib (k + n) +fib-aux-fib-lemma = {!!} -- Come back to this later +``` + +### Exercise 2.1 + +Using `fib-aux-fib-lemma`, **prove** the correctness of `fib-fast`. + +```agda +fib-fast-is-correct : (n : ℕ) → fib-fast n ≡ fib n +fib-fast-is-correct n = {!!} +``` + +*Hints*: +1. The lemma allows you to prove this fairly directly. There is no need to do +induction on `n`. +1. You may also find the `+-base` function from the +[natural-numbers-functions](../natural-numbers-functions.lagda.md) module useful. + +### Exercise 2.2 + +Now **complete** the proof of fundamental lemma `fib-aux-fib-lemma` above. + +*Hint*: You will probably want to use `+-step` from +[natural-numbers-functions](../natural-numbers-functions.lagda.md) at some +point. + +### References + +The exercises and solutions of Part 2 were based on [Neil Mitchell's +proof][mitchell] in the programming language [Idris][idris]. + +[mitchell]: http://neilmitchell.blogspot.com/2017/05/proving-fib-equivalence.html +[idris]: https://en.wikipedia.org/wiki/Idris_(programming_language) diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2-solutions.lagda.md new file mode 100644 index 0000000..e8d6436 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2-solutions.lagda.md @@ -0,0 +1,183 @@ +# Week 2 - Lab Sheet - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab2-solutions where + +open import prelude hiding (_||_) + +_≣_ : Bool → Bool → Type +true ≣ true = 𝟙 +true ≣ false = 𝟘 +false ≣ true = 𝟘 +false ≣ false = 𝟙 + +_==_ : Bool → Bool → Bool +true == true = true +true == false = false +false == true = false +false == false = true + +≡-to-≣ : (x y : Bool) → x ≡ y → x ≣ y +≡-to-≣ true true e = ⋆ +≡-to-≣ false false e = ⋆ + +is-true : Bool → Type +is-true b = b ≡ true + +≣-to-== : (x y : Bool) → x ≣ y → is-true (x == y) +≣-to-== true true e = refl true +≣-to-== false false e = refl true + +==-to-≡ : (x y : Bool) → is-true (x == y) → x ≡ y +==-to-≡ true true e = refl true +==-to-≡ false false e = refl false + +≡-to-== : (x y : Bool) → x ≡ y → is-true (x == y) +≡-to-== x y e = ≣-to-== x y (≡-to-≣ x y e) + +≣-to-≡ : (x y : Bool) → x ≣ y → x ≡ y +≣-to-≡ x y e = ==-to-≡ x y (≣-to-== x y e) + +==-to-≣ : (x y : Bool) → is-true (x == y) → x ≣ y +==-to-≣ x y e = ≡-to-≣ x y (==-to-≡ x y e) + + +data Either (A B : Type) : Type where + left : A → Either A B + right : B → Either A B + +if'_then_else_ : {A B : Type} → Bool → A → B → Either A B +if' true then x else y = left x +if' false then x else y = right y + + +_||_ : Bool → Bool → Bool +true || y = true +false || y = y + +_||'_ : Bool → Bool → Bool +true ||' true = true +true ||' false = true +false ||' true = true +false ||' false = false + + +||-assoc : (a b c : Bool) → a || (b || c) ≡ (a || b) || c +||-assoc true b c = refl true +||-assoc false b c = refl (b || c) + +||'-assoc : (a b c : Bool) → a ||' (b ||' c) ≡ (a ||' b) ||' c +||'-assoc true true true = refl true +||'-assoc true true false = refl true +||'-assoc true false true = refl true +||'-assoc true false false = refl true +||'-assoc false true true = refl true +||'-assoc false true false = refl true +||'-assoc false false true = refl true +||'-assoc false false false = refl false + +||-is-commutative : (a b : Bool) → a || b ≡ b || a +||-is-commutative true true = refl true +||-is-commutative true false = refl true +||-is-commutative false true = refl true +||-is-commutative false false = refl false + +false-left-unit-for-|| : (b : Bool) → false || b ≡ b +false-left-unit-for-|| = refl + +false-right-unit-for-|| : (b : Bool) → b || false ≡ b +false-right-unit-for-|| true = refl true +false-right-unit-for-|| false = refl false + +&&-is-associative : (a b c : Bool) → a && (b && c) ≡ (a && b) && c +&&-is-associative true b c = refl (b && c) +&&-is-associative false b c = refl false + +&&-is-commutative : (a b : Bool) → a && b ≡ b && a +&&-is-commutative true true = refl true +&&-is-commutative true false = refl false +&&-is-commutative false true = refl false +&&-is-commutative false false = refl false + +true-left-unit-for-&& : (b : Bool) → true && b ≡ b +true-left-unit-for-&& = refl + +true-right-unit-for-&& : (b : Bool) → b && true ≡ b +true-right-unit-for-&& true = refl true +true-right-unit-for-&& false = refl false + +&&-distributes-over-|| : (a b c : Bool) → a && (b || c) ≡ (a && b) || (a && c) +&&-distributes-over-|| true true c = refl true +&&-distributes-over-|| true false true = refl true +&&-distributes-over-|| true false false = refl false +&&-distributes-over-|| false b c = refl false + +||-distributes-over-&& : (a b c : Bool) → a || (b && c) ≡ (a || b) && (a || c) +||-distributes-over-&& true b c = refl true +||-distributes-over-&& false true c = refl c +||-distributes-over-&& false false c = refl false + + ++-suc-on-left : (x y : ℕ) → (suc x) + y ≡ suc (x + y) ++-suc-on-left zero y = refl (suc y) ++-suc-on-left (suc x) y = refl (suc (suc x) + y) + +max : ℕ → ℕ → ℕ +max zero n = n +max (suc m) zero = suc m +max (suc m) (suc n) = suc (max m n) + +min : ℕ → ℕ → ℕ +min zero n = zero +min (suc m) zero = zero +min (suc m) (suc n) = suc (min m n) + + ++-0-on-right : (x : ℕ) → x + 0 ≡ x ++-0-on-right zero = refl zero ++-0-on-right (suc x) = ap suc IH + where + IH : x + 0 ≡ x -- IH is short for induction hypothesis + IH = +-0-on-right x + ++-suc-on-right : (x y : ℕ) → x + suc y ≡ suc (x + y) ++-suc-on-right zero y = refl (suc y) ++-suc-on-right (suc x) y = ap suc IH + where + IH : x + suc y ≡ suc (x + y) + IH = +-suc-on-right x y + +max-idempotent : (x : ℕ) → max x x ≡ x +max-idempotent zero = refl zero +max-idempotent (suc x) = ap suc IH + where + IH : max x x ≡ x + IH = max-idempotent x + +max-commutative : (x y : ℕ) → max x y ≡ max y x +max-commutative zero zero = refl zero +max-commutative zero (suc y) = refl (suc y) +max-commutative (suc x) zero = refl (suc x) +max-commutative (suc x) (suc y) = ap suc IH + where + IH : max x y ≡ max y x + IH = max-commutative x y + +min-idempotent : (x : ℕ) → min x x ≡ x +min-idempotent zero = refl zero +min-idempotent (suc x) = ap suc IH + where + IH : min x x ≡ x + IH = min-idempotent x + +min-commutative : (x y : ℕ) → min x y ≡ min y x +min-commutative zero zero = refl zero +min-commutative zero (suc y) = refl zero +min-commutative (suc x) zero = refl zero +min-commutative (suc x) (suc y) = ap suc IH + where + IH : min x y ≡ min y x + IH = min-commutative x y +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2.lagda.md new file mode 100644 index 0000000..efa9e74 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab2.lagda.md @@ -0,0 +1,345 @@ +# Week 2 - Lab Sheet + +```agda +{-# OPTIONS --without-K --safe #-} +module Pool.Lab.lab2 where + +open import prelude hiding (if_then_else_;_||_) +``` + +## Part I: Booleans + +### Section 1: Operations on Booleans + +#### Exercise 1.1 + +We have already defined `if-then-else` like so: + +```agda +if_then_else_ : {A : Type} → Bool → A → A → A +if true then x else y = x +if false then x else y = y +``` + +But this requires our two branches to be of the same type A. Instead, we could +define `if-then-else` to have branches of different types, using the `Either` +datatype + +```agda +data Either (A B : Type) : Type where + left : A → Either A B + right : B → Either A B + +if'_then_else_ : {A B : Type} → Bool → A → B → Either A B +if' b then x else y = {!!} +``` + +**Define** this function. + +#### Exercise 1.2 + +We can define the `_||_` function in (at least) two different ways, depending on +how much pattern-matching we wish to perform: + +```agda +_||_ : Bool → Bool → Bool +true || y = {!!} +false || y = {!!} + +_||'_ : Bool → Bool → Bool +true ||' true = {!!} +true ||' false = {!!} +false ||' true = {!!} +false ||' false = {!!} +``` + +**Complete** the two holes for `_||_` and the four holes for `_||'_`. + +The `_||_` operator is *associative*, meaning `a || (b || c) ≡ (a || b) || c`. + +We can prove this for both of our definitions: + +```agda +||-assoc : (a b c : Bool) → a || (b || c) ≡ (a || b) || c +||-assoc a b c = {!!} + +||'-assoc : (a b c : Bool) → a ||' (b ||' c) ≡ (a ||' b) ||' c +||'-assoc a b c = {!!} +``` + +**Complete** both of these proofs. +(Hint: Because `||` was defined by pattern matching on the first argument, it +may be helpful to use the same strategy.) + +Which of these did you prefer proving, and why? + +#### Exercise 1.3 + +**Complete** the proof that Boolean disjunction is commutative: + +```agda +||-is-commutative : (a b : Bool) → a || b ≡ b || a +||-is-commutative a b = {!!} +``` + +#### Exercise 1.4 + +**Complete** the proof that `false` is the left unit element of `||`: + +```agda +false-left-unit-for-|| : (b : Bool) → false || b ≡ b +false-left-unit-for-|| b = {!!} +``` + +**Complete** the proof that `false` is the right unit element of `||`: + +```agda +false-right-unit-for-|| : (b : Bool) → b || false ≡ b +false-right-unit-for-|| b = {!!} +``` + +#### Exercise 1.5 + +Now prove the analogous properties for `&&`. + +**Complete** the proof that Boolean conjunction is associative: + +```agda +&&-is-associative : (a b c : Bool) → a && (b && c) ≡ (a && b) && c +&&-is-associative a b c = {!!} +``` + +**Complete** the proof that Boolean conjunction is commutative: + +```agda +&&-is-commutative : (a b : Bool) → a && b ≡ b && a +&&-is-commutative a b = {!!} +``` + +**Complete** the proof that that `true` is the left unit element of `&&`: + +```agda +true-left-unit-for-&& : (b : Bool) → true && b ≡ b +true-left-unit-for-&& b = {!!} +``` + +**Complete** the proof that that `true` is the right unit element of `&&`: + +```agda +true-right-unit-for-&& : (b : Bool) → b && true ≡ b +true-right-unit-for-&& b = {!!} +``` + +#### Exercise 1.6 + +An important algebraic property between two operators is *distributivity*. For +example, multiplication distributes over addition since `x * (y + z) = (x * y) + +(x * z)` for every `x, y, z ∈ ℕ`. The same is the case for the Boolean +conjunction and disjunction operators we have defined. + +**Complete** the proof that Boolean conjunction distributes over Boolean +disjunction: + +```agda +&&-distributes-over-|| : (a b c : Bool) → a && (b || c) ≡ (a && b) || (a && c) +&&-distributes-over-|| a b c = {!!} +``` + +```agda +||-distributes-over-&& : (a b c : Bool) → a || (b && c) ≡ (a || b) && (a || c) +||-distributes-over-&& a b c = {!!} +``` + + +### Section 2: Identity type and `Bool` + +With these exercises, you will practise with the identity type and `Bool`. + +#### Exercise 2.1 + +Recall how we defined `_≣_` for [natural +numbers](../introduction.lagda.md#the-empty-type-and-the-unit-type). We could do +the same for the Booleans. + +```agda +_≣_ : Bool → Bool → Type +x ≣ y = {!!} +``` + +**Define** this function. + +Tip: you can type `≣` as `\===`. + +#### Exercise 2.2 + +We can also define a Boolean-valued equality function. + +```agda +_==_ : Bool → Bool → Bool +x == y = {!!} +``` + +**Define** this function. + +#### Exercise 2.3 + +Now we write a conversion function from `x ≡ y` to `x ≣ y`, just like the +function `back` in +[introduction](../introduction.lagda.md#the-identity-type-former-__). + +```agda +≡-to-≣ : (x y : Bool) → x ≡ y → x ≣ y +≡-to-≣ x y e = {!!} +``` + +**Complete** this function. + +#### Exercise 2.4 + +The following helper function makes some functions nicer to read. + +```agda +is-true : Bool → Type +is-true b = b ≡ true +``` + +Now we can consider another conversion function. + +```agda +≣-to-== : (x y : Bool) → x ≣ y → is-true (x == y) +≣-to-== x y e = {!!} +``` + +**Define** this function. + +#### Exercise 2.5 + +And similarly, we have + +```agda +==-to-≡ : (x y : Bool) → is-true (x == y) → x ≡ y +==-to-≡ x y e = {!!} +``` + +**Define** this function. + +#### Exercise 2.6 + +Finally, we can use the above functions to define the three remaining +implications + +```agda +≡-to-== : (x y : Bool) → x ≡ y → is-true (x == y) +≡-to-== x y e = {!!} + +≣-to-≡ : (x y : Bool) → x ≣ y → x ≡ y +≣-to-≡ x y e = {!!} + +==-to-≣ : (x y : Bool) → is-true (x == y) → x ≣ y +==-to-≣ x y e = {!!} +``` + +**Complete** these functions. + + +## Part II: Natural numbers + +### Section 1: Some functions on natural numbers + +#### Exercise 1.1 + +**Complete** the proof of the fact that `(suc x) + y ≡ suc (x + y)` + +```agda ++-suc-on-left : (x y : ℕ) → (suc x) + y ≡ suc (x + y) ++-suc-on-left x y = {!!} +``` + +#### Exercise 1.2 + +We can define `max` operation on natural numbers as follows: + +```agda +max : ℕ → ℕ → ℕ +max zero n = n +max (suc m) zero = suc m +max (suc m) (suc n) = suc (max m n) +``` + +**Define** the function `min` analogously: + +```agda +min : ℕ → ℕ → ℕ +min m n = {!!} +``` + +### Section 2: Natural numbers and proofs using induction + +The function `+-0-on-right` expresses the fact that `0` is a right unit of the +operation `_+_`. Notice how we use an inductive hypothesis (the recursive call) +in the proof. (Spelling this out, `IH` tells us that `x + 0 ≡ x`, so that `ap +suc IH` witnesses that `suc (x + 0) ≡ suc x`.) + +```agda ++-0-on-right : (x : ℕ) → x + 0 ≡ x ++-0-on-right zero = refl zero ++-0-on-right (suc x) = ap suc IH + where + IH : x + 0 ≡ x -- IH is short for induction hypothesis + IH = +-0-on-right x +``` + +Using similar induction hypotheses, try to complete the following exercises. + + +#### Exercise 2.1 + +**Complete** the proof: + +```agda ++-suc-on-right : (x y : ℕ) → x + suc y ≡ suc (x + y) ++-suc-on-right x y = {!!} +``` + +#### Exercise 2.2 + +In algebra, an operator `_*_` is called idempotent iff `x * x = x` for every +`x`. `max` is an example of an idempotent operator: + +**Complete** the hole to prove that `max` is idempotent: + +```agda +max-idempotent : (x : ℕ) → max x x ≡ x +max-idempotent x = {!!} +``` + +#### Exercise 2.3 + +**Complete** the hole by constructing a proof of the fact that `max` is a +commutative operator: + +```agda +max-commutative : (x y : ℕ) → max x y ≡ max y x +max-commutative x y = {!!} +``` + +#### Exercise 2.4 + +Now recall that we defined `min` in Exercise 1.2. Similarly, we can show that +`min` is idempotent and commutative too. + +**Complete** the proof that `min` is idempotent: + +```agda +min-idempotent : (x : ℕ) → min x x ≡ x +min-idempotent x = {!!} +``` + +#### Exercise 2.5 + +**Complete** the proof that `min` is commutative: + +```agda +min-commutative : (x y : ℕ) → min x y ≡ min y x +min-commutative x y = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3-solutions.lagda.md new file mode 100644 index 0000000..075327e --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3-solutions.lagda.md @@ -0,0 +1,122 @@ +# Week 3 - Lab Sheet - Solutions + +```agda +{-# OPTIONS --without-K --safe #-} +module Pool.Lab.lab3-solutions where + +open import prelude hiding (𝟘-nondep-elim) + +¬¬_ : Type → Type +¬¬ A = ¬ (¬ A) + +¬¬¬ : Type → Type +¬¬¬ A = ¬ (¬¬ A) + +∔-introduction-left : {A B : Type} → A → A ∔ B +∔-introduction-left a = inl a + +∔-introduction-right : {A B : Type} → B → A ∔ B +∔-introduction-right b = inr b + +∔-elimination : {A B X : Type} → (A → X) → (B → X) → (A ∔ B → X) +∔-elimination p q (inl a) = p a +∔-elimination p q (inr b) = q b + + +×-elimination-left : {A B : Type} → A × B → A +×-elimination-left (a , b) = a + +×-elimination-right : {A B : Type} → A × B → B +×-elimination-right (a , b) = b + +×-introduction : {A B : Type} → A → B → A × B +×-introduction a b = (a , b) + +×-introduction' : {A B X : Type} → (X → A) → (X → B) → (X → A × B) +×-introduction' p q x = (p x , q x) + + +uncurry : {A B X : Type} → (A → B → X) → (A × B → X) +uncurry p (a , b) = p a b + +curry : {A B X : Type} → (A × B → X) → (A → B → X) +curry p a b = p (a , b) + +→-trans : {A B C : Type} → (A → B) → (B → C) → (A → C) +→-trans p q = λ a → q (p a) + + +𝟘-nondep-elim : {A : Type} → 𝟘 → A +𝟘-nondep-elim = λ () + +¬¬-introduction : {A : Type} → A → ¬¬ A +¬¬-introduction a = λ p → p a + +not-implies-not³ : {A : Type} → ¬ A → ¬¬¬ A +not-implies-not³ p = λ q → q p + +not³-implies-not : {A : Type} → ¬¬¬ A → ¬ A +not³-implies-not p = λ a → p (λ q → q a) + +contraposition : {A B : Type} → (A → B) → ¬ B → ¬ A +contraposition p q = λ a → q (p a) + +¬¬-functor : {A B : Type} → (A → B) → ¬¬ A → ¬¬ B +¬¬-functor p = contraposition (contraposition p) + +¬¬-kleisli : {A B : Type} → (A → ¬¬ B) → ¬¬ A → ¬¬ B +¬¬-kleisli p = contraposition (λ q → λ a → p a q) + + +de-morgan₁ : {A B : Type} → ¬ (A ∔ B) → ¬ A × ¬ B +de-morgan₁ {A} {B} p = ×-introduction p₁ p₂ + where + p₁ : ¬ A + p₁ = p ∘ ∔-introduction-left + p₂ : ¬ B + p₂ = p ∘ ∔-introduction-right + +de-morgan₂ : {A B : Type} → ¬ A ∔ ¬ B → ¬ (A × B) +de-morgan₂ = ∔-elimination (→-trans ×-elimination-left) + (→-trans ×-elimination-right) + +¬-and-+-exercise₁ : {A B : Type} → ¬ A ∔ B → A → B +¬-and-+-exercise₁ (inl p) a = 𝟘-nondep-elim (p a) +¬-and-+-exercise₁ (inr q) _ = q + +¬-and-+-exercise₂ : {A B : Type} → ¬ A ∔ B → ¬ (A × ¬ B) +¬-and-+-exercise₂ (inl p) = λ { (a , _) → p a } +¬-and-+-exercise₂ (inr q) = λ { (a , r) → r q } + +distributivity₁ : {A B C : Type} → (A × B) ∔ C → (A ∔ C) × (B ∔ C) +distributivity₁ {A} {B} {C} = ∔-elimination q₁ q₂ + where + q₁ : A × B → (A ∔ C) × (B ∔ C) + q₁ (a , b) = (inl a , inl b) + q₂ : C → A ∔ C × B ∔ C + q₂ c = (inr c , inr c) + +distributivity₂ : {A B C : Type} → (A ∔ B) × C → (A × C) ∔ (B × C) +distributivity₂ (p , c) = + ∔-elimination (λ a → inl (a , c)) (λ b → inr (b , c)) p + + +Σ-introduction : {A : Type} {B : (A → Type)} → (a : A) → B a → (Σ a ꞉ A , B a) +Σ-introduction a b = (a , b) + +Σ-to-× : {A : Type} {B : (A → Type)} → ((a , _) : (Σ a ꞉ A , B a)) → A × B a +Σ-to-× (a , b) = (a , b) + +Σ-on-Bool : {B : Bool → Type} → (Σ x ꞉ Bool , B x) → B true ∔ B false +Σ-on-Bool (true , b) = inl b +Σ-on-Bool (false , b) = inr b + + +Π-apply : {A : Type} {B : (A → Type)} → ((a : A) → B a) → (a : A) → B a +Π-apply p a = p a + +Π-→ : {A : Type} {B C : A → Type} + → ((a : A) → B a → C a) + → ((a : A) → B a) → ((a : A) → C a) +Π-→ p q = λ a → p a (q a) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3.lagda.md new file mode 100644 index 0000000..b70849c --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab3.lagda.md @@ -0,0 +1,261 @@ +# Week 3 - Lab Sheet + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab3 where + +open import prelude hiding (𝟘-nondep-elim) +``` + +Before we proceed with the exercises, we introduce a some convenient notation +for multiple negations. + +```agda +¬¬_ : Type → Type +¬¬ A = ¬ (¬ A) + +¬¬¬ : Type → Type +¬¬¬ A = ¬ (¬¬ A) +``` + +## Part I: Propositional logic + +### Section 1: Disjunction + +#### Exercise 1.1 + +**Complete** the following proofs involving disjunctions. + +```agda +∔-introduction-left : {A B : Type} → A → A ∔ B +∔-introduction-left = {!!} + +∔-introduction-right : {A B : Type} → B → A ∔ B +∔-introduction-right = {!!} +``` + +#### Exercise 1.2 + +**Complete** the proof about disjunctions. + +```agda +∔-elimination : {A B X : Type} → (A → X) → (B → X) → (A ∔ B → X) +∔-elimination = {!!} +``` + +### Section 2: Conjunction + +#### Exercise 2.1 + +**Complete** the following proofs involving conjunctions. + +```agda +×-elimination-left : {A B : Type} → A × B → A +×-elimination-left = {!!} + +×-elimination-right : {A B : Type} → A × B → B +×-elimination-right = {!!} +``` + +#### Exercise 2.2 + +**Prove** the following: + +```agda +×-introduction : {A B : Type} → A → B → A × B +×-introduction = {!!} + +×-introduction' : {A B X : Type} → (X → A) → (X → B) → (X → A × B) +×-introduction' = {!!} +``` + +### Section 3: Implication + +#### Exercise 3.1 + +**Complete** the following proofs involving implications. + +```agda +uncurry : {A B X : Type} → (A → B → X) → (A × B → X) +uncurry = {!!} + +curry : {A B X : Type} → (A × B → X) → (A → B → X) +curry = {!!} +``` + +You probably already know `curry` and `uncurry` from Haskell, but notice how we +can read them from a logical perspective: `uncurry` says that if `A` implies +that `B` implies `X`, then the conjunction of `A` and `B` implies `X`. + +#### Exercise 3.2 + +**Prove** that implication is transitive. + +``` +→-trans : {A B C : Type} → (A → B) → (B → C) → (A → C) +→-trans = {!!} +``` + +Notice that the proof that implication is transitive is just function +composition. + + +### Section 4: Negation + +The fact that falsity implies everything is known as the [_principle of +explosion_](https://en.wikipedia.org/wiki/Principle_of_explosion) or _ex falso +quodlibet_. + +**Complete** the proof of the principle of explosion in Agda. + +#### Exercise 4.1 + +```agda +𝟘-nondep-elim : {A : Type} → 𝟘 → A +𝟘-nondep-elim = {!!} +``` + +#### Exercise 4.2 + +**Complete** the proof a proposition implies its own double negation. + +```agda +¬¬-introduction : {A : Type} → A → ¬¬ A +¬¬-introduction = {!!} +``` + +#### Exercise 4.3 + +**Prove** that having three negations is the logically equivalent to having a +single negation. + +```agda +not-implies-not³ : {A : Type} → ¬ A → ¬¬¬ A +not-implies-not³ = {!!} + +not³-implies-not : {A : Type} → ¬¬¬ A → ¬ A +not³-implies-not = {!!} +``` + +#### Exercise 4.4 + +**Complete** the proof of contraposition. + +```agda +contraposition : {A B : Type} → (A → B) → ¬ B → ¬ A +contraposition = {!!} +``` + +#### Exercise 4.5 + +Use `contraposition` to **complete** the following two proofs. + +```agda +¬¬-functor : {A B : Type} → (A → B) → ¬¬ A → ¬¬ B +¬¬-functor = {!!} + +¬¬-kleisli : {A B : Type} → (A → ¬¬ B) → ¬¬ A → ¬¬ B +¬¬-kleisli = {!!} +``` +{- + Better HINT: + start with f : A → ¬¬ B + apply ¬¬-functor to get ¬¬ A → ¬¬¬¬ B + Now use ¬¬¬¬ B → ¬¬ B which is a particular case of not-implies-not³ with A = ¬ B + + HINT: + start with f : A → ¬¬ B + apply contraposition to get ¬¬¬ B → ¬A + do this again to get ¬¬ A → ¬¬¬¬ B + Now use ¬¬¬¬ B → ¬¬ B which is a particular case of not-implies-not³ with A = ¬ B +-} + +### Section 5: De Morgan Laws and logical laws + +The De Morgan laws cannot be proved in Agda, though some of the implications +involved in the De Morgan laws _can_ be. The following exercises will involve +proving these (and some other similar laws) for Agda types. + +#### Exercise 5.1 + +**Complete** the proofs. + +```agda +de-morgan₁ : {A B : Type} → ¬ (A ∔ B) → ¬ A × ¬ B +de-morgan₁ = {!!} + +de-morgan₂ : {A B : Type} → ¬ A ∔ ¬ B → ¬ (A × B) +de-morgan₂ = {!!} +``` + +#### Exercise 5.2 + +**Complete** the proofs. + +```agda +¬-and-+-exercise₁ : {A B : Type} → ¬ A ∔ B → A → B +¬-and-+-exercise₁ = {!!} + +¬-and-+-exercise₂ : {A B : Type} → ¬ A ∔ B → ¬ (A × ¬ B) +¬-and-+-exercise₂ = {!!} +``` + +#### Exercise 5.3 + +**Prove** the distributivity laws for `×` and `∔`. + +```agda +distributivity₁ : {A B C : Type} → (A × B) ∔ C → (A ∔ C) × (B ∔ C) +distributivity₁ = {!!} + +distributivity₂ : {A B C : Type} → (A ∔ B) × C → (A × C) ∔ (B × C) +distributivity₂ = {!!} +``` + +## Part II: Logic with quantifiers + +### Section 1: Sums + +#### Exercise 1.1 + +**Complete** the following constructions. + +```agda +Σ-introduction : {A : Type} {B : (A → Type)} → (a : A) → B a → (Σ a ꞉ A , B a) +Σ-introduction = {!!} + +Σ-to-× : {A : Type} {B : (A → Type)} → ((a , _) : (Σ a ꞉ A , B a)) → A × B a +Σ-to-× = {!!} +``` + +#### Exercise 1.2 + +**Complete** the following proof about sums over Booleans. + +```agda +Σ-on-Bool : {B : Bool → Type} → (Σ x ꞉ Bool , B x) → B true ∔ B false +Σ-on-Bool = {!!} +``` + +### Section 2: Products + +#### Exercise 2.1 + +Complete the proof. + +```agda +Π-apply : {A : Type} {B : (A → Type)} → ((a : A) → B a) → (a : A) → B a +Π-apply = {!!} +``` + +#### Exercise 2.2 + +**Prove** the following. + +```agda +Π-→ : {A : Type} {B C : A → Type} + → ((a : A) → B a → C a) + → ((a : A) → B a) → ((a : A) → C a) +Π-→ = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4-solutions.lagda.md new file mode 100644 index 0000000..fc8cb1c --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4-solutions.lagda.md @@ -0,0 +1,89 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab4-solutions where + +open import prelude +open import List-functions + +reverse-lemma : {X : Type} (x : X) (xs : List X) + → x :: reverse xs ≡ reverse (xs ++ [ x ]) +reverse-lemma x [] = refl (x :: []) +reverse-lemma x (y :: xs) = ap (_++ [ y ]) (reverse-lemma x xs) + +reverse-is-involution : {X : Type} → (xs : List X) → xs ≡ reverse (reverse xs) +reverse-is-involution {X} [] = refl [] +reverse-is-involution {X} (x :: xs) + = x :: xs ≡⟨ ap (x ::_) (reverse-is-involution xs) ⟩ + x :: reverse (reverse xs) ≡⟨ reverse-lemma x (reverse xs) ⟩ + reverse (reverse xs ++ [ x ]) ≡⟨ refl (reverse (reverse xs ++ [ x ])) ⟩ + reverse (reverse (x :: xs)) ∎ + ++-assoc : (x y z : ℕ) → x + (y + z) ≡ (x + y) + z ++-assoc zero y z = refl (y + z) ++-assoc (suc x) y z = ap suc (+-assoc x y z) + + +data _≤_ : ℕ → ℕ → Type where + ≤-zero : ( y : ℕ) → 0 ≤ y + ≤-suc : (x y : ℕ) → x ≤ y → suc x ≤ suc y + +_≤'_ : ℕ → ℕ → Type +x ≤' y = Σ k ꞉ ℕ , x + k ≡ y + + +≤'-zero : (y : ℕ) → 0 ≤' y +≤'-zero y = y , refl y + +≤'-suc : (x y : ℕ) → x ≤' y → suc x ≤' suc y +≤'-suc x y (n , p) = n , ap suc p + +≤-prime : (x y : ℕ) → x ≤ y → x ≤' y +≤-prime 0 y (≤-zero y) = ≤'-zero y +≤-prime (suc x) (suc y) (≤-suc x y p) = ≤'-suc x y (≤-prime x y p) + +≤-unprime : (x y : ℕ) → x ≤' y → x ≤ y +≤-unprime zero y (n , p) = ≤-zero y +≤-unprime (suc x) (suc .(x + n)) (n , refl .(suc (x + n))) + = ≤-suc x (x + n) (≤-unprime x (x + n) (n , refl (x + n))) + +≤-trans : (x y z : ℕ) → x ≤ y → y ≤ z → x ≤ z +≤-trans zero y z p q = ≤-zero z +≤-trans (suc x) .(suc y) .(suc z) (≤-suc .x y p) (≤-suc .y z q) + = ≤-suc x z (≤-trans x y z p q) + +≤'-trans : (x y z : ℕ) → x ≤' y → y ≤' z → x ≤' z +≤'-trans x .(x + n) .((x + n) + m) (n , refl .(x + n)) (m , refl .((x + n) + m)) + = (n + m) , +-assoc x n m + +is-decidable : Type → Type +is-decidable A = A ∔ ¬ A + +has-decidable-equality : Type → Type +has-decidable-equality A = (x y : A) → is-decidable (x ≡ y) + +bool-has-decidable-equality : has-decidable-equality Bool +bool-has-decidable-equality true true = inl (refl true) +bool-has-decidable-equality true false = inr (λ ()) +bool-has-decidable-equality false true = inr (λ ()) +bool-has-decidable-equality false false = inl (refl false) + +≤-lemma : (m n : ℕ) → suc m ≤ suc n → m ≤ n +≤-lemma m n (≤-suc m n p) = p + +not-suc-≤-zero : (n : ℕ) → ¬ (suc n ≤ 0) +not-suc-≤-zero n () + +≤-is-decidable : (m n : ℕ) → is-decidable (m ≤ n) +≤-is-decidable zero zero = inl (≤-zero zero) +≤-is-decidable zero (suc n) = inl (≤-zero (suc n)) +≤-is-decidable (suc m) zero = inr (not-suc-≤-zero m) +≤-is-decidable (suc m) (suc n) = ∔-nondep-elim f g IH + where + IH : (m ≤ n) ∔ ¬ (m ≤ n) + IH = ≤-is-decidable m n + f : m ≤ n → is-decidable (suc m ≤ suc n) + f p = inl (≤-suc m n p) + g : ¬ (m ≤ n) → is-decidable (suc m ≤ suc n) + g p = inr (λ q → p (≤-lemma m n q)) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4.lagda.md new file mode 100644 index 0000000..9a3848f --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab4.lagda.md @@ -0,0 +1,210 @@ +# Week 4 - Lab Sheet + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab4 where + +open import prelude +open import List-functions +``` + +## Part I: Reverse is an involution + +We wish to prove that the `reverse` function on lists is an involution; +that is to say that reversing a list twice is the same as doing nothing. + +### Exercise 1.1 + +First, we will prove the following lemma. + +```agda +reverse-lemma : {X : Type} (x : X) (xs : List X) + → x :: reverse xs ≡ reverse (xs ++ [ x ]) +reverse-lemma = {!!} +``` + +**Prove** the lemma about `reverse`. + +### Exercise 1.2 + +The following proof skeleton has been provided for you, using notation for +equational reasoning. + +```agda +reverse-is-involution : {X : Type} → (xs : List X) → xs ≡ reverse (reverse xs) +reverse-is-involution [] = {!!} +reverse-is-involution (x :: xs) + = x :: xs ≡⟨ {!!} ⟩ + x :: reverse (reverse xs) ≡⟨ {!!} ⟩ + reverse (reverse xs ++ [ x ]) ≡⟨ {!!} ⟩ + reverse (reverse (x :: xs)) ∎ +``` + +**Fill** the holes of our proof that `reverse` is an involution. + +## Part II: Associativity of addition of natural numbers + +We wish to prove the associativity of `_+_` on the natural numbers. + +```agda ++-assoc : (x y z : ℕ) → x + (y + z) ≡ (x + y) + z ++-assoc = {!!} +``` + +**Complete** the proof that `_+_` is associative. + +## Part III: Order on the natural numbers + +In this part we will study two ways of expressing that a natural number is less +than or equal to another natural number. + +The first definition is an inductive one. + +```agda +data _≤_ : ℕ → ℕ → Type where + ≤-zero : ( y : ℕ) → 0 ≤ y + ≤-suc : (x y : ℕ) → x ≤ y → suc x ≤ suc y +``` + +It says: +1. that zero is less than or equal to any natural number; +1. if `x` is less than or equal to `y`, then `suc x` is less than or equal to `suc y`. + +The second definition uses a `Σ`-type. + +```agda +_≤'_ : ℕ → ℕ → Type +x ≤' y = Σ k ꞉ ℕ , x + k ≡ y +``` + +It says that `x` is less than or equal to `y` if we have some `k : ℕ` +such that `x + k ≡ y`. + +We will prove that the two definitions are logically equivalent. + +### Exercise 3.1 + +In order to prove that the first definition implies the second, we first +prove two little lemmas about `_≤'_`. + +Note that they amount to the constructors of `_≤_`. + +```agda +≤'-zero : ( y : ℕ) → 0 ≤' y +≤'-zero = {!!} + +≤'-suc : (x y : ℕ) → x ≤' y → suc x ≤' suc y +≤'-suc = {!!} +``` + +**Prove** the two little lemmas about `_≤'_`. + +### Exercise 3.2 + +We now prove that the first definition implies the second. + +```agda +≤-prime : (x y : ℕ) → x ≤ y → x ≤' y +≤-prime = {!!} +``` + +**Prove** that `x ≤ y` implies `x ≤' y` using the little lemmas from 3.1. + +### Exercise 3.3 + +The other direction is slightly trickier. + +```agda +≤-unprime : (x y : ℕ) → x ≤' y → x ≤ y +≤-unprime = {!!} +``` + +**Prove** that `x ≤' y` implies `x ≤ y`. + +*Hint:* The base case only requires pattern matching on `x`, whereas +the inductive case requires further pattern matching. + +### Exercise 3.4 + +The order on the natural numbers is transitive, meaning that if +`x ≤ y` and `y ≤ z` then also `x ≤ z`. We can prove this for +both our definitions of the order. + +```agda +≤-trans : (x y z : ℕ) → x ≤ y → y ≤ z → x ≤ z +≤-trans = {!!} + +≤'-trans : (x y z : ℕ) → x ≤' y → y ≤' z → x ≤' z +≤'-trans = {!!} +``` + +**Complete** the proofs that `_≤_` and `_≤'_` are transitive. + +## Part IV: Decidability and decidable equality + +In last week's exercises, you constructed proofs of logical formulae by writing +Agda programs through the interpretation of Agda types as propositions. We +mentioned, however, that this interpretation does not make sense _a priori_ in +the setting of classical logic and we said that we have to work _constructively_ +to make logical sense of Agda's types. More specifically, we mentioned in +Exercise 2.3 of Homework 3 that the logical interpretation of the law of +excluded middle `(A : Type) → A ∔ ¬ A`, is not provable in Agda. + +Notice, however, that the statement + +```txt + (A : Type) → A ∔ ¬ A +``` + +is to be viewed as asserting that _the law of excluded middle holds for all +types_. Even though this fails to hold in the context of Agda's type system, it +doesn't mean that the law of excluded middle does not hold for _some_ types. In +this exercise, this is precisely the question that we will be interested in; we +will study “types that admit the law of excluded middle”. Such types are called +**decidable types**. We will express this notion through the Agda predicate +`is-decidable`: + +```agda +is-decidable : Type → Type +is-decidable A = A ∔ ¬ A +``` + +To assert `is-decidable A` for some type `A` is to say that type `A` satisfies +the law of excluded middle: we can either construct an inhabitant of `A` or +prove that the existence of an inhabitant for `A` is impossible. + +We can consider the decidability of any type but we will often be interested in +the _decidability of equality types_. A type `A` is said to have _decidable +equality_ iff for any two `x y : A`, the identity type `x ≡ y` is decidable. We +can write this notion down in Agda as follows: + +```agda +has-decidable-equality : Type → Type +has-decidable-equality A = (x y : A) → is-decidable (x ≡ y) +``` + +### Exercise 4.1 + +To familiarise yourself with the notion of decidable equality, **prove** that +the `Bool` type has decidable equality: + +```agda +bool-has-decidable-equality : has-decidable-equality Bool +bool-has-decidable-equality = {!!} +``` + +### Exercise 4.2 + +**Prove** that the `_≤_` relation on `ℕ` is decidable. You will have to use +the following lemma: + +```agda +≤-lemma : (m n : ℕ) → suc m ≤ suc n → m ≤ n +≤-lemma m n (≤-suc m n p) = p +``` + +```agda +≤-is-decidable : (m n : ℕ) → is-decidable (m ≤ n) +≤-is-decidable = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5-solutions.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5-solutions.lagda.md new file mode 100644 index 0000000..fd2acc9 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5-solutions.lagda.md @@ -0,0 +1,219 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab5-solutions where + +open import prelude +open import natural-numbers-functions hiding (_≤_ ; is-even ; +-assoc ; +-comm) +open import List-functions +open import isomorphisms + +open import decidability +open import negation + +open import Pool.Homework.homework4-solutions +open import Pool.Lab.lab4-solutions + + +{- Verbose solution, manually performing and showing many equalities that hold + by definition. -} +map-preserves-length : {A B : Type} (f : A → B) (xs : List A) + → length (map f xs) ≡ length xs +map-preserves-length f [] = refl _ +map-preserves-length f (x :: xs) = + length (map f (x :: xs)) ≡⟨ refl _ ⟩ + length (f x :: (map f xs)) ≡⟨ refl _ ⟩ + 1 + length (map f xs) ≡⟨ ap (1 +_) IH ⟩ + 1 + length xs ∎ + where + IH : length (map f xs) ≡ length xs + IH = map-preserves-length f xs + +{- Alternative, shorter solution that relies on Agda computing the necessary + equalities for us. + We come up with such proofs by looking at the goal, using `C-c C-,`, which + we sometimes normalize (by prefixing the command by `C-u C-u`) to let Agda + do even more work for us. -} +map-preserves-length' : {A B : Type} (f : A → B) (xs : List A) + → length (map f xs) ≡ length xs +map-preserves-length' f [] = refl _ +map-preserves-length' f (x :: xs) = ap suc (map-preserves-length' f xs) + +{- Verbose solution, manually performing and showing many equalities that hold + by definition. -} +length-of-++ : {A : Type} (xs ys : List A) + → length (xs ++ ys) ≡ length xs + length ys +length-of-++ [] ys = refl (length ys) +length-of-++ (x :: xs) ys = + length ((x :: xs) ++ ys) ≡⟨ refl _ ⟩ + length (x :: (xs ++ ys)) ≡⟨ refl _ ⟩ + (1 + length (xs ++ ys)) ≡⟨ ap (1 +_) IH ⟩ + (1 + (length xs + length ys)) ≡⟨ +-assoc 1 (length xs) (length ys) ⟩ + (1 + length xs) + length ys ≡⟨ refl _ ⟩ + length (x :: xs) + length ys ∎ + where + IH : length (xs ++ ys) ≡ length xs + length ys + IH = length-of-++ xs ys + + {- Alternative, shorter solution that relies on Agda computing the necessary + equalities for us. + We come up with such proofs by looking at the goal, using `C-c C-,`, which + we sometimes normalize (by prefixing the command by `C-u C-u`) to let Agda + do even more work for us. -} +length-of-++' : {A : Type} (xs ys : List A) + → length (xs ++ ys) ≡ length xs + length ys +length-of-++' [] ys = refl (length ys) +length-of-++' (x :: xs) ys = ap suc IH + where + IH : length (xs ++ ys) ≡ length xs + length ys + IH = length-of-++ xs ys + +length-of-prefix : {A : Type} (xs ys : List A) + → xs ≼' ys + → length xs ≤' length ys +length-of-prefix xs ys (zs , e) = (length zs , goal) + where + goal = length xs + length zs ≡⟨ sym (length-of-++ xs zs) ⟩ + length (xs ++ zs) ≡⟨ ap length e ⟩ + length ys ∎ + +is-nonempty : {A : Type} → List A → Type +is-nonempty xs = 1 ≤' length xs + +tail : {A : Type} (xs : List A) → is-nonempty xs → List A +tail (x :: xs) p = xs + +head : {A : Type} (xs : List A) → is-nonempty xs → A +head (x :: xs) p = x + +length-of-tail : {A : Type} (xs : List A) (p : is-nonempty xs) + → 1 + length (tail xs p) ≡ length xs +length-of-tail (x :: xs) p = + 1 + length (tail (x :: xs) p) ≡⟨ refl _ ⟩ + 1 + length xs ≡⟨ refl _ ⟩ + length (x :: xs) ∎ + +≤'-suc-lemma : (n : ℕ) → n ≤' (1 + n) +≤'-suc-lemma zero = (1 , refl 1) +≤'-suc-lemma (suc n) = (1 , goal) + where + goal : suc n + 1 ≡ 1 + suc n + goal = +-comm (suc n) 1 + +length-of-tail-decreases : {A : Type} (xs : List A) (p : is-nonempty xs) + → length (tail xs p) ≤' length xs +length-of-tail-decreases (x :: xs) p = goal + where + goal : length xs ≤' (1 + length xs) + goal = ≤'-suc-lemma (length xs) + +×-iso : (X Y : Type) → X × Y ≅ Y × X +×-iso X Y = record { bijection = f ; bijectivity = f-is-bijection } + where + f : X × Y → Y × X + f (x , y) = y , x + + g : Y × X → X × Y + g (y , x) = x , y + + gf : g ∘ f ∼ id + gf (x , y) = refl (x , y) + + fg : f ∘ g ∼ id + fg (y , x) = refl (y , x) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } + ++-iso : (X Y : Type) → X ∔ Y ≅ Y ∔ X ++-iso X Y = record { bijection = f ; bijectivity = f-is-bijection } + where + f : X ∔ Y → Y ∔ X + f (inl x) = inr x + f (inr y) = inl y + + g : Y ∔ X → X ∔ Y + g (inl y) = inr y + g (inr x) = inl x + + gf : g ∘ f ∼ id + gf (inl x) = refl (inl x) + gf (inr y) = refl (inr y) + + fg : f ∘ g ∼ id + fg (inl y) = refl (inl y) + fg (inr x) = refl (inr x) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } + +lists-from-vectors : {A : Type} → List A ≅ (Σ n ꞉ ℕ , Vector A n) +lists-from-vectors {A} + = record { bijection = f ; bijectivity = f-is-bijection } + where + f : List A → Σ n ꞉ ℕ , Vector A n + f [] = 0 , [] + f (x :: xs) = suc (fst (f xs)) , (x :: snd (f xs)) + + g : Σ n ꞉ ℕ , Vector A n → List A + g (0 , [] ) = [] + g (suc n , (x :: xs)) = x :: g (n , xs) + + gf : g ∘ f ∼ id + gf [] = refl [] + gf (x :: xs) = ap (x ::_) (gf xs) + + fg : f ∘ g ∼ id + fg (0 , [] ) = refl (0 , []) + fg (suc n , (x :: xs)) + = ap (λ - → suc (fst -) , (x :: snd -)) (fg (n , xs)) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } + +open _≅_ + +lfv-preserves-length : {A : Type} (xs : List A) + → length xs + ≡ fst (bijection lists-from-vectors xs) +lfv-preserves-length [] = refl 0 +lfv-preserves-length (x :: xs) = ap suc (lfv-preserves-length xs) + +is-even : ℕ → Type +is-even x = Σ y ꞉ ℕ , x ≡ 2 * y + +check-even : ℕ → Bool +check-even zero = true +check-even (suc zero) = false +check-even (suc (suc n)) = check-even n + +evenness-lemma₁ : (n : ℕ) → is-even (2 + n) → is-even n +evenness-lemma₁ n (suc k , p) = k , goal + where + subgoal : suc (suc n) ≡ suc (suc (2 * k)) + subgoal = suc (suc n) ≡⟨ p ⟩ + suc k + suc k ≡⟨ ap suc (+-comm k (suc k)) ⟩ + suc ((suc k) + k) ∎ + + goal : n ≡ 2 * k + goal = suc-is-injective (suc-is-injective subgoal) + +evenness-lemma₂ : (n : ℕ) → is-even n → is-even (2 + n) +evenness-lemma₂ n (k , p) = suc k , goal + where + goal : 2 + n ≡ 2 * suc k + goal = 2 + n ≡⟨ ap (λ - → 2 + -) p ⟩ + 2 + (k + k) ≡⟨ ap suc (+-comm (suc k) k) ⟩ + suc k + suc k ∎ + +even⇒check-even : (n : ℕ) → is-even n → check-even n ≡ true +even⇒check-even zero _ = refl true +even⇒check-even (suc zero) (suc zero , ()) +even⇒check-even (suc zero) (suc (suc _) , ()) +even⇒check-even (suc (suc n)) p = even⇒check-even n (evenness-lemma₁ n p) + +check-even⇒even : (n : ℕ) → check-even n ≡ true → is-even n +check-even⇒even zero (refl true) = zero , refl zero +check-even⇒even (suc zero) () +check-even⇒even (suc (suc n)) p = evenness-lemma₂ n (check-even⇒even n p) +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5.lagda.md b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5.lagda.md new file mode 100644 index 0000000..808e403 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/Lab/lab5.lagda.md @@ -0,0 +1,312 @@ +# Week 5 - Lab Sheet + +```agda +{-# OPTIONS --without-K --safe #-} + +module Pool.Lab.lab5 where + +open import prelude +open import natural-numbers-functions hiding (_≤_ ; is-even ; is-odd ; +-comm) +open import List-functions +open import isomorphisms +``` + +We will also want to use some things from the Lab and Homework sheet of Week 4: + +```agda +open import Pool.Homework.homework4-solutions +open import Pool.Lab.lab4-solutions +``` + +## Part I: Length + +In the file [List-functions.lagda.md](../List-functions.lagda.md), the +function `length` is recursively defined as follows. + +```agdacode +length : {A : Type} → List A → ℕ +length [] = 0 +length (x :: xs) = 1 + length xs +``` + +In the following exercises we will prove some facts involving the length of +lists. In doing so, you will practise with inductive proofs. + +### Exercise 1.1 + +Recall that we defined `map` as follows (see +[List-functions.lagda.md](../List-functions.lagda.md)). + +```agdacode +map : {A B : Type} → (A → B) → List A → List B +map f [] = [] +map f (x :: xs) = f x :: map f xs +``` + +**Prove** that `map` preserves the length of a list. + +```agda +map-preserves-length : {A B : Type} (f : A → B) (xs : List A) + → length (map f xs) ≡ length xs +map-preserves-length = {!!} +``` + +### Exercise 1.2 + +Another useful fact is that the length of two concatenated lists is the sum of +their respective lengths. **Complete** the proof of this fact. + +```agda +length-of-++ : {A : Type} (xs ys : List A) + → length (xs ++ ys) ≡ length xs + length ys +length-of-++ = {!!} +``` + +### Exercise 1.3 + +Recall `≤'` from Lab Sheet 4 and `≼'` from Homework Sheet 4 + +```agdacode +_≤'_ : ℕ → ℕ → Type +x ≤' y = Σ k ꞉ ℕ , x + k ≡ y +``` + +```agdacode +_≼'_ : {X : Type} → List X → List X → Type +_≼'_ {X} xs ys = Σ zs ꞉ List X , xs ++ zs ≡ ys +``` + +and that `x ≤' y` expresses that the natural number `x` is less than or equal to +the natural number `y`, while `xs ≼' ys` expresses that the list `xs` is a +prefix of the list `ys`. + +**Prove** that the length of a prefix of a list `ys` is less than the length of +`ys`, relating the two notions above. + +```agda +length-of-prefix : {A : Type} (xs ys : List A) + → xs ≼' ys + → length xs ≤' length ys +length-of-prefix = {!!} +``` + +### Exercise 1.4 + +A nice use case of the length function is that we are now able to define safe +`head` and `tail` operations on nonempty lists. + +We say that a list is *nonempty* if its length is at least 1. +```agda +is-nonempty : {A : Type} → List A → Type +is-nonempty xs = 1 ≤' length xs +``` + +We can then define `tail` as follows. +```agda +tail : {A : Type} (xs : List A) → is-nonempty xs → List A +tail (x :: xs) p = xs +``` + +Agda accepts this definition and does not complain about missing the `[]`-case, +because it realizes that `[]` does not satisfy `is-nonempty`. + +#### Exercise 1.4a + +```agda +head : {A : Type} (xs : List A) → is-nonempty xs → A +head = {!!} +``` + +**Complete** the definition of `head` yourself. + +#### Exercise 1.4b + +```agda +length-of-tail : {A : Type} (xs : List A) (p : 1 ≤' length xs) + → 1 + length (tail xs p) ≡ length xs +length-of-tail = {!!} +``` + +**Prove** that the length of a list is obtained by adding 1 to the length of the +tail. + +#### Exercise 1.4c + +```agda +≤'-suc-lemma : (n : ℕ) → n ≤' (1 + n) +≤'-suc-lemma = {!!} + +length-of-tail-decreases : {A : Type} (xs : List A) (p : 1 ≤' length xs) + → length (tail xs p) ≤' length xs +length-of-tail-decreases = {!!} +``` + +**Complete** the proof of the following lemma and use it to prove that the +length of the tail of a list is less than the length of the full list. + +## Part II: Isomorphisms + +Make sure you have read the [file on isomorphisms](../isomorphisms.lagda.md), +where we define ismorphisms and show that `Bool ≅ 𝟚`. + +You will now give three more isomorphisms. In each case, you should think +about *why* and *how* each pair of types are isomorphic before attemping to +formalise the isomorphism. + +### Exercise 2.1 + +```agda +×-iso : (X Y : Type) → X × Y ≅ Y × X +×-iso X Y = record { bijection = f ; bijectivity = f-is-bijection } + where + f : X × Y → Y × X + f = {!!} + + g : Y × X → X × Y + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +**Show** that X × Y is isomorphic to Y × X using the above template. + +### Exercise 2.2 + +```agda ++-iso : (X Y : Type) → X ∔ Y ≅ Y ∔ X ++-iso X Y = record { bijection = f ; bijectivity = f-is-bijection } + where + f : X ∔ Y → Y ∔ X + f = {!!} + + g : Y ∔ X → X ∔ Y + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +### Exercise 2.3 + +```agda +lists-from-vectors : {A : Type} → List A ≅ (Σ n ꞉ ℕ , Vector A n) +lists-from-vectors {A} + = record { bijection = f ; bijectivity = f-is-bijection } + where + f : List A → Σ n ꞉ ℕ , Vector A n + f = {!!} + + g : Σ n ꞉ ℕ , Vector A n → List A + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +**Show** that the the type `List A` is isomorphic to the type `Σ (Vector A)`. + +(Note that this is the first question in [this question +sheet](../vector-and-list-isomorphisms.lagda.md)). + +Hint: The statement of Exercise 2.3b may help you. + +### Exercise 2.3b + +```agda +open _≅_ + +lfv-preserves-length : {A : Type} (xs : List A) + → fst (bijection lists-from-vectors xs) + ≡ length xs +lfv-preserves-length = {!!} +``` + +Notice how `bijection` extracts the function `f` you defined in +`lists-from-vectors`. + +**Prove** that for any `xs : List A`, the length of `xs` is the same as the +first projection of `f xs : Σ (Vector A)` (where `f : ℕ → List 𝟙` is as you +defined in Exercise 4a). + +## Part III: Evenness + +In the lecture notes, you have seen the predicates `is-even` and `is-odd`: + +```agda +is-even is-odd : ℕ → Type +is-even x = Σ y ꞉ ℕ , x ≡ 2 * y +is-odd x = Σ y ꞉ ℕ , x ≡ 1 + 2 * y +``` + +In these exercises, we will define a Boolean-valued version of the `is-even` +predicate and prove that the two versions are _logically_ equivalent: + +```agda +check-even : ℕ → Bool +check-even zero = true +check-even (suc zero) = false +check-even (suc (suc n)) = check-even n +``` + +### Exercise 3.1 + +First, we will have to prove two lemmas that we will use in Exercise 3.2, where +we will prove our main result. + +```agda +evenness-lemma₁ : (n : ℕ) → is-even (2 + n) → is-even n +evenness-lemma₁ n (suc k , p) = k , goal + where + subgoal : suc (suc n) ≡ suc (suc (2 * k)) + subgoal = suc (suc n) ≡⟨ {!!} ⟩ + suc k + suc k ≡⟨ {!!} ⟩ + suc ((suc k) + k) ∎ + + goal : n ≡ 2 * k + goal = suc-is-injective (suc-is-injective subgoal) + +evenness-lemma₂ : (n : ℕ) → is-even n → is-even (2 + n) +evenness-lemma₂ n (k , p) = suc k , goal + where + goal : 2 + n ≡ 2 * suc k + goal = 2 + n ≡⟨ {!!} ⟩ + 2 + (k + k) ≡⟨ {!!} ⟩ + suc k + suc k ∎ +``` + +**Complete** the above proofs. + +### Exercise 3.2 + +**Prove** that if `is-even n` is inhabited, then `check-even n ≡ true`. + +```agda +even⇒check-even : (n : ℕ) → is-even n → check-even n ≡ true +even⇒check-even = {!!} +``` + +**Prove** that if `check-even n ≡ true` then `is-even n` is inhabited: + +```agda +check-even⇒even : (n : ℕ) → check-even n ≡ true → is-even n +check-even⇒even = {!!} +``` diff --git a/src/HoTTEST/Agda/Exercises/Pool/README.md b/src/HoTTEST/Agda/Exercises/Pool/README.md new file mode 100644 index 0000000..b917f90 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/Pool/README.md @@ -0,0 +1,19 @@ +Tom de Jong. + +A pool of Agda exercises originally intended for lab sessions and homework in +the module "Advanced Functional Programming" at the [School of Computer Science](https://www.birmingham.ac.uk/schools/computer-science/index.aspx) +of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + +These exercises were created together with +[Ayberk Tosun](https://www.cs.bham.ac.uk/~axt978/) and +[Todd Waugh Ambridge](https://www.cs.bham.ac.uk/~txw467/). + + +In general (but not always) the [Homework](Homework) exercises were designed +to be a little more challenging than the exercises selected for the [Lab](Lab). +The numbers in the file names, e.g. `Lab4.lagda.md`, mark the week numbers of +the original module. Thus, a higher number typically indicates more advanced +exercises. + +All exercises are solved in a corresponding file marked with the suffix +`-solutions.lagda.md`. diff --git a/src/HoTTEST/Agda/Exercises/README.md b/src/HoTTEST/Agda/Exercises/README.md new file mode 100644 index 0000000..d614f56 --- /dev/null +++ b/src/HoTTEST/Agda/Exercises/README.md @@ -0,0 +1,21 @@ +# Links to the Agda Exercise Sheets + +1. Week 01: [Exercises](01-Exercises.lagda.md); + [Solutions](01-Solutions.lagda.md). +1. Week 02: [Exercises](02-Exercises.lagda.md); + [Solutions](02-Solutions.lagda.md). +1. Week 03: [Exercises](03-Exercises.lagda.md); + [Solutions](03-Solutions.lagda.md). +1. Week 04: [Exercises](../HITs/Exercises4.lagda.md); + [Solutions](../HITs/Solutions4.lagda.md). +1. Week 05: [Exercises](../HITs/Exercises5.lagda.md); + [Solutions (by Dan Licata)](../HITs/Solutions5-dan.lagda.md); + [Solutions (by Tom de Jong)](../HITs/Solutions5-tom.lagda.md). +1. Week 06: [Exercises](../HITs/Exercises6.lagda.md); + [Solutions](../HITs/Solutions6.lagda.md). +1. Week 07: [Exercises](../Cubical/Exercises7.lagda.md); + [Solutions](../Cubical/Solutions7.lagda.md). +1. Week 08: [Exercises](../Cubical/Exercises8.lagda.md); + [Solutions](../Cubical/Solutions8.lagda.md). +1. Week 09: [Exercises](../Cubical/Exercises9.lagda.md); + [Solutions](../Cubical/Solutions9.lagda.md). diff --git a/src/HoTTEST/Agda/HITs/Exercises4.lagda.md b/src/HoTTEST/Agda/HITs/Exercises4.lagda.md new file mode 100644 index 0000000..65dbdd2 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Exercises4.lagda.md @@ -0,0 +1,152 @@ +```agda +{-# OPTIONS --without-K --rewriting --allow-unsolved-metas #-} + +open import new-prelude +open import Lecture4-notes + +module Exercises4 where +``` + +# Constructing homotopies between paths + +(⋆) Give two "different" path-between-paths/homotopies between (loop ∙ ! +loop) ∙ loop and loop. They should be different in the sense that one +should cancel the !loop with the first loop, and one with the second +loop. But these aren't really *really* different, in that there will be +a path-between-paths-between-paths between the two! + +```agda +homotopy1 : (loop ∙ ! loop) ∙ loop ≡ loop +homotopy1 = {!!} + +homotopy2 : (loop ∙ ! loop) ∙ loop ≡ loop +homotopy2 = {!!} +``` + +(Harder exercise (🌶️): give a path between homotopy1 and +homotopy2! I'd recommend saving this until later though, because there +is a trick to it that we haven't covered in lecture yet.) + +```agda +path-between-paths-between-paths : homotopy1 ≡ homotopy2 +path-between-paths-between-paths = {!!} +``` + +# Functions are group homomorphism + +(⋆⋆) State and prove a general lemma about what ap of a function on the +inverse of a path (! p) does (see ap-∙ for inspiration). + +State and prove a general lemma about what ! (p ∙ q) is. + +Use them to prove that the double function takes loop-inverse to +loop-inverse concatenated with itself. + +```agda +double-!loop : ap double (! loop) ≡ ! loop ∙ ! loop +double-!loop = {!!} +``` + +(⋆) Define a function invert : S1 → S1 such that (ap invert) inverts a path +on the circle, i.e. sends the n-fold loop to the -n-fold loop. + +```agda +invert : S1 → S1 +invert = {!!} +``` + +# Circles equivalence + +See the maps between the 1 point circle and the 2 point circle in the +lecture code. Check that the composite map S1 → S1 +is homotopic to the identity on base and loop: + +(⋆) + +```agda +to-from-base : from (to base) ≡ base +to-from-base = {!!} +``` + +(⋆⋆⋆) + +``` +to-from-loop : ap from (ap to loop) ≡ loop +to-from-loop = {!!} +``` + +Note: the problems below here are progressively more optional, so if you +run out of time/energy at some point that is totally fine. + +# Torus to circles + +The torus is equivalent to the product of two circles S1 × S1. The idea +for the map from the torus to S1 × S1 that is part of this equivalence +is that one loop on on the torus is sent to to the circle loop in one +copy of S1, and the other loop on the torus to the loop in the other +copy of the circle. Define this map. + +Hint: for the image of the square, you might want a lemma saying how +paths in product types compose (⋆⋆⋆): + +```agda +compose-pair≡ : {A B : Type} {x1 x2 x3 : A} {y1 y2 y3 : B} + (p12 : x1 ≡ x2) (p23 : x2 ≡ x3) + (q12 : y1 ≡ y2) (q23 : y2 ≡ y3) + → ((pair≡ p12 q12) ∙ (pair≡ p23 q23)) ≡ {!!} [ (x1 , y1) ≡ (x3 , y3) [ A × B ] ] +compose-pair≡ = {!!} +``` + +(🌶️) +``` +torus-to-circles : Torus → S1 × S1 +torus-to-circles = {!!} +``` + +# Suspensions (🌶️) + +Find a type X such that the two-point circle Circle2 is equivalent to +the suspension Susp X. Check your answer by defining a logical +equivalence (functions back and forth), since we haven't seen how to +prove that such functions are inverse yet. + +```agda +c2s : Circle2 → Susp {!!} +c2s = {!!} + +s2c : Susp {!!} → Circle2 +s2c = {!!} +``` + +Suspension is a functor from types, which means that it acts on +functions as well as types. Define the action of Susp on functions: + +```agda +susp-func : {X Y : Type} → (f : X → Y) → Susp X → Susp Y +susp-func f = {!!} +``` + +To really prove that Susp is a functor, we should check that this action +on functions preserves the identity function and composition of +functions. But to do that we would need the dependent elimination rule +for suspensions, which we haven't talked about yet. + +# Pushouts (🌶️) + +Fix a type X. Find types A,B,C with functions f : C → A and g : C → B +such that the suspension Susp X is equivalent to the Pushout C A B f g. +Check your answer by defining a logical equivalence (functions back and +forth), since we haven't seen how to prove that such functions are +inverse yet. + +```agda +SuspFromPush : Type → Type +SuspFromPush A = {!!} + +s2p : (A : Type) → Susp A → SuspFromPush A +s2p A = {!!} + +p2s : (A : Type) → SuspFromPush A → Susp A +p2s A = {!!} +``` + diff --git a/src/HoTTEST/Agda/HITs/Exercises5.lagda.md b/src/HoTTEST/Agda/HITs/Exercises5.lagda.md new file mode 100644 index 0000000..fd7b108 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Exercises5.lagda.md @@ -0,0 +1,207 @@ + +```agda +{-# OPTIONS --rewriting --without-K --allow-unsolved-metas #-} + +open import new-prelude +open import Lecture5-notes +open import Solutions4 using (ap-!; to-from-base; to-from-loop; s2c; c2s; susp-func) + +module Exercises5 where +``` + +# 1 point and 2 point circles are equivalent (⋆) + +In lecture, we defined maps between the one point circle (S1) and the +two point circle (Circle2) and checked that the round-trip on Circle2 is +the identity. + +Prove that the round-trip on S1 is the identity (use to-from-base +and to-from-loop from the Lecture 4 exercises), and package all of +these items up as an equivalence S1 ≃ Circle2. + +```agda +to-from : (x : S1) → from (to x) ≡ x +to-from = {!!} + +circles-equivalent : S1 ≃ Circle2 +circles-equivalent = {!!} +``` + +# Reversing the circle (⋆⋆) + +Define a map S1 → S1 that "reverses the orientation of the circle", +i.e. sends loop to ! loop. + +```agda +rev : S1 → S1 +rev = {!!} +``` + +Prove that rev is an equivalence. Hint: you will need to state and prove +one new generalized "path algebra" lemma and to use one of the lemmas from +the "Functions are group homomorphism" section of Lecture 4's exercises. +```agda +rev-equiv : is-equiv rev +rev-equiv = {!!} +``` + + +# Circles to torus (⋆⋆) + +In the Lecture 4 exercises, you defined a map from the Torus to S1 × S1. +In this problem, you will define a converse map. The goal is for these +two maps to be part of an equivalence, but we won't prove that in these +exercises. + +You will need to state and prove a lemma characterizing a path over a +path in a path fibration. Then, to define the map S1 × S1 → Torus, you +will want to curry it and use S1-rec and/or S1-elim on each circle. + +```agda +PathOver-path≡ : ∀ {A B : Type} {g : A → B} {f : A → B} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ (g a)} + {r : (f a') ≡ (g a')} + → {!!} + → q ≡ r [ (\ x → (f x) ≡ (g x)) ↓ p ] +PathOver-path≡ {A}{B}{g}{f}{a}{a'}{p}{q}{r} h = {!!} + +circles-to-torus : S1 → (S1 → Torus) +circles-to-torus = {!!} + +circles-to-torus' : S1 × S1 → Torus +circles-to-torus' (x , y) = circles-to-torus x y +``` + +**Below are some "extra credit" exercise if you want more to do. These +are (even more) optional: nothing in the next lecture will depend on you +understanding them. The next section (H space) is harder code but uses +only the circle, whereas the following sections are a bit easier code +but require understanding the suspension type, which we haven't talked +about too much yet.** + +# H space + +The multiplication operation on the circle discussed in lecture is part +of what is called an "H space" structure on the circle. One part of +this structure is that the circle's basepoint is a unit element for +multiplication. + +(⋆) Show that base is a left unit. +```agda +mult-unit-l : (y : S1) → mult base y ≡ y +mult-unit-l y = {!!} +``` + +(⋆) Because we'll need it in a second, show that ap distributes over +function composition: +```agda +ap-∘ : ∀ {l1 l2 l3 : Level} {A : Type l1} {B : Type l2} {C : Type l3} + (f : A → B) (g : B → C) + {a a' : A} + (p : a ≡ a') + → ap (g ∘ f) p ≡ ap g (ap f p) +ap-∘ = {!!} +``` + +(⋆⋆) Suppose we have a curried function f : S1 → A → B. Under the +equivalence between paths in S1 × A and pairs of paths discussed in +Lecture 3, we can then "apply" (the uncurried version of) f to a pair of +paths (p : x ≡ y [ S1 ] , q : z ≡ w [ A ]) to get a path (f x z ≡ f y w +[ B ]). In the special case where q is reflexivity on a, this +application to p and q can be simplified to ap (\ x → f x a) p : f x a ≡ +f y a [ B ]. + +Now, suppose that f is defined by circle recursion. We would expect +some kind of reduction for applying f to the pair of paths (loop , q) --- +i.e. we should have reductions for *nested* pattern matching on HITs. +In the case where q is reflexivity, applying f to the pair (loop , refl) +can reduce like this: +```agda +S1-rec-loop-1 : ∀ {A B : Type} {f : A → B} {h : f ≡ f} {a : A} + → ap (\ x → S1-rec f h x a) loop ≡ app≡ h a +S1-rec-loop-1 {A}{B}{f}{h}{a} = {!!} +``` +Prove this reduction using ap-∘ and the reduction rule for S1-rec on the loop. + +(⋆⋆⋆) Show that base is a right unit for multiplication. You will need +a slightly different path-over lemma than before. + +```agda +PathOver-endo≡ : ∀ {A : Type} {f : A → A} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ a} + {r : (f a') ≡ a'} + → the Type {! !} + → q ≡ r [ (\ x → f x ≡ x) ↓ p ] +PathOver-endo≡ {p = (refl _)} {q = q} {r} h = {!!} + +mult-unit-r : (x : S1) → mult x base ≡ x +mult-unit-r = {!!} +``` + +# Suspensions and the 2-point circle + +(⋆) Postulate the computation rules for the non-dependent susp-rec and +declare rewrites for the reduction rules on the point constructors. +```agda +postulate + Susp-rec-north : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m northS ≡ {!!} + Susp-rec-south : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m southS ≡ {!!} +-- {-# REWRITE Susp-rec-north #-} +-- {-# REWRITE Susp-rec-south #-} +postulate + Susp-rec-merid : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → (x : A) → ap (Susp-rec n s m) (merid x) ≡ {!!} +``` + +(⋆) Postulate the dependent elimination rule for suspensions: + +```agda +postulate + Susp-elim : {l : Level} {A : Type} (P : Susp A → Type l) + → (n : {!!}) + → (s : {!!}) + → (m : {!!}) + → (x : Susp A) → P x +``` + +(⋆⋆) Show that the maps s2c and c2s from the Lecture 4 exercises are mutually inverse: + +```agda +c2s2c : (x : Circle2) → s2c (c2s x) ≡ x +c2s2c = {!!} + +s2c2s : (x : Susp Bool) → c2s (s2c x) ≡ x +s2c2s = {!!} +``` + +(⋆) Conclude that Circle2 is equivalent to Susp Bool: + +```agda +Circle2-Susp-Bool : Circle2 ≃ Susp Bool +Circle2-Susp-Bool = {!!} +``` + +# Functoriality of suspension (⋆⋆) + +In the Lecture 4 exercises, we defined functoriality for the suspension +type, which given a function X → Y gives a function Σ X → Σ Y. Show +that this operation is functorial, meaning that it preserves identity +and composition of functions: +```agda +susp-func-id : ∀ {X : Type} → susp-func {X} id ∼ id +susp-func-id = {!!} + +susp-func-∘ : ∀ {X Y Z : Type} (f : X → Y) (g : Y → Z) + → susp-func {X} (g ∘ f) ∼ susp-func g ∘ susp-func f +susp-func-∘ f g = {!!} +``` + + + diff --git a/src/HoTTEST/Agda/HITs/Exercises6.lagda.md b/src/HoTTEST/Agda/HITs/Exercises6.lagda.md new file mode 100644 index 0000000..c4e590f --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Exercises6.lagda.md @@ -0,0 +1,292 @@ + +```agda + +{-# OPTIONS --rewriting --without-K --allow-unsolved-metas #-} + +open import new-prelude + +open import Lecture6-notes +open import Lecture5-notes + +module Exercises6 where +``` + +In this problem set, you will look at a variation on the circle, a +higher inductive type for a "bowtie", i.e. two loops at a point. +(Unscaffolded harder exercise: do these problems for a "wedge of k +circles" for any natural number k.) + +# HIT recursion from induction + +In general, the dependent elimination rule for a higher inductive type +implies the simple/non-dependent elimination rule. In this problem, you +will show this for the bowtie. We could have done this for the circles +in the past lectures, but I wanted to introduce the non-dependent +elimination rule first, and then left both as postulates. + +Note that this problem has a bit of a "metamathematical" flavor (showing +that a set of axioms is implied by a shorter set). If you prefer to +jump right to the more "mathematical" problem of characterizing the loop +space of the bowtie below, I recommend turning Bowtie-rec and its +associated reductions into postulates like we have done for previous +higher inductive types, and adding a rewrite for the reduction on the +base point. This will make Agda display things in a more easy to read +way (otherwise, it will display Bowtie-rec as a meta-variable). + +Here is the definition of the bowtie and its dependent elimination rule: + +```agda +postulate + Bowtie : Set + baseB : Bowtie + loop1 : baseB ≡ baseB + loop2 : baseB ≡ baseB + Bowtie-elim : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → (x : Bowtie) → X x + Bowtie-elim-base : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → Bowtie-elim X x p q baseB ≡ x +{-# REWRITE Bowtie-elim-base #-} + +postulate + Bowtie-elim-loop1 : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → apd (Bowtie-elim X x p q) loop1 ≡ p + Bowtie-elim-loop2 : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → apd (Bowtie-elim X x p q) loop2 ≡ q +``` + +Next, we will prove the non-dependent elim/"recursion principle" from +these. First, we need some lemmas. + +(⋆) Paths over a path in a constant fibration are equivalent to paths. +It is simple to prove this by path induction. + +```agda +PathOver-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 + → b1 ≡ b2 [ (\ _ → B) ↓ p ] +PathOver-constant = {!!} + +PathOver-constant-inverse : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 [ (\ _ → B) ↓ p ] + → b1 ≡ b2 +PathOver-constant-inverse = {!!} + +PathOver-constant-inverse-cancel1 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2) + → PathOver-constant-inverse p (PathOver-constant p q) ≡ q +PathOver-constant-inverse-cancel1 = {!!} + +PathOver-constant-inverse-cancel2 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2 [ _ ↓ p ]) + → PathOver-constant p (PathOver-constant-inverse p q) ≡ q +PathOver-constant-inverse-cancel2 = {!!} + +PathOver-constant-equiv : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (b1 ≡ b2) ≃ (b1 ≡ b2 [ (\ _ → B) ↓ p ]) +PathOver-constant-equiv p = improve (Isomorphism (PathOver-constant p) + (Inverse (PathOver-constant-inverse p) + (PathOver-constant-inverse-cancel1 p) + (PathOver-constant-inverse-cancel2 p))) + +``` + +(⋆) Next, for a non-dependent function f, there is an annoying mismatch between +ap f and apd f, which we can reconcile as follows: + +```agda +ap-apd-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → (f : A → B) + → ap f p ≡ PathOver-constant-inverse _ (apd f p) +ap-apd-constant = {!!} +``` + +(⋆) Define Bowtie-rec and prove the reduction for base: + +```agda +Bowtie-rec : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → (Bowtie) → X +Bowtie-rec {_} {X} x p q = {!!} + +Bowtie-rec-base : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → Bowtie-rec x p q baseB ≡ x +Bowtie-rec-base _ _ _ = {!!} +``` + +(⋆⋆) Prove the reductions for loop: + +```agda +Bowtie-rec-loop1 : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → ap (Bowtie-rec x p q) loop1 ≡ p [ x ≡ x ] +Bowtie-rec-loop1 x p q = {!!} + +Bowtie-rec-loop2 : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → ap (Bowtie-rec x p q) loop2 ≡ q [ x ≡ x ] +Bowtie-rec-loop2 x p q = {!!} +``` + +# Loop space of the bowtie + +In this problem, you will show that the loop space of the bowtie is the +"free group on two generators", which we will write in Agda as F2. The +point of this problem is mainly for you to read and really understand +the proof that the loop space of the circle is ℤ. All of the code is +essentially a rearrangement of code from that proof. I'd suggest +trying the proof yourself, and looking at the analogous bits of the +Circle proof if you get stuck. + +## Some lemmas (⋆⋆) + +In the Circle proof in lecture, I inlined a couple of things that +can be proved more generally. You might want to prove these general +versions in advance and use them in your proof, or, if that seems +confusing, you might first do the proof without these lemmas +to motivate them. + +```agda +concat-equiv : ∀ {A : Type} (a : A) {a' a'' : A} + → (p : a' ≡ a'') + → (a ≡ a') ≃ (a ≡ a'') +concat-equiv = {!!} + +concat-equiv-map : ∀ {A : Type} {a a' a'' : A} + → (p : a' ≡ a'') + → fwd (concat-equiv a p) ≡ \ q → q ∙ p +concat-equiv-map = {!!} +``` + +(Note: you could also write out all of the components, but this was easier.) + +```agda +transport-∙ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + {a1 a2 a3 : A} (p : a1 ≡ a2) (q : a2 ≡ a3) + → transport B (p ∙ q) ∼ transport B q ∘ transport B p +transport-∙ = {!!} +``` +## Calculating the loop space + +First, we will assume a type F2 representing the free group on 2 +generators. + +ℤ is the free group on one generator, with 0 as the neutral element and +succℤ corresponding to "addition" with the one generator. succℤ is an +equivalence, with the inverse representing "addition" with -1. + +For other groups, it is somewhat more common to think of the group +operation as "multiplication" rather than "addition", so we will name +the neutral element as "1" and the action of the elements as +"multiplication". Thus, we assume a type with an element 1F, and two +equivalences, which we think of as "multiplication with generator 1" and +"multiplication with generator 2". + +Unscaffolded hard exercise: You can implement F2 as lists whose +elements are a four-element type g1, g2, g1⁻¹, g2⁻¹ with no adjacent +pairs of inverse elements. Then the forward directions of mult1/mult2 +will be implement by cons'ing g1/g2 on and "reducing" if that creates +two adjacent inverses, the backwards directions by consing g1⁻¹ and g2⁻¹ +on and reducing, and the inverse laws will hold because the reduction +cancels the inverses. + +For this problem, we will simply assume the nice universal property for +this type: that it maps uniquely into any other type with a point and +two equivalences, and that it is a set. + +```agda +module AssumeF2 + (F2 : Type) + (1F : F2) + (mult1 : F2 ≃ F2) + (mult2 : F2 ≃ F2) + (F2-rec : {X : Type} + (o : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → F2 → X) + (F2-rec-1 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → F2-rec z m1 m2 1F ≡ z) + (F2-rec-mult1 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + (a : F2) → F2-rec z m1 m2 (fwd mult1 a) ≡ fwd m1 (F2-rec z m1 m2 a)) + (F2-rec-mult2 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + (a : F2) → F2-rec z m1 m2 (fwd mult2 a) ≡ fwd m2 (F2-rec z m1 m2 a)) + (F2-rec-unique : {X : Type} + (f : F2 → X) + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → f 1F ≡ z + → ((f ∘ fwd mult1) ∼ (fwd m1 ∘ f)) + → ((f ∘ fwd mult2) ∼ (fwd m2 ∘ f)) + → (x : F2) → f x ≡ F2-rec z m1 m2 x) + (hSetF : is-set F2) where +``` + +(⋆⋆⋆) Prove that the loop space of the Bowtie is F2. Each bit of the +proof will be analogous to the corresponding part of the Circle proof. + +```agda + Cover : Bowtie → Type + Cover = {!!} + + encode : (x : Bowtie) → baseB ≡ x → Cover x + encode = {!!} + + decode : (x : Bowtie) → Cover x → baseB ≡ x + decode = {!!} + + encode-decode : (x : Bowtie) (p : baseB ≡ x) → decode x (encode x p) ≡ p + encode-decode = {!!} + + decode-encode : (x : Bowtie) (p : Cover x) → encode x (decode x p) ≡ p + decode-encode = {!!} + +``` + diff --git a/src/HoTTEST/Agda/HITs/Lecture4-live.lagda.md b/src/HoTTEST/Agda/HITs/Lecture4-live.lagda.md new file mode 100644 index 0000000..84054b7 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture4-live.lagda.md @@ -0,0 +1,137 @@ +```agda + +{-# OPTIONS --without-K --rewriting #-} + +open import new-prelude + +module Lecture4-live where + + {- + postulate + Bool : Type + true : Bool + false : Bool + -} + + postulate + S1 : Type + base : S1 + loop : base ≡ base [ S1 ] + + example-path : base ≡ base + example-path = loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop + + -- groupoid laws + + ∙assoc : {A : Type} {x y z w : A} + (p : x ≡ y) + (q : y ≡ z) + (r : z ≡ w) + → (p ∙ (q ∙ r)) ≡ ((p ∙ q) ∙ r) [ x ≡ w [ A ] ] + ∙assoc p q (refl _) = refl _ + + ∙unit-l : {A : Type} {x y : A} + → (p : x ≡ y) + → (refl _ ∙ p) ≡ p + ∙unit-l (refl _) = refl _ + + ∙unit-r : {A : Type} {x y : A} + → (p : x ≡ y) → (p ∙ refl _) ≡ p + ∙unit-r p = refl p + + !-inv-l : {A : Type} {x y : A} + → (p : x ≡ y) + → (! p ∙ p) ≡ refl _ + !-inv-l (refl _) = refl _ + + !-inv-r : {A : Type} {x y : A} → (p : x ≡ y) + → (p ∙ ! p) ≡ refl _ + !-inv-r (refl _) = refl _ + + example : loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop + ≡ loop + example = loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop ≡⟨ refl _ ⟩ + (((loop ∙ ! loop) ∙ loop) ∙ ! loop) ∙ loop ≡⟨ ap (\ H → H ∙ loop ∙ ! loop ∙ loop) (!-inv-r loop) ⟩ + ((refl _ ∙ loop) ∙ ! loop) ∙ loop ≡⟨ ap (λ H → H ∙ ! loop ∙ loop) (∙unit-l loop) ⟩ + ((loop ∙ ! loop) ∙ loop) ≡⟨ ! (∙assoc _ _ loop) ⟩ + (loop ∙ (! loop ∙ loop)) ≡⟨ ap (\ H → loop ∙ H) (!-inv-l loop) ⟩ + (loop ∙ refl _) ≡⟨ ∙unit-r loop ⟩ + loop ∎ + + -- Circle recursion + postulate + S1-rec : {X : Type} (x : X) (l : x ≡ x [ X ]) → (S1 → X) + S1-rec-base : {X : Type} (x : X) (l : x ≡ x [ X ]) + → (S1-rec x l) base ≡ x + + {-# BUILTIN REWRITE _≡_ #-} + {-# REWRITE S1-rec-base #-} + + postulate + S1-rec-loop : {X : Type} (x : X) (l : x ≡ x [ X ]) + → ap (S1-rec x l) loop ≡ l + + double : S1 → S1 + double = S1-rec base (loop ∙ loop) + + double-loop : ap double loop ≡ loop ∙ loop + double-loop = S1-rec-loop _ _ + + ap-∙ : {A B : Type} {f : A → B} {x y z : A} + (p : x ≡ y) + (q : y ≡ z) + → ap f (p ∙ q) ≡ ap f p ∙ ap f q + ap-∙ (refl _) (refl _) = refl _ + + double-2-loops : ap double (loop ∙ loop) ≡ (loop ∙ loop) ∙ (loop ∙ loop) + double-2-loops = + ap double (loop ∙ loop) ≡⟨ ap-∙ loop loop ⟩ + ap double loop ∙ ap double loop ≡⟨ ap₂ (\ p q → p ∙ q) + (S1-rec-loop _ _) + (S1-rec-loop _ _) ⟩ + (loop ∙ loop) ∙ (loop ∙ loop) ∎ + + postulate + Circle2 : Type + north : Circle2 + south : Circle2 + west : north ≡ south + east : north ≡ south + Circle2-rec : {X : Type} + (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2 → X + Circle2-rec-north : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2-rec n s w e north ≡ n + Circle2-rec-south : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2-rec n s w e south ≡ s + {-# REWRITE Circle2-rec-north #-} + {-# REWRITE Circle2-rec-south #-} + + postulate + Circle2-rec-west : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → ap (Circle2-rec n s w e) west ≡ w + Circle2-rec-east : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → ap (Circle2-rec n s w e) east ≡ e + + to : S1 → Circle2 + to = S1-rec north (east ∙ ! west) + + from : Circle2 → S1 + from = Circle2-rec base base (refl _) loop + + from-to-north : to (from north) ≡ north + from-to-north = refl _ + + from-to-south : to (from south) ≡ south + from-to-south = west + + from-to-west : ap to (ap from west) ∙ from-to-south ≡ west + from-to-west = ap to (ap from west) ∙ west ≡⟨ ap (\ H → ap to H ∙ west) (Circle2-rec-west _ _ _ _) ⟩ + ap to (refl base) ∙ west ≡⟨ ∙unit-l west ⟩ + west ∎ + + + -- Q: is loop distinct from refl? + -- A: so far, maybe + +``` diff --git a/src/HoTTEST/Agda/HITs/Lecture4-notes.lagda.md b/src/HoTTEST/Agda/HITs/Lecture4-notes.lagda.md new file mode 100644 index 0000000..706b365 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture4-notes.lagda.md @@ -0,0 +1,349 @@ +```agda + +{-# OPTIONS --without-K --rewriting #-} + +open import new-prelude + +module Lecture4-notes where +``` + +# The circle + +In general, a type should be thought of as an infinity-groupoid: it has +points, paths (or identifications, equalities, depending on what you +want to call them), and then paths between paths, and so on. The +*regular* inductive types that we've, like booleans or natural numbers, +have constructors only for points. *Higher* inductive types (HITs) also have +constructors for paths, paths between paths, etc. + +A good first example is the circle, which we will write as S1. As as a +higher inductive type, it is presented as "the free infinity groupoid +generated by a point and a loop at that point". + +```agda +postulate + S1 : Type + base : S1 + loop : base ≡ base +``` + +base is a point constructor, like zero for the natural numbers or true +and false for the booleans. loop is a *path* constructor, which means +that instead of constructing an element of S1, it constructs a path in S1 +from base to base. We represent this using the identity type ≡. Since +Agda doesn't have HITs built-in (until we get to Cubical Agda in Lecture +7), we will "postulate" these, which means to assume constants of these +types for the rest of the development. A postulate is roughly the same +as making the whole rest of the codebase depend on a variable of that +type. + +Recall from previous lectures that for any type we can define operations +of path concatenation and path inverses. We will write concatenation as +p ∙ q and inverse as ! p. (This was written (-)⁻¹ in previous lectures, +but I find a prefix notation easier to read when Agda displays terms.) +You can review the definitions in the [prelude](new-prelude.lagda.md). + +Using these we can construct other paths on the circle from loop: + +```agda +example-path : base ≡ base +example-path = loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop +``` + +There should be a homotopy between example (forwards, backwards, +forwards, backwards, forwards) and loop (forwards). To define it, +we can generalize and prove the groupoid laws for paths. + +# Groupoid laws for paths + +If you've seen some abstract algebra, you know that the *group laws* say +that the "multiplication" of the group is associative and has the +neutral element as a unit on both sides, and that the inverse of an +element multiplied with that element is the neutral element. The +*groupoid laws* for paths are a "typed" version of these. They say that +the identity path (refl) is a unit for path concatenation, path +concatenation is associative, and the inverse of a path composed with +that path is the identity. Let's state and prove them. + +```agda +∙unit-r : {A : Type} {x y : A} → (p : x ≡ y) → (p ∙ refl _) ≡ p +∙unit-r p = refl _ + +∙unit-l : {A : Type} {x y : A} → (p : x ≡ y) → (refl _ ∙ p) ≡ p +∙unit-l (refl _) = refl _ + +∙assoc : {A : Type} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙ (q ∙ r) ≡ (p ∙ q) ∙ r +∙assoc p (refl _) (refl _) = (refl _) + +!-inv-l : {A : Type} {x y : A} → (p : x ≡ y) → (! p ∙ p) ≡ refl _ +!-inv-l (refl _) = refl _ + +!-inv-r : {A : Type} {x y : A} → (p : x ≡ y) → (p ∙ ! p) ≡ refl _ +!-inv-r (refl _) = refl _ +``` + +It's kind of amazing that all of this structure emerges from a very +simple axiom (path induction!). + +OK, now we can use these to construct a path/homotopy between the path +above and loop: + +```agda +example : (loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop) ≡ loop [ base ≡ base ] +example = loop ∙ ! loop ∙ loop ∙ ! loop ∙ loop ≡⟨ refl _ ⟩ + (((loop ∙ ! loop) ∙ loop) ∙ ! loop) ∙ loop ≡⟨ ap (\ H → H ∙ loop ∙ ! loop ∙ loop) (!-inv-r loop) ⟩ + refl _ ∙ loop ∙ ! loop ∙ loop ≡⟨ ap (\ H → H ∙ ! loop ∙ loop) (∙unit-l (loop)) ⟩ + loop ∙ ! loop ∙ loop ≡⟨ ! (∙assoc _ (! loop) loop) ⟩ + loop ∙ (! loop ∙ loop) ≡⟨ ap (\ H → loop ∙ H) (!-inv-l loop) ⟩ + loop ∙ (refl _) ≡⟨ refl _ ⟩ + loop ∎ +``` + +We'll use the notation x ≡ y [ A ] if we want to make the type a path is +in explicit. We'll sometimes write (\ x → e) for (λ x → e) because it +easier to type and de-emphasizes the function a little (since ap almost +always needs a lambda abstraction as input, it would IMO be nicer to have +a syntax with a bound variable like ap (H → ...) p). + +Up to homotopy, there should be ℤ many paths on the circle: +... , ! loop ∙ ! loop, ! loop, refl, loop, loop ∙ loop, ... + +# Circle "recursion" + +Terminology note: for an inductive type like booleans, natural numbers, +lists, etc., it is common to call the dependent elimination principle +the "elimination" or "induction" principle, and the non-dependent +elimination principle the "recursion" principle. This makes perfect +sense for natural numbers, but the analogy gets a little strained for +types like booleans. Since bool isn't properly inductive (there are no +constructors for bool that take a boolean as input, like successor for +natural numbers), the "recursion" principle is just if-then-else, and +the "induction" principle is just proof by case analysis. Nonetheless, +for uniformity, we will refer to non-dependent elimination rules as +"recursion principles" even when the type is not properly inductive. + +With this in mind, the "recursion" principle for the circle expresses +that it is the *initial* type with a point and a loop at that point: +given any other type with a point and a loop, we can map the circle into +it (in a unique way, but we'll get to that later). + +```agda +postulate + S1-rec : {l : Level} {X : Type l} (x : X) (p : x ≡ x) → S1 → X + S1-rec-base : {l : Level} {X : Type l} (x : X) (p : x ≡ x) + → S1-rec {l} {X} x p base ≡ x + S1-rec-loop' : {l : Level} {X : Type l} (x : X) (p : x ≡ x) + → ap (S1-rec x p) loop ≡ (S1-rec-base x p) ∙ p ∙ ! (S1-rec-base x p) +``` +The base and loop postulates are the reduction rules for S1-rec. +Note that we use ap to apply S1-rec to the loop. + +It's common to assume that at least the point reductions like +S1-rec-base are definitional equalities. We can do this in Agda like +this: +```agda +{-# BUILTIN REWRITE _≡_ #-} +{-# REWRITE S1-rec-base #-} +``` +Then, for example, we don't need to explicitly mention S1-rec-base to state the loop reduction. +```agda +postulate + S1-rec-loop : {l : Level} {X : Type l} (x : X) (p : x ≡ x) + → ap (S1-rec x p) loop ≡ p +``` + +This is mainly for convenience, but it's a big convenience! When we get +to Cubical Agda, both of these reductions will be definitional, but for +axiomatic HoTT we only put in the point one. + +## Doubling function + +As a first example of circle recursion, we can define the function from +the circle to the circle that sends the loop to the concatenation of two +loops. Thinking of the paths as integers, this sends x to 2*x. + +```agda +double : S1 → S1 +double = S1-rec base (loop ∙ loop) + +calculate-double-loop : ap double loop ≡ (loop ∙ loop) +calculate-double-loop = S1-rec-loop _ _ +``` + +How do we "reduce" this function on two loops? We need a general fact +that functions are functors, which in particular means that they are +group homomorphisms on paths, seen as groups with composition as the +group operation. + +```agda +ap-∙ : {A B : Type} {f : A → B} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → ap f (p ∙ q) ≡ ap f p ∙ ap f q +ap-∙ (refl _) (refl _) = refl _ + +calculate-double-2-loops : ap double (loop ∙ loop) ≡ loop ∙ loop ∙ loop ∙ loop +calculate-double-2-loops = + ap double (loop ∙ loop) ≡⟨ ap-∙ loop loop ⟩ + ap double loop ∙ ap double loop ≡⟨ ap₂ (\ p q → p ∙ q) (S1-rec-loop _ _) (S1-rec-loop _ _) ⟩ + (loop ∙ loop) ∙ (loop ∙ loop) ≡⟨ (∙assoc (loop ∙ loop) loop loop ) ⟩ + ((loop ∙ loop) ∙ loop) ∙ loop ∎ +``` + +# 2 point version of the circle + +We can also "subdivide" the circle into two points with two paths +between them: + +```agda +postulate + Circle2 : Type + north : Circle2 + south : Circle2 + west : north ≡ south + east : north ≡ south + Circle2-rec : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2 → X + Circle2-rec-north : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2-rec n s w e north ≡ n + Circle2-rec-south : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → Circle2-rec n s w e south ≡ s +{-# REWRITE Circle2-rec-north #-} +{-# REWRITE Circle2-rec-south #-} + +postulate + Circle2-rec-west : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → ap (Circle2-rec n s w e) west ≡ w + Circle2-rec-east : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + → ap (Circle2-rec n s w e) east ≡ e +``` + +Let's work towards showing that these two definitions of the circle give +equivalent types. This means defining functions back and forth whose +composites are *homotopic* to the identity function. But inside type +theory, this looks just like giving a "isomorphism" in the sense defined +in the Lecture 3 exercises. + +First, to map the 1-point circle to the 2-point circle, we need a point +and a loop. Of course, there are lots of choices, but let's pick +```agda +to : S1 → Circle2 +to = S1-rec north (east ∙ ! west) +``` + +In the other direction, we need to map north and south each to some +point of the circle, but we've only got one point, base. For the paths, +we can go "twice as fast" on one and "stand still" on the other. +```agda +from : Circle2 → S1 +from = Circle2-rec base base (refl base) loop +``` + +We can check that on the constructors, the composite from-then-to is the identity: + +```agda +from-to-north : to (from north) ≡ north +from-to-north = refl _ + +from-to-south : to (from south) ≡ south +from-to-south = west + +from-to-west : ap to (ap from west) ∙ from-to-south ≡ west +from-to-west = ap to (ap from west) ∙ west ≡⟨ ap (\ H → ap to H ∙ west) (Circle2-rec-west _ _ _ _) ⟩ + ap to (refl base) ∙ west ≡⟨ ∙unit-l west ⟩ + west ∎ + +from-to-east : ap to (ap from east) ∙ from-to-south ≡ east +from-to-east = ap to (ap from east) ∙ west ≡⟨ ap (\ H → ap to H ∙ west) (Circle2-rec-east _ _ _ _) ⟩ + ap to loop ∙ west ≡⟨ ap (\ H → H ∙ west) (S1-rec-loop _ _) ⟩ + east ∙ ! west ∙ west ≡⟨ ! (∙assoc _ (! west) west) ⟩ + east ∙ (! west ∙ west) ≡⟨ ap (\ H → east ∙ H) (!-inv-l west) ⟩ + east ∎ +``` + +Note that it doesn't even type check to ask that ap to (ap from west) ≡ +west, because to (from south) ≡ north. So we need to compare them "up +to" a path. We'll talk about why (- ∙ from-to-south) is the right path +to pick next time. + +# The torus + +The torus is a nice example of a type with a path-between-paths constructor: + +```agda +postulate + Torus : Type + baseT : Torus + pT : baseT ≡ baseT + qT : baseT ≡ baseT + sT : pT ∙ qT ≡ qT ∙ pT + T-rec : {l : Level} {X : Type l} (x : X) (p : x ≡ x) (q : x ≡ x) (s : p ∙ q ≡ q ∙ p) → Torus → X + T-rec-base : {l : Level} {X : Type l} (x : X) (p : x ≡ x) (q : x ≡ x) (s : p ∙ q ≡ q ∙ p) + → T-rec {l} {X} x p q s baseT ≡ x +{-# REWRITE T-rec-base #-} +postulate + T-rec-pT : {l : Level} {X : Type l} (x : X) (p : x ≡ x) (q : x ≡ x) (s : p ∙ q ≡ q ∙ p) + → ap (T-rec x p q s) pT ≡ p + T-rec-qT : {l : Level} {X : Type l} (x : X) (p : x ≡ x) (q : x ≡ x) (s : p ∙ q ≡ q ∙ p) + → ap (T-rec x p q s) qT ≡ q + -- there needs to be an analogous reduction for sT, but it's a little tricky to state it, + -- so we will defer it for now +``` + +Lots of other important concepts can be defined with HITs, including +quotients, suspensions, pushouts, truncations. + +# Suspensions + +```agda +postulate + Susp : Type → Type + northS : {A : Type} → Susp A + southS : {A : Type} → Susp A + merid : {A : Type} → A → northS ≡ southS [ Susp A ] + Susp-rec : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp A → X +``` + +# Pushouts + +```agda +postulate + Pushout : (C : Type) (A : Type) (B : Type) (f : C → A) (g : C → B) → Type + +module _ {C : Type} {A : Type} {B : Type} {f : C → A} {g : C → B} where + + postulate + inl : A → Pushout C A B f g + inr : B → Pushout C A B f g + glue : (c : C) → inl (f c) ≡ inr (g c) + Push-rec : {X : Type} (l : A → X) (r : B → X) (gl : (c : C) → l (f c) ≡ r (g c)) + → Pushout C A B f g → X +``` + +# Relation quotient + +Suppressing some details, think of this as the quotient of A by the +equivalence relation closure of R (technically, this description is +precisely right when R(x,y) is a proposition and A is a set; otherwise, +you need to think about the higher homotopies induced by the +constructors). + +```agda +postulate + _/_ : (A : Type) → (A → A → Type) → Type + i : {A : Type} {R : A → A → Type} → A → A / R + q : {A : Type} {R : A → A → Type} (x y : A) → R x y → i x ≡ i y [ A / R ] + +``` + +# Propositional truncation + +The universal way of making a type into a -1-type. This is a nice +example of a (properly) inductive HIT. +```agda +postulate + ∥_∥₋₁ : Type → Type + ∣_∣ : {A : Type} → A → ∥ A ∥₋₁ + trunc : {A : Type} → is-prop ∥ A ∥₋₁ +``` diff --git a/src/HoTTEST/Agda/HITs/Lecture4-pictures.pdf b/src/HoTTEST/Agda/HITs/Lecture4-pictures.pdf new file mode 100644 index 0000000..6fc738d Binary files /dev/null and b/src/HoTTEST/Agda/HITs/Lecture4-pictures.pdf differ diff --git a/src/HoTTEST/Agda/HITs/Lecture5-live.lagda.md b/src/HoTTEST/Agda/HITs/Lecture5-live.lagda.md new file mode 100644 index 0000000..6f2013d --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture5-live.lagda.md @@ -0,0 +1,195 @@ +```agda +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +module Lecture5-live where + +open import Lecture4-notes hiding (to; from; from-to-north; from-to-south; from-to-west; from-to-east; q) public + +record is-equiv {l1 l2 : Level} {A : Type l1} {B : Type l2} (f : A → B) : Type (l1 ⊔ l2) where + constructor + Inverse + field + inverse : B → A + η : inverse ∘ f ∼ id + inverse2 : B → A + ε : f ∘ inverse2 ∼ id + +record _≃_ {l1 l2 : Level} (A : Type l1) (B : Type l2) : Type (l1 ⊔ l2) where + constructor + Equivalence + field + map : A → B + is-equivalence : is-equiv map + +fwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≃ B → A → B +fwd e = _≃_.map e + +bwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≃ B → B → A +bwd e = is-equiv.inverse (_≃_.is-equivalence e) + +improve : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≅ B → A ≃ B +improve (Isomorphism f (Inverse g gf fg)) = + Equivalence f (Inverse g gf g fg) + +data PathOver {l1 l2 : Level} + {A : Type l1} (B : A → Type l2) : + {a1 a2 : A} (p : a1 ≡ a2) + (b1 : B a1) (b2 : B a2) → Type (l1 ⊔ l2) where + reflo : {x : A} {y : B x} → PathOver B (refl x) y y + +-- b1 ≡ b2 [ B a2 ] doesn't type check + +syntax PathOver B p b1 b2 = b1 ≡ b2 [ B ↓ p ] + +transport-to-pathover : {l1 l2 : Level} {A : Type l1} (B : A → Type l2) + {a1 a2 : A} (p : a1 ≡ a2) + (b1 : B a1) (b2 : B a2) + → (transport B p b1 ≡ b2 [ B a2 ]) ≃ (b1 ≡ b2 [ B ↓ p ]) +transport-to-pathover B (refl _) b1 b2 = + Equivalence (λ { (refl _) → reflo }) + (Inverse (\ { reflo → refl _}) + (\ { (refl _) → refl _} ) + ((\ { reflo → refl _})) + (\ { reflo → refl _})) + +path-to-pathover : ∀ {A : Type} {B : A → Type} + → {a : A} {x y : B a} + → (p : x ≡ y) + → x ≡ y [ B ↓ refl a ] +path-to-pathover = fwd (transport-to-pathover _ _ _ _) + +apd : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + (f : (x : A) → B x) {a1 a2 : A} (p : a1 ≡ a2) + → f a1 ≡ f a2 [ B ↓ p ] +apd f (refl _) = reflo + +postulate + + -- Circle2-rec : {X : Type} (n : X) (s : X) (w : n ≡ s) (e : n ≡ s) + -- → Circle2 → X + + Circle2-elim : (X : Circle2 → Type) + (n : X north) + (s : X south) + (w : n ≡ s [ X ↓ west ]) + (e : n ≡ s [ X ↓ east ]) + → (x : Circle2) → X x + + Circle2-elim-north : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → Circle2-elim X n s w e north ≡ n + Circle2-elim-south : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → Circle2-elim X n s w e south ≡ s +{-# REWRITE Circle2-elim-north #-} +{-# REWRITE Circle2-elim-south #-} +postulate + Circle2-elim-west : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → apd (Circle2-elim X n s w e) west + ≡ w + Circle2-elim-east : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → apd (Circle2-elim X n s w e) east ≡ e + + +module RememberTheseFromLastTime where + + to : S1 → Circle2 + to = S1-rec north (east ∙ ! west) + + from : Circle2 → S1 + from = Circle2-rec base base (refl base) loop + + from-to-north : to (from north) ≡ north + from-to-north = refl _ + + from-to-south : to (from south) ≡ south + from-to-south = west + + + from-to-west : ap to (ap from west) ∙ from-to-south ≡ west + from-to-west = ap to (ap from west) ∙ from-to-south ≡⟨ ap (\ H → ap to H ∙ from-to-south) (Circle2-rec-west _ _ _ _) ⟩ + ap to (refl base) ∙ from-to-south ≡⟨ ∙unit-l from-to-south ⟩ + from-to-south ≡⟨ refl _ ⟩ + west ∎ + + from-to-east : ap to (ap from east) ∙ from-to-south ≡ east + from-to-east = ap to (ap from east) ∙ from-to-south ≡⟨ ap (\ H → ap to H ∙ from-to-south) (Circle2-rec-east _ _ _ _) ⟩ + ap to loop ∙ from-to-south ≡⟨ ap (\ H → H ∙ from-to-south) (S1-rec-loop _ _) ⟩ + east ∙ ! west ∙ from-to-south ≡⟨ ! (∙assoc _ (! west) from-to-south) ⟩ + east ∙ (! west ∙ from-to-south) ≡⟨ ap (\ H → east ∙ H) (!-inv-l west) ⟩ + east ∎ + +open RememberTheseFromLastTime public + +PathOver-roundtrip≡ : ∀ {A B : Type} (g : B → A) (f : A → B) + {a a' : A} (p : a ≡ a') + {q : g (f a) ≡ a} + {r : g (f a') ≡ a'} + → ! q ∙ ap g (ap f p) ∙ r ≡ p + → q ≡ r [ (\ x → g (f x) ≡ x) ↓ p ] +PathOver-roundtrip≡ g f (refl _) h = path-to-pathover (coh _ _ h) where + coh : ∀ {A : Type} {a1 a2 : A} (q : a1 ≡ a2) (r : a1 ≡ a2) + → ! q ∙ r ≡ refl _ + → q ≡ r + coh (refl _) q h = ! h ∙ ∙unit-l q + +from-to : (y : Circle2) → to (from y) ≡ y +from-to = Circle2-elim _ + from-to-north + from-to-south + (PathOver-roundtrip≡ to from west {q = from-to-north} {r = from-to-south} + ( ! (∙assoc _ _ from-to-south) ∙ (∙unit-l _ ∙ from-to-west ))) + ((PathOver-roundtrip≡ to from east (! (∙assoc _ _ from-to-south) ∙ ∙unit-l _ ∙ from-to-east))) + +postulate + S1-elim : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ] ) + → (x : S1) → X x + + S1-elim-base : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + → S1-elim X x p base ≡ x + +{-# REWRITE S1-elim-base #-} +postulate + S1-elim-loop : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + → apd (S1-elim X x p) loop ≡ p + + + +PathOver-path-loop : ∀ {A : Type} + {a a' : A} {p : a ≡ a'} + {q : a ≡ a} + {r : a' ≡ a'} + → q ∙ p ≡ p ∙ r + → q ≡ r [ (\ x → x ≡ x) ↓ p ] +PathOver-path-loop {p = refl _} (refl .(refl _ ∙ _)) = path-to-pathover (∙unit-l _) + +mult : S1 → (S1 → S1) +mult = S1-rec ((\ x → x)) + (λ≡ (S1-elim _ + loop + (PathOver-path-loop (refl _) ))) + +-- full funext + +app≡ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + {f g : (x : A) → B x} + → f ≡ g → f ∼ g +app≡ p x = ap (\ f → f x) p + +postulate + λ≡βinv : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} + → (h : f ∼ g) + → app≡ (λ≡ h) ≡ h + + +``` diff --git a/src/HoTTEST/Agda/HITs/Lecture5-notes.lagda.md b/src/HoTTEST/Agda/HITs/Lecture5-notes.lagda.md new file mode 100644 index 0000000..6ca77b5 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture5-notes.lagda.md @@ -0,0 +1,299 @@ +```agda +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +module Lecture5-notes where + +open import Lecture4-notes hiding (to; from; from-to-north; from-to-south; from-to-west; from-to-east; q) public + +``` + +# Equivalences + +In this lecture, we will start proving some type equivalences, so we +need to code a definition of equivalence. In the Lecture 3 exercies, we +saw the definition of a bijection/isomorphism/quasi-equivalence in Agda: +a record consisting of maps back and forth with homotopies showing that +they compose to the identity. In the HoTT track, we saw the definition +of equivalence of types as a bi-invertible map. We can code this +similarly in Agda: + +```agda +record is-equiv {l1 l2 : Level} {A : Type l1} {B : Type l2} (f : A → B) : Type (l1 ⊔ l2) where + constructor Inverse + field + post-inverse : B → A + is-post-inverse : post-inverse ∘ f ∼ id + pre-inverse : B → A + is-pre-inverse : f ∘ pre-inverse ∼ id + +record _≃_ {l1 l2 : Level} (A : Type l1) (B : Type l2) : Type (l1 ⊔ l2) where + constructor + Equivalence + field + map : A → B + is-equivalence : is-equiv map +``` + +Here are some short names for projecting the the forward and (one of +the) backward maps: + +```agda +fwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≃ B → A → B +fwd e = _≃_.map e + +bwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≃ B → B → A +bwd e = is-equiv.post-inverse (_≃_.is-equivalence e) +``` + +An isomorphism can be improved to an equivalence by using the same +function as the pre-inverse and post-inverse: + +```agda +improve : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} → A ≅ B → A ≃ B +improve (Isomorphism f (Inverse g gf fg)) = Equivalence f (Inverse g gf g fg) + +``` + +# Path over a path + +Next, we will work towards stating the dependent elimination rules / +induction principles for HITs. To do this, we will need the notion of a +"dependent path" or "path over a path". + +Suppose you have a type A with elements a1 and a2 and a type family B : +A → Type with elements b1 : B a1 and b2 : B a2. It doesn't type check +to ask if b1 ≡ b2 because b1 and b2 have different types. But if we +also have a path p : a1 ≡ a2 then we can ask whether b1 ≡ b2 "up to" +that path. + +Using topological terminology, we can think about the *total space* of +B, type Σ x ∶ A , B, which has a map pr₁ : (Σ x ∶ A , B) → A. When b : +B x, we say that b is "in the fiber over x", because the element (x , b) +: (Σ x ∶ A , B) projects to x by pr₁. One way to compare elements in +different fibres is to ask for a path in the total space, i.e. a path +(a1 , b1) ≡ (a2 , b2) [ Σ x ∶ A , B ]. Using the characterization of +paths in Σ types from Lecture 3 as pairs of paths, such a path has a +first component path a1 ≡ a2 [ A ]. We are often interested in paths in +the total space whose first component is a specified path p : a1 ≡ a2 [ +A ]. This can be represented by a pair of a path q : (a1 , b1) ≡ (a2 , +b2) [ Σ x ∶ A , B ] with a path fst (from-Σ-≡ q) ≡ p [ a1 ≡ a2 ]. +However, there are a couple of more direct ways to represent this in +type theory. + +The first way to represent a path over p is with a new inductive definition: + +```agda +data PathOver {l1 l2 : Level} {A : Type l1} (B : A → Type l2) : + {a1 a2 : A} (p : a1 ≡ a2) + (b1 : B a1) (b2 : B a2) → Type (l1 ⊔ l2) where + reflo : {x : A} {y : B x} → PathOver B (refl x) y y + +syntax PathOver B p b1 b2 = b1 ≡ b2 [ B ↓ p ] +``` + +This is an "inductive family" with one constructor that says that any y +: B x is equal to itself over reflexivity. + +Another way to represent a path over p is as a homogeneous path +transport B p b1 ≡ b2 [ B a2 ]. Semantically, this is because dependent +types are fibrations, which includes that any path-over b1 ≡ b2 [ B ↓ p +] factors uniquely into a heterogeneous path b1 ≡ transport B p b1 [ B ↓ +p ] composed with a homogeneous path transport B p b1 ≡ b2 [ B a2 ]. + +In type theory, we can prove that these are equivalent by path induction: + +``` +transport-to-pathover : {l1 l2 : Level} {A : Type l1} (B : A → Type l2) + {a1 a2 : A} (p : a1 ≡ a2) + (b1 : B a1) (b2 : B a2) + → (transport B p b1 ≡ b2) ≃ (b1 ≡ b2 [ B ↓ p ]) +transport-to-pathover B (refl _) b1 b2 = improve (Isomorphism ((λ { (refl _) → reflo })) + (Inverse (\ { reflo → refl _}) + (\ {(refl _) → refl _}) + (\ {(reflo) → refl _}))) + +path-to-pathover : ∀ {A : Type} {B : A → Type} + → {a : A} {x y : B a} + → (p : x ≡ y) + → x ≡ y [ B ↓ refl a ] +path-to-pathover p = fwd (transport-to-pathover _ (refl _) _ _) p +``` + +ap of a dependent function naturally creates a path-over: +```agda +apd : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + (f : (x : A) → B x) {a1 a2 : A} (p : a1 ≡ a2) + → f a1 ≡ f a2 [ B ↓ p ] +apd f (refl _) = reflo +``` + +We'll use the inductive family definition mainly because it will be +closer to what you'll see in Lectures 7,8,9 on Cubical Agda. + +# Dependent elims/induction for HITs + +We're finally in a position to state the dependent elimination rules for +HITs. Let's start with Circle2 because it's a little easier to see +what's going on. + +```agda +postulate + Circle2-elim : (X : Circle2 → Type) + (n : X north) + (s : X south) + (w : n ≡ s [ X ↓ west ]) + (e : n ≡ s [ X ↓ east ]) + → (x : Circle2) → X x + Circle2-elim-north : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → Circle2-elim X n s w e north ≡ n + Circle2-elim-south : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → Circle2-elim X n s w e south ≡ s +{-# REWRITE Circle2-elim-north #-} +{-# REWRITE Circle2-elim-south #-} +postulate + Circle2-elim-west : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → apd (Circle2-elim X n s w e) west ≡ w + Circle2-elim-east : (X : Circle2 → Type) (n : X north) (s : X south) + (w : n ≡ s [ X ↓ west ]) (e : n ≡ s [ X ↓ east ]) + → apd (Circle2-elim X n s w e) east ≡ e +``` + +# Proving equivalences + +One common reason dependent elims get used is to prove that functions +are mutually inverse. + +```agda +module RememberTheseFromLastTime where + + to : S1 → Circle2 + to = S1-rec north (east ∙ ! west) + + from : Circle2 → S1 + from = Circle2-rec base base (refl base) loop + + from-to-north : to (from north) ≡ north + from-to-north = refl _ + + from-to-south : to (from south) ≡ south + from-to-south = west + + from-to-west : ap to (ap from west) ∙ from-to-south ≡ west + from-to-west = ap to (ap from west) ∙ west ≡⟨ ap (\ H → ap to H ∙ west) (Circle2-rec-west _ _ _ _) ⟩ + ap to (refl base) ∙ west ≡⟨ ∙unit-l west ⟩ + west ∎ + + from-to-east : ap to (ap from east) ∙ from-to-south ≡ east + from-to-east = ap to (ap from east) ∙ west ≡⟨ ap (\ H → ap to H ∙ west) (Circle2-rec-east _ _ _ _) ⟩ + ap to loop ∙ west ≡⟨ ap (\ H → H ∙ west) (S1-rec-loop _ _) ⟩ + east ∙ ! west ∙ west ≡⟨ ! (∙assoc _ (! west) west) ⟩ + east ∙ (! west ∙ west) ≡⟨ ap (\ H → east ∙ H) (!-inv-l west) ⟩ + east ∎ +open RememberTheseFromLastTime public +``` + +```agda +PathOver-roundtrip≡ : ∀ {A B : Type} (g : B → A) (f : A → B) + {a a' : A} (p : a ≡ a') + {q : g (f a) ≡ a} + {r : g (f a') ≡ a'} + → ! q ∙ ((ap g (ap f p)) ∙ r) ≡ p + → q ≡ r [ (\ x → g (f x) ≡ x) ↓ p ] +PathOver-roundtrip≡ g f (refl _) {q = q} {r} h = + path-to-pathover (ap (\ H → q ∙ H) (! h) ∙ + ( ∙assoc _ _ (refl _ ∙ r) ∙ + (ap (\ H → H ∙ (refl _ ∙ r)) (!-inv-r q) ∙ + (∙unit-l (refl _ ∙ r) ∙ ∙unit-l r )) )) + +from-to : (y : Circle2) → to (from y) ≡ y +from-to = Circle2-elim _ + from-to-north + from-to-south + (PathOver-roundtrip≡ to from west (∙unit-l _ ∙ from-to-west)) + (PathOver-roundtrip≡ to from east (∙unit-l _ ∙ from-to-east)) +``` + +To do the other direction you'll need + +```agda +postulate + S1-elim : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + → (x : S1) → X x + + S1-elim-base : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + → S1-elim X x p base ≡ x + +{-# REWRITE S1-elim-base #-} +postulate + S1-elim-loop : (X : S1 → Type) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + → apd (S1-elim X x p) loop ≡ p +``` + +# Multiplication on the circle + +Classically, if you think of the points on the circle as complex +numbers, you can multiply two points, and the result will represent a +point on the circle. There is a synthetic version of this operation in +type theory. Defining it uses S1-elim and *function extensionality*, +which says that paths between functions are homotopies. + +```agda +PathOver-path-loop : ∀ {A : Type} + {a a' : A} {p : a ≡ a'} + {q : a ≡ a} + {r : a' ≡ a'} + → q ∙ p ≡ p ∙ r + → q ≡ r [ (\ x → x ≡ x) ↓ p ] +PathOver-path-loop {p = (refl _)} h = path-to-pathover (h ∙ (∙unit-l _)) + +mult : S1 → S1 → S1 +mult = S1-rec ((\ y → y)) (λ≡ (S1-elim (λ z → z ≡ z) loop (PathOver-path-loop (refl _)))) +``` + +Note that it is also possible to do this without funext by binding the +second input before doing the S1-rec on the first input (thanks Ulrik!): +```agda +mult-nofunext : S1 → S1 → S1 +mult-nofunext x y = S1-rec y (S1-elim (λ z → z ≡ z) loop (PathOver-path-loop (refl _)) y) x +``` + +Above, we used the main part of function extensionality: a homotopy +induces a path between functions. The full form of the function +extensionality axiom (mentioned briefly in HoTT Lecture 5) is that +homotopies are equivalent to paths between functions, and in particular +that the map from paths to homotopies that you can define by path +induction is an equivalence. + +```agda +app≡ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} → f ≡ g → f ∼ g +app≡ p x = ap (\ f → f x) p + +postulate + λ≡βinv : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} → (h : f ∼ g) + → app≡ (λ≡ h) ≡ h + +-- it's often more useful to have this as a homotopy +λ≡β : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} → (h : f ∼ g) + → app≡ (λ≡ h) ∼ h +λ≡β h = app≡ (λ≡βinv h) + +full-funext : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} + → is-equiv {A = f ≡ g} {B = f ∼ g} app≡ +full-funext = Inverse λ≡' λ≡η λ≡ λ≡βinv where + postulate + λ≡' : _ + λ≡η : _ +``` + + diff --git a/src/HoTTEST/Agda/HITs/Lecture5-pictures.pdf b/src/HoTTEST/Agda/HITs/Lecture5-pictures.pdf new file mode 100644 index 0000000..011842b Binary files /dev/null and b/src/HoTTEST/Agda/HITs/Lecture5-pictures.pdf differ diff --git a/src/HoTTEST/Agda/HITs/Lecture6-live.lagda.md b/src/HoTTEST/Agda/HITs/Lecture6-live.lagda.md new file mode 100644 index 0000000..f2a3088 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture6-live.lagda.md @@ -0,0 +1,182 @@ +```agda + +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +open import Lecture4-notes +open import Lecture5-notes +open import Solutions5-dan using (PathOver-path≡) + +module Lecture6-live where + +{- remember the definition of S1 from Lectures 4 and 5 + + S1 : Type + base : S1 + loop : base ≡ base + + For {l : Level} {X : Type l} (x : X) (p : x ≡ x) + S1-rec : S1 → X + S1-rec-base : S1-rec x p base ≡ x + S1-rec-loop : ap (S1-rec x p) loop ≡ p + + For {l : Level} (X : S1 → Type l) + (x : X base) + (p : x ≡ x [ X ↓ loop ]) + S1-elim : (x : S1) → X x + S1-elim-base : S1-elim X x p base ≡ x + S1-elim-loop : apd (S1-elim X x p) loop ≡ p + +-} + +Ω¹S1 : Type +Ω¹S1 = base ≡ base [ S1 ] + +Ω²S1 : Type +Ω²S1 = refl base ≡ refl base [ base ≡ base [ S1 ] ] + +{- +Correspondence + +-2 ! loop ∙ ! loop +-1 ! loop + 0 refl _ , loop ∙ ! loop + 1 loop , loop ∙ loop ∙ ! loop + 2 loop ∙ loop + + refl _ =? loop NO, but we have to prove it + +-} + +data ℤ : Type where + Pos : ℕ → ℤ -- Pos 0 = 1 + Zero : ℤ -- 0 + Neg : ℕ → ℤ -- Neg 0 = -1 + +module AssumeInts + (ℤ : Type) + + (0ℤ : ℤ) + + (succℤ : ℤ ≃ ℤ) -- successor and predecessor + + + (ℤ-rec : {X : Type} + (z : X) + (s : X ≃ X) + → ℤ → X) + + + (ℤ-rec-0ℤ : {X : Type} + (z : X) + (s : X ≃ X) + → ℤ-rec z s 0ℤ ≡ z ) + (ℤ-rec-succℤ : {X : Type} + (z : X) + (s : X ≃ X) + → (x : ℤ) → ℤ-rec z s (fwd succℤ x) ≡ fwd s (ℤ-rec z s x)) + + (ℤ-rec-unique : {X : Type} + (f : ℤ → X) + (z : X) + (s : X ≃ X) + → f 0ℤ ≡ z + → ((x : ℤ) → f (fwd succℤ x) ≡ fwd s (f x)) + → (x : ℤ) → f x ≡ ℤ-rec z s x) + where + + + loop^ : ℤ → base ≡ base + loop^ = ℤ-rec (refl _) + (improve (Isomorphism (\ p → p ∙ loop) + (Inverse (\ q → q ∙ ! loop) + (\ p → ! (∙assoc _ loop (! loop)) + ∙ ap (\ H → p ∙ H) (!-inv-r loop)) + (\ p → ! (∙assoc _ (! loop) (loop)) + ∙ ap (\ H → p ∙ H) (!-inv-l loop))))) + + postulate + ua : ∀ {l : Level} {X Y : Type l} → X ≃ Y → X ≡ Y + uaβ : ∀ {l : Level} {X Y : Type l} (e : X ≃ Y) {x : X} + → transport (\ X → X) (ua e) x ≡ fwd e x + + Cover : S1 → Type + Cover = S1-rec ℤ (ua succℤ) + + transport-ap-assoc : {A : Type} (C : A → Type) {a a' : A} (p : a ≡ a') {x : C a} + → transport C p x ≡ + transport (\ X → X) (ap C p) x + transport-ap-assoc C (refl _) = refl _ + + + transport-Cover-loop : (x : ℤ) → transport Cover loop x ≡ fwd succℤ x + transport-Cover-loop x = transport-ap-assoc Cover loop ∙ + (ap (\ H → transport id H x) (S1-rec-loop _ _) + ∙ uaβ _) + + transport-Cover-then-loop : {x : S1} (p : x ≡ base) (y : Cover x) + → transport Cover (p ∙ loop) y ≡ fwd succℤ (transport Cover p y) + transport-Cover-then-loop (refl _) y = ap (\ Z → transport Cover (Z) y) (∙unit-l loop) ∙ + transport-Cover-loop _ + + encode : (x : S1) → base ≡ x → Cover x + encode x p = transport Cover p 0ℤ + -- .base (refl base) = 0ℤ + + encode-firsttry : base ≡ base → ℤ + encode-firsttry = encode base + + endo-ℤ-is-id : (f : ℤ → ℤ) + → f 0ℤ ≡ 0ℤ + → ((x : ℤ) → f (fwd succℤ x) ≡ fwd succℤ (f x)) + → f ∼ id + endo-ℤ-is-id f f0 fsucc x = ℤ-rec-unique f 0ℤ succℤ f0 fsucc x ∙ + ! (ℤ-rec-unique (\x → x) 0ℤ succℤ (refl _) (\ _ → refl _) x) + + encode-loop^ : (x : ℤ) → encode-firsttry (loop^ x) ≡ x + encode-loop^ = endo-ℤ-is-id (encode-firsttry ∘ loop^) + (ap encode-firsttry (ℤ-rec-0ℤ _ _) ∙ + refl _) + \ x → ap (\ H → encode base H) (ℤ-rec-succℤ _ _ x) ∙ + goal x where + goal : (x : _) → transport Cover (loop^ x ∙ loop) 0ℤ ≡ + fwd succℤ (transport Cover (loop^ x) 0ℤ) + goal x = transport-Cover-then-loop (loop^ x) 0ℤ + + postulate + PathOver-→ : {X : Type} + {A : X → Type} + {B : X → Type} + {x x' : X} {p : x ≡ x'} + {f1 : A x → B x} + {f2 : A x' → B x'} + → ((a : A x) → f1 a ≡ f2 (transport A p a) [ B ↓ p ]) + → f1 ≡ f2 [ (\ z → A z → B z) ↓ p ] + -- this can be proved from function extensionality, see the Lecture 6 notes + + PathOver-path-to : ∀ {A : Type} + {a0 a a' : A} {p : a ≡ a'} + {q : a0 ≡ a} + {r : a0 ≡ a'} + → q ∙ p ≡ r + → q ≡ r [ (\ x → a0 ≡ x) ↓ p ] + PathOver-path-to {p = refl _} (refl _) = reflo + + decode : (x : S1) → Cover x → base ≡ x + decode = S1-elim _ loop^ + (PathOver-→ \ z → PathOver-path-to + (! (ap loop^ (transport-Cover-loop z) ∙ + ℤ-rec-succℤ _ _ _ ))) + + encode-decode : (x : S1) (p : base ≡ x) → decode x (encode x p) ≡ p + encode-decode .base (refl base) = ℤ-rec-0ℤ _ _ + + loop^-encode : (p : base ≡ base) → loop^ (encode-firsttry p) ≡ p + loop^-encode = encode-decode base + + theorem : (base ≡ base) ≃ ℤ + theorem = improve (Isomorphism (encode base) (Inverse loop^ loop^-encode encode-loop^)) + +``` + diff --git a/src/HoTTEST/Agda/HITs/Lecture6-notes.lagda.md b/src/HoTTEST/Agda/HITs/Lecture6-notes.lagda.md new file mode 100644 index 0000000..840e6c0 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Lecture6-notes.lagda.md @@ -0,0 +1,447 @@ +```agda + +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +open import Lecture4-notes +open import Lecture5-notes +open import Solutions5-dan using (PathOver-path≡) + +module Lecture6-notes where + +``` + +In this class, we will go over the proof that the fundamental group of +the circle is the integers. Mike Shulman first proved this theorem in +HoTT; the proof below is my simplified version of his proof. + +# Univalence + +Univalence says that paths in the universe are equivalences. We'll +start with just the ability to turn an equivalence into a path. + +```agda +postulate + ua : ∀ {l : Level} {X Y : Type l} → X ≃ Y → X ≡ Y +``` + +There is already a converse to this map: to turn a path X ≡ Y into and +equivalence, we can do path induction to contract X to Y, and then +return the identity equivalence X ≃ X. Call this path-to-equiv: + +```agda +id≃ : ∀ {A : Type} → A ≃ A +id≃ = Equivalence ((\ x → x)) (Inverse (\ x → x) (\ _ → refl _) (\ x → x) (\ _ → refl _)) + +path-to-equiv : ∀ {A B : Type} → A ≡ B → A ≃ B +path-to-equiv (refl _) = id≃ +``` + +Note that fwd (path-to-equiv p) is equal to transport (\ X → X) p): +```agda +fwd-path-to-equiv : ∀ {A B : Type} (p : A ≡ B) → fwd (path-to-equiv p) ≡ transport (\ X → X) p +fwd-path-to-equiv (refl _) = refl _ +``` + +In full, the univalence axiom says that this map path-to-equiv is an +equivalence between equivalences and paths. For our purposes today, we +will only need one more bit of that equivalence, which says that +transport (\ X → X) after ua is the identity: + +```agda +postulate + uaβ : ∀{l : Level} {X Y : Type l} (e : X ≃ Y) {x : X} + → transport (\ X → X) (ua e) x ≡ fwd e x +``` +(This extra bit actually turns out to imply the full univalence axiom, +using the fact that is-equiv is a proposition, and the fundamental +theorem of identity types (see the HoTT track).) + +# Lemma library + +Here are a bunch of general facts that are provable from path induction. +In class, we will prove these as they come up in the proof, but they +need to be lifted outside of the AssumeInts module for the exercises +code to work, since they don't actually depend on those assumptions. + +```agda +transport-ap-assoc : {A : Type} (C : A → Type) {a a' : A} (p : a ≡ a') {x : C a} + → transport C p x ≡ transport (\ X → X) (ap C p) x +transport-ap-assoc C (refl _) = refl _ + +transport-→ : {X : Type} + {A : X → Type} + {B : X → Type} + {x x' : X} (p : x ≡ x') + {f : A x → B x} + → transport (λ z → (y : A z) → B (z)) p f ≡ \ a → transport B p (f (transport A (! p) a)) +transport-→ (refl _) = refl _ + +transport-inv-r : {X : Type} + {A : X → Type} + {x x' : X} (p : x ≡ x') (a : A x') + → transport A p (transport A (! p) a) ≡ a +transport-inv-r (refl _) a = refl _ + +PathOver-→ : {X : Type} + {A : X → Type} + {B : X → Type} + {x x' : X} {p : x ≡ x'} + {f1 : A x → B x} + {f2 : A x' → B x'} + → ((a : A x) → f1 a ≡ f2 (transport A p a) [ B ↓ p ]) + → f1 ≡ f2 [ (\ z → A z → B z) ↓ p ] +PathOver-→ {A = A} {B} {p = p} {f2 = f2} h = + fwd (transport-to-pathover _ _ _ _) + (transport-→ p ∙ λ≡ \ a → bwd (transport-to-pathover _ _ _ _) + (h (transport A (! p) a)) ∙ ap f2 (transport-inv-r p a)) + +pair≡d : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + {a a' : A} (p : a ≡ a') + {b : B a} {b' : B a'} (q : b ≡ b' [ B ↓ p ]) + → (a , b) ≡ (a' , b') [ Σ B ] +pair≡d (refl _) reflo = refl _ + +fill-left : {X : Type} + {A : X → Type} + {x x' : X} (p : x ≡ x') (a : A x') + → (transport A (! p) a) ≡ a [ A ↓ p ] +fill-left (refl _) a = reflo + +transport-Π : {X : Type} + {A : X → Type} + {B : Σ A → Type} + {x x' : X} (p : x ≡ x') + {f : (y : A x) → B (x , y)} + → transport (λ z → (y : A z) → B (z , y)) p f ≡ \ a → transport B (pair≡d p (fill-left p a)) (f (transport A (! p) a)) +transport-Π (refl _) = refl _ + +PathOver-Π : {X : Type} + {A : X → Type} + {B : Σ A → Type} + {x x' : X} {p : x ≡ x'} + {f1 : (y : A x) → B (x , y)} + {f2 : (y' : A x') → B (x' , y')} + → ({a : A x} {a' : A x'} (q : a ≡ a' [ A ↓ p ]) → f1 a ≡ f2 a' [ B ↓ pair≡d p q ]) + → f1 ≡ f2 [ (\ z → (y : A z) → B (z , y)) ↓ p ] +PathOver-Π {A = A} {B} {p = p} {f1 = f1} {f2} h = + fwd (transport-to-pathover _ _ _ _) + ((transport-Π p) ∙ λ≡ \ a → + bwd (transport-to-pathover B (pair≡d p (fill-left p a)) _ _) (h _)) + +PathOver-path-to : ∀ {A : Type} + {a0 a a' : A} {p : a ≡ a'} + {q : a0 ≡ a} + {r : a0 ≡ a'} + → q ∙ p ≡ r + → q ≡ r [ (\ x → a0 ≡ x) ↓ p ] +PathOver-path-to {p = refl _} {q = refl _} (refl _) = reflo +``` + +# Fundamental group of the circle + +Next we will work towards characterizing the fundamental group of the +circle. The homotopy groups of a space are (roughly) the groups of +paths, paths-between-paths, etc. with path concatenation as the group +operation. The homotopy groups are almost defined as iterated identity +types like this: + +```agda +Ω¹S1 : Type +Ω¹S1 = base ≡ base + +Ω²S1 : Type +Ω²S1 = refl base ≡ refl base [ base ≡ base ] +``` + +The almost is because you also need to use truncations (which we haven't +talked about yet) to remove the higher structure of these path types. +The iterated identity types represent what is called the n-fold loop +space of a type, Ω^n A, which is the whole space of loops, +loops-between-refls, etc. + +"Calculating a loops space (or homotopy group)" means proving an +equivalence between the specified loop space and some more explicit +description of a group. E.g. Ω¹S1 turns out to be the integers ℤ. + +We can enumerate ℤ many paths on the circle: … -2 (! loop ∙ ! loop), -1 +(! loop), 0 (refl _), 1 (loop), 2 (loop ∙ loop), … Of course, there are +others, like loop ∙ ! loop, but that has a path-between-paths to +refl --- we're counting paths "up to homotopy" or "up to the higher +identity types". So, intuitively, every loop on the circle is homotopic +to one of the above form. The other way this equivalence could fail is if +e.g. loop ≡ refl _, so there wouldn't be as many paths as it looks like. +But we didn't add any path-between-path constructors to S1, +so intuitively in the "least" type generated by base and loop, +there are no such identifications. + +To prove this in HoTT, we will use the universal cover of the circle, +which is a fibration (type family) over the circle whose fiber over each +point is ℤ, and where going around the loop in the base goes up one +level. This can be pictured as a helix. But in type theory we just +define it by saying what the fibers are and what happens when you go +around the loop. To do this, we need a definition of the integers. + +For now, we'll just assume a definition of the integers that supplies +exactly what we need to do this proof. Much of an implementation is +below. + +First, we have a type ℤ with 0 and successor functions. Just that would +define ℕ; to get ℤ, we also want successor to be an equivalence -- with +the inverse being the predecessor. + +Next, the universal property of the integers is that it's the +"least"/"initial" type with a point and an equivalence: you can map into +any other type with a point and an equivalence. +Intuitively, this sends 0 to b, +n to (fwd s)^n b +and -b to (bwd s)^n b. Moreover, this map is unique, in the sense that +any other map that sends 0 to some point z and successor to some equivalence s +is equal to the map determined by ℤ-rec. + +```agda +module AssumeInts + (ℤ : Type) + (0ℤ : ℤ) + (succℤ : ℤ ≃ ℤ) + (ℤ-rec : {X : Type} + (b : X) + (s : X ≃ X) + → ℤ → X) + (ℤ-rec-0ℤ : {X : Type} + (b : X) + (s : X ≃ X) + → ℤ-rec b s 0ℤ ≡ b) + (ℤ-rec-succℤ : {X : Type} + (b : X) + (s : X ≃ X) + (a : ℤ) → ℤ-rec b s (fwd succℤ a) ≡ fwd s (ℤ-rec b s a)) + (ℤ-rec-unique : {X : Type} + (f : ℤ → X) + (z : X) + (s : X ≃ X) + → f 0ℤ ≡ z + → ((f ∘ fwd succℤ) ∼ (fwd s ∘ f)) + → (x : ℤ) → f x ≡ ℤ-rec z s x) + (hSetℤ : is-set ℤ) where +``` + +Using ℤ-rec, we define a map "loop to the nth power", which maps an integer n to +a loop on the circle as indicated above: +… -2 (! loop ∙ ! loop), -1 (! loop), 0 (refl _), 1 (loop), 2 (loop ∙ loop), … + +```agda + loop^ : ℤ → base ≡ base + loop^ = ℤ-rec (refl _) + (improve (Isomorphism (\ p → p ∙ loop) + (Inverse (\ p → p ∙ (! loop)) + (\ p → ! (∙assoc _ loop (! loop)) ∙ + ap (\ H → p ∙ H) (!-inv-r loop) ) + (\ p → ! (∙assoc _ (! loop) loop) ∙ + ap (\ H → p ∙ H) (!-inv-l loop))))) +``` + +But mapping loops to ints is harder. The trick is to define a type +family/fibrations over S1, where transporting in that type family +converts paths to ints. This is where the universal cover of the +circle/the helix comes in: + +```agda + Cover : S1 → Type + Cover = S1-rec ℤ (ua succℤ) +``` + +For example, we can calculate that transporting in the Cover on the loop +adds one: + +```agda + transport-Cover-loop : (x : ℤ) → transport Cover loop x ≡ fwd succℤ x + transport-Cover-loop x = transport-ap-assoc Cover loop ∙ + ap (\ H → transport id H x) (S1-rec-loop _ _) ∙ + (uaβ succℤ) + + PathOver-Cover-loop : (x : ℤ) → x ≡ fwd succℤ x [ Cover ↓ loop ] + PathOver-Cover-loop x = fwd (transport-to-pathover _ _ _ _) (transport-Cover-loop x) +``` + +Now we can define + +```agda + encode : (x : S1) → base ≡ x → Cover x + encode x p = transport Cover p 0ℤ +``` + +For the other direction, we need to use S1-elim: + +```agda + decode : (x : S1) → Cover x → base ≡ x + decode = S1-elim _ + loop^ + (PathOver-→ (\ a → + PathOver-path-to (! (ℤ-rec-succℤ _ _ a) ∙ + ! (ap loop^ (transport-Cover-loop _))))) +``` + +One composite is easy by path induction: + +```agda + encode-decode : (x : S1) (p : base ≡ x) → decode x (encode x p) ≡ p + encode-decode .base (refl base) = ℤ-rec-0ℤ _ _ +``` + +(This composite can actually be abstracted into a general lemma; see the +fundamental theorem of identity types in the HoTT track.) + +For the other composite, we do circle induction, use the uniqueness +principle for maps out of ℤ, and use the fact that ℤ is a set to easily +finish the final goal: + +```agda + endo-ℤ-is-id : (f : ℤ → ℤ) + → f 0ℤ ≡ 0ℤ + → (f ∘ fwd succℤ) ∼ (fwd succℤ ∘ f) + → f ∼ id + endo-ℤ-is-id f f0 fsucc x = ℤ-rec-unique f 0ℤ succℤ f0 fsucc x ∙ + ! (ℤ-rec-unique (\ x → x) 0ℤ succℤ (refl _) (\ _ → refl _) x) + + transport-Cover-then-loop : {x : S1} (p : x ≡ base) (y : Cover x) + → transport Cover (p ∙ loop) y ≡ fwd succℤ (transport Cover p y) + transport-Cover-then-loop (refl _) y = ap (\ Z → transport Cover (Z) y) (∙unit-l loop) ∙ + transport-Cover-loop _ + + decode-encode-base : (x : ℤ) → encode base (loop^ x) ≡ x + decode-encode-base x = endo-ℤ-is-id encode-loop^ encode-loop^-zero encode-loop^-succ x where + encode-loop^ : ℤ → ℤ + encode-loop^ x = encode base (loop^ x) + + encode-loop^-zero : encode-loop^ 0ℤ ≡ 0ℤ + encode-loop^-zero = ap (\ H → transport Cover H 0ℤ) (ℤ-rec-0ℤ _ _) + + encode-loop^-succ : (encode-loop^ ∘ fwd succℤ) ∼ (fwd succℤ ∘ encode-loop^) + encode-loop^-succ x = ap (\ H → encode base H) (ℤ-rec-succℤ _ _ x) ∙ + transport-Cover-then-loop (loop^ x) 0ℤ + + + decode-encode : (x : S1) (p : Cover x) → encode x (decode x p) ≡ p + decode-encode = S1-elim _ + decode-encode-base + (PathOver-Π \ aa' → fwd (transport-to-pathover _ _ _ _) (hSetℤ _ _ _ _)) +``` + +Here's most of an implementation of integers: + +```agda +module ImplementInts where + + fix : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} + → A ≃ B + → A ≃ B + fix (Equivalence f (Inverse g fg g' fg')) = + Equivalence f (Inverse g fg g + (\x → ! (ap f (ap g (fg' x))) ∙ (ap f (fg (g' x)) ∙ fg' x))) + + fwd-bwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} + → (e : A ≃ B) + → fwd e ∘ bwd e ∼ id + fwd-bwd e = is-equiv.is-pre-inverse (_≃_.is-equivalence (fix e)) + + bwd-fwd : ∀ {l1 l2 : Level} {A : Type l1} {B : Type l2} + → (e : A ≃ B) + → bwd e ∘ fwd e ∼ id + bwd-fwd e = is-equiv.is-post-inverse (_≃_.is-equivalence (fix e)) + + data ℤ : Type where + Pos : ℕ → ℤ + Zero : ℤ + Neg : ℕ → ℤ + + 0ℤ : ℤ + 0ℤ = Zero + + succ-fn : ℤ → ℤ + succ-fn (Pos x) = Pos (suc x) + succ-fn Zero = Pos zero + succ-fn (Neg zero) = Zero + succ-fn (Neg (suc x)) = Neg x + + pred-fn : ℤ → ℤ + pred-fn (Pos zero) = Zero + pred-fn (Pos (suc x)) = Pos x + pred-fn Zero = Neg zero + pred-fn (Neg x) = Neg (suc x) + + succ-pred : (x : ℤ) → succ-fn (pred-fn x) ≡ x + succ-pred (Pos zero) = refl _ + succ-pred (Pos (suc x)) = refl _ + succ-pred Zero = refl _ + succ-pred (Neg zero) = refl _ + succ-pred (Neg (suc x)) = refl _ + + pred-succ : (x : ℤ) → pred-fn (succ-fn x) ≡ x + pred-succ (Pos zero) = refl _ + pred-succ (Pos (suc x)) = refl _ + pred-succ Zero = refl _ + pred-succ (Neg zero) = refl _ + pred-succ (Neg (suc x)) = refl _ + + succℤ : ℤ ≃ ℤ + succℤ = improve (Isomorphism succ-fn (Inverse pred-fn pred-succ succ-pred)) + + ℤ-rec : {X : Type} + (b : X) + (s : X ≃ X) + → ℤ → X + ℤ-rec b s (Pos zero) = fwd s b + ℤ-rec b s (Pos (suc x)) = fwd s (ℤ-rec b s (Pos x)) + ℤ-rec b s Zero = b + ℤ-rec b s (Neg zero) = bwd s b + ℤ-rec b s (Neg (suc x)) = bwd s (ℤ-rec b s (Neg x)) + + ℤ-rec-0ℤ : {X : Type} + (b : X) + (s : X ≃ X) + → ℤ-rec b s 0ℤ ≡ b + ℤ-rec-0ℤ b s = refl _ + + ℤ-rec-succℤ : {X : Type} + (b : X) + (s : X ≃ X) + (a : ℤ) → ℤ-rec b s (fwd succℤ a) ≡ fwd s (ℤ-rec b s a) + ℤ-rec-succℤ b s (Pos x) = refl _ + ℤ-rec-succℤ b s Zero = refl _ + ℤ-rec-succℤ b s (Neg zero) = ! (fwd-bwd s b) + ℤ-rec-succℤ b s (Neg (suc zero)) = ! (fwd-bwd s _) + ℤ-rec-succℤ b s (Neg (suc (suc x))) = ! (fwd-bwd s _) + + f-pred : {X : Type} + (f : ℤ → X) + (s : X ≃ X) + → ((f ∘ fwd succℤ) ∼ (fwd s ∘ f)) + → f ∘ bwd succℤ ∼ bwd s ∘ f + f-pred f s h x = ! (bwd-fwd s _) ∙ + ! (ap (bwd s) (h (bwd succℤ x))) ∙ + ap (bwd s) (ap f (fwd-bwd succℤ x)) + + ℤ-rec-unique : {X : Type} + (f : ℤ → X) + (z : X) + (s : X ≃ X) + → f 0ℤ ≡ z + → ((f ∘ fwd succℤ) ∼ (fwd s ∘ f)) + → (x : ℤ) → f x ≡ ℤ-rec z s x + ℤ-rec-unique f z s p q (Pos zero) = q 0ℤ ∙ ap (fwd s) p + ℤ-rec-unique f z s p q (Pos (suc x)) = q (Pos x) ∙ ap (fwd s) (ℤ-rec-unique f z s p q (Pos x)) + ℤ-rec-unique f z s p q Zero = p + ℤ-rec-unique f z s p q (Neg zero) = f-pred f s q Zero ∙ ap (bwd s) p + ℤ-rec-unique f z s p q (Neg (suc x)) = f-pred f s q (Neg x) ∙ ap (bwd s) ((ℤ-rec-unique f z s p q (Neg x))) + + -- the fact that ℤ is a set can be proved using Hedberg's theorem, + -- see ../Lecture-Notes/Hedbergs-Theorem.lagda.md + -- and ../Lecture-Notes/decidability.lagda.md + -- ℤ has decidable equality because it is equivalent to the coproduct ℕ + 1 + ℕ + -- and ℕ has decidable equality and coproducts preserve decidable equality + + + module Instantiate = AssumeInts ℤ 0ℤ succℤ ℤ-rec ℤ-rec-0ℤ ℤ-rec-succℤ ℤ-rec-unique + +``` diff --git a/src/HoTTEST/Agda/HITs/Solutions4.lagda.md b/src/HoTTEST/Agda/HITs/Solutions4.lagda.md new file mode 100644 index 0000000..58a04eb --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions4.lagda.md @@ -0,0 +1,179 @@ +```agda +{-# OPTIONS --without-K --rewriting #-} + +open import new-prelude +open import Lecture4-notes + +module Solutions4 where +``` + +# Constructing homotopies between paths + +Give two "different" path-between-paths/homotopies between (loop ∙ ! +loop) ∙ loop and loop. They should be different in the sense that one +should cancel the !loop with the first loop, and one with the second +loop. But these aren't really *really* different, in that there will be +a path-between-paths-between-paths between the two! (Harder exercise +that we haven't really prepared for: prove this!) + +```agda +homotopy1 : (loop ∙ ! loop) ∙ loop ≡ loop +homotopy1 = (loop ∙ ! loop) ∙ loop ≡⟨ ap ( \ H → H ∙ loop) (!-inv-r loop) ⟩ + refl _ ∙ loop ≡⟨ ∙unit-l loop ⟩ + loop ∎ + +homotopy2 : (loop ∙ ! loop) ∙ loop ≡ loop +homotopy2 = (loop ∙ ! loop) ∙ loop ≡⟨ ! (∙assoc loop (! loop) loop) ⟩ + loop ∙ (! loop ∙ loop) ≡⟨ ap (\ H → loop ∙ H) (!-inv-l loop) ⟩ + loop ∎ +``` + +For fun, here's the proof they're the same. The above proofs work for +any path p, so we can generalize the goal and then do path induction. +in practice, it would be better to define homotopy1 and homotopy2 for +a general p in the first place and then instantiate them to loop so +that you don't need to copy and paste their definitions into the goal +here, but I think it's helpful to be concrete when first practicing +these path algebra steps. + + +```agda +path-between-paths-between-paths : homotopy1 ≡ homotopy2 +path-between-paths-between-paths = gen loop where + gen : ∀ {A : Type} {x y : A} (p : x ≡ y) + → (ap ( \ H → H ∙ p) (!-inv-r p) ∙ ∙unit-l p) + ≡ + (! (∙assoc p (! p) p) ∙ ap (\ H → p ∙ H) (!-inv-l p)) + [ ((p ∙ ! p) ∙ p) ≡ p [ x ≡ y [ A ] ] ] + gen (refl _) = refl _ + +``` + +# Functions are group homomorphism + +State and prove a general lemma about what ap of a function on the +inverse of a path (! p) does (see ap-∙ for inspiration). + +State and prove a general lemma about what ! (p ∙ q) is. + +Us them to prove that the double function takes loop-inverse to +loop-inverse concatenated with itself. + +```agda +!-∙ : {A : Type} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → ! (p ∙ q) ≡ ! q ∙ ! p +!-∙ (refl _) (refl _) = refl _ + +ap-! : {A B : Type} {f : A → B} {x y : A} (p : x ≡ y) + → ap f (! p) ≡ ! (ap f p) +ap-! (refl _) = refl _ + +double-!loop : ap double (! loop) ≡ ! loop ∙ ! loop +double-!loop = ap double (! loop) ≡⟨ ap-! loop ⟩ + ! (ap double loop) ≡⟨ ap ! (S1-rec-loop _ _) ⟩ + ! (loop ∙ loop) ≡⟨ !-∙ loop _ ⟩ + ! loop ∙ ! loop ∎ +``` + +```agda +invert : S1 → S1 +invert = S1-rec base (! loop) +``` + +# Circles equivalence + +See the maps between the 1 point circle and the 2 point circle in the +lecture code. Check that the composite map S1 → S1 is +homotopic to the identity on base and loop. + +```agda +to-from-base : from (to base) ≡ base +to-from-base = refl _ + +to-from-loop : ap from (ap to loop) ≡ loop +to-from-loop = ap from (ap to loop) ≡⟨ ap (ap from) (S1-rec-loop _ _) ⟩ + ap from (east ∙ ! west) ≡⟨ ap-∙ east (! west) ⟩ + ap from east ∙ ap from (! west) ≡⟨ ap₂ _∙_ (Circle2-rec-east _ _ _ _) + (ap-! west ∙ ap ! (Circle2-rec-west _ _ _ _)) ⟩ + loop ∙ ! (refl base) ≡⟨ refl _ ⟩ + loop ∎ +``` + +Note: the problems below here are progressively more optional, so if you +run out of time/energy at some point that is totally fine. + +# Torus to circles + +The torus is equivalent to the product of two circles S1 × S1. The idea +for the map from the torus to S1 × S1 that is part of this equivalence +is that one loop on on the torus is sent to to the circle loop in one +copy of S1, and the other loop on the torus to the loop in the other +copy of the circle. Define this map. + +Hint: for the image of the square, you might want a lemma saying how +paths in product types compose: + +```agda +compose-pair≡ : {A B : Type} {x1 x2 x3 : A} {y1 y2 y3 : B} + (p12 : x1 ≡ x2) (p23 : x2 ≡ x3) + (q12 : y1 ≡ y2) (q23 : y2 ≡ y3) + → (pair≡ p12 q12) ∙ (pair≡ p23 q23) ≡ pair≡ (p12 ∙ p23) (q12 ∙ q23) +compose-pair≡ (refl _) (refl _) (refl _) (refl _) = refl _ + +torus-to-circles : Torus → S1 × S1 +torus-to-circles = T-rec (base , base) + (pair≡ (refl base) loop ) + (pair≡ loop (refl base )) + (compose-pair≡ (refl _) loop loop (refl _) ∙ + ap₂ pair≡ (∙unit-l loop) (! (∙unit-l loop)) ∙ + ! (compose-pair≡ loop (refl _) (refl _) loop)) +``` + +# Suspensions + +Find a type X such that the two-point circle Circle2 is equivalent to +the suspension Susp X. Check your answer by defining a logical +equivalence (functions back and forth), since we haven't seen how to +prove that such functions are inverse yet. + +```agda +c2s : Circle2 → Susp Bool +c2s = Circle2-rec northS southS (merid true) (merid false) + +s2c : Susp Bool → Circle2 +s2c = Susp-rec north south (\ { true → west ; false → east }) +``` + +Suspension is a functor from types, which means that it acts on +functions as well as types. Define the action of Susp on functions: + +```agda +susp-func : {X Y : Type} → (f : X → Y) → Susp X → Susp Y +susp-func f = Susp-rec northS southS (\ x → merid (f x) ) +``` + +To really prove that Susp is a functor, we should check that this action +on functions preserves the identity function and composition of +functions. But to do that we would need the dependent elimination rule +for suspensions, which we haven't talked about yet. + + +# Pushouts + +Fix a type X. Find types A,B,C with functions f : C → A and g : C → B +such that the suspension Susp X is equivalent to the Pushout C A B f g. +Check your answer by defining a logical equivalence (functions back and +forth), since we haven't seen how to prove that such functions are +inverse yet. + +```agda +SuspFromPush : Type → Type +SuspFromPush A = Pushout A 𝟙 𝟙 (\ _ → ⋆) (\ _ → ⋆) + +s2p : (A : Type) → Susp A → SuspFromPush A +s2p A = Susp-rec (inl ⋆) (inr ⋆) glue + +p2s : (A : Type) → SuspFromPush A → Susp A +p2s A = Push-rec (\ _ → northS) (\ _ → southS) merid +``` + diff --git a/src/HoTTEST/Agda/HITs/Solutions5-dan.lagda.md b/src/HoTTEST/Agda/HITs/Solutions5-dan.lagda.md new file mode 100644 index 0000000..08ebe0b --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions5-dan.lagda.md @@ -0,0 +1,284 @@ + +Note: for this week we have two solution files for the exercises. Dan +Licata's solutions are below. Tom de Jong's solutions in are in +Solutions5-tom.lagda.md. We are providing both to illustrate slightly +different ways of using Agda. Tom's are a little more verbose and +easier for a person to read non-interactively. Dan's are a little more +inlined and meant to show you roughly the minimum text you need to get +Agda to check the proofs, but are harder to read unless you do so +interactively (by putting holes around sub-expressions and getting Agda +to tell you what the types are supposed to be). It's sometimes nice to +code in the terse style when you are coding for yourself, and then to +clean it up more like the more verbose solution when writing it up for +other people. + +Note: Agda will get confused if you make a single file that imports both +our solutions and your solutions in Exercises5. The reason is that +rewrite rules are installed globally, and both files declare rewrites +for the reduction rules for suspension types (if you get that far), and +Agda won't let you install a rewrite for something that already reduces. +(For a similar reason, Solutions5-tom has the rewrites commented out to +make our build scripts for github work, so you will need to uncomment +those in your copy to load the file.) + +```agda +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude +open import Lecture5-notes +open import Solutions4 using (ap-!; to-from-base; to-from-loop; s2c; c2s; susp-func) + +module Solutions5-dan where +``` + +# 1 point and 2 point circles are equivalent (⋆) + +In lecture, we defined maps between the one point circle (S1) and the +two point circle (Circle2) and checked that the round-trip on Circle2 is +the identity. + +Prove that the round-trip on S1 is the identity (use to-from-base +and to-from-loop from the Lecture 4 exercises), and package all of +these items up as an equivalence S1 ≃ Circle2. + +```agda +to-from : (x : S1) → from (to x) ≡ x +to-from = S1-elim _ + to-from-base + (PathOver-roundtrip≡ from to loop (∙unit-l _ ∙ to-from-loop)) + +circles-equivalent : S1 ≃ Circle2 +circles-equivalent = improve (Isomorphism to (Inverse from to-from from-to)) +``` + +# Reversing the circle (⋆⋆) + +Define a map S1 → S1 that "reverses the orientation of the circle", +i.e. sends loop to ! loop. + +```agda +rev : S1 → S1 +rev = S1-rec base (! loop) +``` + +Prove that rev is an equivalence. Hint: you will need to state and prove +one new generalized "path algebra" lemma and to use one of the lemmas from +the "Functions are group homomorphism" section of Lecture 4's exercises. +``` +!-invol : {l : Level} {A : Type l} {x y : A} → (p : x ≡ y) → (! (! p)) ≡ p +!-invol (refl _) = refl _ + +rev-equiv : is-equiv rev +rev-equiv = Inverse rev invol rev invol where + invol : rev ∘ rev ∼ id + invol = S1-elim _ + (refl _) + (PathOver-roundtrip≡ rev rev _ + (∙unit-l _ ∙ + ap (ap rev) (S1-rec-loop _ _) ∙ + ap-! loop ∙ + ap ! (S1-rec-loop _ _) ∙ + !-invol loop)) +``` + + +# Circles to torus (⋆⋆) + +In the Lecture 4 exercises, you defined a map from the Torus to S1 × S1. +In this problem, you will define a converse map. The goal is for these +two maps to be part of an equivalence, but we won't prove that in these +exercises. + +You will need to state and prove a lemma characterizing a path over a +path in a path fibration. Then, to define the map S1 × S1 → Torus, you +will want to curry it and use S1-rec and/or S1-elim on each circle. + +```agda +PathOver-path≡ : ∀ {A B : Type} {g : A → B} {f : A → B} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ (g a)} + {r : (f a') ≡ (g a')} + → q ∙ ap g p ≡ ap f p ∙ r + → q ≡ r [ (\ x → (f x) ≡ (g x)) ↓ p ] +PathOver-path≡ {p = (refl _)} h = path-to-pathover (h ∙ ∙unit-l _) + +circles-to-torus : S1 → (S1 → Torus) +circles-to-torus = S1-rec (S1-rec baseT qT) + (λ≡ (S1-elim _ + pT + (PathOver-path≡ (ap (\ H → pT ∙ H) (S1-rec-loop _ _) ∙ + sT ∙ + ap (\ H → H ∙ pT) (! (S1-rec-loop _ _)))))) + +circles-to-torus-nofunext : S1 → (S1 → Torus) +circles-to-torus-nofunext x y = S1-rec (S1-rec baseT qT y) + ( (S1-elim (\ y → S1-rec baseT qT y ≡ S1-rec baseT qT y) + pT + (PathOver-path≡ (ap (\ H → pT ∙ H) (S1-rec-loop _ _) ∙ + sT ∙ + ap (\ H → H ∙ pT) (! (S1-rec-loop _ _))))) y ) + x + +circles-to-torus' : S1 × S1 → Torus +circles-to-torus' (x , y) = circles-to-torus x y +``` + +**Below are some "extra credit" exercises if you want more to do. These +are (even more) optional: nothing in the next lecture will depend on you +understanding them. The next section (H space) is harder code but uses +only the circle, whereas the following sections are a bit easier code +but require understanding the suspension type, which we haven't talked +about too much yet.** + +# H space + +The multiplication operation on the circle discussed in lecture is part +of what is called an "H space" structure on the circle. One part of +this structure is that the circle's basepoint is a unit element for +multiplication. + +(⋆) Show that base is a left unit. +```agda +mult-unit-l : (y : S1) → mult base y ≡ y +mult-unit-l y = refl _ +``` + +(⋆) Because we'll need it in a second, show that ap distributes over +function composition: +```agda +ap-∘ : ∀ {l1 l2 l3 : Level} {A : Type l1} {B : Type l2} {C : Type l3} + (f : A → B) (g : B → C) + {a a' : A} + (p : a ≡ a') + → ap (g ∘ f) p ≡ ap g (ap f p) +ap-∘ _ _ (refl _) = refl _ +``` + +(⋆⋆) Suppose we have a curried function f : S1 → A → B. Under the +equivalence between paths in S1 × A and pairs of paths discussed in +Lecture 3, we can then "apply" (the uncurried version of) f to a pair of +paths (p : x ≡ y [ S1 ] , q : z ≡ w [ A ]) to get a path (f x z ≡ f y w +[ B ]). In the special case where q is reflexivity on a, this +application to p and q can be simplified to ap (\ x → f x a) p : f x a ≡ +f y a [ B ]. + +Now, suppose that f is defined by circle recursion. We would expect +some kind of reduction for applying f to the pair of paths (loop , q) --- +i.e. we should have reductions for *nested* pattern matching on HITs. +In the case where q is reflexivity, applying f to the pair (loop , refl) +can reduce like this: +```agda +S1-rec-loop-1 : ∀ {A B : Type} {f : A → B} {h : f ≡ f} {a : A} + → ap (\ x → S1-rec f h x a) loop ≡ app≡ h a +S1-rec-loop-1 {f = f}{h}{a} = (ap-∘ (S1-rec f h) (\ z → z a) loop) ∙ ap (ap (\ z → z a)) (S1-rec-loop _ _) +``` +Prove this reduction using ap-∘ and the reduction rule for S1-rec on the loop. + +(⋆⋆⋆) Show that base is a right unit for multiplication. You will need +a slightly different path-over lemma than before. + +```agda +PathOver-endo≡ : ∀ {A : Type} {f : A → A} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ a} + {r : (f a') ≡ a'} + → ! q ∙ ((ap f p) ∙ r) ≡ p + → q ≡ r [ (\ x → f x ≡ x) ↓ p ] +PathOver-endo≡ {p = (refl _)} {q = q} {r} h = + path-to-pathover (ap (\ H → q ∙ H) (! h) ∙ + ( ∙assoc _ _ (refl _ ∙ r) ∙ + (ap (\ H → H ∙ (refl _ ∙ r)) (!-inv-r q) ∙ + (∙unit-l (refl _ ∙ r) ∙ ∙unit-l r )) )) + +mult-unit-r : (x : S1) → mult x base ≡ x +mult-unit-r = S1-elim _ + (refl _) + (PathOver-endo≡ + ((∙unit-l _) ∙ (S1-rec-loop-1 ∙ ( λ≡β _ _)) )) +``` + +# Suspensions and the 2-point circle + +(⋆) Postulate the computation rules for the non-dependent susp-rec and +declare rewrites for the point reduction rules on the point constructors. +```agda +postulate + Susp-rec-north : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m northS ≡ n + Susp-rec-south : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m southS ≡ s +{-# REWRITE Susp-rec-north #-} +{-# REWRITE Susp-rec-south #-} +postulate + Susp-rec-merid : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → (x : A) → ap (Susp-rec n s m) (merid x) ≡ m x +``` + +(⋆) Postulate the dependent elimination rule for suspensions: + +```agda +postulate + Susp-elim : {A : Type} (P : Susp A → Type) + → (n : P northS) + → (s : P southS) + → (m : (x : A) → n ≡ s [ P ↓ merid x ]) + → (x : Susp A) → P x +``` + +(⋆⋆) Show that the maps s2c and c2s from the Lecture 4 exercises are mutually inverse: + +```agda +c2s2c : (x : Circle2) → s2c (c2s x) ≡ x +c2s2c = Circle2-elim _ (refl _) (refl _) + (PathOver-roundtrip≡ s2c c2s _ + (∙unit-l _ ∙ (ap (ap s2c) (Circle2-rec-west _ _ _ _) ∙ Susp-rec-merid _ _ _ _))) + (PathOver-roundtrip≡ s2c c2s _ + (∙unit-l _ ∙ (ap (ap s2c) (Circle2-rec-east _ _ _ _) ∙ Susp-rec-merid _ _ _ _))) + +s2c2s : (x : Susp Bool) → c2s (s2c x) ≡ x +s2c2s = Susp-elim _ (refl _) (refl _) \ { true → PathOver-roundtrip≡ c2s s2c _ + (∙unit-l _ ∙ + (ap (ap c2s) (Susp-rec-merid _ _ _ _) + ∙ Circle2-rec-west _ _ _ _)) ; + false → PathOver-roundtrip≡ c2s s2c _ + (∙unit-l _ ∙ + (ap (ap c2s) (Susp-rec-merid _ _ _ _) + ∙ Circle2-rec-east _ _ _ _))} +``` + +(⋆) Conclude that Circle2 is equivalent to Susp Bool: + +```agda +Circle2-Susp-Bool : Circle2 ≃ Susp Bool +Circle2-Susp-Bool = improve (Isomorphism c2s (Inverse s2c c2s2c s2c2s)) +``` + +# Functoriality of suspension (⋆⋆) + +In the Lecture 4 exercises, we defined functoriality for the suspension +type, which given a function X → Y gives a function Σ X → Σ Y. Show +that this operation is functorial, meaning that it preserves identity +and composition of functions: +```agda +susp-func-id : ∀ {X : Type} → susp-func {X} id ∼ id +susp-func-id = Susp-elim _ (refl _) + (refl _) + (\x → PathOver-endo≡ (∙unit-l _ ∙ (Susp-rec-merid _ _ _ _)) ) + +susp-func-∘ : ∀ {X Y Z : Type} (f : X → Y) (g : Y → Z) + → susp-func {X} (g ∘ f) ∼ susp-func g ∘ susp-func f +susp-func-∘ f g = Susp-elim _ + (refl _) + (refl _) + (\x → PathOver-path≡ (∙unit-l _ ∙ + ap-∘ (susp-func f) (susp-func g) _ ∙ + ap (ap (susp-func g)) (Susp-rec-merid _ _ _ _) ∙ + Susp-rec-merid _ _ _ _ ∙ + ! (Susp-rec-merid _ _ _ _))) +``` + + + diff --git a/src/HoTTEST/Agda/HITs/Solutions5-tom.lagda.md b/src/HoTTEST/Agda/HITs/Solutions5-tom.lagda.md new file mode 100644 index 0000000..7be99f2 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions5-tom.lagda.md @@ -0,0 +1,354 @@ + + +**To make this file typecheck, you will need to uncomment two REWRITE + rules, see (†) below. (They are commented because of idiosyncrasies in + building corresponding GitHub pages.)** + +For this week we have two solution files for the exercises. Tom de +Jong's solutions below. Dan Licata's solutions are in +Solutions5-dan.lagda.md. We are providing both to illustrate slightly +different ways of using Agda. Tom's are a little more verbose and +easier for a person to read non-interactively. Dan's are a little more +inlined and meant to show you roughly the minimum text you need to get +Agda to check the proofs, but are harder to read unless you do so +interactively (by putting holes around sub-expressions and getting Agda +to tell you what the types are supposed to be). It's sometimes nice to +code in the terse style when you are coding for yourself, and then to +clean it up more like the more verbose solution when writing it up for +other people. + +Note: Agda will get confused if you make a single file that imports both +our solutions and your solutions in Exercises5. The reason is that +rewrite rules are installed globally, and both files declare rewrites +for the reduction rules for suspension types (if you get that far), and +Agda won't let you install a rewrite for something that already reduces. +(For a similar reason, Solutions5-tom has the rewrites commented out to +make our build scripts for github work, so you will need to uncomment +those in your copy to load the file.) + + +```agda +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude +open import Lecture5-notes +open import Solutions4 using (ap-!; to-from-base; to-from-loop; s2c; c2s; susp-func) + +module Solutions5-tom where +``` + +# 1 point and 2 point circles are equivalent (⋆) + +In lecture, we defined maps between the one point circle (S1) and the +two point circle (Circle2) and checked that the round-trip on Circle2 is +the identity. + +Prove that the round-trip on S1 is the identity (use to-from-base +and to-from-loop from the Lecture 4 exercises), and package all of +these items up as an equivalence S1 ≃ Circle2. + +```agda +to-from : (x : S1) → from (to x) ≡ x +to-from = S1-elim (\x → from (to x) ≡ x) to-from-base p + where + p : to-from-base ≡ to-from-base [ (\ x → from (to x) ≡ x) ↓ loop ] + p = PathOver-roundtrip≡ from to loop q + where + q : ! to-from-base ∙ (ap from (ap to loop) ∙ to-from-base) + ≡ loop + q = ! to-from-base ∙ (ap from (ap to loop) ∙ to-from-base) ≡⟨ refl _ ⟩ + refl _ ∙ (ap from (ap to loop)) ∙ refl _ ≡⟨ refl _ ⟩ + refl _ ∙ (ap from (ap to loop)) ≡⟨ ∙unit-l (ap from (ap to loop)) ⟩ + ap from (ap to loop) ≡⟨ to-from-loop ⟩ + loop ∎ + +circles-equivalent : S1 ≃ Circle2 +circles-equivalent = improve (Isomorphism to (Inverse from to-from from-to)) +``` + +# Reversing the circle (⋆⋆) + +Define a map S1 → S1 that "reverses the orientation of the circle", +i.e. sends loop to ! loop. + +```agda +rev : S1 → S1 +rev = S1-rec base (! loop) +``` + +Prove that rev is an equivalence. Hint: you will need to state and prove +one new generalized "path algebra" lemma and to use one of the lemmas from +the "Functions are group homomorphism" section of Lecture 4's exercises. +```agda +rev-loop : ap rev loop ≡ ! loop +rev-loop = S1-rec-loop base (! loop) + +!-involutive : {X : Type} {x y : X} (p : x ≡ y) → ! (! p) ≡ p +!-involutive (refl _) = refl _ + +rev-equiv : is-equiv rev +rev-equiv = Inverse rev h rev h + where + h : rev ∘ rev ∼ id + h = S1-elim (\ x → rev (rev x) ≡ x) (refl base) p + where + p : refl base ≡ refl base [ (\x -> rev (rev x) ≡ x) ↓ loop ] + p = PathOver-roundtrip≡ rev rev loop q + where + q = refl base ∙ ap rev (ap rev loop) ≡⟨ ∙unit-l _ ⟩ + ap rev (ap rev loop) ≡⟨ ap (ap rev) rev-loop ⟩ + ap rev (! loop) ≡⟨ ap-! loop ⟩ + ! (ap rev loop) ≡⟨ ap ! rev-loop ⟩ + ! (! loop) ≡⟨ !-involutive loop ⟩ + loop ∎ +``` + +# Circles to torus (⋆⋆) + +In the Lecture 4 exercises, you defined a map from the Torus to S1 × S1. +In this problem, you will define a converse map. The goal is for these +two maps to be part of an equivalence, but we won't prove that in these +exercises. + +You will need to state and prove a lemma characterizing a path over a +path in a path fibration. Then, to define the map S1 × S1 → Torus, you +will want to curry it and use S1-rec and/or S1-elim on each circle. + +```agda +PathOver-path≡ : ∀ {A B : Type} {g : A → B} {f : A → B} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ (g a)} + {r : (f a') ≡ (g a')} + → q ∙ ap g p ≡ ap f p ∙ r + → q ≡ r [ (\ x → (f x) ≡ (g x)) ↓ p ] +PathOver-path≡ {g = g} {f = f} {p = refl a} {q} {r} e = to-prove + where + to-prove : q ≡ r [ (\ x → (f x) ≡ (g x)) ↓ (refl a) ] + to-prove = path-to-pathover (q ≡⟨ e ⟩ + refl (f a) ∙ r ≡⟨ ∙unit-l r ⟩ + r ∎) + +circles-to-torus : S1 → (S1 → Torus) +circles-to-torus = S1-rec f e + where + f : S1 → Torus + f = S1-rec baseT pT + e : f ≡ f + e = λ≡ (S1-elim (λ x → f x ≡ f x) qT p) + where + p : qT ≡ qT [ (\ x → f x ≡ f x) ↓ loop ] + p = PathOver-path≡ (qT ∙ ap f loop ≡⟨ ap (qT ∙_) lemma ⟩ + qT ∙ pT ≡⟨ ! sT ⟩ + pT ∙ qT ≡⟨ ap (_∙ qT) (! lemma) ⟩ + ap f loop ∙ qT ∎) + where + lemma : ap f loop ≡ pT + lemma = S1-rec-loop baseT pT + +circles-to-torus' : S1 × S1 → Torus +circles-to-torus' (x , y) = circles-to-torus x y +``` + +**Below are some "extra credit" exercise if you want more to do. These +are (even more) optional: nothing in the next lecture will depend on you +understanding them. The next section (H space) is harder code but uses +only the circle, whereas the following sections are a bit easier code +but require understanding the suspension type, which we haven't talked +about too much yet.** + +# H space + +The multiplication operation on the circle discussed in lecture is part +of what is called an "H space" structure on the circle. One part of +this structure is that the circle's basepoint is a unit element for +multiplication. + +(⋆) Show that base is a left unit. +```agda +mult-unit-l : (y : S1) → mult base y ≡ y +mult-unit-l y = refl _ -- Choosing (refl _) here helps later on. +``` + +(⋆) Because we'll need it in a second, show that ap distributes over +function composition: +```agda +ap-∘ : ∀ {l1 l2 l3 : Level} {A : Type l1} {B : Type l2} {C : Type l3} + (f : A → B) (g : B → C) + {a a' : A} + (p : a ≡ a') + → ap (g ∘ f) p ≡ ap g (ap f p) +ap-∘ _ _ (refl _) = refl _ +``` + +(⋆⋆) Suppose we have a curried function f : S1 → A → B. Under the +equivalence between paths in S1 × A and pairs of paths discussed in +Lecture 3, we can then "apply" (the uncurried version of) f to a pair of +paths (p : x ≡ y [ S1 ] , q : z ≡ w [ A ]) to get a path (f x z ≡ f y w +[ B ]). In the special case where q is reflexivity on a, this +application to p and q can be simplified to ap (\ x → f x a) p : f x a ≡ +f y a [ B ]. + +Now, suppose that f is defined by circle recursion. We would expect +some kind of reduction for applying f to the pair of paths (loop , q) --- +i.e. we should have reductions for *nested* pattern matching on HITs. +In the case where q is reflexivity, applying f to the pair (loop , refl) +can reduce like this: +```agda +S1-rec-loop-1 : ∀ {A B : Type} {f : A → B} {h : f ≡ f} {a : A} + → ap (\ x → S1-rec f h x a) loop ≡ app≡ h a +S1-rec-loop-1 {f = f} {h} {a} = ap (\ x → S1-rec f h x a) loop ≡⟨ ap-∘ (S1-rec f h) (\ f → f a) loop ⟩ + ap (\ f → f a) (ap (S1-rec f h) loop) ≡⟨ ap (ap (\ f → f a)) (S1-rec-loop f h) ⟩ + ap (\ f → f a) h ≡⟨ refl _ ⟩ + app≡ h a ∎ +``` +Prove this reduction using ap-∘ and the reduction rule for S1-rec on the loop. + +(⋆⋆⋆) Show that base is a right unit for multiplication. You will need +a slightly different path-over lemma than before. +```agda +PathOver-endo≡ : ∀ {A : Type} {f : A → A} + {a a' : A} {p : a ≡ a'} + {q : (f a) ≡ a} + {r : (f a') ≡ a'} + → ! q ∙ ((ap f p) ∙ r) ≡ p + → q ≡ r [ (\ x → f x ≡ x) ↓ p ] +PathOver-endo≡ {f = f} {p = (refl a)} {q = q} {r} h = to-show + where + to-show : q ≡ r [ (\ x → f x ≡ x) ↓ refl a ] + to-show = path-to-pathover (q ≡⟨ refl _ ⟩ + q ∙ refl a ≡⟨ ap (q ∙_) (! h) ⟩ + q ∙ (! q ∙ (refl _ ∙ r)) ≡⟨ ap (\ - → q ∙ (! q ∙ -)) (∙unit-l r) ⟩ + q ∙ (! q ∙ r) ≡⟨ ∙assoc q (! q) r ⟩ + (q ∙ ! q) ∙ r ≡⟨ ap (_∙ r) (!-inv-r q) ⟩ + refl _ ∙ r ≡⟨ ∙unit-l r ⟩ + r ∎) + +mult-unit-r : (x : S1) → mult x base ≡ x +mult-unit-r = S1-elim (\ x → mult x base ≡ x) (mult-unit-l base) p + where + p : mult-unit-l base ≡ mult-unit-l base [ (\ x → mult x base ≡ x) ↓ loop ] + p = PathOver-endo≡ q + where + q = ! (mult-unit-l base) ∙ (ap (\ x → mult x base) loop ∙ mult-unit-l base) ≡⟨ refl _ ⟩ + ! (refl _) ∙ (ap (\ x → mult x base) loop ∙ refl _) ≡⟨ refl _ ⟩ + ! (refl _) ∙ ap (\ x → mult x base) loop ≡⟨ ∙unit-l _ ⟩ + ap (\ x → mult x base) loop ≡⟨ S1-rec-loop-1 ⟩ + app≡ (λ≡ m) base ≡⟨ λ≡β m base ⟩ + loop ∎ + where + m : (x : S1) → x ≡ x + m = S1-elim _ loop (PathOver-path-loop (refl (loop ∙ loop))) +``` + +# Suspensions and the 2-point circle + +(⋆) Postulate the computation rules for the non-dependent susp-rec and +declare rewrites for the point reduction rules on the point constructors. +```agda +postulate + Susp-rec-north : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m northS ≡ n + Susp-rec-south : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → Susp-rec n s m southS ≡ s +``` + +**(†) Uncomment these REWRITE rules to make the file typecheck. + (They are commented because of idiosyncrasies in building corresponding + GitHub pages.)** +``` +-- {-# REWRITE Susp-rec-north #-} +-- {-# REWRITE Susp-rec-south #-} + +postulate + Susp-rec-merid : {l : Level} {A : Type} {X : Type l} + (n : X) (s : X) (m : A → n ≡ s) + → (x : A) → ap (Susp-rec n s m) (merid x) ≡ m x +``` + +(⋆) Postulate the dependent elimination rule for suspensions: + +```agda +postulate + Susp-elim : {A : Type} (P : Susp A → Type) + → (n : P northS) + → (s : P southS) + → (m : (x : A) → n ≡ s [ P ↓ merid x ]) + → (x : Susp A) → P x +``` + +(⋆⋆) Show that the maps s2c and c2s from the Lecture 4 exercises are mutually inverse: + +```agda +c2s2c : (x : Circle2) → s2c (c2s x) ≡ x +c2s2c = Circle2-elim _ pₙ pₛ pʷ pᵉ + where + pₙ : s2c (c2s north) ≡ north + pₙ = refl _ + pₛ : s2c (c2s south) ≡ south + pₛ = refl _ + pʷ : pₙ ≡ pₛ [ (\ x → s2c (c2s x) ≡ x) ↓ west ] + pʷ = PathOver-roundtrip≡ s2c c2s west + (refl _ ∙ ap s2c (ap c2s west) ≡⟨ ∙unit-l _ ⟩ + ap s2c (ap c2s west) ≡⟨ ap (ap s2c) (Circle2-rec-west _ _ _ _) ⟩ + ap s2c (merid true) ≡⟨ Susp-rec-merid _ _ _ _ ⟩ + west ∎) + pᵉ : pₙ ≡ pₛ [ (\ x → s2c (c2s x) ≡ x) ↓ east ] + pᵉ = PathOver-roundtrip≡ s2c c2s east + (refl _ ∙ ap s2c (ap c2s east) ≡⟨ ∙unit-l _ ⟩ + ap s2c (ap c2s east) ≡⟨ ap (ap s2c) (Circle2-rec-east _ _ _ _) ⟩ + ap s2c (merid false) ≡⟨ Susp-rec-merid _ _ _ _ ⟩ + east ∎) + +s2c2s : (x : Susp Bool) → c2s (s2c x) ≡ x +s2c2s = Susp-elim _ pₙ pₛ q + where + pₙ : c2s (s2c northS) ≡ northS + pₙ = refl _ + pₛ : c2s (s2c southS) ≡ southS + pₛ = refl _ + q : (b : Bool) → PathOver (\ x → c2s (s2c x) ≡ x) (merid b) pₙ pₛ + q true = PathOver-roundtrip≡ c2s s2c (merid true) + (refl _ ∙ ap c2s (ap s2c (merid true)) ≡⟨ ∙unit-l _ ⟩ + ap c2s (ap s2c (merid true)) ≡⟨ ap (ap c2s) (Susp-rec-merid _ _ _ _) ⟩ + ap c2s west ≡⟨ Circle2-rec-west _ _ _ _ ⟩ + merid true ∎) + q false = PathOver-roundtrip≡ c2s s2c (merid false) + (refl _ ∙ ap c2s (ap s2c (merid false)) ≡⟨ ∙unit-l _ ⟩ + ap c2s (ap s2c (merid false)) ≡⟨ ap (ap c2s) (Susp-rec-merid _ _ _ _) ⟩ + ap c2s east ≡⟨ Circle2-rec-east _ _ _ _ ⟩ + merid false ∎) +``` + +(⋆) Conclude that Circle2 is equivalent to Susp Bool: + +```agda +Circle2-Susp-Bool : Circle2 ≃ Susp Bool +Circle2-Susp-Bool = improve (Isomorphism c2s (Inverse s2c c2s2c s2c2s)) +``` + +# Functoriality of suspension (⋆⋆) + +In the Lecture 4 exercises, we defined functoriality for the suspension +type, which given a function X → Y gives a function Σ X → Σ Y. Show +that this operation is functorial, meaning that it preserves identity +and composition of functions: +```agda +susp-func-id : ∀ {X : Type} → susp-func {X} id ∼ id +susp-func-id = Susp-elim _ (refl _) (refl _) + (\x → PathOver-endo≡ (∙unit-l _ ∙ (Susp-rec-merid _ _ _ _)) ) + +susp-func-∘ : ∀ {X Y Z : Type} (f : X → Y) (g : Y → Z) + → susp-func {X} (g ∘ f) ∼ susp-func g ∘ susp-func f +susp-func-∘ {X = X} f g = Susp-elim _ (refl _) (refl _) h + where + h : (x : X) + → refl _ ≡ refl _ [ (\s → susp-func (g ∘ f) s ≡ (susp-func g ∘ susp-func f) s) ↓ merid x ] + h x = PathOver-path≡ (refl _ ∙ ap (susp-func g ∘ susp-func f) (merid x) ≡⟨ ∙unit-l _ ⟩ + ap (susp-func g ∘ susp-func f) (merid x) ≡⟨ ap-∘ (susp-func f) (susp-func g) (merid x) ⟩ + ap (susp-func g) (ap (susp-func f) (merid x)) ≡⟨ ap (ap (susp-func g)) (Susp-rec-merid _ _ _ x) ⟩ + ap (susp-func g) (merid (f x)) ≡⟨ Susp-rec-merid _ _ _ (f x) ⟩ + merid (g (f x)) ≡⟨ ! (Susp-rec-merid _ _ _ x) ⟩ + ap (susp-func (g ∘ f)) (merid x) ∎) +``` diff --git a/src/HoTTEST/Agda/HITs/Solutions6-Astra.agda b/src/HoTTEST/Agda/HITs/Solutions6-Astra.agda new file mode 100644 index 0000000..9a9a598 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions6-Astra.agda @@ -0,0 +1,224 @@ +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +open import Lecture6-notes +open import Lecture5-notes + +module Solutions6-Astra where + +private + variable + ℓ ℓ₁ ℓ₂ : Level + +postulate + 𝑓𝑖𝑔₈ : Type + 𝑝𝑡 : 𝑓𝑖𝑔₈ + 𝑙₁ : 𝑝𝑡 ≡ 𝑝𝑡 + 𝑙₂ : 𝑝𝑡 ≡ 𝑝𝑡 + 𝑓𝑖𝑔₈-elim : (X : 𝑓𝑖𝑔₈ → Type ℓ) + (x : X 𝑝𝑡) (p : x ≡ x [ X ↓ 𝑙₁ ]) (q : x ≡ x [ X ↓ 𝑙₂ ]) + (x : 𝑓𝑖𝑔₈) → X x + 𝑓𝑖𝑔₈-elim-𝑝𝑡 : (X : 𝑓𝑖𝑔₈ → Type ℓ) + (x : X 𝑝𝑡) (p : x ≡ x [ X ↓ 𝑙₁ ]) (q : x ≡ x [ X ↓ 𝑙₂ ]) → + 𝑓𝑖𝑔₈-elim X x p q 𝑝𝑡 ≡ x +{-# REWRITE 𝑓𝑖𝑔₈-elim-𝑝𝑡 #-} + +postulate + 𝑓𝑖𝑔₈-elim-𝑙₁ : (X : 𝑓𝑖𝑔₈ → Type ℓ) + (x : X 𝑝𝑡) (p : x ≡ x [ X ↓ 𝑙₁ ]) (q : x ≡ x [ X ↓ 𝑙₂ ]) → + apd (𝑓𝑖𝑔₈-elim X x p q) 𝑙₁ ≡ p + 𝑓𝑖𝑔₈-elim-𝑙₂ : (X : 𝑓𝑖𝑔₈ → Type ℓ) + (x : X 𝑝𝑡) (p : x ≡ x [ X ↓ 𝑙₁ ]) (q : x ≡ x [ X ↓ 𝑙₂ ]) → + apd (𝑓𝑖𝑔₈-elim X x p q) 𝑙₂ ≡ q + +Path→PathP : {A : Type ℓ₁} {B : Type ℓ₂} {a₁ a₂ : A} {b₁ b₂ : B} + (p : a₁ ≡ a₂) → b₁ ≡ b₂ → b₁ ≡ b₂ [ (λ _ → B) ↓ p ] +Path→PathP (refl _) (refl _) = reflo + +PathP→Path : {A : Type ℓ₁} {B : Type ℓ₂} {a₁ a₂ : A} + {b₁ b₂ : B} (p : a₁ ≡ a₂) → b₁ ≡ b₂ [ (λ _ → B) ↓ p ] → b₁ ≡ b₂ +PathP→Path (refl _) reflo = refl _ + +Path-η : {A : Type ℓ₁} {B : Type ℓ₂} + {a1 a2 : A} {b1 b2 : B} (p : a1 ≡ a2) (q : b1 ≡ b2) → + PathP→Path p (Path→PathP p q) ≡ q +Path-η (refl _) (refl _) = refl _ + +ap-apd : {A : Type ℓ₁} {B : Type ℓ₂} {a1 a2 : A} + (p : a1 ≡ a2) (f : A → B) → + Path→PathP p (ap f p) ≡ apd f p +ap-apd (refl _) f = refl reflo + +𝑓𝑖𝑔₈-rec : {X : Type ℓ} (x : X) (p : x ≡ x [ X ]) (q : x ≡ x [ X ]) → + 𝑓𝑖𝑔₈ → X +𝑓𝑖𝑔₈-rec {X = X} x p q = + 𝑓𝑖𝑔₈-elim (λ _ → X) x (Path→PathP 𝑙₁ p) (Path→PathP 𝑙₂ q) + +𝑓𝑖𝑔₈-rec-𝑝𝑡 : {X : Type ℓ} (x : X) (p : x ≡ x [ X ]) (q : x ≡ x [ X ]) → + 𝑓𝑖𝑔₈-rec x p q 𝑝𝑡 ≡ x +𝑓𝑖𝑔₈-rec-𝑝𝑡 _ _ _ = refl _ + +𝑓𝑖𝑔₈-rec-𝑙₁ : {X : Type ℓ} (x : X) (p : x ≡ x [ X ]) (q : x ≡ x [ X ]) → + ap (𝑓𝑖𝑔₈-rec x p q) 𝑙₁ ≡ p +𝑓𝑖𝑔₈-rec-𝑙₁ {X = X} x p q = + ! (Path-η 𝑙₁ (ap (𝑓𝑖𝑔₈-rec x p q) 𝑙₁)) + ∙ ap (PathP→Path 𝑙₁) + (ap-apd 𝑙₁ (𝑓𝑖𝑔₈-rec x p q) + ∙ 𝑓𝑖𝑔₈-elim-𝑙₁ (λ _ → X) x (Path→PathP 𝑙₁ p) (Path→PathP 𝑙₂ q)) + ∙ Path-η 𝑙₁ p + +𝑓𝑖𝑔₈-rec-𝑙₂ : {X : Type ℓ} (x : X) (p : x ≡ x [ X ]) (q : x ≡ x [ X ]) → + ap (𝑓𝑖𝑔₈-rec x p q) 𝑙₂ ≡ q +𝑓𝑖𝑔₈-rec-𝑙₂ {X = X} x p q = + ! (Path-η 𝑙₂ (ap (𝑓𝑖𝑔₈-rec x p q) 𝑙₂)) + ∙ ap (PathP→Path 𝑙₂) + (ap-apd 𝑙₂ (𝑓𝑖𝑔₈-rec x p q) + ∙ 𝑓𝑖𝑔₈-elim-𝑙₂ (λ _ → X) x (Path→PathP 𝑙₁ p) (Path→PathP 𝑙₂ q)) + ∙ Path-η 𝑙₂ q + +concat-equiv : {A : Type} {x y z : A} → + y ≡ z → (x ≡ y) ≃ (x ≡ z) +concat-equiv p = + Equivalence + (λ q → q ∙ p) + (Inverse + (λ q → q ∙ ! p) + (λ q → + ! (∙assoc q p (! p)) ∙ ap (q ∙_) (!-inv-r p)) + (λ q → q ∙ ! p) + (λ q → + ! (∙assoc q (! p) p) ∙ ap (q ∙_) (!-inv-l p))) + +transport-∙ : {A : Type ℓ₁} {B : A → Type ℓ₂} + {x y z : A} (p : x ≡ y) (q : y ≡ z) → + transport B (p ∙ q) ∼ transport B q ∘ transport B p +transport-∙ (refl _) (refl _) α = refl α + +module AssumeF₂ + (F₂ : Type) + (𝑒 : F₂) + (𝑠ℎ₁ : F₂ ≃ F₂) + (𝑠ℎ₂ : F₂ ≃ F₂) + (F₂-rec : {ℓ : Level} {X : Type ℓ} (x : X) (m1 : X ≃ X) (m2 : X ≃ X) → + F₂ → X) + (F₂-rec-𝑒 : {ℓ : Level} {X : Type ℓ} (x : X) (m1 : X ≃ X) (m2 : X ≃ X) → + F₂-rec x m1 m2 𝑒 ≡ x) + (F₂-rec-𝑠ℎ₁ : {ℓ : Level} {X : Type ℓ} (x : X) (m1 : X ≃ X) (m2 : X ≃ X) + (a : F₂) → F₂-rec x m1 m2 (fwd 𝑠ℎ₁ a) ≡ fwd m1 (F₂-rec x m1 m2 a)) + (F₂-rec-𝑠ℎ₂ : {ℓ : Level} {X : Type ℓ} (x : X) (m1 : X ≃ X) (m2 : X ≃ X) + (a : F₂) → F₂-rec x m1 m2 (fwd 𝑠ℎ₂ a) ≡ fwd m2 (F₂-rec x m1 m2 a)) + (F₂-rec-unique : {ℓ : Level} {X : Type ℓ} (f : F₂ → X) (x : X) + (m1 : X ≃ X) (m2 : X ≃ X) → f 𝑒 ≡ x → + ((f ∘ fwd 𝑠ℎ₁) ∼ (fwd m1 ∘ f)) → ((f ∘ fwd 𝑠ℎ₂) ∼ (fwd m2 ∘ f)) → + (z : F₂) → f z ≡ F₂-rec x m1 m2 z) + (hSetF : is-set F₂) where + + Cover : 𝑓𝑖𝑔₈ → Type + Cover = 𝑓𝑖𝑔₈-rec F₂ (ua 𝑠ℎ₁) (ua 𝑠ℎ₂) + + encode : {x : 𝑓𝑖𝑔₈} → 𝑝𝑡 ≡ x → Cover x + encode p = transport Cover p 𝑒 + + loopify : F₂ → 𝑝𝑡 ≡ 𝑝𝑡 + loopify = F₂-rec (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) + + tr-Cover-𝑙₁ : (α : F₂) → transport Cover 𝑙₁ α ≡ fwd 𝑠ℎ₁ α + tr-Cover-𝑙₁ α = + transport Cover 𝑙₁ α + ≡⟨ transport-ap-assoc Cover 𝑙₁ ⟩ + transport (λ X → X) (ap Cover 𝑙₁) α + ≡⟨ ap (λ ϕ → transport (λ X → X) ϕ α) + (𝑓𝑖𝑔₈-rec-𝑙₁ F₂ (ua 𝑠ℎ₁) (ua 𝑠ℎ₂)) ⟩ + transport (λ X → X) (ua 𝑠ℎ₁) α + ≡⟨ uaβ 𝑠ℎ₁ ⟩ + fwd 𝑠ℎ₁ α + ∎ + + tr-Cover-𝑙₂ : (α : F₂) → transport Cover 𝑙₂ α ≡ fwd 𝑠ℎ₂ α + tr-Cover-𝑙₂ α = + transport Cover 𝑙₂ α + ≡⟨ transport-ap-assoc Cover 𝑙₂ ⟩ + transport (λ X → X) (ap Cover 𝑙₂) α + ≡⟨ ap (λ ϕ → transport (λ X → X) ϕ α) + (𝑓𝑖𝑔₈-rec-𝑙₂ F₂ (ua 𝑠ℎ₁) (ua 𝑠ℎ₂)) ⟩ + transport (λ X → X) (ua 𝑠ℎ₂) α + ≡⟨ uaβ 𝑠ℎ₂ ⟩ + fwd 𝑠ℎ₂ α + ∎ + + loopify-𝑙₁ : (α : F₂) → + loopify (transport Cover 𝑙₁ α) ≡ loopify α ∙ 𝑙₁ + loopify-𝑙₁ α = + loopify (transport Cover 𝑙₁ α) + ≡⟨ ap loopify (tr-Cover-𝑙₁ α) ⟩ + loopify (fwd 𝑠ℎ₁ α) + ≡⟨ F₂-rec-𝑠ℎ₁ (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) α ⟩ + loopify α ∙ 𝑙₁ + ∎ + + loopify-𝑙₂ : (α : F₂) → + loopify (transport Cover 𝑙₂ α) ≡ loopify α ∙ 𝑙₂ + loopify-𝑙₂ α = + loopify (transport Cover 𝑙₂ α) + ≡⟨ ap loopify (tr-Cover-𝑙₂ α) ⟩ + loopify (fwd 𝑠ℎ₂ α) + ≡⟨ F₂-rec-𝑠ℎ₂ (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) α ⟩ + loopify α ∙ 𝑙₂ + ∎ + + decode : {x : 𝑓𝑖𝑔₈} → Cover x → 𝑝𝑡 ≡ x + decode {x} = + 𝑓𝑖𝑔₈-elim (λ α → Cover α → 𝑝𝑡 ≡ α) + loopify + (PathOver-→ (λ α → PathOver-path-to (! (loopify-𝑙₁ α)))) + (PathOver-→ (λ α → PathOver-path-to (! (loopify-𝑙₂ α)))) + x + + encode-decode : {x : 𝑓𝑖𝑔₈} (p : 𝑝𝑡 ≡ x) → decode (encode p) ≡ p + encode-decode (refl .𝑝𝑡) = + F₂-rec-𝑒 (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) + + 𝑠ℎ₁-lem : (α : F₂) → + encode (loopify (fwd 𝑠ℎ₁ α)) ≡ fwd 𝑠ℎ₁ (encode (loopify α)) + 𝑠ℎ₁-lem α = + encode (loopify (fwd 𝑠ℎ₁ α)) + ≡⟨ ap encode + (F₂-rec-𝑠ℎ₁ (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) α) ⟩ + encode (loopify α ∙ 𝑙₁) + ≡⟨ transport-∙ (loopify α) 𝑙₁ 𝑒 ⟩ + transport Cover 𝑙₁ (transport Cover (loopify α) 𝑒) + ≡⟨ tr-Cover-𝑙₁ (transport Cover (loopify α) 𝑒) ⟩ + fwd 𝑠ℎ₁ (encode (loopify α)) + ∎ + + 𝑠ℎ₂-lem : (α : F₂) → + encode (loopify (fwd 𝑠ℎ₂ α)) ≡ fwd 𝑠ℎ₂ (encode (loopify α)) + 𝑠ℎ₂-lem α = + encode (loopify (fwd 𝑠ℎ₂ α)) + ≡⟨ ap encode + (F₂-rec-𝑠ℎ₂ (refl 𝑝𝑡) (concat-equiv 𝑙₁) (concat-equiv 𝑙₂) α) ⟩ + encode (loopify α ∙ 𝑙₂) + ≡⟨ transport-∙ (loopify α) 𝑙₂ 𝑒 ⟩ + transport Cover 𝑙₂ (transport Cover (loopify α) 𝑒) + ≡⟨ tr-Cover-𝑙₂ (transport Cover (loopify α) 𝑒) ⟩ + fwd 𝑠ℎ₂ (encode (loopify α)) + ∎ + + encode-loopify : (α : F₂) → encode (loopify α) ≡ α + encode-loopify α = + F₂-rec-unique (encode ∘ loopify) 𝑒 𝑠ℎ₁ 𝑠ℎ₂ + (ap encode (encode-decode (refl 𝑝𝑡))) 𝑠ℎ₁-lem 𝑠ℎ₂-lem α + ∙ ! (F₂-rec-unique (λ β → β) 𝑒 𝑠ℎ₁ 𝑠ℎ₂ + (refl 𝑒) (λ _ → refl _) (λ _ → refl _) α) + + decode-encode : {x : 𝑓𝑖𝑔₈} (p : Cover x) → encode (decode {x} p) ≡ p + decode-encode {x} = + 𝑓𝑖𝑔₈-elim (λ x → (p : Cover x) → encode (decode {x} p) ≡ p) + encode-loopify + (PathOver-Π (λ {y} {z} q → + fwd (transport-to-pathover _ _ _ _) (hSetF _ _ _ _))) + (PathOver-Π (λ {y} {z} q → + fwd (transport-to-pathover _ _ _ _) (hSetF _ _ _ _))) + x + diff --git a/src/HoTTEST/Agda/HITs/Solutions6-wedge-of-circles.lagda.md b/src/HoTTEST/Agda/HITs/Solutions6-wedge-of-circles.lagda.md new file mode 100644 index 0000000..dd94074 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions6-wedge-of-circles.lagda.md @@ -0,0 +1,267 @@ +```agda + +{-# OPTIONS --rewriting --without-K --allow-unsolved-metas #-} + +open import new-prelude + +open import Lecture6-notes +open import Lecture5-notes hiding (S1; base; loop) + +module Solutions6-wedge-of-circles where +``` + +We need finite types + +```agda +data Fin : ℕ → Type where + zero : {n : ℕ} → Fin (suc n) + suc : {n : ℕ} → Fin n → Fin (suc n) +``` + +We postulate the HIT representing the wedge of k circles + +```agda +postulate + Wedge-S1 : ℕ → Type + +module _ {n : ℕ} where + + postulate + base : Wedge-S1 n + loop : Fin n → base ≡ base + Wedge-S1-elim : {l : Level} + → (A : Wedge-S1 n → Type l) + → (a : A base) + → ((k : Fin n) → a ≡ a [ A ↓ loop k ]) + → (x : Wedge-S1 n) → A x + Wedge-S1-elim-base : {l : Level} + → {A : Wedge-S1 n → Type l} + → {a : A base} + → {p : (k : Fin n) → a ≡ a [ A ↓ loop k ]} + → Wedge-S1-elim A a p base ≡ a + {-# REWRITE Wedge-S1-elim-base #-} + + postulate + Wedge-S1-elim-loop : {l : Level} + → {A : Wedge-S1 n → Type l} + → {a : A base} + → {p : (k : Fin n) → a ≡ a [ A ↓ loop k ]} + → (k : Fin n) → apd (Wedge-S1-elim A a p) (loop k) ≡ p k +``` + +We prove the recursion principle + +```agda +PathOver-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 + → b1 ≡ b2 [ (\ _ → B) ↓ p ] +PathOver-constant (refl _) (refl _) = reflo + +PathOver-constant-inverse : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 [ (\ _ → B) ↓ p ] + → b1 ≡ b2 +PathOver-constant-inverse (refl _) reflo = refl _ + +PathOver-constant-inverse-cancel1 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2) + → PathOver-constant-inverse p (PathOver-constant p q) ≡ q +PathOver-constant-inverse-cancel1 (refl _) (refl _) = refl _ + +PathOver-constant-inverse-cancel2 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2 [ _ ↓ p ]) + → PathOver-constant p (PathOver-constant-inverse p q) ≡ q +PathOver-constant-inverse-cancel2 (refl _) reflo = refl _ + +PathOver-constant-equiv : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (b1 ≡ b2) ≃ (b1 ≡ b2 [ (\ _ → B) ↓ p ]) +PathOver-constant-equiv p = improve (Isomorphism (PathOver-constant p) + (Inverse (PathOver-constant-inverse p) + (PathOver-constant-inverse-cancel1 p) + (PathOver-constant-inverse-cancel2 p))) + +ap-apd-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → (f : A → B) + → ap f p ≡ PathOver-constant-inverse p (apd f p) +ap-apd-constant (refl _) f = refl _ +``` + +```agda +module _ {n : ℕ} where + + Wedge-S1-rec : {l : Level} + → {A : Type l} + → (a : A) + → ((k : Fin n) → a ≡ a) + → Wedge-S1 n → A + Wedge-S1-rec {A = A} a p = + Wedge-S1-elim + (λ _ → A) + (a) + (λ k → PathOver-constant (loop k) (p k)) + + Wedge-S1-rec-loop : {l : Level} + → {A : Type l} + → {a : A} + → {p : (k : Fin n) → a ≡ a} + → (k : Fin n) → ap (Wedge-S1-rec a p) (loop k) ≡ p k + Wedge-S1-rec-loop {A = A} {a} {p} k = + ap (Wedge-S1-rec a p) (loop k) ≡⟨ ap-apd-constant (loop k) (Wedge-S1-rec a p) ⟩ + PathOver-constant-inverse (loop k) (apd (Wedge-S1-rec a p) (loop k)) ≡⟨ ap (PathOver-constant-inverse (loop k)) + (Wedge-S1-elim-loop k) ⟩ + PathOver-constant-inverse (loop k) (PathOver-constant (loop k) (p k)) ≡⟨ PathOver-constant-inverse-cancel1 (loop k) (p k) ⟩ + p k ∎ +``` + +We will now calculate the loop space of the wedge of circles + +```agda +concat-equiv : ∀ {A : Type} (a : A) {a' a'' : A} + → (p : a' ≡ a'') + → (a ≡ a') ≃ (a ≡ a'') +concat-equiv a p = + improve + (Isomorphism (_∙ p) + (Inverse + (_∙ ! p) + (λ q → ! (∙assoc q p (! p)) ∙ ap (q ∙_) (!-inv-r p)) + (λ q → ! (∙assoc q (! p) p) ∙ ap (q ∙_) (!-inv-l p)))) + +transport-∙ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + → {a1 a2 a3 : A} (p : a1 ≡ a2) (q : a2 ≡ a3) + → transport B (p ∙ q) ∼ transport B q ∘ transport B p +transport-∙ (refl _) (refl _) _ = refl _ +``` + +We assume we have the free group on k generators + +```agda +module AssumeFreeGroup (n : ℕ) + (FreeGroup : Type) + (one : FreeGroup) + (mult : (k : Fin n) → FreeGroup ≃ FreeGroup) + (FreeGroup-rec : {A : Type} + → (a : A) + → ((k : Fin n) → A ≃ A) + → FreeGroup → A) + (FreeGroup-rec-one : {A : Type} + → {a : A} + → {p : (k : Fin n) → A ≃ A} + → FreeGroup-rec a p one ≡ a) + (FreeGroup-rec-mult : {A : Type} + → {a : A} + → {p : (k : Fin n) → A ≃ A} + → (k : Fin n) → FreeGroup-rec a p ∘ fwd (mult k) ∼ fwd (p k) ∘ FreeGroup-rec a p) + (FreeGroup-rec-unique : {A : Type} + → {a : A} + → {p : (k : Fin n) → A ≃ A} + → (f : FreeGroup → A) + → f one ≡ a + → ((k : Fin n) → f ∘ fwd (mult k) ∼ fwd (p k) ∘ f) + → f ∼ FreeGroup-rec a p) + (FreeGroup-is-set : is-set FreeGroup) where + + Cover : Wedge-S1 n → Type + Cover = Wedge-S1-rec FreeGroup (λ k → ua (mult k)) + + encode : (x : Wedge-S1 n) → base ≡ x → Cover x + encode x p = transport Cover p one + + decode-base : FreeGroup → base ≡ base [ Wedge-S1 n ] + decode-base = FreeGroup-rec (refl base) (λ k → concat-equiv base (loop k)) + + transport-Cover-loop : (k : Fin n) (x : FreeGroup) → transport Cover (loop k) x ≡ fwd (mult k) x + transport-Cover-loop k x = + transport Cover (loop k) x ≡⟨ transport-ap-assoc Cover (loop k) ⟩ + transport (λ X → X) (ap Cover (loop k)) x ≡⟨ ap (λ P → transport (λ X → X) P x) (Wedge-S1-rec-loop k) ⟩ + transport (λ X → X) (ua (mult k)) x ≡⟨ uaβ (mult k) ⟩ + fwd (mult k) x ∎ + + decode-loop : (k : Fin n) → decode-base ≡ decode-base [ (λ x → Cover x → base ≡ x) ↓ loop k ] + decode-loop k = PathOver-→ (λ x → PathOver-path-to (decode-loop-lma x)) where + + decode-loop-lma : (x : FreeGroup) → decode-base x ∙ loop k ≡ decode-base (transport (λ z → Cover z) (loop k) x) + decode-loop-lma x = + decode-base x ∙ loop k ≡⟨ refl _ ⟩ + fwd (concat-equiv base (loop k)) (decode-base x) ≡⟨ ! (FreeGroup-rec-mult k x) ⟩ + decode-base (fwd (mult k) x) ≡⟨ ap decode-base (! (transport-Cover-loop k x)) ⟩ + decode-base (transport Cover (loop k) x) ∎ + + decode : (x : Wedge-S1 n) → Cover x → base ≡ x + decode = + Wedge-S1-elim + (λ x → Cover x → base ≡ x) + decode-base + decode-loop + + encode-decode : (x : Wedge-S1 n) (p : base ≡ x) → decode x (encode x p) ≡ p + encode-decode .base (refl base) = + decode base (encode base (refl base)) ≡⟨ refl _ ⟩ + decode base one ≡⟨ FreeGroup-rec-one ⟩ + refl base ∎ + + endo-FreeGroup-is-id : (f : FreeGroup → FreeGroup) + → f one ≡ one + → ((k : Fin n) → f ∘ fwd (mult k) ∼ fwd (mult k) ∘ f) + → f ∼ id + endo-FreeGroup-is-id f fone=one f∘mult∼mult∘f x = + f x ≡⟨ FreeGroup-rec-unique f fone=one f∘mult∼mult∘f x ⟩ + FreeGroup-rec one mult x ≡⟨ ! (FreeGroup-rec-unique id (refl one) (λ k x' → refl _) x) ⟩ + x ∎ + + decode-encode-base : (p : FreeGroup) → encode base (decode base p) ≡ p + decode-encode-base = + endo-FreeGroup-is-id + (encode base ∘ decode base) + decode-encode-base-one + decode-encode-base-mult where + + decode-encode-base-one : encode base (decode base one) ≡ one + decode-encode-base-one = + encode base (decode base one) ≡⟨ ap (encode base) FreeGroup-rec-one ⟩ + encode base (refl base) ≡⟨ refl _ ⟩ + one ∎ + + decode-encode-base-mult : (k : Fin n) (x : FreeGroup) + → encode base (decode base (fwd (mult k) x)) + ≡ fwd (mult k) (encode base (decode base x)) + decode-encode-base-mult k x = + encode base (decode base (fwd (mult k) x)) ≡⟨ ap (encode base) (FreeGroup-rec-mult k x) ⟩ + encode base (fwd (concat-equiv base (loop k)) (decode base x)) ≡⟨ refl _ ⟩ + encode base (decode base x ∙ loop k) ≡⟨ refl _ ⟩ + transport Cover (decode base x ∙ loop k) one ≡⟨ transport-∙ (decode base x) (loop k) one ⟩ + transport Cover (loop k) (transport Cover (decode base x) one) ≡⟨ refl _ ⟩ + transport Cover (loop k) (encode base (decode base x)) ≡⟨ transport-Cover-loop k (encode base (decode base x)) ⟩ + fwd (mult k) (encode base (decode base x)) ∎ + + decode-encode-loop : (k : Fin n) + → decode-encode-base ≡ decode-encode-base [ (λ x → (p : Cover x) → encode x (decode x p) ≡ p) ↓ loop k ] + decode-encode-loop k = PathOver-Π (λ q → fwd (transport-to-pathover _ _ _ _) (FreeGroup-is-set _ _ _ _)) + + decode-encode : (x : Wedge-S1 n) (p : Cover x) → encode x (decode x p) ≡ p + decode-encode = + Wedge-S1-elim + (λ x → (p : Cover x) → encode x (decode x p) ≡ p) + decode-encode-base + decode-encode-loop + + ΩWedge-S1≃FreeGroup : (base ≡ base [ Wedge-S1 n ]) ≃ FreeGroup + ΩWedge-S1≃FreeGroup = improve (Isomorphism (encode base) (Inverse (decode base) (encode-decode base) (decode-encode base))) +``` + diff --git a/src/HoTTEST/Agda/HITs/Solutions6.lagda.md b/src/HoTTEST/Agda/HITs/Solutions6.lagda.md new file mode 100644 index 0000000..5138d4c --- /dev/null +++ b/src/HoTTEST/Agda/HITs/Solutions6.lagda.md @@ -0,0 +1,336 @@ + +```agda + +{-# OPTIONS --rewriting --without-K #-} + +open import new-prelude + +open import Lecture6-notes +open import Lecture5-notes + +module Solutions6 where +``` + +In this problem set, you will look at a variation on the circle, a +higher inductive type for a "bowtie", i.e. two loops at a point. +(Unscaffolded harder exercise: do these problems for a "wedge of k +circles" for any natural number k.) + +# HIT recursion from induction + +In general, the dependent elimination rule for a higher inductive type +implies the simple/non-dependent elimination rule. In this problem, you +will show this for the bowtie. We could have done this for the circles +in the past lectures, but I wanted to introduce the non-dependent +elimination rule first, and then left both as postulates. + +Note that this problem has a bit of a "metamathematical" flavor (showing +that a set of axioms is implied by a shorter set). If you prefer to +jump right to the more "mathematical" problem of characterizing the loop +space of the bowtie below, I recommend turning Bowtie-rec and its +associated reductions into postulates like we have done for previous +higher inductive types, and adding a rewrite for the reduction on the +base point. This will make Agda display things in a more easy to read +way (otherwise, it will display Bowtie-rec as a meta-variable). + +Here is the definition of the bowtie and its dependent elimination rule: + +```agda +postulate + Bowtie : Set + baseB : Bowtie + loop1 : baseB ≡ baseB + loop2 : baseB ≡ baseB + Bowtie-elim : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → (x : Bowtie) → X x + Bowtie-elim-base : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → Bowtie-elim X x p q baseB ≡ x +{-# REWRITE Bowtie-elim-base #-} + +postulate + Bowtie-elim-loop1 : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → apd (Bowtie-elim X x p q) loop1 ≡ p + Bowtie-elim-loop2 : {l : Level} (X : Bowtie → Type l) + (x : X baseB) + (p : x ≡ x [ X ↓ loop1 ]) + (q : x ≡ x [ X ↓ loop2 ]) + → apd (Bowtie-elim X x p q) loop2 ≡ q +``` + +Next, we will prove the non-dependent elim/"recursion principle" from +these. First, we need some lemmas. + +(⋆) Paths over a path in a constant fibration are equivalent to paths. +It is simple to prove this by + +```agda +PathOver-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 + → b1 ≡ b2 [ (\ _ → B) ↓ p ] +PathOver-constant (refl _) (refl _) = reflo + +PathOver-constant-inverse : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → b1 ≡ b2 [ (\ _ → B) ↓ p ] + → b1 ≡ b2 +PathOver-constant-inverse .(refl _) reflo = refl _ + +PathOver-constant-inverse-cancel1 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2) + → PathOver-constant-inverse p (PathOver-constant p q) ≡ q +PathOver-constant-inverse-cancel1 (refl _) (refl _) = refl _ + +PathOver-constant-inverse-cancel2 : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (q : b1 ≡ b2 [ _ ↓ p ]) + → PathOver-constant p (PathOver-constant-inverse p q) ≡ q +PathOver-constant-inverse-cancel2 (refl _) (reflo) = refl _ + +PathOver-constant-equiv : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → {b1 b2 : B} + → (b1 ≡ b2) ≃ (b1 ≡ b2 [ (\ _ → B) ↓ p ]) +PathOver-constant-equiv p = improve (Isomorphism (PathOver-constant p) + (Inverse (PathOver-constant-inverse p) + (PathOver-constant-inverse-cancel1 p) + (PathOver-constant-inverse-cancel2 p))) + +``` + +(⋆) Next, for a non-dependent function f, there is an annoying mismatch between +ap f and apd f, which we can reconcile as follows: + +```agda +ap-apd-constant : {l1 l2 : Level} {A : Type l1} {B : Type l2} + → {a1 a2 : A} + → (p : a1 ≡ a2) + → (f : A → B) + → ap f p ≡ PathOver-constant-inverse _ (apd f p) +ap-apd-constant (refl _) f = refl _ +``` + +(⋆) Define Bowtie-rec and prove the reduction for base: + +```agda +Bowtie-rec : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → (Bowtie) → X +Bowtie-rec {_} {X} x p q = Bowtie-elim (\ _ → X) x (PathOver-constant _ p) (PathOver-constant _ q) + +Bowtie-rec-base : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → Bowtie-rec x p q baseB ≡ x +Bowtie-rec-base _ _ _ = refl _ +``` + +(⋆⋆) Prove the reductions for loop: + +```agda +Bowtie-rec-loop1 : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → ap (Bowtie-rec x p q) loop1 ≡ p +Bowtie-rec-loop1 x p q = ap-apd-constant _ _ ∙ + ap (PathOver-constant-inverse loop1) (Bowtie-elim-loop1 _ x (PathOver-constant _ p) (PathOver-constant _ q)) ∙ + PathOver-constant-inverse-cancel1 _ _ + +Bowtie-rec-loop2 : {l : Level} {X : Type l} + (x : X) + (p : x ≡ x [ X ]) + (q : x ≡ x [ X ]) + → ap (Bowtie-rec x p q) loop2 ≡ q +Bowtie-rec-loop2 x p q = ap-apd-constant _ _ ∙ + ap (PathOver-constant-inverse loop2) (Bowtie-elim-loop2 _ x (PathOver-constant _ p) (PathOver-constant _ q)) ∙ + PathOver-constant-inverse-cancel1 _ _ +``` + +# Loop space of the bowtie + +In this problem, you will show that the loop space of the bowtie is the +"free group on two generators", which we will write in Agda as F2. The +point of this problem is mainly for you to read and really understand +the proof that the loop space of the circle is ℤ. All of the code is +essentially a rearrangement of code from that proof. I'd suggest +trying the proof yourself, and looking at the analogous bits of the +Circle proof if you get stuck. + +## Some lemmas (⋆⋆) + +In the Circle proof in lecture, I inlined a couple of things that +can be proved more generally. You might want to prove these general +versions in advance and use them in your proof, or, if that seems +confusing, you might first do the proof without these lemmas +to motivate them. + +```agda +concat-equiv : ∀ {A : Type} (a : A) {a' a'' : A} + → (p : a' ≡ a'') + → (a ≡ a') ≃ (a ≡ a'') +concat-equiv _ (refl _) = id≃ + +concat-equiv-map : ∀ {A : Type} {a a' a'' : A} + → (p : a' ≡ a'') + → fwd (concat-equiv a p) ≡ \ q → q ∙ p +concat-equiv-map (refl _) = refl _ +``` + +(Note: you could also write out all of the components, but this was easier.) + +```agda +transport-∙ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} + {a1 a2 a3 : A} (p : a1 ≡ a2) (q : a2 ≡ a3) + → transport B (p ∙ q) ∼ transport B q ∘ transport B p +transport-∙ (refl _) (refl _) _ = refl _ +``` +## Calculating the loop space + +First, we will assume a type F2 representing the free group on 2 +generators. + +ℤ is the free group on one generator, with 0 as the neutral element and +succℤ corresponding to "addition" with the one generator. succℤ is an +equivalence, with the inverse representing "addition" with -1. + +For other groups, it is somewhat more common to think of the group +operation as "multiplication" rather than "addition", so we will name +the neutral element as "1" and the action of the elements as +"multiplication". Thus, we assume a type with an element 1F, and two +equivalences, which we think of as "multiplication with generator 1" and +"multiplication with generator 2". + +Unscaffolded hard exercise: You can implement F2 as lists whose +elements are a four-element type g1, g2, g1⁻¹, g2⁻¹ with no adjacent +pairs of inverse elements. Then the forward directions of mult1/mult2 +will be implement by cons'ing g1/g2 on and "reducing" if that creates +two adjacent inverses, the backwards directions by consing g1⁻¹ and g2⁻¹ +on and reducing, and the inverse laws will hold because the reduction +cancels the inverses. + +For this problem, we will simply assume the nice universal property for +this type: that it maps uniquely into any other type with a point and +two equivalences, and that it is a set. + +```agda +module AssumeF2 + (F2 : Type) + (1F : F2) + (mult1 : F2 ≃ F2) + (mult2 : F2 ≃ F2) + (F2-rec : {X : Type} + (o : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → F2 → X) + (F2-rec-1 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → F2-rec z m1 m2 1F ≡ z) + (F2-rec-mult1 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + (a : F2) → F2-rec z m1 m2 (fwd mult1 a) ≡ fwd m1 (F2-rec z m1 m2 a)) + (F2-rec-mult2 : {X : Type} + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + (a : F2) → F2-rec z m1 m2 (fwd mult2 a) ≡ fwd m2 (F2-rec z m1 m2 a)) + (F2-rec-unique : {X : Type} + (f : F2 → X) + (z : X) + (m1 : X ≃ X) + (m2 : X ≃ X) + → f 1F ≡ z + → ((f ∘ fwd mult1) ∼ (fwd m1 ∘ f)) + → ((f ∘ fwd mult2) ∼ (fwd m2 ∘ f)) + → (x : F2) → f x ≡ F2-rec z m1 m2 x) + (hSetF : is-set F2) where +``` + +(⋆⋆⋆) Prove that the loop space of the Bowtie is F2. Each bit of the +proof will be analogous to the corresponding part of the Circle proof. + +```agda + Cover : Bowtie → Type + Cover = Bowtie-rec F2 + (ua mult1) + (ua mult2) + + encode : (x : Bowtie) → baseB ≡ x → Cover x + encode x p = transport Cover p 1F + + transport-Cover-loop1 : (x : F2) → transport Cover loop1 (x) ≡ (fwd mult1 x) + transport-Cover-loop1 x = transport-ap-assoc _ loop1 ∙ + (ap (\ H → transport (\ X → X) H x) (Bowtie-rec-loop1 _ _ _) ∙ + uaβ _) + + transport-Cover-loop2 : (x : F2) → transport Cover loop2 (x) ≡ (fwd mult2 x) + transport-Cover-loop2 x = transport-ap-assoc _ loop2 ∙ + (ap (\ H → transport (\ X → X) H x) (Bowtie-rec-loop2 _ _ _) ∙ + uaβ _) + + decode-F2 : F2 → baseB ≡ baseB [ Bowtie ] + decode-F2 = F2-rec (refl baseB) + (concat-equiv _ loop1) + (concat-equiv _ loop2) + + decode-F2-mult1 : (x : F2) → decode-F2 (fwd mult1 x) ≡ decode-F2 x ∙ loop1 + decode-F2-mult1 x = F2-rec-mult1 _ _ _ _ ∙ app≡ (concat-equiv-map loop1) _ + + decode-F2-mult2 : (x : F2) → decode-F2 (fwd mult2 x) ≡ decode-F2 x ∙ loop2 + decode-F2-mult2 x = F2-rec-mult2 _ _ _ _ ∙ app≡ (concat-equiv-map loop2) _ + + decode : (x : Bowtie) → Cover x → baseB ≡ x + decode = Bowtie-elim _ + decode-F2 + (PathOver-→ \ a → PathOver-path-to (! (decode-F2-mult1 a) ∙ ap decode-F2 (! (transport-Cover-loop1 _)))) + (PathOver-→ \ a → PathOver-path-to (! (decode-F2-mult2 a) ∙ ap decode-F2 (! (transport-Cover-loop2 _)))) + + encode-decode : (x : Bowtie) (p : baseB ≡ x) → decode x (encode x p) ≡ p + encode-decode .baseB (refl baseB) = F2-rec-1 _ _ _ + + endo-F2-is-id : (f : F2 → F2) + → f 1F ≡ 1F + → (f ∘ fwd mult1) ∼ (fwd mult1 ∘ f) + → (f ∘ fwd mult2) ∼ (fwd mult2 ∘ f) + → f ∼ id + endo-F2-is-id f f1 fmult1 fmult2 x = F2-rec-unique f 1F mult1 mult2 f1 fmult1 fmult2 x ∙ + ! (F2-rec-unique (\ x → x) 1F mult1 mult2 (refl _) (\ _ → refl _) (\ _ → refl _) x) + + decode-encode : (x : Bowtie) (p : Cover x) → encode x (decode x p) ≡ p + decode-encode = Bowtie-elim _ + (endo-F2-is-id (λ z → encode baseB (decode-F2 z)) + (ap (\ H → encode baseB H) (F2-rec-1 _ _ _)) + (\ x → ap (encode baseB) (decode-F2-mult1 x) ∙ (transport-∙ _ loop1 _ ∙ (transport-Cover-loop1 _)) ) + (\ x → ap (encode baseB) (decode-F2-mult2 x) ∙ (transport-∙ _ loop2 _ ∙ (transport-Cover-loop2 _)) )) + (PathOver-Π \ a → fwd (transport-to-pathover _ _ _ _) (hSetF _ _ _ _)) + (PathOver-Π \ a → fwd (transport-to-pathover _ _ _ _) (hSetF _ _ _ _)) + +``` + diff --git a/src/HoTTEST/Agda/HITs/hits-lectures.agda-lib b/src/HoTTEST/Agda/HITs/hits-lectures.agda-lib new file mode 100644 index 0000000..b58a822 --- /dev/null +++ b/src/HoTTEST/Agda/HITs/hits-lectures.agda-lib @@ -0,0 +1,5 @@ +name: hits-lectures +include: . +flags: + --cubical-compatible + -WnoUnsupportedIndexedMatch diff --git a/src/HoTTEST/Agda/HITs/new-prelude.lagda.md b/src/HoTTEST/Agda/HITs/new-prelude.lagda.md new file mode 100644 index 0000000..6e39eaa --- /dev/null +++ b/src/HoTTEST/Agda/HITs/new-prelude.lagda.md @@ -0,0 +1,149 @@ +This file contains selected definitions from the Lecture 1, 2, 3 code +(see ../Lecture-Notes) that we will use in Lectures 4, 5, 6. Some +definitions have been made "universe polymorphic" (see Lecture 3) so +that we can use them for more than one universe, because we will need +this soon. + + +```agda + +{-# OPTIONS --without-K #-} + +module new-prelude where + +open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) renaming (Set to Type) public + +data _≡_ {l : Level} {A : Type l} : A → A → Type l where + refl : (x : A) → x ≡ x + +Path : {l : Level} (A : Type l) → A → A → Type l +Path A x y = x ≡ y + +syntax Path A x y = x ≡ y [ A ] + +infix 0 _≡_ + +_∙_ : {l : Level} {A : Type l} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +p ∙ (refl y) = p + +infixl 7 _∙_ + +-- path inverses/symmetry was called (-)⁻¹ in previous lectures, but I prefer a prefix +-- rather than post-fix notation +! : {l : Level} {A : Type l} {x y : A} → x ≡ y → y ≡ x +! (refl x) = refl x + +ap : {l1 l2 : Level} {A : Type l1} {B : Type l2} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f (refl x) = refl (f x) + +ap₂ : {l1 l2 l3 : Level} {A : Type l1} {B : Type l2} {C : Type l3} (f : A → B → C) {x x' : A} {y y' : B} + → x ≡ x' → y ≡ y' → f x y ≡ f x' y' +ap₂ f (refl x) (refl y) = refl (f x y) + +transport : {l1 l2 : Level} {X : Type l1} (A : X → Type l2) + → {x y : X} → x ≡ y → A x → A y +transport A (refl x) a = a + + +_∼_ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} → ((x : A) → B x) → ((x : A) → B x) → Type (l1 ⊔ l2) +f ∼ g = ∀ x → f x ≡ g x + +infix 0 _∼_ + +_≡⟨_⟩_ : {l : Level} {X : Type l} (x : X) {y z : X} → x ≡ y → y ≡ z → x ≡ z +x ≡⟨ p ⟩ q = p ∙ q + +_∎ : {l : Level} {X : Type l} (x : X) → x ≡ x +x ∎ = refl x + +infixr 0 _≡⟨_⟩_ +infix 1 _∎ + +record Σ {l1 l2 : Level} {A : Type l1 } (B : A → Type l2) : Type (l1 ⊔ l2) where + constructor + _,_ + field + pr₁ : A + pr₂ : B pr₁ + +open Σ public +infixr 0 _,_ + +Sigma : {l1 l2 : Level} (A : Type l1) (B : A → Type l2) → Type (l1 ⊔ l2) +Sigma A B = Σ B + +syntax Sigma A (λ x → b) = Σ x ꞉ A , b + +infix -1 Sigma + +id : {l : Level} {A : Type l} → A → A +id x = x + +_∘_ : {l1 l2 l3 : Level} {A : Type l1} {B : Type l2} {C : B → Type l3} + → ((y : B) → C y) + → (f : A → B) + → (x : A) → C (f x) +g ∘ f = λ x → g (f x) + +_×_ : ∀ {l1 l2} → Type l1 → Type l2 → Type (l1 ⊔ l2) +A × B = Sigma A (\ _ → B) + +infixr 2 _×_ + +-- sums-equality.to-×-≡ in ../Lecture-Notes/sums-equality +pair≡ : {l1 l2 : Level} {A : Type l1} {B : Type l2} {x x' : A} {y y' : B} + → x ≡ x' + → y ≡ y' + → _≡_{_}{A × B} (x , y) (x' , y') +pair≡ (refl _) (refl _) = refl _ + +postulate + λ≡ : {l1 l2 : Level} {A : Type l1} {B : A → Type l2} {f g : (x : A) → B x} → f ∼ g → f ≡ g + +record is-bijection {l1 l2 : Level} {A : Type l1} {B : Type l2} (f : A → B) : Type (l1 ⊔ l2) where + constructor + Inverse + field + inverse : B → A + η : inverse ∘ f ∼ id + ε : f ∘ inverse ∼ id + +record _≅_ {l1 l2 : Level} (A : Type l1) (B : Type l2) : Type (l1 ⊔ l2) where + constructor + Isomorphism + field + bijection : A → B + bijectivity : is-bijection bijection + +infix 0 _≅_ + +is-prop : {l : Level} → Type l → Type l +is-prop X = (x y : X) → x ≡ y + +is-set : {l : Level} → Type l → Type l +is-set X = (x y : X) → is-prop (x ≡ y) + +data BoolL {l : Level} : Type l where + true false : BoolL + +Bool : Type +Bool = BoolL {lzero} + +if_then_else_ : {l : Level} {A : Type l} → Bool → A → A → A +if true then x else y = x +if false then x else y = y + +record Unit {l : Level} : Type l where + constructor + ⋆ + +𝟙 : Type +𝟙 = Unit {lzero} + +data ℕ : Type where + zero : ℕ + suc : ℕ → ℕ + +the : ∀ {l} (A : Type l) → A → A +the _ x = x +``` diff --git a/src/HoTTEST/Agda/Lecture-Log/lecture1.lagda.md b/src/HoTTEST/Agda/Lecture-Log/lecture1.lagda.md new file mode 100644 index 0000000..2463089 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Log/lecture1.lagda.md @@ -0,0 +1,173 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module lecture1 where +import introduction + +Type = Set +Type₁ = Set₁ + +data Bool : Type where + true false : Bool + +-- Type "→" as "\to" or "\->" +not : Bool → Bool +not true = false +not false = true + +idBool : Bool → Bool +idBool x = x + +idBool' : Bool → Bool +idBool' = λ (x : Bool) → x + +-- The following is a Π type +id' : (X : Type) → X → X +id' X x = x + +-- Implicit +id : {X : Type} → X → X +id x = x + +idBool'' : Bool → Bool +idBool'' = id' _ + +-- "propositions as types" "mathematical statements as types" +example : {P Q : Type} → P → (Q → P) +example {P} {Q} p = f + where + f : Q → P + f _ = p + +example' : {P Q : Type} → P → (Q → P) +example' p = λ q → p + +open import binary-products + +-- "×" is "and" in propositions as types +ex : {P Q : Type} → P × Q → Q × P +ex (p , q) = (q , p) + +-- \bN +data ℕ : Type where + zero : ℕ + suc : ℕ → ℕ + +three : ℕ +three = suc (suc (suc zero)) + +-- {-# BUILTIN NATURAL ℕ #-} +-- for technical reasons we can not have this binding here. however +-- there is already a type ℕ in the introduction module which /does/ +-- support numeral notation. we are careful not to use the ℕ type +-- defined in this module again. + +three' : introduction.ℕ +three' = 3 -- synonym for the above + +D : Bool → Type +D true = introduction.ℕ +D false = Bool + +-- "mix-fix" operator (3rd sense of "_" in Agda) +-- b x y +if_then_else_ : {X : Type} → Bool → X → X → X +if true then x else y = x +if false then x else y = y + +if[_]_then_else_ : (X : Bool → Type) + → (b : Bool) + → X true + → X false + → X b +if[ X ] true then x else y = x +if[ X ] false then x else y = y + +-- Π (b : Bool), D b +unfamiliar : (b : Bool) → D b +unfamiliar b = if[ D ] b then 3 else false + +data List (A : Type) : Type where + [] : List A -- empty list + _::_ : A → List A → List A -- if xs is a list then x :: xs is list + +infixr 10 _::_ + +ff : Type → Type +ff = List + +sample-list₀ : List introduction.ℕ +sample-list₀ = 0 :: 1 :: 2 :: [] + +length : {X : Type} → List X → ℕ +length [] = zero +length (x :: xs) = suc (length xs) + +-- Principle of induction on ℕ +ℕ-elim : {A : ℕ → Type} + → A zero -- base case + → ((k : ℕ) → A k → A (suc k)) -- induction step + → (n : ℕ) → A n +ℕ-elim {A} a₀ f = h + where + h : (n : ℕ) → A n + h zero = a₀ + h (suc n) = f n IH + where + IH : A n + IH = h n + +ℕ-elim' : {A : ℕ → Type} + → A zero -- base case + → ((k : ℕ) → A k → A (suc k)) -- induction step + → (n : ℕ) → A n +ℕ-elim' {A} a₀ f zero = a₀ +ℕ-elim' {A} a₀ f (suc n) = f n (ℕ-elim' a₀ f n) + +List-elim : {X : Type} (A : List X → Type) + → A [] -- base + → ((x : X) (xs : List X) → A xs → A (x :: xs)) -- step + → (xs : List X) → A xs +List-elim {X} A a f = h + where + h : (xs : List X) → A xs + h [] = a -- base + h (x :: xs) = f x xs (h xs) -- step + +-- \b0 +data 𝟘 : Type where + +-- \b1 +data 𝟙 : Type where + ⋆ : 𝟙 -- \star + +_≣_ : ℕ → ℕ → Type +zero ≣ zero = 𝟙 +zero ≣ suc y = 𝟘 +suc x ≣ zero = 𝟘 +suc x ≣ suc y = x ≣ y + +infix 0 _≣_ + +ℕ-refl : (x : ℕ) → x ≣ x +ℕ-refl zero = ⋆ +ℕ-refl (suc x) = ℕ-refl x + +_+_ : ℕ → ℕ → ℕ +zero + y = y +suc x + y = suc (x + y) + +infixr 20 _+_ + +_++_ : {A : Type} → List A → List A → List A +[] ++ ys = ys +(x :: xs) ++ ys = x :: (xs ++ ys) + +lh : {X : Type} (xs ys : List X) + → length (xs ++ ys) ≣ length xs + length ys +lh [] ys = ℕ-refl (length ys) +lh (x :: xs) ys = IH + where + IH : length (xs ++ ys) ≣ (length xs + length ys) + IH = lh xs ys +``` diff --git a/src/HoTTEST/Agda/Lecture-Log/lecture2.lagda.md b/src/HoTTEST/Agda/Lecture-Log/lecture2.lagda.md new file mode 100644 index 0000000..63e106d --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Log/lecture2.lagda.md @@ -0,0 +1,348 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module lecture2 where + +-- lecture 2 +-- Plan: basic MLTT types, including their elimination principles. +-- + +open import lecture1 hiding (𝟘 ; 𝟙 ; D ; ℕ ; _+_) +open import introduction using (ℕ ; zero ; suc ; _+_) + +-- empty type +data 𝟘 : Type where + +-- Π x ꞉ X , A x +-- (X → B) =. Π x ꞉ X , B + +𝟘-elim : {A : 𝟘 → Type} (x : 𝟘) → A x +𝟘-elim () + +-- 𝟘 interpreted as "false" + +¬_ : Type → Type +¬ A = A → 𝟘 + +infix 1000 ¬_ + +𝟘-nondep-elim : {B : Type} → 𝟘 → B +𝟘-nondep-elim {B} = 𝟘-elim {λ _ → B} + +-- A := λ _ → B + + +is-empty : Type → Type +is-empty A = A → 𝟘 + +-- A ≡ 𝟘 or A ≃ 𝟘. + +𝟘-is-empty'' : is-empty 𝟘 +𝟘-is-empty'' = λ () + +𝟘-is-empty : is-empty 𝟘 +𝟘-is-empty = 𝟘-nondep-elim + +𝟘-is-empty' : is-empty 𝟘 +𝟘-is-empty' = id + +-- Unit type: + +-- Record definitions satisfy a certain "η" rule. + +record 𝟙 : Type where + constructor + ⋆ + +open 𝟙 public + +𝟙-is-nonempty' : ¬ is-empty 𝟙 +𝟙-is-nonempty' = λ (f : 𝟙 → 𝟘) → f ⋆ + +𝟙-is-nonempty : ¬ is-empty 𝟙 +𝟙-is-nonempty f = f ⋆ + +𝟙-elim : {A : 𝟙 → Type} + → A ⋆ + → (x : 𝟙) → A x +𝟙-elim a x = a + +𝟙-nondep-elim : {A : Type} + → A + → 𝟙 → A +𝟙-nondep-elim {A} = 𝟙-elim {λ _ → A} + +-- Type of binary digits: + +data 𝟚 : Type where + 𝟎 𝟏 : 𝟚 + +𝟚-elim : {A : 𝟚 → Type} + → A 𝟎 + → A 𝟏 + → (x : 𝟚) → A x +𝟚-elim a₀ a₁ 𝟎 = a₀ +𝟚-elim a₀ a₁ 𝟏 = a₁ + +𝟚-nondep-elim : {A : Type} + → A + → A + → 𝟚 → A +𝟚-nondep-elim {A} = 𝟚-elim {λ _ → A} + +-- Π types in Agda are primitive. +-- +-- We have that Π x ꞉ X , A x is written +-- +-- (x : X) → A x, or +-- ∀ (x : X) → A x, or +-- ∀ x → A x (if Agda can infer X). +-- +-- We can introduce Π-syntax if we wish: + +Pi : (A : Type) (B : A → Type) → Type +Pi A B = (x : A) → B x + +syntax Pi A (λ x → b) = Π x ꞉ A , b +-- ↑ +-- this is typed "\:4" in emacs mode and is not the same as ":". +-- (we can't use the normal one unfortunately.) + + +-- Function composition. + +-- The usual one found in mathematics: + +module _ where + private + _∘_ : {A B C : Type} → (B → C) → (A → B) → (A → C) + (g ∘ f) x = g (f x) + +-- A more general version: + +_∘_ : {A B : Type} {C : B → Type} + → ((y : B) → C y) + → (f : A → B) + → (x : A) → C (f x) +(g ∘ f) x = g (f x) + +-- The types-as-mathematical-statements reading of dependent function composition is: +-- +-- If (for all y : B we have that C y) and f : A → B is any function, then +-- for all x : A we have that C (f x). +-- +-- The proof is function composition. + + +-- Σ-types: + +module _ where + private + + data Σ {A : Type } (B : A → Type) : Type where + _,_ : (x : A) (y : B x) → Σ {A} B + + pr₁ : {A : Type} {B : A → Type} → Σ B → A + pr₁ (x , y) = x + + pr₂ : {A : Type} {B : A → Type} → (z : Σ B) → B (pr₁ z) + pr₂ (x , y) = y + +-- Our preferred definition: + +record Σ {A : Type } (B : A → Type) : Type where + constructor + _,_ + field + pr₁ : A + pr₂ : B pr₁ + +open Σ public +infixr 0 _,_ + +pr₁-again : {A : Type} {B : A → Type} → Σ B → A +pr₁-again = pr₁ + +pr₂-again : {A : Type} {B : A → Type} ((x , y) : Σ B) → B x +pr₂-again = pr₂ + + +-- This satisfies the η-rule z = (pr₁ z , pr₂ z), which the definition using `data` doesn't. + + +Sigma : (A : Type) (B : A → Type) → Type +Sigma A B = Σ {A} B + +syntax Sigma A (λ x → b) = Σ x ꞉ A , b + +infix -1 Sigma + +-- Recall that we defined D as follows in the first lecture: + +D : Bool → Type +D true = ℕ +D false = Bool + +-- Example + +Σ-example₁ Σ-example₂ : Σ b ꞉ Bool , D b +Σ-example₁ = (true , 17) +Σ-example₂ = (false , true) + +-- Σ-elim is "curry": + +Σ-elim : {A : Type } {B : A → Type} {C : (Σ x ꞉ A , B x) → Type} + → ((x : A) (y : B x) → C (x , y)) + → (z : Σ x ꞉ A , B x) → C z +Σ-elim f (x , y) = f x y + +Σ-uncurry : {A : Type } {B : A → Type} {C : (Σ x ꞉ A , B x) → Type} + → ((z : Σ x ꞉ A , B x) → C z) + → (x : A) (y : B x) → C (x , y) +Σ-uncurry g x y = g (x , y) + +_×_ : Type → Type → Type +A × B = Σ x ꞉ A , B + +-- (x : X) → A x +-- (x : X) × A x + +infixr 2 _×_ + +-- We will have that A₀ × A₁ ≅ Π (n : 𝟚) , A n ≅ ((n : 𝟚) → A n) +-- where A 𝟎 = A₀ +-- A 𝟏 = A₁ +-- A : 𝟚 → Type +-- f ↦ (f 𝟎 , f 𝟏) +-- (a₀ , a₁) ↦ 𝟚-elim a₀ a₁ +-- But we need function extensionality to prove that this works. +-- Binary products are special cases of products. + +``` + +Various uses of Σ: + + * + * + * + +```agda + +-- Binary sums _+_ ∔ + +data _∔_ (A B : Type) : Type where + inl : A → A ∔ B + inr : B → A ∔ B + +-- Mathematically A ∔ B is (disjoint) union. +-- Logically, it is "or" (disjunction). +-- ∥ A ∔ B ∥. + +infixr 20 _∔_ + +∔-elim : {A B : Type} (C : A ∔ B → Type) + → ((x : A) → C (inl x)) -- f + → ((y : B) → C (inr y)) -- g + → (z : A ∔ B) → C z +∔-elim C f g (inl x) = f x +∔-elim C f g (inr y) = g y + +∔-nondep-elim : {A B C : Type} + → (A → C) + → (B → C) + → (A ∔ B → C) +∔-nondep-elim {A} {B} {C} = ∔-elim (λ z → C) + +-- We will have that A₀ ∔ A₁ ≅ Σ (n : 𝟚) , A n +-- where A 𝟎 = A₀ +-- A 𝟏 = A₁ +-- inl a₀ ↦ (𝟎 , a₀) +-- inr a₁ ↦ (𝟏 , a₁) +-- Binary sums are special cases of sums. + + +-- We call an element of the identity type x ≡ y an +-- "identification". The terminology "path" is also used. +-- I prefer the former. + +data _≡_ {A : Type} : A → A → Type where + refl : (x : A) → x ≡ x + +-- refl x : proof that x is equal to itself. + +infix 0 _≡_ + +-- The following is also called "J": + +≡-elim : {X : Type} (A : (x y : X) → x ≡ y → Type) + → ((x : X) → A x x (refl x)) + → (x y : X) (p : x ≡ y) → A x y p +≡-elim A f x x (refl x) = f x + +-- To conclude that a property A x y p of identifications p of +-- elements x and y holds for all x, y and p, it is enough to show +-- that A x x (refl x) holds for all x. + +≡-nondep-elim : {X : Type} (A : X → X → Type) + → ((x : X) → A x x) + → (x y : X) → x ≡ y → A x y +≡-nondep-elim A = ≡-elim (λ x y _ → A x y) + +-- We finished lecture 2 here. So we'll start lecture 3 here. + +trans : {A : Type} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +trans p (refl y) = p + +sym : {A : Type} {x y : A} → x ≡ y → y ≡ x +sym (refl x) = refl x + +ap : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f (refl x) = refl (f x) + +ap₂ : {A B C : Type} (f : A → B → C) {x x' : A} {y y' : B} + → x ≡ x' → y ≡ y' → f x y ≡ f x' y' +ap₂ f (refl x) (refl y) = refl (f x y) + +transport : {X : Type} (A : X → Type) + → {x y : X} → x ≡ y → A x → A y +transport A (refl x) a = a + +_∙_ : {A : Type} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ = trans + +infixl 7 _∙_ + +_⁻¹ : {A : Type} {x y : A} → x ≡ y → y ≡ x +_⁻¹ = sym + +infix 40 _⁻¹ + +_≤_ : ℕ → ℕ → Type +0 ≤ y = 𝟙 +suc x ≤ 0 = 𝟘 +suc x ≤ suc y = x ≤ y + +_≥_ : ℕ → ℕ → Type +x ≥ y = y ≤ x + +_*_ : ℕ → ℕ → ℕ +0 * y = 0 +suc x * y = x * y + y + +infixr 30 _*_ + +_divides_ : ℕ → ℕ → Type +x divides y = Σ z ꞉ ℕ , x * z ≡ y + +is-prime : ℕ → Type +is-prime p = (p ≥ 2) × ((n : ℕ) → n divides p → (n ≡ 1) ∔ (n ≡ p)) + + +twin-prime-conjecture : Type +twin-prime-conjecture = (n : ℕ) → Σ p ꞉ ℕ , (p ≥ n) + × is-prime p + × is-prime (p + 2) + +there-are-infinitely-many-primes : Type +there-are-infinitely-many-primes = (n : ℕ) → Σ p ꞉ ℕ , (p ≥ n) × is-prime p +``` diff --git a/src/HoTTEST/Agda/Lecture-Log/lecture3.lagda.md b/src/HoTTEST/Agda/Lecture-Log/lecture3.lagda.md new file mode 100644 index 0000000..f1fb575 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Log/lecture3.lagda.md @@ -0,0 +1,355 @@ +```agda +{-# OPTIONS --without-K --safe #-} + +module lecture3 where + +-- lecture 3 +-- Plan: Complete last lecture. +-- Generalize some definitions to use universe levels. +-- Uses of Sigma, including examples like monoids. +-- Use of universes to prove that ¬ (false ≡ true). +-- Characterization of equality in Σ types. + + +open import lecture1 hiding (𝟘 ; 𝟙 ; ⋆ ; D ; _≣_ ; ℕ) +open import lecture2 using (is-prime ; _*_ ; 𝟘 ; 𝟙 ; ⋆ ; _≥_) +open import introduction using (ℕ ; zero ; suc ; _+_) + +-- Give Σ a universe-polymorphic type + +open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) + renaming (Set to 𝓤) + public + +variable i j k : Level + +record Σ {A : 𝓤 i} (B : A → 𝓤 j) : 𝓤 (i ⊔ j) where + constructor + _,_ + field + pr₁ : A + pr₂ : B pr₁ + +open Σ public +infixr 1 _,_ + +Sigma : (A : 𝓤 i) (B : A → 𝓤 j) → 𝓤 (i ⊔ j) +Sigma {i} {j} A B = Σ {i} {j} {A} B + +syntax Sigma A (λ x → b) = Σ x ꞉ A , b + +infix -1 Sigma + +_×_ : 𝓤 i → 𝓤 j → 𝓤 (i ⊔ j) +A × B = Σ x ꞉ A , B + +-- (x : X) → A x +-- (x : X) × A x + +infixr 2 _×_ + +-- More general type of negation: + +¬_ : 𝓤 i → 𝓤 i +¬ A = A → 𝟘 + +-- Give the identity type more general universe assignments: + +data _≡_ {X : 𝓤 i} : X → X → 𝓤 i where + refl : (x : X) → x ≡ x + +_≢_ : {X : 𝓤 i} → X → X → 𝓤 i +x ≢ y = ¬ (x ≡ y) + +infix 0 _≡_ + +≡-elim : {X : 𝓤 i} (A : (x y : X) → x ≡ y → 𝓤 j) + → ((x : X) → A x x (refl x)) + → (x y : X) (p : x ≡ y) → A x y p +≡-elim A f x x (refl x) = f x + +≡-nondep-elim : {X : 𝓤 i} (A : X → X → 𝓤 j) + → ((x : X) → A x x) + → (x y : X) → x ≡ y → A x y +≡-nondep-elim A = ≡-elim (λ x y _ → A x y) + +trans : {A : 𝓤 i} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +trans p (refl y) = p + +sym : {A : 𝓤 i} {x y : A} → x ≡ y → y ≡ x +sym (refl x) = refl x + +ap : {A : 𝓤 i} {B : 𝓤 j} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f (refl x) = refl (f x) + +ap₂ : {A : 𝓤 i} {B : 𝓤 j} {C : 𝓤 k} (f : A → B → C) {x x' : A} {y y' : B} + → x ≡ x' → y ≡ y' → f x y ≡ f x' y' +ap₂ f (refl x) (refl y) = refl (f x y) + +transport : {X : 𝓤 i} (A : X → 𝓤 j) + → {x y : X} → x ≡ y → A x → A y +transport A (refl x) a = a + +_∙_ : {A : 𝓤 i} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ = trans + +infixl 7 _∙_ + +_⁻¹ : {A : 𝓤 i} {x y : A} → x ≡ y → y ≡ x +_⁻¹ = sym + +infix 40 _⁻¹ + +-- The (sub)type of prime numbers + +ℙ : 𝓤₀ +ℙ = Σ p ꞉ ℕ , is-prime p + +ℙ-inclusion : ℙ → ℕ +ℙ-inclusion = pr₁ + +-- We can prove that this map is left-cancellable, i.e. it satisfies +-- ℙ-inclusion u ≡ ℙ-inclusion u → u ≡ v. +-- Moreover, this map is an embedding (we haven't defined this concept yet). + +-- Not quite the type of composite numbers: + +CN : 𝓤 +CN = Σ x ꞉ ℕ , Σ (y , z) ꞉ ℕ × ℕ , x ≡ y * z + +CN' : 𝓤 +CN' = Σ x ꞉ ℕ , Σ (y , z) ꞉ ℕ × ℕ , (y ≥ 2) × (z ≥ 2) × (x ≡ y * z) + +CN-projection : CN → ℕ +CN-projection = pr₁ + +-- This map is not left-cancellable, and hence can't be considered to +-- be an an inclusion. + +counter-example : CN-projection (6 , (3 , 2) , refl 6) + ≡ CN-projection (6 , (2 , 3) , refl 6) +counter-example = refl 6 + +-- But how do we prove that these two tuples are *different*? They +-- certainly do look different. We'll do this later. + +-- We will need to define +-- +-- CN = Σ x ꞉ ℕ , ∥ Σ (y , z) ꞉ ℕ × ℕ , x ≡ y * z ∥, or equivalently +-- CN = Σ x ꞉ ℕ , ∃ (y , z) ꞉ ℕ × ℕ , x ≡ y * z ∥ +-- +-- to really get a *subtype* of composite numbers. + + +-- Another use of Σ. +-- The type of monoids. + +is-prop : 𝓤 i → 𝓤 i +is-prop X = (x y : X) → x ≡ y + +is-set : 𝓤 i → 𝓤 i +is-set X = (x y : X) → is-prop (x ≡ y) + +Mon : 𝓤 (lsuc i) +Mon {i} = Σ X ꞉ 𝓤 i -- data + , is-set X -- property (we show that) + × (Σ 𝟏 ꞉ X , -- data (but...) + Σ _·_ ꞉ (X → X → X) -- data + , (((x : X) → (x · 𝟏 ≡ x)) -- (1) property + × ((x : X) → (𝟏 · x ≡ x)) -- (2) property + × ((x y z : X) → (x · (y · z)) ≡ ((x · y) · z)))) -- (3) property + +-- This can be defined using a record in Agda: + +record Mon' : 𝓤 (lsuc i) where + constructor mon + field + carrier : 𝓤 i -- X + carrier-is-set : is-set carrier + 𝟏 : carrier + _·_ : carrier → carrier → carrier + left-unit-law : (x : carrier) → x · 𝟏 ≡ x + right-unit-law : (x : carrier) → 𝟏 · x ≡ x + assoc-law : (x y z : carrier) → (x · (y · z)) ≡ ((x · y) · z) + +α : Mon {i} → Mon' {i} +α (X , X-is-set , 𝟏 , _·_ , l , r , a) = mon X X-is-set 𝟏 _·_ l r a + +β : Mon' {i} → Mon {i} +β (mon X X-is-set 𝟏 _·_ l r a) = (X , X-is-set , 𝟏 , _·_ , l , r , a) + +βα : (M : Mon {i}) → β (α M) ≡ M +βα = refl + +αβ : (M : Mon' {i}) → α (β M) ≡ M +αβ = refl + +-- This kind of proof doesn't belong to the realm of MLTT: + +false-is-not-true[not-an-MLTT-proof] : false ≢ true +false-is-not-true[not-an-MLTT-proof] () + +-- Proof in MLTT, which requires a universe (Cf. Ulrik's 2nd HoTT +-- lecture): + +_≣_ : Bool → Bool → 𝓤₀ +true ≣ true = 𝟙 +true ≣ false = 𝟘 +false ≣ true = 𝟘 +false ≣ false = 𝟙 + +≡-gives-≣ : {x y : Bool} → x ≡ y → x ≣ y +≡-gives-≣ (refl true) = ⋆ +≡-gives-≣ (refl false) = ⋆ + +false-is-not-true : ¬ (false ≡ true) +false-is-not-true p = II + where + I : false ≣ true + I = ≡-gives-≣ p + + II : 𝟘 + II = I + +false-is-not-true' : ¬ (false ≡ true) +false-is-not-true' = ≡-gives-≣ + +-- Notice that this proof is different from the one given by Ulrik in +-- the HoTT track. Exercise: implement Ulrik's proof in Agda. + +-- Exercise: prove that ¬ (0 ≡ 1) in the natural numbers in MLTT style +-- without using `()`. + +-- contrapositives. + +contrapositive : {A : 𝓤 i} {B : 𝓤 j} → (A → B) → (¬ B → ¬ A) +contrapositive f g a = g (f a) + +Π-¬-gives-¬-Σ : {X : 𝓤 i} {A : X → 𝓤 j} + → ((x : X) → ¬ A x) + → ¬ (Σ x ꞉ X , A x) +Π-¬-gives-¬-Σ ϕ (x , a) = ϕ x a + +¬-Σ-gives-Π-¬ : {X : 𝓤 i} {A : X → 𝓤 j} + → ¬ (Σ x ꞉ X , A x) + → ((x : X) → ¬ A x) +¬-Σ-gives-Π-¬ γ x a = γ (x , a) + + +-- Equality in Σ types. + +from-Σ-≡' : {X : 𝓤 i} {A : X → 𝓤 j} + {(x , a) (y , b) : Σ A} + → (x , a) ≡ (y , b) + → Σ p ꞉ (x ≡ y) , (transport A p a ≡ b) +from-Σ-≡' (refl (x , a)) = (refl x , refl a) + +to-Σ-≡' : {X : 𝓤 i} {A : X → 𝓤 j} + {(x , a) (y , b) : Σ A} + → (Σ p ꞉ (x ≡ y) , (transport A p a ≡ b)) + → (x , a) ≡ (y , b) +to-Σ-≡' (refl x , refl a) = refl (x , a) + +module _ {X : 𝓤 i} {A : 𝓤 j} + {(x , a) (y , b) : X × A} where + + from-×-≡ : (x , a) ≡ (y , b) + → (x ≡ y) × (a ≡ b) + from-×-≡ (refl (x , a)) = refl x , refl a + + + to-×-≡ : (x ≡ y) × (a ≡ b) + → (x , a) ≡ (y , b) + to-×-≡ (refl x , refl a) = refl (x , a) + +module _ {X : 𝓤 i} {A : X → 𝓤 j} + {(x , a) (y , b) : Σ A} where + + -- x y : X + -- a : A x + -- b : A y + + from-Σ-≡ : (x , a) ≡ (y , b) + → Σ p ꞉ (x ≡ y) , transport A p a ≡ b + from-Σ-≡ (refl (x , a)) = refl x , refl a + + + to-Σ-≡ : (Σ p ꞉ (x ≡ y) , (transport A p a ≡ b)) + → (x , a) ≡ (y , b) + to-Σ-≡ (refl x , refl a) = refl (x , a) + + + contra-from-Σ-≡ : ¬ (Σ p ꞉ (x ≡ y) , (transport A p a ≡ b)) + → (x , a) ≢ (y , b) + contra-from-Σ-≡ = contrapositive from-Σ-≡ + + contra-to-Σ-≡ : (x , a) ≢ (y , b) + → ¬ (Σ p ꞉ (x ≡ y) , (transport A p a ≡ b)) + contra-to-Σ-≡ = contrapositive to-Σ-≡ + + to-Σ-≢ : ((p : x ≡ y) → transport A p a ≢ b) + → (x , a) ≢ (y , b) + to-Σ-≢ u = contra-from-Σ-≡ (Π-¬-gives-¬-Σ u) + + from-Σ-≢ : (x , a) ≢ (y , b) + → ((p : x ≡ y) → transport A p a ≢ b) + from-Σ-≢ v = ¬-Σ-gives-Π-¬ (contra-to-Σ-≡ v) +``` + +We now revisit the example above. How do we prove that aa and bb are +different? It's not easy. We use the above lemmas. + +```agda +aa bb : CN +aa = (6 , (3 , 2) , refl 6) +bb = (6 , (2 , 3) , refl 6) +``` + +To prove that aa ≢ bb, we need to know that ℕ is a set! And this is +difficult. See the module +[Hedbergs-Theorem](../Lecture-Notes/files/Hedbergs-Theorem.lagda.md) for +a complete proof. + +For the moment we just assume that ℕ is a set, and prove that 3 ≢ 2 by +cheating (produce a genuine MLTT proof as an exercise). + +```agda + +3-is-not-2 : 3 ≢ 2 +3-is-not-2 () + +example-revisited : is-set ℕ → aa ≢ bb +example-revisited ℕ-is-set = I + where + A : ℕ → 𝓤₀ + A x = Σ (y , z) ꞉ ℕ × ℕ , x ≡ y * z + + II : (p : 6 ≡ 6) → transport A p ((3 , 2) , refl 6) ≢ ((2 , 3) , refl 6) + II p = VIII + where + III : p ≡ refl 6 + III = ℕ-is-set 6 6 p (refl 6) + + IV : transport A p ((3 , 2) , refl 6) ≡ ((3 , 2) , refl 6) + IV = ap (λ - → transport A - ((3 , 2) , refl 6)) III + + V : ((3 , 2) , refl 6) ≢ ((2 , 3) , refl 6) + V q = 3-is-not-2 VII + where + VI : (3 , 2) ≡ (2 , 3) + VI = ap pr₁ q + + VII : 3 ≡ 2 + VII = ap pr₁ VI + + VIII : transport A p ((3 , 2) , refl 6) ≢ ((2 , 3) , refl 6) + VIII r = V IX + where + IX : ((3 , 2) , refl 6) ≡ ((2 , 3) , refl 6) + IX = trans (IV ⁻¹) r + + I : aa ≢ bb + I = to-Σ-≢ II +``` + +If there is time, I will do some isomorphisms. diff --git a/src/HoTTEST/Agda/Lecture-Notes/README.md b/src/HoTTEST/Agda/Lecture-Notes/README.md new file mode 100644 index 0000000..388ef74 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/README.md @@ -0,0 +1,51 @@ +[Martin Escardo](https://www.cs.bham.ac.uk/~mhe). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + +## [Preferred link to see these lecture notes](https://martinescardo.github.io/HoTTEST-Summer-School/) + +## [Lecture Notes](./files/) + + 1. [Introduction](files/introduction.lagda.md) + 1. Propositions as types and basic Martin-Löf type theory + 1. [Basic Martin-Löf Type Theory](files/curry-howard.lagda.md) + 1. [Empty type](files/empty-type.lagda.md) `𝟘` + 1. [Unit type](files/unit-type.lagda.md) `𝟙` + 1. [Binary type](files/binary-type.lagda.md) `𝟚` + 1. [Product types](files/products.lagda.md) `Π` + 1. [Sum types](files/sums.lagda.md) `Σ` + 1. [Binary products](files/binary-products.lagda.md) `_×_` + 1. [Binary sums](files/binary-sums.lagda.md) `_∔_` + 1. [Identity types](files/identity-type.lagda.md) `_≡_` + 1. [Natural numbers](files/natural-numbers-type.lagda.md) `ℕ` + 1. [Negation](files/negation.lagda.md) `¬` + 1. [Function extensionality](files/function-extensionality.lagda.md) + 1. [Propositions as types versus propositions as booleans](files/decidability.lagda.md) + 1. Isomorphisms `≅` + 1. [Definition and basic examples](files/isomorphisms.lagda.md) + 1. [Binary sums as a special case of sums](files/binary-sums-as-sums.lagda.md) + 1. [Binary products as a special case of products](files/binary-products-as-products.lagda.md) + 1. More types, their elimination principles, and their isomorphism with Basic MLTT types + 1. [Booleans](files/Bool.lagda.md) and [their manifestation as a Basic MLTT type](files/Bool-functions.lagda.md) + 1. [Finite types](files/Fin.lagda.md) and [their manifestation as a Basic MLTT type](files/Fin-functions.lagda.md) + 1. [Lists](files/List.lagda.md) + 1. [Vectors](files/Vector.lagda.md) + 1. [List and vector isomorphisms](files/vector-and-list-isomorphisms.lagda.md) + 1. More constructions and proofs with the above types: + 1. [Natural numbers](files/natural-numbers-functions.lagda.md) + 1. [Booleans](files/Bool-functions.lagda.md) + 1. [Finite types](files/Fin-functions.lagda.md) + 1. [Lists](files/List-functions.lagda.md) + 1. [Vectors](files/Vector-functions.lagda.md) + 1. [Hedberg's Theorem](files/Hedbergs-Theorem.lagda.md) + +## [Index of all Agda files](files/all.lagda.md) + +## Advanced further reading + + 1. [Introduction to Homotopy Type Theory and Univalent Foundations (HoTT/UF) with Agda](https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#contents). + +## Follow the links + +Each of the links above will take you to a different Agda module. You can use this to interactively explore the lecture notes: the Agda code is all type-checked (with the exception of some exercises which are left as holes, `{! !}`), and presented as hypertext. Click on an identifier to be taken to its definition. + +## [HoTTEST School github repository](https://github.com/martinescardo/HoTTEST-Summer-School) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Bool-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Bool-functions.lagda.md new file mode 100644 index 0000000..89f39c8 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Bool-functions.lagda.md @@ -0,0 +1,45 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The booleans + +The booleans are isomorphic to a Basic MLTT type: + +```agda +open import isomorphisms + +Bool-isomorphism : Bool ≅ 𝟙 ∔ 𝟙 +Bool-isomorphism = record { bijection = f ; bijectivity = f-is-bijection } + where + f : Bool → 𝟙 ∔ 𝟙 + f false = inl ⋆ + f true = inr ⋆ + + g : 𝟙 ∔ 𝟙 → Bool + g (inl ⋆) = false + g (inr ⋆) = true + + gf : g ∘ f ∼ id + gf true = refl true + gf false = refl false + + fg : f ∘ g ∼ id + fg (inl ⋆) = refl (inl ⋆) + fg (inr ⋆) = refl (inr ⋆) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Bool.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Bool.lagda.md new file mode 100644 index 0000000..0a10e39 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Bool.lagda.md @@ -0,0 +1,180 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The booleans + +We discuss the elimination principle for the booleans. The booleans +are defined by constructors `true` and `false`: +```agda +data Bool : Type where + true false : Bool +``` +## `if-then-else` + +The non-dependent eliminator of the type of booleans amounts to `if-then-else` +```agda +if_then_else_ : {A : Type} → Bool → A → A → A +if true then x else y = x +if false then x else y = y +``` +In general, the non-dependent elimination principle of a type explains how to "get out of the type", whereas the constructors tell us how to "get into the type". + +## Dependent `if-then-else` + +Notice that both `x` (the `then` branch) and `y` (the `else` branch) have the same type, name `A`. Using dependent type, we can have different types in the dependent version of the eliminator. We make the type `A` depend on the boolean condition of the `if-then-else`. This means that now we will have `A : Bool → Type` instead of `A : Bool`. This is a function that given a boolean `b : Bool`, returns a type `A b`. Functions whose return value is a type are also called *type families*. Also `A b` is called a *dependent type*. It depends on the value of the boolean `b`. Here is an example, which we make private to this module. +```agda +private + open import natural-numbers-type + A-example : Bool → Type + A-example true = ℕ + A-example false = Bool +``` +Using this idea, we have the following dependently-typed version of `if_then_else_`, which now has four explicit arguments. We make `A` explicit this time, because Agda can hardly ever infer it. +```agda +dependent-on_if_then_else_ : (A : Bool → Type) → (b : Bool) → A true → A false → A b +dependent-on A if true then x else y = x +dependent-on A if false then x else y = y +``` +Notice that the return type `A b` depends on the second argument `b` of the function. +Notice also that `x : A true` and `y : A false`. +```agda +private + example₀ : (b : Bool) → A-example b + example₀ b = dependent-on A-example if b then 3 else true +``` +This works because `3 : A-example true` and `true : A-example false`. So the dependent version of `if-then-else` allows the `then` and `else` branches have different types, which depend on the type of the condition. + +## The official definition of the eliminator of the type of booleans + +Traditionally the argument of the type we want to eliminate (the booleans in our case) is written last: +```agda +Bool-elim : (A : Bool → Type) + → A true + → A false + → (b : Bool) → A b +Bool-elim A x y true = x +Bool-elim A x y false = y +``` +The type of `Bool-elim` says that if we provide elements of the type `A true` and `A false`, we get a function `(b : Bool) → A b`. + +The non-dependent version is a particular case of the dependent version, by considering the constant type family `λ _ → A` for a given `A : Type`. This time we make the first argument `A` implicit: +```agda +Bool-nondep-elim : {A : Type} + → A + → A + → Bool → A +Bool-nondep-elim {A} = Bool-elim (λ _ → A) +``` +This produces a function `Bool → A` from two given elements of the type `A`. + +## Logical reading of the eliminator + +The *conclusion* of `Bool-elim` is `(b : Bool) → A b`, which under *propositions as types* has the logical reading "for all `b : Bool`, the proposition `A b` holds". The *hypotheses* `A true` and `A false` are all is needed to reach this conclusion. + +Thus the logical reading of `Bool-elim` is: + + * In order to prove that "for all `b : Bool`, the proposition `A b` holds" + +it is enough to prove that + + * the proposition `A true` holds, and + + * the proposition `A false` holds, + +which should be intuitively clear. + +## Examples of proofs using the eliminator + +First define +```agda +not : Bool → Bool +not true = false +not false = true +``` +Then we can prove that `not` can be expressed using `if-then-else`: +```agda +open import identity-type +not-defined-with-if : (b : Bool) → not b ≡ if b then false else true +not-defined-with-if true = refl false +not-defined-with-if false = refl true +``` +Using the eliminator, this can be proved as follows: +```agda +not-defined-with-if₀ : (b : Bool) → not b ≡ if b then false else true +not-defined-with-if₀ = Bool-elim A x y + where + A : Bool → Type + A b = not b ≡ if b then false else true + + x : A true + x = refl false + + y : A false + y = refl true +``` +Of course, we can "in-line" the definitions of the `where` clause: +```agda +not-defined-with-if₁ : (b : Bool) → not b ≡ if b then false else true +not-defined-with-if₁ = Bool-elim + (λ b → not b ≡ if b then false else true) + (refl false) + (refl true) +``` +This is shorter but probably less readable. The following is even shorter, using the fact that Agda can infer the property `A : Bool → Type` we want to prove automatically. We use `_` to tell Agda "please figure out yourself what this argument in the function has to be": +```agda +not-defined-with-if₂ : (b : Bool) → not b ≡ if b then false else true +not-defined-with-if₂ = Bool-elim _ (refl false) (refl true) +``` +In situations where we try to use `_` but Agda can't determine that there is a *unique* answer to what `_` should be, the colour yellow is used to indicate this in the syntax highlighting, accompanied by an error message. To give another example, we first define the notion of an [involution](https://en.wikipedia.org/wiki/Involution_(mathematics)), or involutive function: +```agda +is-involution : {X : Type} → (X → X) → Type +is-involution {X} f = (x : X) → f (f x) ≡ x +``` +For example, list reversal is an involution. Another example is boolean negation: +```agda +not-is-involution : is-involution not +not-is-involution = Bool-elim _ (refl true) (refl false) +``` +A proof without mentioning `is-involution` and without using the eliminator is also possible, of course: +```agda +not-is-involution' : (b : Bool) → not (not b) ≡ b +not-is-involution' true = refl true +not-is-involution' false = refl false +``` +Very often we will give definitions by pattern-matching as above instead of +`Bool-elim`. But the concept of eliminator for a type remains useful, and it is what `MLTT` (Martin-Löf Type Theory), the foundation of our programming language Agda, uses to define types. Types are defined by formation rules, construtors, eliminators, and equations explaining how the constructors interact with the eliminators. Pattern-matching can be considered as "syntax sugar" for definitions using eliminators. Usually definitions using pattern matching are more readable than definitions using eliminators, but they are equivalent to definitions using eliminators. + +Notice that in the definition of `is-involution` we needed to explicitly indicate the implicit argument `X` using curly brackets. Agda allows the notation `∀` in order to be able to omit the type `X`, provided it can be inferred automaticaly, which it can in our situation: +```agda +is-involution' : {X : Type} → (X → X) → Type +is-involution' f = ∀ x → f (f x) ≡ x +``` + +## Some useful functions + +```agda +_&&_ : Bool → Bool → Bool +true && y = y +false && y = false + +_||_ : Bool → Bool → Bool +true || y = true +false || y = y + +infixr 20 _||_ +infixr 30 _&&_ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Fin-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Fin-functions.lagda.md new file mode 100644 index 0000000..878face --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Fin-functions.lagda.md @@ -0,0 +1,74 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +# Isomorphism of Fin n with a Basic MLTT type + +```agda +Fin' : ℕ → Type +Fin' 0 = 𝟘 +Fin' (suc n) = 𝟙 ∔ Fin' n + +zero' : {n : ℕ} → Fin' (suc n) +zero' = inl ⋆ + +suc' : {n : ℕ} → Fin' n → Fin' (suc n) +suc' = inr + +open import Fin +open import isomorphisms + +Fin-isomorphism : (n : ℕ) → Fin n ≅ Fin' n +Fin-isomorphism n = record { bijection = f n ; bijectivity = f-is-bijection n } + where + f : (n : ℕ) → Fin n → Fin' n + f (suc n) zero = zero' + f (suc n) (suc k) = suc' (f n k) + + g : (n : ℕ) → Fin' n → Fin n + g (suc n) (inl ⋆) = zero + g (suc n) (inr k) = suc (g n k) + + gf : (n : ℕ) → g n ∘ f n ∼ id + gf (suc n) zero = refl zero + gf (suc n) (suc k) = γ + where + IH : g n (f n k) ≡ k + IH = gf n k + + γ = g (suc n) (f (suc n) (suc k)) ≡⟨ refl _ ⟩ + g (suc n) (suc' (f n k)) ≡⟨ refl _ ⟩ + suc (g n (f n k)) ≡⟨ ap suc IH ⟩ + suc k ∎ + + fg : (n : ℕ) → f n ∘ g n ∼ id + fg (suc n) (inl ⋆) = refl (inl ⋆) + fg (suc n) (inr k) = γ + where + IH : f n (g n k) ≡ k + IH = fg n k + + γ = f (suc n) (g (suc n) (suc' k)) ≡⟨ refl _ ⟩ + f (suc n) (suc (g n k)) ≡⟨ refl _ ⟩ + suc' (f n (g n k)) ≡⟨ ap suc' IH ⟩ + suc' k ∎ + + f-is-bijection : (n : ℕ) → is-bijection (f n) + f-is-bijection n = record { inverse = g n ; η = gf n ; ε = fg n} +``` + +**Exercise.** Show that the type `Fin n` is isomorphic to the type `Σ k : ℕ , k < n`. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Fin.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Fin.lagda.md new file mode 100644 index 0000000..0b81413 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Fin.lagda.md @@ -0,0 +1,77 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Finite types + +We now define a type `Fin n` which has exactly `n` elements, where `n : ℕ` is a natural number. + +```agda +open import natural-numbers-type public + +data Fin : ℕ → Type where + zero : {n : ℕ} → Fin (suc n) + suc : {n : ℕ} → Fin n → Fin (suc n) +``` +Examples: +```agda +private + x₀ x₁ x₂ : Fin 3 + x₀ = zero + x₁ = suc zero + x₂ = suc (suc zero) + + y₀ y₁ y₂ : Fin 3 + y₀ = zero {2} + y₁ = suc {2} (zero {1}) + y₂ = suc {2} (suc {1} (zero {0})) +``` +And these are all the elements of `Fin 3`. Notice that `Fin 0` is empty: +```agda + +open import empty-type public + +Fin-0-is-empty : is-empty (Fin 0) +Fin-0-is-empty () +``` +Here `()` is a pseudo-pattern that tells that there is no constructor for the type. +```agda +Fin-suc-is-nonempty : {n : ℕ} → ¬ is-empty (Fin (suc n)) +Fin-suc-is-nonempty f = f zero +``` + +## Elimination principle + +```agda +Fin-elim : (A : {n : ℕ} → Fin n → Type) + → ({n : ℕ} → A {suc n} zero) + → ({n : ℕ} (k : Fin n) → A k → A (suc k)) + → {n : ℕ} (k : Fin n) → A k +Fin-elim A a f = h + where + h : {n : ℕ} (k : Fin n) → A k + h zero = a + h (suc k) = f k (h k) +``` +So this again looks like [primitive recursion](natural-numbers-type.lagda.md). And it gives an induction principle for `Fin`. + +## Every element of `Fin n` can be regarded as an element of `ℕ` + +```agda +ι : {n : ℕ} → Fin n → ℕ +ι zero = zero +ι (suc x) = suc (ι x) +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Hedbergs-Theorem.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Hedbergs-Theorem.lagda.md new file mode 100644 index 0000000..3df34da --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Hedbergs-Theorem.lagda.md @@ -0,0 +1,116 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +# Hedberg's Theorem + +Sometimes we wish to know that the identity type `x ≡ y` has at most one element for all elements `x` and `y` of a type. In this case we say that the type is a set. Alternatively, one says that the type satisfies *uniqueness of identity proofs* (UIP), or that it satisfies the *K axiom*. Perhaps surprisingly, this can't be proved for all types. But it can be proved for quite a few types, including the booleans, natural numbers, and functions `ℕ → ℕ`, among many others. + +Hedberg's Theorem, whose proof is short, but quite difficult to understand, even for experts in Martin-Löf type theory, is a main tool for establishing that some types are sets. +Agda has the axiom `K` discussed above enabled by default. We are disabling it in all modules, including this. The reason is that towards the end of term we intend to give examples of types that are not sets, and explain why they are interesting. + +```agda +{-# OPTIONS --without-K --safe #-} + +module Hedbergs-Theorem where + +open import prelude +open import negation + +is-prop : Type → Type +is-prop X = (x y : X) → x ≡ y + +𝟘-is-prop : is-prop 𝟘 +𝟘-is-prop () () + +𝟙-is-prop : is-prop 𝟙 +𝟙-is-prop ⋆ ⋆ = refl ⋆ + +is-set : Type → Type +is-set X = (x y : X) → is-prop (x ≡ y) + +is-constant : {X Y : Type} → (X → Y) → Type +is-constant {X} f = (x x' : X) → f x ≡ f x' + +has-constant-endofunction : Type → Type +has-constant-endofunction X = Σ f ꞉ (X → X), is-constant f + +⁻¹-left∙ : {X : Type} {x y : X} (p : x ≡ y) + → p ⁻¹ ∙ p ≡ refl y +⁻¹-left∙ (refl x) = refl (refl x) + +⁻¹-right∙ : {X : Type} {x y : X} (p : x ≡ y) + → p ∙ p ⁻¹ ≡ refl x +⁻¹-right∙ (refl x) = refl (refl x) + +Hedbergs-Lemma : {X : Type} (x : X) + → ((y : X) → has-constant-endofunction (x ≡ y)) + → (y : X) → is-prop (x ≡ y) +Hedbergs-Lemma {X} x c y p q = II + where + f : (y : X) → x ≡ y → x ≡ y + f y = pr₁ (c y) + + κ : (y : X) (p q : x ≡ y) → f y p ≡ f y q + κ y = pr₂ (c y) + + I : (y : X) (p : x ≡ y) → (f x (refl x))⁻¹ ∙ f y p ≡ p + I x (refl x) = r + where + r : (f x (refl x)) ⁻¹ ∙ f x (refl x) ≡ refl x + r = ⁻¹-left∙ (f x (refl x)) + + II = p ≡⟨ (I y p)⁻¹ ⟩ + (f x (refl x))⁻¹ ∙ f y p ≡⟨ ap ((f x (refl x))⁻¹ ∙_) (κ y p q) ⟩ + (f x (refl x))⁻¹ ∙ f y q ≡⟨ I y q ⟩ + q ∎ + + +is-Hedberg-type : Type → Type +is-Hedberg-type X = (x y : X) → has-constant-endofunction (x ≡ y) + +Hedberg-types-are-sets : (X : Type) + → is-Hedberg-type X + → is-set X +Hedberg-types-are-sets X c x = Hedbergs-Lemma x + (λ y → pr₁ (c x y) , pr₂ (c x y)) + +pointed-types-have-constant-endofunction : {X : Type} + → X + → has-constant-endofunction X +pointed-types-have-constant-endofunction x = ((λ y → x) , (λ y y' → refl x)) + +empty-types-have-constant-endofunction : {X : Type} + → is-empty X + → has-constant-endofunction X +empty-types-have-constant-endofunction e = (id , (λ x x' → 𝟘-elim (e x))) + +open import decidability + +decidable-types-have-constant-endofunctions : {X : Type} + → is-decidable X + → has-constant-endofunction X +decidable-types-have-constant-endofunctions (inl x) = + pointed-types-have-constant-endofunction x +decidable-types-have-constant-endofunctions (inr e) = + empty-types-have-constant-endofunction e + +types-with-decidable-equality-are-Hedberg : {X : Type} + → has-decidable-equality X + → is-Hedberg-type X +types-with-decidable-equality-are-Hedberg d x y = + decidable-types-have-constant-endofunctions (d x y) + +Hedbergs-Theorem : {X : Type} → has-decidable-equality X → is-set X +Hedbergs-Theorem {X} d = Hedberg-types-are-sets X + (types-with-decidable-equality-are-Hedberg d) + +ℕ-is-set : is-set ℕ +ℕ-is-set = Hedbergs-Theorem ℕ-has-decidable-equality + +Bool-is-set : is-set Bool +Bool-is-set = Hedbergs-Theorem Bool-has-decidable-equality +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/List-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/List-functions.lagda.md new file mode 100644 index 0000000..4117f0b --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/List-functions.lagda.md @@ -0,0 +1,117 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Some functions on lists + +## Length, concatenation, map and singleton lists + +```agda +open import List +open import natural-numbers-type + +length : {A : Type} → List A → ℕ +length [] = 0 +length (x :: xs) = 1 + length xs + +_++_ : {A : Type} → List A → List A → List A +[] ++ ys = ys +(x :: xs) ++ ys = x :: (xs ++ ys) + +infixr 20 _++_ + +map : {A B : Type} → (A → B) → List A → List B +map f [] = [] +map f (x :: xs) = f x :: map f xs + +[_] : {A : Type} → A → List A +[ x ] = x :: [] +``` + +## Properties of list concatenation + +```agda +[]-left-neutral : {X : Type} (xs : List X) → [] ++ xs ≡ xs +[]-left-neutral = refl + +[]-right-neutral : {X : Type} (xs : List X) → xs ++ [] ≡ xs +[]-right-neutral [] = refl [] +[]-right-neutral (x :: xs) = + (x :: xs) ++ [] ≡⟨ refl _ ⟩ + x :: (xs ++ []) ≡⟨ ap (x ::_) ([]-right-neutral xs) ⟩ + x :: xs ∎ + +++-assoc : {A : Type} (xs ys zs : List A) → (xs ++ ys) ++ zs ≡ xs ++ (ys ++ zs) +++-assoc [] ys zs = refl (ys ++ zs) +++-assoc (x :: xs) ys zs = + ((x :: xs) ++ ys) ++ zs ≡⟨ refl _ ⟩ + x :: ((xs ++ ys) ++ zs) ≡⟨ ap (x ::_) (++-assoc xs ys zs) ⟩ + x :: (xs ++ (ys ++ zs)) ≡⟨ refl _ ⟩ + (x :: xs) ++ ys ++ zs ∎ +``` + +Short proofs: +```agda +[]-right-neutral' : {X : Type} (xs : List X) → xs ++ [] ≡ xs +[]-right-neutral' [] = refl [] +[]-right-neutral' (x :: xs) = ap (x ::_) ([]-right-neutral' xs) + +++-assoc' : {A : Type} (xs ys zs : List A) → (xs ++ ys) ++ zs ≡ xs ++ (ys ++ zs) +++-assoc' [] ys zs = refl (ys ++ zs) +++-assoc' (x :: xs) ys zs = ap (x ::_) (++-assoc' xs ys zs) +``` + +## List reversal + +First an obvious, but slow reversal algorithm: +```agda +reverse : {A : Type} → List A → List A +reverse [] = [] +reverse (x :: xs) = reverse xs ++ [ x ] +``` +This is quadratic time. To get a linear-time algorithm, we use the following auxiliary function: +```agda +rev-append : {A : Type} → List A → List A → List A +rev-append [] ys = ys +rev-append (x :: xs) ys = rev-append xs (x :: ys) +``` +The intended behaviour of `rev-append` is that `rev-append xs ys = reverse xs ++ ys`. +Using this fact, the linear time algorithm is the following: +```agda +rev : {A : Type} → List A → List A +rev xs = rev-append xs [] +``` +We now want to show that `rev` and `reverse` behave in the same way. To do this, we first show that the auxiliary function does behave as intended: +```agda +rev-append-behaviour : {A : Type} (xs ys : List A) + → rev-append xs ys ≡ reverse xs ++ ys +rev-append-behaviour [] ys = refl ys +rev-append-behaviour (x :: xs) ys = + rev-append (x :: xs) ys ≡⟨ refl _ ⟩ + rev-append xs (x :: ys) ≡⟨ rev-append-behaviour xs (x :: ys) ⟩ + reverse xs ++ (x :: ys) ≡⟨ refl _ ⟩ + reverse xs ++ ([ x ] ++ ys) ≡⟨ sym (++-assoc (reverse xs) [ x ] ys) ⟩ + (reverse xs ++ [ x ]) ++ ys ≡⟨ refl _ ⟩ + reverse (x :: xs) ++ ys ∎ +``` +We state this as follows in Agda: +```agda +rev-correct : {A : Type} (xs : List A) → rev xs ≡ reverse xs +rev-correct xs = + rev xs ≡⟨ refl _ ⟩ + rev-append xs [] ≡⟨ rev-append-behaviour xs [] ⟩ + reverse xs ++ [] ≡⟨ []-right-neutral (reverse xs) ⟩ + reverse xs ∎ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/List.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/List.lagda.md new file mode 100644 index 0000000..f6a5a52 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/List.lagda.md @@ -0,0 +1,52 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Lists + +This type has already been briefly discussed in the introduction. +```agda +data List (A : Type) : Type where + [] : List A + _::_ : A → List A → List A + +infixr 10 _::_ +``` + +## Elimination principle + +```agda +List-elim : {X : Type} (A : List X → Type) + → A [] + → ((x : X) (xs : List X) → A xs → A (x :: xs)) + → (xs : List X) → A xs +List-elim {X} A a f = h + where + h : (xs : List X) → A xs + h [] = a + h (x :: xs) = f x xs (h xs) +``` +Notice that the definition of `h` is the same as that of the usual `fold` function for lists, if you know this, but the type is more general. In fact, the `fold` function is just the non-dependent version of the eliminator +```agda +List-nondep-elim : {X A : Type} + → A + → (X → List X → A → A) + → List X → A +List-nondep-elim {X} {A} a f = List-elim {X} (λ _ → A) a f +``` + +## Induction on lists + +In terms of logic, the elimination principle gives an induction principle for proving properties of lists. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Vector-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Vector-functions.lagda.md new file mode 100644 index 0000000..8f01920 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Vector-functions.lagda.md @@ -0,0 +1,114 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Some functions on vectors + +As mentioned in the [introduction](introduction.lagda.md), vectors allow us to have safe `head` and `tail` functions. +```agda +head : {A : Type} {n : ℕ} → Vector A (suc n) → A +head (x :: xs) = x + +tail : {A : Type} {n : ℕ} → Vector A (suc n) → Vector A n +tail (x :: xs) = xs +``` + +We can also define a safe indexing function on vectors using [finite types](Fin.lagda.md) as follows. +```agda +open import Fin + +_!!_ : ∀ {X n} → Vector X n → Fin n → X +(x :: xs) !! zero = x +(x :: xs) !! suc n = xs !! n +``` +Alternatively, this can be defined as follows: +```agda +_!!'_ : {X : Type} {n : ℕ} → Vector X n → Fin n → X +[] !!' () +(x :: xs) !!' zero = x +(x :: xs) !!' suc n = xs !!' n +``` + +We can also do vector concatenation: +```agda +_++_ : {X : Type} {m n : ℕ} → Vector X m → Vector X n → Vector X (m + n) +[] ++ ys = ys +(x :: xs) ++ ys = x :: (xs ++ ys) + +infixr 20 _++_ +``` + +## Vectors represented as a Basic MLTT type + +```agda +open import unit-type +open import binary-products + +Vector' : Type → ℕ → Type +Vector' A 0 = 𝟙 +Vector' A (suc n) = A × Vector' A n + +[]' : {A : Type} → Vector' A 0 +[]' = ⋆ + +_::'_ : {A : Type} {n : ℕ} → A → Vector' A n → Vector' A (suc n) +x ::' xs = x , xs + +infixr 10 _::'_ + +private + example₀ : Vector' ℕ 3 + example₀ = 1 ::' 2 ::' 3 ::' ([]' {ℕ}) + + example₁ : example₀ ≡ (1 , 2 , 3 , ⋆) + example₁ = refl _ + +open import isomorphisms + +Vector-iso : {A : Type} {n : ℕ} → Vector A n ≅ Vector' A n +Vector-iso {A} {n} = record { bijection = f n ; bijectivity = f-is-bijection n } + where + f : (n : ℕ) → Vector A n → Vector' A n + f 0 [] = []' {A} + f (suc n) (x :: xs) = x ::' f n xs + + g : (n : ℕ) → Vector' A n → Vector A n + g 0 ⋆ = [] + g (suc n) (x , xs) = x :: g n xs + + gf : (n : ℕ) → g n ∘ f n ∼ id + gf 0 [] = refl [] + gf (suc n) (x :: xs) = ap (x ::_) (gf n xs) + + fg : (n : ℕ) → f n ∘ g n ∼ id + fg 0 ⋆ = refl ⋆ + fg (suc n) (x , xs) = ap (x ,_) (fg n xs) + + f-is-bijection : (n : ℕ) → is-bijection (f n) + f-is-bijection n = record { inverse = g n ; η = gf n ; ε = fg n } + +private + open _≅_ + open is-bijection + + example₂ : bijection Vector-iso (1 :: 2 :: 3 :: []) ≡ (1 , 2 , 3 , ⋆) + example₂ = refl _ + + example₄ : Vector ℕ 3 + example₄ = inverse (bijectivity Vector-iso) (1 , 2 , 3 , ⋆) + + example₅ : example₄ ≡ 1 :: 2 :: 3 :: [] + example₅ = refl _ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/Vector.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/Vector.lagda.md new file mode 100644 index 0000000..6809ea3 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/Vector.lagda.md @@ -0,0 +1,67 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Vectors + +This type has already been briefly discussed in the [introduction](introduction.lagda.md). +```agda +open import natural-numbers-type + +data Vector (A : Type) : ℕ → Type where + [] : Vector A 0 + _::_ : {n : ℕ} → A → Vector A n → Vector A (suc n) + +infixr 10 _::_ +``` + +## Elimination principle + +```agda +Vector-elim : {X : Type} (A : (n : ℕ) → Vector X n → Type) + → A 0 [] + → ((x : X) (n : ℕ) (xs : Vector X n) → A n xs → A (suc n) (x :: xs)) + → (n : ℕ) (xs : Vector X n) → A n xs +Vector-elim {X} A a f = h + where + h : (n : ℕ) (xs : Vector X n) → A n xs + h 0 [] = a + h (suc n) (x :: xs) = f x n xs (h n xs) +``` +It is better, in practice, to make the parameter `n` implicit, because it can be inferred from the type of `xs`, and so we get less clutter: +```agda +Vector-elim' : {X : Type} (A : {n : ℕ} → Vector X n → Type) + → A [] + → ((x : X) {n : ℕ} (xs : Vector X n) → A xs → A (x :: xs)) + → {n : ℕ} (xs : Vector X n) → A xs +Vector-elim' {X} A a f = h + where + h : {n : ℕ} (xs : Vector X n) → A xs + h [] = a + h (x :: xs) = f x xs (h xs) +``` +Again, the non-dependent version gives a fold function for vectors: +```agda +Vector-nondep-elim' : {X A : Type} + → A + → (X → {n : ℕ} → Vector X n → A → A) + → {n : ℕ} → Vector X n → A +Vector-nondep-elim' {X} {A} = Vector-elim' {X} (λ {_} _ → A) +``` + +## Induction on vectors + +As for lists, it is given by the proposition-as-types reading of the type of `Vector-elim`. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/all.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/all.lagda.md new file mode 100644 index 0000000..6aee2a3 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/all.lagda.md @@ -0,0 +1,44 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +```agda +{-# OPTIONS --without-K --safe #-} + +module all where + +import Bool +import Bool-functions +import Fin +import Fin-functions +import Hedbergs-Theorem +import List +import List-functions +import Vector +import Vector-functions +import binary-products +import binary-products-as-products +import binary-sums +import binary-sums-as-sums +import binary-type +import decidability +import empty-type +import function-extensionality +import general-notation +import identity-type +import isomorphism-functions +import isomorphisms +import natural-numbers-functions +import natural-numbers-type +import negation +import prelude +import products +import searchability +import sums +import sums-equality +import unit-type +-- import vector-and-list-isomorphisms +``` + +[Go back to the table of contents](../README.md) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/binary-products-as-products.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/binary-products-as-products.lagda.md new file mode 100644 index 0000000..8d6e864 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/binary-products-as-products.lagda.md @@ -0,0 +1,66 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +## Binary products as a special case of arbitrary products + +Using the [binary type](binary-type.lagda.md) `𝟚`, binary products can be seen as a special case of arbitrary products as follows: +```agda +open import binary-type + +_×'_ : Type → Type → Type +A₀ ×' A₁ = Π n ꞉ 𝟚 , A n + where + A : 𝟚 → Type + A 𝟎 = A₀ + A 𝟏 = A₁ + +infixr 2 _×'_ +``` +We could have written the type `Π n ꞉ 𝟚 , A n` as simply `(n : 𝟚) → A n`, but we wanted to emphasize that binary products `_×_` are special cases of arbitrary products `Π`. + +To justify this claim, we establish an [isomorphism](isomorphisms.lagda.md). But we need to assume [function extensionality](function-extensionality.lagda.md) for this purpose. +```agda +open import isomorphisms +open import function-extensionality + +binary-product-isomorphism : FunExt → (A₀ A₁ : Type) → A₀ × A₁ ≅ A₀ ×' A₁ +binary-product-isomorphism funext A₀ A₁ = record { bijection = f ; bijectivity = f-is-bijection } + where + f : A₀ × A₁ → A₀ ×' A₁ + f (x₀ , x₁) 𝟎 = x₀ + f (x₀ , x₁) 𝟏 = x₁ + + g : A₀ ×' A₁ → A₀ × A₁ + g h = h 𝟎 , h 𝟏 + + gf : g ∘ f ∼ id + gf (x₀ , x₁) = refl (x₀ , x₁) + + fg : f ∘ g ∼ id + fg h = γ + where + p : f (g h) ∼ h + p 𝟎 = refl (h 𝟎) + p 𝟏 = refl (h 𝟏) + + γ : f (g h) ≡ h + γ = funext p + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` +Notice that the above argument, in Agda, not only *shows* that the types are indeed isomorphic, but also explains *how* and *why* they are isomorphic. Thus, logical arguments coded in Agda are useful not only to get confidence in correctness, but also to gain understanding. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/binary-products.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/binary-products.lagda.md new file mode 100644 index 0000000..5174604 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/binary-products.lagda.md @@ -0,0 +1,56 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The cartesian-product type former `_×_` + +As [discussed before](sums.lagda.md), the cartesian product `A × B` is simply `Σ {A} (λ x → B)` which means that for `(x , y) : A × B` the type of `y` is always `B`, independently of what `x` is. Using our special syntax for `Σ` this can be defined as follows in Agda: +```agda +open import sums public + +_×_ : Type → Type → Type +A × B = Σ x ꞉ A , B + +infixr 2 _×_ +``` +So the elements of `A × B` are of the form `(x , y)` with `x : A` and `y : B`. + +## Logical conjunction ("and") + +The logical interpretation of `A × B` is "A and B". This is because in order to show that the proposition "A and B" holds, we have to show that each of A and B holds. To show that A holds we exhibit an element `x` of A, and to show that B holds we exhibit an element `y` of B, and so to show that "A and B" holds we exhibit a pair `(x , y)` with `x : A` and `b : B` + +## Logical equivalence + +We now can define bi-implication, or logical equivalence, as follows: +```agda +_⇔_ : Type → Type → Type +A ⇔ B = (A → B) × (B → A) + +infix -2 _⇔_ +``` +The symbol `⇔` is often pronounced "if and only if". + +We use the first and second projections to extract the left-to-right implication and the right-to-left implication: +```agda +lr-implication : {A B : Type} → (A ⇔ B) → (A → B) +lr-implication = pr₁ + +rl-implication : {A B : Type} → (A ⇔ B) → (B → A) +rl-implication = pr₂ +``` + +## Alternative definition of `_×_` + +There is [another way to define binary products](binary-products-as-products.lagda.md) as a special case of `Π` rather than `Σ`. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums-as-sums.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums-as-sums.lagda.md new file mode 100644 index 0000000..9af2930 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums-as-sums.lagda.md @@ -0,0 +1,57 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +## Binary sums as a special case of arbitrary sums + +Using the binary type `𝟚`, binary sums can be seen as a special case of arbitrary sums as follows: +```agda +open import binary-type + +_∔'_ : Type → Type → Type +A₀ ∔' A₁ = Σ n ꞉ 𝟚 , A n + where + A : 𝟚 → Type + A 𝟎 = A₀ + A 𝟏 = A₁ +``` + +To justify this claim, we establish an isomorphism. +```agda +open import isomorphisms + +binary-sum-isomorphism : (A₀ A₁ : Type) → A₀ ∔ A₁ ≅ A₀ ∔' A₁ +binary-sum-isomorphism A₀ A₁ = record { bijection = f ; bijectivity = f-is-bijection } + where + f : A₀ ∔ A₁ → A₀ ∔' A₁ + f (inl x₀) = 𝟎 , x₀ + f (inr x₁) = 𝟏 , x₁ + + g : A₀ ∔' A₁ → A₀ ∔ A₁ + g (𝟎 , x₀) = inl x₀ + g (𝟏 , x₁) = inr x₁ + + gf : g ∘ f ∼ id + gf (inl x₀) = refl (inl x₀) + gf (inr x₁) = refl (inr x₁) + + fg : f ∘ g ∼ id + fg (𝟎 , x₀) = refl (𝟎 , x₀) + fg (𝟏 , x₁) = refl (𝟏 , x₁) + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums.lagda.md new file mode 100644 index 0000000..2b6f3db --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/binary-sums.lagda.md @@ -0,0 +1,88 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +# The binary-sum type former `_∔_` + +This is the same as (or, more precisely, [isomorphic](isomorphisms.lagda.md) to) the `Either` type defined earlier (you can try this as an exercise). The notation in type theory is `_+_`, but we want to reserve this for addition of natural numbers, and hence we use the same symbol with a dot on top: +```agda +data _∔_ (A B : Type) : Type where + inl : A → A ∔ B + inr : B → A ∔ B + +infixr 20 _∔_ +``` + +The type `A ∔ B` is called the coproduct of `A` and `B`, or the sum of `A` and `B`, or the disjoint union of `A` and `B`. The elements of `A ∔ B` are of the form `inl x` with `x : A` and `inr y` with `y : B`. + +In terms of computation, we use the type `A ∔ B` when we want to put the two types together into a single type. It is also possible to write `A ∔ A`, in which case we will have two copies of the type `A`, so that now every element `x` of `A` has two different incarnations `inl a` and `inr a` in the type `A ∔ A`. For example, the [unit type](unit-type.lagda.md) `𝟙` has exactly one element, namely `⋆ : 𝟙`, and hence the type `𝟙 ∔ 𝟙` has precisely two elements, namely `inl ⋆` and `inr ⋆`. + +## Logical disjunction ("or") + +In terms of logic, we use the type `A ∔ B` to express "A or B". This is because in order for "A or B" to hold, at least one of A and B must hold. The type `A → A ∔ B` of the function `inl` is interpreted as saying that if A holds then so does "A or B", and similarly, the type of B → A ∔ B of the function `inr` says that if B holds then so does "A or B". In other words, if `x : A` is a proof of `A`, then `inl x : A + B` is a proof of `A or B`, and if `y : B` is a proof of B, them `inr y : A + B` is a proof of "A or B". Here when we said "proof" we meant "program" after the propositions-as-types and proofs-as-programs paradigm. + +## Logical disjunction in HoTT/UF + +In HoTT/UF it useful to have an alternative disjunction operation `A ∨ B` defined to be `∥ A ∔ B ∥` where `∥_∥` is a certain *propositional truncation* operation. + +## Elimination principle + +Now suppose we want to define a dependent function `(z : A ∔ B) → C z`. How can we do that? If we have two functions `f : (x : A) → C (inl x)` and `g : (y : B) → C (inr y)`, then, given `z : A ∔ B`, we can inspect whether `z` is of the form `inl x` with `x : A` or of the form `inr y` with `y : B`, and the respectively apply `f` or `g` to get an element of `C z`. This procedure is called the *elimination* principle for the type former `_∔_` and can be written in Agda as follows: + +```agda +∔-elim : {A B : Type} (C : A ∔ B → Type) + → ((x : A) → C (inl x)) + → ((y : B) → C (inr y)) + → (z : A ∔ B) → C z +∔-elim C f g (inl x) = f x +∔-elim C f g (inr y) = g y +``` +So the eliminator amounts to simply definition by cases. In terms of logic, it says that in order to show that "for all z : A ∔ B we have that C z holds" it is enough to show two things: (1) "for all x : A it is the case that C (inl x) holds", and (2) "forall y : B it is the case that C (inr y) holds". This is not only sufficient, but also necessary: +```agda +open import binary-products + +∔-split : {A B : Type} (C : A ∔ B → Type) + → ((z : A ∔ B) → C z) + → ((x : A) → C (inl x)) × ((y : B) → C (inr y)) +∔-split {A} {B} C h = f , g + where + f : (x : A) → C (inl x) + f x = h (inl x) + + g : (y : B) → C (inr y) + g y = h (inr y) +``` + +There is also a version of the eliminator in which `C` doesn't depend on `z : A ∔ B` and is always the same: +```agda +∔-nondep-elim : {A B C : Type} + → (A → C) + → (B → C) + → (A ∔ B → C) +∔-nondep-elim {A} {B} {C} = ∔-elim (λ z → C) +``` +In terms of logic, this means that in order to show that "A or B implies C", it is enough to show that both "A implies C" and "B implies C". This also can be inverted: +```agda +∔-nondep-split : {A B C : Type} + → (A ∔ B → C) + → (A → C) × (B → C) +∔-nondep-split {A} {B} {C} = ∔-split (λ z → C) +``` +In terms of logic, this means that if `A or B implies C` then both `A implies C` and `B implies C`. + +## Alternative definition of `_∔_` + +There is [another way to define binary sums](binary-sums-as-sums.lagda.md) as a special case of `Σ`. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/binary-type.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/binary-type.lagda.md new file mode 100644 index 0000000..ed5672a --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/binary-type.lagda.md @@ -0,0 +1,46 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The binary type 𝟚 + +This type can be defined to be `𝟙 ∔ 𝟙` using [binary sums](binary-sums.lagda.md), but we give a direct definition which will allow us to discuss some relationships between the various type formers of Basic MLTT. + +```agda +data 𝟚 : Type where + 𝟎 𝟏 : 𝟚 +``` +This type is not only [isomorphic to `𝟙 ∔ 𝟙`](isomorphisms.lagda.md) but also to the type [`Bool`](Bool.lagda.md) of booleans. +Its elimination principle is as follows: +```agda +𝟚-elim : {A : 𝟚 → Type} + → A 𝟎 + → A 𝟏 + → (x : 𝟚) → A x +𝟚-elim x₀ x₁ 𝟎 = x₀ +𝟚-elim x₀ x₁ 𝟏 = x₁ +``` +In logical terms, this says that it order to show that a property `A` of elements of the binary type `𝟚` holds for all elements of the type `𝟚`, it is enough to show that it holds for `𝟎` and for `𝟏`. The non-dependent version of the eliminator is the following: +```agda +𝟚-nondep-elim : {A : Type} + → A + → A + → 𝟚 → A +𝟚-nondep-elim {A} = 𝟚-elim {λ _ → A} +``` +Notice that the non-dependent version is like `if-then-else`, if we think of one of the elements of `A` as `true` and the other as `false`. + +The dependent version generalizes the *type* of the non-dependent +version, with the same definition of the function. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/curry-howard.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/curry-howard.lagda.md new file mode 100644 index 0000000..6c3b8c9 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/curry-howard.lagda.md @@ -0,0 +1,38 @@ +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + +```agda +{-# OPTIONS --without-K #-} +module curry-howard where +``` + + +# Propositions as types and basic Martin-Löf type theory + +We now complete the proposition-as-types interpretation of logic. + +| Logic | English | Type theory | Agda | Handouts | Alternative terminology | +|-----------------|----------------------------|---------------|---------------|-----------------------------------------------------|--------------------------------------------| +| ⊥ | false | ℕ₀ | 𝟘 | [empty type](empty-type.lagda.md) | | +| ⊤ | true (*) | ℕ₁ | 𝟙 | [unit type](unit-type.lagda.md) | | +| A ∧ B | A and B | A × B | A × B | [binary product](binary-products.lagda.md) | cartesian product | +| A ∨ B | A or B | A + B | A ∔ B | [binary sum](binary-sums.lagda.md) | coproduct,
binary disjoint union | +| A → B | A implies B | A → B | A → B | [function type](products.lagda.md) | non-dependent function type | +| ¬ A | not A | A → ℕ₀ | A → 𝟘 | [negation](negation.lagda.md) | | +| ∀ x : A, B x | for all x:A, B x | Π x : A , B x | (x : A) → B x | [product](products.lagda.md) | dependent function type | +| ∃ x : A, B x | there is x:A such that B x | Σ x ꞉ A , B x | Σ x ꞉ A , B x | [sum](sums.lagda.md) | disjoint union,
dependent pair type | +| x = y | x equals y | Id x y | x ≡ y | [identity type](identity-type.lagda.md) | equality type,
propositional equality | + +## Remarks + + * (*) Not only the unit type 𝟙, but also any type with at least one element can be regarded as "true" in the propositions-as-types interpretation. + + * In Agda we can use any notation we like, of course, and people do use slightly different notatations for e.g. `𝟘`, `𝟏`, `+` and `Σ`. + + * We will refer to the above type theory together with the type `ℕ` of natural numbers as *Basic Martin-Löf Type Theory*. + + * As we will see, this type theory is very expressive and allows us to construct rather sophisticated types, including e.g. lists, vectors and trees. + + * All types in MLTT ([Martin-Löf type theory](http://archive-pml.github.io/martin-lof/pdfs/Bibliopolis-Book-1984.pdf)) come with *introduction* and *elimination* principles. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/decidability.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/decidability.lagda.md new file mode 100644 index 0000000..50aeabb --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/decidability.lagda.md @@ -0,0 +1,360 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Propositions as types versus propositions as booleans + +When programming in conventional programming languages such as Haskell, C, Java, Python, etc., we use *booleans* rather than *types* to encode logical propositions. But this works only because we restrict ourselves to *decidable* propositions, as we'll see below. + +We now discuss *why* we use *types* to encode logical propositions, and +*when* we can use *booleans* instead. It is not always. + +## Discussion and motivation + +**Examples.** We *can automatically check* equality of booleans, integers, strings and much more, using algorithms. + +**Counter-example.** We *can't check* equality of functions of type `ℕ → ℕ`, for instance. Intuitively, to check that two functions `f` and `g` of this type are equal, we need to check infinitely many cases, namely `f x = g x` for all `x : ℕ`. But, we are afraid, intuition is not enough. This has to be proved. Luckily in our case, [Alan Turing](https://en.wikipedia.org/wiki/Alan_Turing) provided the basis to prove that. He showed that the [Halting Problem](https://en.wikipedia.org/wiki/Halting_problem) can't be solved by an algorithm in any programming language. It follows from this that we can't check whether two such functions `f` and `g` are equal or not using an algorithm. + +The above examples and counter-examples show that sometimes we can decide equality with an algorithm, and sometimes we can't. However, for example, the identity type `_≡_` applies to *all* types, whether they have decidable equality or not, and this is why it is useful. We can think about equality, not only in our heads but also in Agda, without worrying whether it can be *checked* to be true or not by a computer. The identity type is not about *checking* equality. It is about asserting that two things are equal, and then proving this ourselves. In fact, equality is very often not checkable by the computer. It is instead about *stating* and *proving* or *disproving* equalities, where the proving and disproving is done by people (the lecturers and the students in this case), by hard, intelligent work. + +## Decidable propositions + +Motivated by the above discussion, we define when a logical proposition is decidable under the understanding of propositions as types: +```agda +is-decidable : Type → Type +is-decidable A = A ∔ ¬ A +``` +This means that we can produce an element of `A` or show that no such element can be found. + +Although it is not possible in general to write a program of type `¬¬ A → A`, this is possible when `A` is decidable: +```agda +¬¬-elim : {A : Type} → is-decidable A → ¬¬ A → A +¬¬-elim (inl x) f = x +¬¬-elim (inr g) f = 𝟘-elim (f g) +``` + +## Decidable propositions as booleans + +The following shows that a type `A` is decidable if and only if there is `b : Bool` such that `A` holds if and only if the boolean `b` is `true`. + +For the purposes of this handout, understanding the following proof is not important at a first reading. What is important is to understand *what* the type of the following function is saying, which is what we explained above. +```agda +decidability-with-booleans : (A : Type) → is-decidable A ⇔ Σ b ꞉ Bool , (A ⇔ b ≡ true) +decidability-with-booleans A = f , g + where + f : is-decidable A → Σ b ꞉ Bool , (A ⇔ b ≡ true) + f (inl x) = true , (α , β) + where + α : A → true ≡ true + α _ = refl true + + β : true ≡ true → A + β _ = x + + f (inr ν) = false , (α , β) + where + α : A → false ≡ true + α x = 𝟘-elim (ν x) + + β : false ≡ true → A + β () + + g : (Σ b ꞉ Bool , (A ⇔ b ≡ true)) → is-decidable A + g (true , α , β) = inl (β (refl true)) + g (false , α , β) = inr (λ x → false-is-not-true (α x)) +``` + +## Decidable predicates as boolean-valued functions + +Consider the logical statement "x is even". This is decidable, because +there is an easy algorithm that tells whether a natural number `x` is +even or not. In programming languages we write this algorithm as a +procedure that returns a boolean. But an equally valid definition is to say that `x` is even if there is a natural number `y` such that `x = 2 * y`. This definition doesn't automatically give an algorithm to check whether or not `x` is even. + +```agda + is-even : ℕ → Type + is-even x = Σ y ꞉ ℕ , x ≡ 2 * y +``` +This says what to be even *means*. But it doesn't say how we *check* with a computer program whether a number is even or not, which would be given by a function `check-even : ℕ → Bool`. +```agda + check-even : ℕ → Bool + check-even 0 = true + check-even (suc x) = not (check-even x) +``` + +For this function to be correct, it has to be the case that + + > `is-even x ⇔ check-even x ≡ true` + +**Exercise.** We believe you have learned enough Agda, try this. + +This is possible because + + > `(x : X) → is-decidable (is-even x)`. + +The following generalizes the above discussion and implements it in Agda. + +First we define what it means for a predicate, such as `A = is-even`, to be decidable: +```agda +is-decidable-predicate : {X : Type} → (X → Type) → Type +is-decidable-predicate {X} A = (x : X) → is-decidable (A x) + +``` +In our example, this means that we can tell whether a number is even or not. + +Next we show that a predicate `A` is decidable if and only if there is a boolean valued function `α` such that `A x` holds if and only if `α x ≡ true`. In the above example, `A` plays the role of `is-even` and `alpha` plays the role of `check-even`. + +Again, what is important at a first reading of this handout is not to understand the proof but what the type of the function is saying. But we will discuss the proof in lectures. + +```agda +predicate-decidability-with-booleans : {X : Type} (A : X → Type) + → is-decidable-predicate A + ⇔ Σ α ꞉ (X → Bool) , ((x : X) → A x ⇔ α x ≡ true) +predicate-decidability-with-booleans {X} A = f , g + where + f : is-decidable-predicate A → Σ α ꞉ (X → Bool) , ((x : X) → A x ⇔ α x ≡ true) + f d = α , β + where + α : X → Bool + α x = pr₁ (lr-implication I (d x)) + where + I : is-decidable (A x) ⇔ Σ b ꞉ Bool , (A x ⇔ b ≡ true) + I = decidability-with-booleans (A x) + + β : (x : X) → A x ⇔ α x ≡ true + β x = ϕ , γ + where + I : is-decidable (A x) → Σ b ꞉ Bool , (A x ⇔ b ≡ true) + I = lr-implication (decidability-with-booleans (A x)) + + II : Σ b ꞉ Bool , (A x ⇔ b ≡ true) + II = I (d x) + + ϕ : A x → α x ≡ true + ϕ = lr-implication (pr₂ II) + + γ : α x ≡ true → A x + γ = rl-implication (pr₂ II) + + g : (Σ α ꞉ (X → Bool) , ((x : X) → A x ⇔ α x ≡ true)) → is-decidable-predicate A + g (α , ϕ) = d + where + d : is-decidable-predicate A + d x = III + where + I : (Σ b ꞉ Bool , (A x ⇔ b ≡ true)) → is-decidable (A x) + I = rl-implication (decidability-with-booleans (A x)) + + II : Σ b ꞉ Bool , (A x ⇔ b ≡ true) + II = (α x , ϕ x) + + III : is-decidable (A x) + III = I II +``` + +Although boolean-valued predicates are fine, we prefer to use type-valued predicates for the sake of uniformity, because we can always define type valued predicates, but only on special circumstances can we define boolean-valued predicates. It is better to define all predicates in the same way, and then write Agda code for predicates that happen to be decidable. + +## Preservation of decidability + +If `A` and `B` are logically equivalent, then `A` is decidable if and only if `B` is decidable. We first prove one direction. +```agda +map-decidable : {A B : Type} → (A → B) → (B → A) → is-decidable A → is-decidable B +map-decidable f g (inl x) = inl (f x) +map-decidable f g (inr h) = inr (λ y → h (g y)) + +map-decidable-corollary : {A B : Type} → (A ⇔ B) → (is-decidable A ⇔ is-decidable B) +map-decidable-corollary (f , g) = map-decidable f g , map-decidable g f +``` +Variation: +```agda +map-decidable' : {A B : Type} → (A → ¬ B) → (¬ A → B) → is-decidable A → is-decidable B +map-decidable' f g (inl x) = inr (f x) +map-decidable' f g (inr h) = inl (g h) +``` + +```agda +pointed-types-are-decidable : {A : Type} → A → is-decidable A +pointed-types-are-decidable = inl + +empty-types-are-decidable : {A : Type} → ¬ A → is-decidable A +empty-types-are-decidable = inr + +𝟙-is-decidable : is-decidable 𝟙 +𝟙-is-decidable = pointed-types-are-decidable ⋆ + +𝟘-is-decidable : is-decidable 𝟘 +𝟘-is-decidable = empty-types-are-decidable 𝟘-is-empty + +∔-preserves-decidability : {A B : Type} + → is-decidable A + → is-decidable B + → is-decidable (A ∔ B) +∔-preserves-decidability (inl x) _ = inl (inl x) +∔-preserves-decidability (inr _) (inl y) = inl (inr y) +∔-preserves-decidability (inr h) (inr k) = inr (∔-nondep-elim h k) + +×-preserves-decidability : {A B : Type} + → is-decidable A + → is-decidable B + → is-decidable (A × B) +×-preserves-decidability (inl x) (inl y) = inl (x , y) +×-preserves-decidability (inl _) (inr k) = inr (λ (x , y) → k y) +×-preserves-decidability (inr h) _ = inr (λ (x , y) → h x) + +→-preserves-decidability : {A B : Type} + → is-decidable A + → is-decidable B + → is-decidable (A → B) +→-preserves-decidability _ (inl y) = inl (λ _ → y) +→-preserves-decidability (inl x) (inr k) = inr (λ f → k (f x)) +→-preserves-decidability (inr h) (inr k) = inl (λ x → 𝟘-elim (h x)) + +¬-preserves-decidability : {A : Type} + → is-decidable A + → is-decidable (¬ A) +¬-preserves-decidability d = →-preserves-decidability d 𝟘-is-decidable +``` + +## Exhaustively searchable types + +We will explain in a future lecture why we need to use `Type₁` rather than `Type` in the following definition. For the moment we just mention that because the definition mentions `Type`, there would be a circularity if the type of the definition where again `Type`. Such circular definitions are not allowed because if they were then we would be able to prove that `0=1`. We have that `Type : Type₁` (the type of `Type` is `Type₁`) but we can't have `Type : Type`. +```agda +is-exhaustively-searchable : Type → Type₁ +is-exhaustively-searchable X = (A : X → Type) + → is-decidable-predicate A + → is-decidable (Σ x ꞉ X , A x) +``` +**Exercise**. Show, in Agda, that the types `𝟘`, `𝟙` , `Bool` and `Fin n`, for any `n : ℕ`, are exhaustively searchable. The idea is that we check whether or not `A x` holds for each `x : A`, and if we find at least one, we conclude that `Σ x ꞉ X , A x`, and otherwise we conclude that `¬ (Σ x ꞉ X , A x)`. This is possible because these types are finite. + +**Exercise**. Think why there can't be any program of type `is-exhaustively-searchable ℕ`, by reduction to the Halting Problem. No Agda code is asked in this question. In fact, the question is asking you to think why such Agda code can't exist. This is very different from proving, in Agda, that `¬ is-exhaustively-searchable ℕ`. Interestingly, this is also not provable in Agda, but explaining why this is the case is beyond the scope of this module. In any case, this is an example of a statement `A` such that neither `A` nor `¬ A` are provable in Agda. Such statements are called *independent*. It must be remarked that the reason why there isn't an Agda program of type `is-exhaustively-searchable ℕ` is *not* merely that `ℕ` is infinite, because there are, perhaps surprisingly, infinite types `A` such that a program of type `is-exhastively-searchable A` can be coded in Agda. One really does an argument such as reduction to the Halting Problem to show that there is no program that can exaustively search the set `ℕ` of natural numbers. + +```agda +Π-exhaustibility : (X : Type) + → is-exhaustively-searchable X + → (A : X → Type) + → is-decidable-predicate A + → is-decidable (Π x ꞉ X , A x) +Π-exhaustibility X s A d = VI + where + I : is-decidable-predicate (λ x → ¬ (A x)) + I x = ¬-preserves-decidability (d x) + + II : is-decidable (Σ x ꞉ X , ¬ (A x)) + II = s (λ x → ¬ (A x)) I + + III : (Σ x ꞉ X , ¬ (A x)) → ¬ (Π x ꞉ X , A x) + III (x , f) g = f (g x) + + IV : ¬ (Σ x ꞉ X , ¬ (A x)) → (Π x ꞉ X , A x) + IV h x = ii + where + i : ¬¬ A x + i f = h (x , f) + + ii : A x + ii = ¬¬-elim (d x) i + + V : is-decidable (Σ x ꞉ X , ¬ (A x)) → is-decidable (Π x ꞉ X , A x) + V = map-decidable' III IV + + VI : is-decidable (Π x ꞉ X , A x) + VI = V II +``` +**Exercises.** If two types `A` and `B` are exhaustively searchable types, then so are the types `A × B` and `A + B`. Moreover, if `X` is an exhaustively searchable type and `A : X → Type` is a family of types, and the type `A x` is exhaustively searchable for each `x : X`, then the type `Σ x ꞉ X , A x` is exhaustively searchable. + +## Decidable equality + +A particular case of interest regarding the above discussion is the notion of a type having decidable equality, which can be written in Agda as follows. + +```agda +has-decidable-equality : Type → Type +has-decidable-equality X = (x y : X) → is-decidable (x ≡ y) +``` +**Exercise.** Show, in Agda, that a type `X` has decidable equality if and only if there is a function `X → X → Bool` that checks whether two elements of `X` are equal or not. + +Some examples: +```agda +Bool-has-decidable-equality : has-decidable-equality Bool +Bool-has-decidable-equality true true = inl (refl true) +Bool-has-decidable-equality true false = inr true-is-not-false +Bool-has-decidable-equality false true = inr false-is-not-true +Bool-has-decidable-equality false false = inl (refl false) + +open import natural-numbers-functions + +ℕ-has-decidable-equality : has-decidable-equality ℕ +ℕ-has-decidable-equality 0 0 = inl (refl zero) +ℕ-has-decidable-equality 0 (suc y) = inr zero-is-not-suc +ℕ-has-decidable-equality (suc x) 0 = inr suc-is-not-zero +ℕ-has-decidable-equality (suc x) (suc y) = III + where + IH : is-decidable (x ≡ y) + IH = ℕ-has-decidable-equality x y + + I : x ≡ y → suc x ≡ suc y + I = ap suc + + II : suc x ≡ suc y → x ≡ y + II = suc-is-injective + + III : is-decidable (suc x ≡ suc y) + III = map-decidable I II IH +``` + +## Equality of functions + +As discussed above, it is not possible to decide whether or not we have `f ∼ g` for two functions `f` and `g`, for example of type `ℕ → ℕ`. However, sometimes we can *prove* or *disprove* this for *particular* functions. Here are some examples: + +```agda +private + f g h : ℕ → ℕ + + f x = x + + g 0 = 0 + g (suc x) = suc (g x) + + h x = suc x + + f-equals-g : f ∼ g + f-equals-g 0 = refl (f 0) + f-equals-g (suc x) = γ + where + IH : f x ≡ g x + IH = f-equals-g x + + γ : f (suc x) ≡ g (suc x) + γ = f (suc x) ≡⟨ refl _ ⟩ + suc x ≡⟨ refl _ ⟩ + suc (f x) ≡⟨ ap suc IH ⟩ + suc (g x) ≡⟨ refl _ ⟩ + g (suc x) ∎ + + f-not-equals-h : ¬ (f ∼ h) + f-not-equals-h e = contradiction d + where + d : 0 ≡ 1 + d = e 0 + + contradiction : ¬ (0 ≡ 1) + contradiction () +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/empty-type.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/empty-type.lagda.md new file mode 100644 index 0000000..4994b30 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/empty-type.lagda.md @@ -0,0 +1,92 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The empty type 𝟘 + +It is convenient to have an empty type `𝟘`, with no elements at all. For example, this can be used in order to define [negation](negation.lagda.md), among other things. + +```agda +data 𝟘 : Type where +``` +And that's the complete definition. The list of constructors that define the type `𝟘` is empty. + +Perhaps counter-intuitively, there is one function `𝟘 → 𝟘`, namely the [identity function](products.lagda.md). So although the type `𝟘` is empty, the type `𝟘 → 𝟘` is non-empty. In fact, the non-dependent elimination principle generalizes that. + +## Proposition as types interpretation + +The empty type is used to interpret "false". Because there is no way to prove the statement `false`, we use the empty type to represent `false` as a type. + +In logic, in order to prove that a proposition is false, we assume it is true and use this assumption to reach a contradiction, such as `0 = 1`. With proofs as programs, in order to show that a statement represented by a type `A` is false, we assume a hypothetical element `x : A`, and from this we try to build a (necessarily impossible) element of the type `𝟘`, which is the desired contradiction. Because of this, in logic the negation of `A` is defined as `A implies false` or `A implies a contradiction`. Hence in type theory we define the negation of a type `A` to be the function type `A → 𝟘`: +``` +¬_ : Type → Type +¬ A = A → 𝟘 + +infix 1000 ¬_ +``` + + +## Elimination principle + +```agda +𝟘-elim : {A : 𝟘 → Type} (x : 𝟘) → A x +𝟘-elim () +``` +The [absurd pattern](https://agda.readthedocs.io/en/latest/language/function-definitions.html#absurd-patterns) `()` +expresses the fact that there is no pattern available because the type is empty. So in the same way that we define the type by giving an empty list of constructors, we define the function `𝟘-elim` by giving an empty list of equations. But we indicate this explicitly with the absurd pattern. + +In terms of logic, this says that in order to show that a property `A` of elements of the empty type holds for all `x : 𝟘`, we have to do nothing, because there is no element to check, and by doing nothing we exhaust all possibilities. This is called [vacuous truth](https://en.wikipedia.org/wiki/Vacuous_truth). + +It is important to notice that this is not a mere technicality. We'll see practical examples in due course. + +The non-dependent version of the eliminator says that there is a function from the empty type to any type: +```agda +𝟘-nondep-elim : {A : Type} → 𝟘 → A +𝟘-nondep-elim {A} = 𝟘-elim {λ _ → A} +``` + +## Definition of emptiness + +On the other hand, there is a function `f : A → 𝟘` if and only if `A` +has no elements, that is, if `A` is also empty. This is because if +`x : A`, there is no element `y : 𝟘` we can choose +in order to define `f x` to be `y`. In fact, we make this observation into our +definition of emptiness: +```agda +is-empty : Type → Type +is-empty A = A → 𝟘 +``` +So notice that this is the same definition as that of negation. + +Here is another example of a type that is empty. In the [introduction](introduction.lagda.md) we defined the identity type former `_≡_`, which [we will revisit](identity-type.lagda.md), and we have that, for example, the type `3 ≡ 4` is empty, whereas the type `3 ≡ 3` has an element `refl 3`. Here are some examples coded in Agda: +```agda +𝟘-is-empty : is-empty 𝟘 +𝟘-is-empty = 𝟘-nondep-elim + +open import unit-type + +𝟙-is-nonempty : ¬ is-empty 𝟙 +𝟙-is-nonempty f = f ⋆ +``` + +The last function works as follows. First we unfold the definition of `¬ is-empty 𝟙` to get `is-empty 𝟙 → 𝟘`. Unfolding again, we get the type `(𝟙 → 𝟘) → 𝟘`. So, given a hypothetical function `f : 𝟙 → 𝟘`, which of course cannot exist (and this what we are trying to conclude), we need to produce an element of `𝟘`. We do this by simply applying the mythical `f` to `⋆ : 𝟙`. We can actually incorporate this discussion in the Agda code, if we want: +```agda +𝟙-is-nonempty' : ¬ is-empty 𝟙 +𝟙-is-nonempty' = γ + where + γ : (𝟙 → 𝟘) → 𝟘 + γ f = f ⋆ +``` +Agda accepts this second version because it automatically unfolds definitions, just as we have done above, to check whether what we have written makes sense. In this case, Agda knows that `¬ is-empty 𝟙` is exactly the same thing as `(𝟙 → 𝟘) → 𝟘` *by definition* of `¬` and `is-empty`. More examples are given in the file [negation](negation.lagda.md). + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/function-extensionality.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/function-extensionality.lagda.md new file mode 100644 index 0000000..497fc58 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/function-extensionality.lagda.md @@ -0,0 +1,28 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +# Function extensionality + +Recall that we defined pointwise equality `f ∼ g` of functions in the [identity type handout](identity-type.lagda.md). +The principle of function extensionality says that pointwise equal functions are equal and is given by the following type `FunExt`: +```agda +open import identity-type + +FunExt = {A : Type} {B : A → Type} {f g : (x : A) → B x} → f ∼ g → f ≡ g +``` +Unfortunately, this principle is not provable or disprovable in Agda or MLTT (we say that it is [independent](https://en.wikipedia.org/wiki/Independence_(mathematical_logic))). +But it is provable in [Cubical Agda](https://agda.readthedocs.io/en/latest/language/cubical.html), which is based on Cubical Type Theory and is outside the scope of these lecture notes. Sometimes we will use function extensionality as an explicit assumption. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/general-notation.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/general-notation.lagda.md new file mode 100644 index 0000000..9da3955 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/general-notation.lagda.md @@ -0,0 +1,43 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +# General notation + +The option `--without-K` will be explained later (it concerns the identity type). The option `--safe` disables experimental features of Agda that may make it *inconsistent*, that is, allow to prove false statements such as `0 = 1`. + +```agda +{-# OPTIONS --without-K --safe #-} +``` + +Any Agda file is a module with the same name as the file. + +```agda +module general-notation where +``` + +Agda calls types `sets` by default, but we prefer to refer to them by their traditional name. + +```agda +Type = Set +Type₁ = Set₁ +``` +The following functions allow us to extract the type of an element of a type, and the domain and codomain of a function. We don't need them very often, but they are handy when Agda can't infer types from the context, or simply when the names of the types in question are not available in the scope of a definition: +```agda +type-of : {X : Type} → X → Type +type-of {X} x = X + +domain : {X Y : Type} → (X → Y) → Type +domain {X} {Y} f = X + +codomain : {X Y : Type} → (X → Y) → Type +codomain {X} {Y} f = Y + +_$_ : {A B : Type} → (A → B) → A → B +f $ x = f x + +infixr 0 _$_ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/identity-type.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/identity-type.lagda.md new file mode 100644 index 0000000..7fefbc7 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/identity-type.lagda.md @@ -0,0 +1,125 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The identity type former `_≡_` + +The original and main terminology for the following type is *identity type*, but sometimes it is also called the *equality type*. Sometimes this is also called *propositional equality*, but we will avoid this terminology as it sometimes leads to confusion. +```agda +data _≡_ {A : Type} : A → A → Type where + refl : (x : A) → x ≡ x + +infix 0 _≡_ +``` + +## Elimination principle + +The elimination principle for this type, also called `J` in the literature, is defined as follows: +```agda +≡-elim : {X : Type} (A : (x y : X) → x ≡ y → Type) + → ((x : X) → A x x (refl x)) + → (x y : X) (p : x ≡ y) → A x y p +≡-elim A f x x (refl x) = f x +``` +This says that in order to show that `A x y p` holds for all `x y : X` and `p : x ≡ y`, it is enough to show that `A x x (refl x)` holds for all `x : X`. +In the literature, this elimination principle is called `J`. Again, we are not going to use it explicitly, because we can use definitions by pattern matching on `refl`, just as we did for defining it. + +We also have the non-dependent version of the eliminator: +```agda +≡-nondep-elim : {X : Type} (A : X → X → Type) + → ((x : X) → A x x) + → (x y : X) → x ≡ y → A x y +≡-nondep-elim A = ≡-elim (λ x y _ → A x y) +``` +A property of two variables like `A` above is referred to as a *relation*. The assumption `(x : X) → A x x` says that the relation is reflexive. Then the non-dependent version of the principle says that the reflexive relation given by the identity type `_≡_` can always be mapped to any reflexive relation, or we may say that `_≡_` is the smallest reflexive relation on the type `X`. + +## Fundamental constructions with the identity type + +As an exercise, you may try to rewrite the following definitions to use `≡-nondep-elim` instead of pattern matching on `refl`: +```agda +trans : {A : Type} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +trans p (refl y) = p + +sym : {A : Type} {x y : A} → x ≡ y → y ≡ x +sym (refl x) = refl x + +ap : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f (refl x) = refl (f x) + +ap₂ : {A B C : Type} (f : A → B → C) {x x' : A} {y y' : B} + → x ≡ x' → y ≡ y' → f x y ≡ f x' y' +ap₂ f (refl x) (refl y) = refl (f x y) + +transport : {X : Type} (A : X → Type) + → {x y : X} → x ≡ y → A x → A y +transport A (refl x) a = a +``` +We have already seen the first three. In the literature, `ap` is often called `cong`. In logical terms, the last one, often called `subst` in the literature, says that if `x` is equal `y` and `A x` holds, then so does `A y`. That is, we can substitute equals for equals in logical statements. + +## HoTT/UF notation + +```agda +_∙_ : {A : Type} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +_∙_ = trans + +infixl 7 _∙_ + +_⁻¹ : {A : Type} {x y : A} → x ≡ y → y ≡ x +_⁻¹ = sym + +infix 40 _⁻¹ + +``` + + +## Pointwise equality of functions + +It is often convenient to work with *pointwise equality* of functions, defined as follows: +```agda +_∼_ : {A : Type} {B : A → Type} → ((x : A) → B x) → ((x : A) → B x) → Type +f ∼ g = ∀ x → f x ≡ g x + +infix 0 _∼_ +``` + +Unfortunately, it is not provable or disprovable in Agda that pointwise equal functions are equal, that is, that `f ∼ g` implies `f ≡ g` (but it is provable in [Cubical Agda](https://agda.readthedocs.io/en/latest/language/cubical.html), which is outside the scope of these lecture notes). This principle is very useful and is called [function extensionality](function-extensionality.lagda.md). + +## Notation for equality reasoning + +When writing `trans p q` we lose type information of the +identifications `p : x ≡ y` and `q : y ≡ z`, which makes some definitions using `trans` hard to read. We now +introduce notation to be able to write e.g. + + > `x ≡⟨ p ⟩` + + > `y ≡⟨ q ⟩` + + > `z ≡⟨ r ⟩` + + > `t ∎` + +rather than the more unreadable `trans p (trans q r) : x ≡ t`. + +```agda +_≡⟨_⟩_ : {X : Type} (x : X) {y z : X} → x ≡ y → y ≡ z → x ≡ z +x ≡⟨ p ⟩ q = trans p q + +_∎ : {X : Type} (x : X) → x ≡ x +x ∎ = refl x + +infixr 0 _≡⟨_⟩_ +infix 1 _∎ +``` +We'll see examples of uses of this in other handouts. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/index.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/index.lagda.md new file mode 100644 index 0000000..f267ad5 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/index.lagda.md @@ -0,0 +1,42 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +```agda +{-# OPTIONS --without-K --safe #-} + +module index where + +import Bool +import Bool-functions +import Fin +import Fin-functions +import Hedbergs-Theorem +import List +import List-functions +import Vector +import Vector-functions +import binary-products +import binary-products-as-products +import binary-sums +import binary-sums-as-sums +import binary-type +import decidability +import empty-type +import function-extensionality +import general-notation +import identity-type +import isomorphism-functions +import isomorphisms +import natural-numbers-functions +import natural-numbers-type +import negation +import prelude +import products +import searchability +import sums +import sums-equality +import unit-type +-- import vector-and-list-isomorphisms +``` diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/introduction.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/introduction.lagda.md new file mode 100644 index 0000000..81a6266 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/introduction.lagda.md @@ -0,0 +1,233 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +# Introduction to Agda + +Everything defined and briefly discussed in this introduction will be redefined and discussed in more detail in other handouts. + +## Initial examples of types in Agda + + + +Here are some examples of types: + + +```agda +data Bool : Type where + true false : Bool + +data ℕ : Type where + zero : ℕ + suc : ℕ → ℕ + +{-# BUILTIN NATURAL ℕ #-} + +data List (A : Type) : Type where + [] : List A + _::_ : A → List A → List A + +infixr 10 _::_ +``` + +The pragma `BUILTIN NATURAL` is to get syntax sugar to be able to write 0,1,2,3,... rather than the more verbose + + * zero + * suc zero + * suc (suc zero) + * suc (suc (suc zero)) + * ⋯ + +We pronounce `suc` as "successor". + +## Examples of definitions using the above types + +```agda +not : Bool → Bool +not true = false +not false = true + +_&&_ : Bool → Bool → Bool +true && y = y +false && y = false + +_||_ : Bool → Bool → Bool +true || y = true +false || y = y + +infixr 20 _||_ +infixr 30 _&&_ + +if_then_else_ : {A : Type} → Bool → A → A → A +if true then x else y = x +if false then x else y = y + +_+_ : ℕ → ℕ → ℕ +zero + y = y +suc x + y = suc (x + y) + +_*_ : ℕ → ℕ → ℕ +zero * y = 0 +suc x * y = x * y + y + +infixr 20 _+_ +infixr 30 _*_ + +sample-list₀ : List ℕ +sample-list₀ = 1 :: 2 :: 3 :: [] + +sample-list₁ : List Bool +sample-list₁ = true || false && true :: false :: true :: true :: [] + +length : {A : Type} → List A → ℕ +length [] = 0 +length (x :: xs) = 1 + length xs + +_++_ : {A : Type} → List A → List A → List A +[] ++ ys = ys +(x :: xs) ++ ys = x :: (xs ++ ys) + +infixr 20 _++_ +``` + +## More sophisticated examples of types in Agda + +Sometimes we may wish to consider lists over a type `A` of a given length `n : ℕ`. The elements of this type, written `Vector A n`, are called *vectors*, and the type can be defined as follows: + +```agda +data Vector (A : Type) : ℕ → Type where + [] : Vector A 0 + _::_ : {n : ℕ} → A → Vector A n → Vector A (suc n) +``` +This is called a *dependent type* because it is a type that depends on *elements* `n` of another type, namely `ℕ`. + +In Agda, we can't define the `head` and `tail` functions on lists, because they are undefined for the empty list, and functions must be total in Agda. Vectors solve this problem: + +```agda +head : {A : Type} {n : ℕ} → Vector A (suc n) → A +head (x :: xs) = x + +tail : {A : Type} {n : ℕ} → Vector A (suc n) → Vector A n +tail (x :: xs) = xs +``` +Agda accepts the above definitions because it knows that the input vector has at least one element, and hence does have a head and a tail. Here is another example. + +Dependent types are pervasive in Agda. + +## The empty type and the unit type + +A type with no elements can be defined as follows: +```agda +data 𝟘 : Type where +``` +We will also need the type with precisely one element, which we define as follows: +```agda +data 𝟙 : Type where + ⋆ : 𝟙 +``` + +Here is an example of a dependent type defined using the above types: +```agda +_≣_ : ℕ → ℕ → Type +0 ≣ 0 = 𝟙 +0 ≣ suc y = 𝟘 +suc x ≣ 0 = 𝟘 +suc x ≣ suc y = x ≣ y + +infix 0 _≣_ +``` +The idea of the above definition is that `x ≣ y` is a type which either has precisely one element, if `x` and `y` are the same natural number, or else is empty, if `x` and `y` are different. +The following definition says that for any natural number `x` we can find an element of the type `x ≣ x`. +```agda +ℕ-refl : (x : ℕ) → x ≣ x +ℕ-refl 0 = ⋆ +ℕ-refl (suc x) = ℕ-refl x +``` +## The identity type former `_≡_` + +It is possible to generalize the above definition +for any type in place of that of natural numbers as follows: +```agda +data _≡_ {A : Type} : A → A → Type where + refl : (x : A) → x ≡ x + +infix 0 _≡_ +``` +Here are some functions we can define with the identity type: +```agda +trans : {A : Type} {x y z : A} → x ≡ y → y ≡ z → x ≡ z +trans (refl x) (refl x) = refl x + +sym : {A : Type} {x y : A} → x ≡ y → y ≡ x +sym (refl x) = refl x + +ap : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y +ap f (refl x) = refl (f x) +``` + +The identity type is a little bit subtle and there is a lot to say about it. +For the moment, let's convince ourselves that we can convert back and forth between the types `x ≡ y` and `x ≣ y`, in the case that `A` is the type of natural numbers, as follows: + +```agda +back : (x y : ℕ) → x ≡ y → x ≣ y +back x x (refl x) = ℕ-refl x + +forth : (x y : ℕ) → x ≣ y → x ≡ y +forth 0 0 ⋆ = refl 0 +forth (suc x) (suc y) p = I + where + IH : x ≡ y + IH = forth x y p + + I : suc x ≡ suc y + I = ap suc IH +``` + +## Propositions as types + +The [CurryHoward interpretation of logic](https://en.wikipedia.org/wiki/Curry%E2%80%93Howard_correspondence), after [Haskell Curry](https://en.wikipedia.org/wiki/Haskell_Curry) and [William Howard](https://en.wikipedia.org/wiki/William_Alvin_Howard), interprets logical statements, also known as propositions, as *types*. [Per Martin-Löf](https://en.wikipedia.org/wiki/Per_Martin-L%C3%B6f) extended this interpretation of propositions as types with equality, by introducing the identity type discussed above. + +An incomplete table of the CurryHoward--Martin-Löf interpretation of logical propositions is the following: + +| Proposition | Type | +| - | --- | +| A implies B | function type A → B | +| ∀ x : A, B x | dependent function type (x : A) → B x | +| equality | identity type `_≡_` | + +This was enough for our examples above. + +For more complex examples of reasoning about programs, we need to complete the following table: + +| Logic | English | Type | +| - | --- | --- | +| ¬ A | not A | ? | +| A ∧ B | A and B | ? | +| A ∨ B | A or B | ? | +| A → B | A implies B | A → B | +| ∀ x : A, B x | for all x:A, B x | (x : A) → B x | +| ∃ x : A, B x | there is x:A such that B x | ? | +| x = y | x equals y | x ≡ y | + +This will be the subject of future handouts. + +## Agda manual + +Please study the [Agda manual](https://agda.readthedocs.io/en/latest/) as a complement to these lecture notes. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-functions.lagda.md new file mode 100644 index 0000000..f4ab5a3 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-functions.lagda.md @@ -0,0 +1,49 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Some constructions with isomorphisms + +```agda +≃-refl : (X : Type) → X ≅ X +≃-refl X = Isomorphism id (Inverse id refl refl) + +≅-sym : {X Y : Type} → X ≅ Y → Y ≅ X +≅-sym (Isomorphism f (Inverse g η ε)) = Isomorphism g (Inverse f ε η) + +≅-trans : {X Y Z : Type} → X ≅ Y → Y ≅ Z → X ≅ Z +≅-trans (Isomorphism f (Inverse g η ε)) + (Isomorphism f' (Inverse g' η' ε')) + = Isomorphism (f' ∘ f) + (Inverse (g ∘ g') + (λ x → g (g' (f' (f x))) ≡⟨ ap g (η' (f x)) ⟩ + g (f x) ≡⟨ η x ⟩ + x ∎) + (λ y → f' (f (g (g' y))) ≡⟨ ap f' (ε (g' y)) ⟩ + f' (g' y) ≡⟨ ε' y ⟩ + y ∎)) +``` + +Notation for chains of isomorphisms (like chains of equalities): + +```agda +_≃⟨_⟩_ : (X {Y} {Z} : Type) → X ≅ Y → Y ≅ Z → X ≅ Z +_ ≃⟨ d ⟩ e = ≅-trans d e + +_■ : (X : Type) → X ≅ X +_■ = ≃-refl + +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-template.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-template.lagda.md new file mode 100644 index 0000000..a5d68c5 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphism-template.lagda.md @@ -0,0 +1,17 @@ +iso : ? ≅ ? +iso = record { bijection = f ; bijectivity = f-is-bijection } + where + f : ? → ? + f = ? + + g : ? → ? + g = ? + + gf : g ∘ f ∼ id + gf = ? + + fg : f ∘ g ∼ id + fg = ? + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/isomorphisms.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphisms.lagda.md new file mode 100644 index 0000000..5054a1f --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/isomorphisms.lagda.md @@ -0,0 +1,115 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Type isomorphisms + +A function `f : A → B` is called a *bijection* if there is a function `g : B → A` in the opposite direction such that `g ∘ f ∼ id` and `f ∘ g ∼ id`. Recall that `_∼_` is [pointwise equality](identity-type.lagda.md) and that `id` is the [identity function](products.lagda.md). This means that we can convert back and forth between the types `A` and `B` landing at the same element with started with, either from `A` or from `B`. In this case, we say that the types `A` and `B` are *isomorphic*, and we write `A ≅ B`. Bijections are also called type *isomorphisms*. We can define these concepts in Agda using [sum types](sums.lagda.md) or [records](https://agda.readthedocs.io/en/latest/language/record-types.html). We will adopt the latter, but we include both definitions for the sake of illustration. Recall that we [defined](general-notation.lagda.md) the domain of a function `f : A → B` to be `A` and its codomain to be `B`. + +Here we apply the proposition-as-types interpretation of logic to define the concepts: + +```agda + is-bijection : {A B : Type} → (A → B) → Type + is-bijection f = Σ g ꞉ (codomain f → domain f) , ((g ∘ f ∼ id) × (f ∘ g ∼ id)) + + _≅_ : Type → Type → Type + A ≅ B = Σ f ꞉ (A → B) , is-bijection f +``` +And here we give an equivalent definition which uses records and is usually more convenient in practice and is the one we adopt: +```agda +record is-bijection {A B : Type} (f : A → B) : Type where + constructor + Inverse + field + inverse : B → A + η : inverse ∘ f ∼ id + ε : f ∘ inverse ∼ id + +record _≅_ (A B : Type) : Type where + constructor + Isomorphism + field + bijection : A → B + bijectivity : is-bijection bijection + +infix 0 _≅_ +``` +The definition with `Σ` is probably more intuitive, but, as discussed above, the definition with a record is often easier to work with, because we can easily extract the components of the definitions using the names of the fields. It also often allows Agda to infer more types, and to give us more sensible goals in the interactive development of Agda programs and proofs. + +Notice that `inverse` plays the role of `g`. The reason we use `inverse` is that then we can use the word `inverse` to extract the inverse of a bijection. Similarly we use `bijection` for `f`, as we can use the word `bijection` to extract the bijection from a record. + +## Equivalences in HoTT/UF + +In HoTT/UF one uses a refinement of the notion of isomorphism defined above, not discussed in these lecture notes. + +## Some basic examples + +```agda +open import binary-type +open import Bool + +Bool-𝟚-isomorphism : Bool ≅ 𝟚 +Bool-𝟚-isomorphism = record { bijection = f ; bijectivity = f-is-bijection } + where + f : Bool → 𝟚 + f false = 𝟎 + f true = 𝟏 + + g : 𝟚 → Bool + g 𝟎 = false + g 𝟏 = true + + gf : g ∘ f ∼ id + gf true = refl true + gf false = refl false + + fg : f ∘ g ∼ id + fg 𝟎 = refl 𝟎 + fg 𝟏 = refl 𝟏 + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` +But there is also a different isomorphism: +```agda +Bool-𝟚-isomorphism' : Bool ≅ 𝟚 +Bool-𝟚-isomorphism' = record { bijection = f ; bijectivity = f-is-bijection } + where + f : Bool → 𝟚 + f false = 𝟏 + f true = 𝟎 + + g : 𝟚 → Bool + g 𝟎 = true + g 𝟏 = false + + gf : g ∘ f ∼ id + gf true = refl true + gf false = refl false + + fg : f ∘ g ∼ id + fg 𝟎 = refl 𝟎 + fg 𝟏 = refl 𝟏 + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` +And these are the only two isomorphisms (you could try to prove this in Agda as a rather advanced exercise). More advanced examples are in other files. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-functions.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-functions.lagda.md new file mode 100644 index 0000000..8d39f43 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-functions.lagda.md @@ -0,0 +1,449 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +# Natural numbers functions, relations and properties + +## Some general properties + +```agda +suc-is-not-zero : {x : ℕ} → suc x ≢ 0 +suc-is-not-zero () + +zero-is-not-suc : {x : ℕ} → 0 ≢ suc x +zero-is-not-suc () + +pred : ℕ → ℕ +pred 0 = 0 +pred (suc n) = n + +suc-is-injective : {x y : ℕ} → suc x ≡ suc y → x ≡ y +suc-is-injective = ap pred + +plus-on-paths : {m n k : ℕ} → n ≡ m → n + k ≡ m + k +plus-on-paths {_} {_} {k} = ap (_+ k) + +``` + +## Order relation _≤_ + +The less-than order on natural numbers can be defined in a number of +equivalent ways. The first one says that `x ≤ y` iff `x + z ≡ y` for +some `z`: +```agda +_≤₀_ : ℕ → ℕ → Type +x ≤₀ y = Σ a ꞉ ℕ , x + a ≡ y +``` +The second one is defined by recursion: +```agda +_≤₁_ : ℕ → ℕ → Type +0 ≤₁ y = 𝟙 +suc x ≤₁ 0 = 𝟘 +suc x ≤₁ suc y = x ≤₁ y +``` +The third one, which we will adopt as the official one, is defined *by induction* using `data`: +```agda +data _≤_ : ℕ → ℕ → Type where + 0-smallest : {y : ℕ} → 0 ≤ y + suc-preserves-≤ : {x y : ℕ} → x ≤ y → suc x ≤ suc y + +_≥_ : ℕ → ℕ → Type +x ≥ y = y ≤ x + +infix 0 _≤_ +infix 0 _≥_ +``` + +We will now show some properties of these relations. +```agda +suc-reflects-≤ : {x y : ℕ} → suc x ≤ suc y → x ≤ y +suc-reflects-≤ {x} {y} (suc-preserves-≤ l) = l + +suc-preserves-≤₀ : {x y : ℕ} → x ≤₀ y → suc x ≤₀ suc y +suc-preserves-≤₀ {x} {y} (a , p) = γ + where + q : suc (x + a) ≡ suc y + q = ap suc p + + γ : suc x ≤₀ suc y + γ = (a , q) + +≤₀-implies-≤₁ : {x y : ℕ} → x ≤₀ y → x ≤₁ y +≤₀-implies-≤₁ {zero} {y} (a , p) = ⋆ +≤₀-implies-≤₁ {suc x} {suc y} (a , p) = IH + where + q : x + a ≡ y + q = suc-is-injective p + + γ : x ≤₀ y + γ = (a , q) + + IH : x ≤₁ y + IH = ≤₀-implies-≤₁ {x} {y} γ + +≤₁-implies-≤ : {x y : ℕ} → x ≤₁ y → x ≤ y +≤₁-implies-≤ {zero} {y} l = 0-smallest +≤₁-implies-≤ {suc x} {suc y} l = γ + where + IH : x ≤ y + IH = ≤₁-implies-≤ l + + γ : suc x ≤ suc y + γ = suc-preserves-≤ IH + +≤-implies-≤₀ : {x y : ℕ} → x ≤ y → x ≤₀ y +≤-implies-≤₀ {0} {y} 0-smallest = (y , refl y) +≤-implies-≤₀ {suc x} {suc y} (suc-preserves-≤ l) = γ + where + IH : x ≤₀ y + IH = ≤-implies-≤₀ {x} {y} l + + γ : suc x ≤₀ suc y + γ = suc-preserves-≤₀ IH +``` + +## Exponential function + +```agda +_^_ : ℕ → ℕ → ℕ +y ^ 0 = 1 +y ^ suc x = y * y ^ x + +infix 40 _^_ +``` + +## Maximum and minimum + +```agda +max : ℕ → ℕ → ℕ +max 0 y = y +max (suc x) 0 = suc x +max (suc x) (suc y) = suc (max x y) + +min : ℕ → ℕ → ℕ +min 0 y = 0 +min (suc x) 0 = 0 +min (suc x) (suc y) = suc (min x y) +``` + +## No natural number is its own successor + +We now show that there is no natural number `x` such that `x = suc x`. +```agda +every-number-is-not-its-own-successor : (x : ℕ) → x ≢ suc x +every-number-is-not-its-own-successor 0 e = zero-is-not-suc e +every-number-is-not-its-own-successor (suc x) e = γ + where + IH : x ≢ suc x + IH = every-number-is-not-its-own-successor x + + e' : x ≡ suc x + e' = suc-is-injective e + + γ : 𝟘 + γ = IH e' + +there-is-no-number-which-is-its-own-successor : ¬ (Σ x ꞉ ℕ , x ≡ suc x) +there-is-no-number-which-is-its-own-successor (x , e) = every-number-is-not-its-own-successor x e +``` + +## Prime numbers + +```agda +is-prime : ℕ → Type +is-prime n = (n ≥ 2) × ((x y : ℕ) → x * y ≡ n → (x ≡ 1) ∔ (x ≡ n)) +``` +**Exercise.** Show that `is-prime n` is [decidable](decidability.lagda.md) for every `n : ℕ`. Hard. + +The following is a conjecture that so far mathematicians haven't been able to prove or disprove. But we can still say what the conjecture is in Agda: +```agda +twin-prime-conjecture : Type +twin-prime-conjecture = (n : ℕ) → Σ p ꞉ ℕ , (p ≥ n) + × is-prime p + × is-prime (p + 2) +``` + +## Properties of addition + +```agda ++-base : (x : ℕ) → x + 0 ≡ x ++-base 0 = 0 + 0 ≡⟨ refl _ ⟩ + 0 ∎ + ++-base (suc x) = suc (x + 0) ≡⟨ ap suc (+-base x) ⟩ + suc x ∎ + ++-step : (x y : ℕ) → x + suc y ≡ suc (x + y) ++-step 0 y = 0 + suc y ≡⟨ refl _ ⟩ + suc y ∎ + ++-step (suc x) y = suc x + suc y ≡⟨ refl _ ⟩ + suc (x + suc y) ≡⟨ ap suc (+-step x y) ⟩ + suc (suc (x + y)) ≡⟨ refl _ ⟩ + suc (suc x + y) ∎ + ++-comm : (x y : ℕ) → x + y ≡ y + x ++-comm 0 y = 0 + y ≡⟨ refl _ ⟩ + y ≡⟨ sym (+-base y) ⟩ + y + 0 ∎ + ++-comm (suc x) y = suc x + y ≡⟨ refl _ ⟩ + suc (x + y) ≡⟨ ap suc (+-comm x y) ⟩ + suc (y + x) ≡⟨ refl _ ⟩ + suc y + x ≡⟨ sym (+-step y x) ⟩ + y + suc x ∎ + + +plus-is-injective : {n m k : ℕ} → (n + k ≡ m + k) → n ≡ m +plus-is-injective {n} {m} {zero} p = n ≡⟨ sym (+-base n) ⟩ n + zero ≡⟨ p ⟩ m + zero ≡⟨ +-base m ⟩ m ∎ +plus-is-injective {n} {m} {suc k} p = goal + where + IH : (n + k ≡ m + k) → n ≡ m + IH = plus-is-injective {n} {m} {k} + goal : n ≡ m + goal = IH (suc-is-injective (suc (n + k) ≡⟨ sym (+-step n k) ⟩ n + suc k ≡⟨ p ⟩ m + suc k ≡⟨ +-step m k ⟩ suc(m + k) ∎)) +``` + +## Associativity of addition + +```agda ++-assoc : (x y z : ℕ) → (x + y) + z ≡ x + (y + z) ++-assoc 0 y z = refl (y + z) ++-assoc (suc x) y z = + (suc x + y) + z ≡⟨ refl _ ⟩ + suc (x + y) + z ≡⟨ refl _ ⟩ + suc ((x + y) + z) ≡⟨ ap suc (+-assoc x y z) ⟩ + suc (x + (y + z)) ≡⟨ refl _ ⟩ + suc x + (y + z) ∎ + ++-assoc' : (x y z : ℕ) → (x + y) + z ≡ x + (y + z) ++-assoc' 0 y z = refl (y + z) ++-assoc' (suc x) y z = ap suc (+-assoc' x y z) +``` + +## 1 is a neutral element of multiplication + +```agda +1-*-left-neutral : (x : ℕ) → 1 * x ≡ x +1-*-left-neutral x = refl x + +1-*-right-neutral : (x : ℕ) → x * 1 ≡ x +1-*-right-neutral 0 = refl 0 +1-*-right-neutral (suc x) = + suc x * 1 ≡⟨ refl _ ⟩ + x * 1 + 1 ≡⟨ ap (_+ 1) (1-*-right-neutral x) ⟩ + x + 1 ≡⟨ +-comm x 1 ⟩ + 1 + x ≡⟨ refl _ ⟩ + suc x ∎ +``` + +## Multiplication distributes over addition: + +```agda +*-+-distrib : (x y z : ℕ) → x * (y + z) ≡ x * y + x * z +*-+-distrib 0 y z = refl 0 +*-+-distrib (suc x) y z = γ + where + IH : x * (y + z) ≡ x * y + x * z + IH = *-+-distrib x y z + + γ : suc x * (y + z) ≡ suc x * y + suc x * z + γ = suc x * (y + z) ≡⟨ refl _ ⟩ + x * (y + z) + (y + z) ≡⟨ ap (_+ y + z) IH ⟩ + (x * y + x * z) + y + z ≡⟨ +-assoc (x * y) (x * z) (y + z) ⟩ + x * y + x * z + y + z ≡⟨ ap (x * y +_) (sym (+-assoc (x * z) y z)) ⟩ + x * y + (x * z + y) + z ≡⟨ ap (λ - → x * y + - + z) (+-comm (x * z) y) ⟩ + x * y + (y + x * z) + z ≡⟨ ap (x * y +_) (+-assoc y (x * z) z) ⟩ + x * y + y + x * z + z ≡⟨ sym (+-assoc (x * y) y (x * z + z)) ⟩ + (x * y + y) + x * z + z ≡⟨ refl _ ⟩ + suc x * y + suc x * z ∎ +``` + +## Commutativity of multiplication + +```agda +*-base : (x : ℕ) → x * 0 ≡ 0 +*-base 0 = refl 0 +*-base (suc x) = + suc x * 0 ≡⟨ refl _ ⟩ + x * 0 + 0 ≡⟨ ap (_+ 0) (*-base x) ⟩ + 0 + 0 ≡⟨ refl _ ⟩ + 0 ∎ + +*-comm : (x y : ℕ) → x * y ≡ y * x +*-comm 0 y = sym (*-base y) +*-comm (suc x) y = + suc x * y ≡⟨ refl _ ⟩ + x * y + y ≡⟨ +-comm (x * y) y ⟩ + y + x * y ≡⟨ ap (y +_) (*-comm x y) ⟩ + y + y * x ≡⟨ ap (_+ (y * x)) (sym (1-*-right-neutral y)) ⟩ + y * 1 + y * x ≡⟨ sym (*-+-distrib y 1 x) ⟩ + y * (1 + x) ≡⟨ refl _ ⟩ + y * suc x ∎ + +``` + +## Associativity of multiplication + +```agda +*-assoc : (x y z : ℕ) → (x * y) * z ≡ x * (y * z) +*-assoc zero y z = refl _ +*-assoc (suc x) y z = (x * y + y) * z ≡⟨ *-comm (x * y + y) z ⟩ + z * (x * y + y) ≡⟨ *-+-distrib z (x * y) y ⟩ + z * (x * y) + z * y ≡⟨ ap (z * x * y +_) (*-comm z y) ⟩ + z * (x * y) + y * z ≡⟨ ap (_+ y * z) (*-comm z (x * y)) ⟩ + (x * y) * z + y * z ≡⟨ ap (_+ y * z) (*-assoc x y z) ⟩ + x * y * z + y * z ∎ +``` + + +## Even and odd numbers + +```agda +is-even is-odd : ℕ → Type +is-even x = Σ y ꞉ ℕ , x ≡ 2 * y +is-odd x = Σ y ꞉ ℕ , x ≡ 1 + 2 * y + +zero-is-even : is-even 0 +zero-is-even = 0 , refl 0 + +ten-is-even : is-even 10 +ten-is-even = 5 , refl _ + +zero-is-not-odd : ¬ is-odd 0 +zero-is-not-odd () + +one-is-not-even : ¬ is-even 1 +one-is-not-even (0 , ()) +one-is-not-even (suc (suc x) , ()) + +one-is-not-even' : ¬ is-even 1 +one-is-not-even' (suc zero , ()) + +one-is-odd : is-odd 1 +one-is-odd = 0 , refl 1 + +even-gives-odd-suc : (x : ℕ) → is-even x → is-odd (suc x) +even-gives-odd-suc x (y , e) = γ + where + e' : suc x ≡ 1 + 2 * y + e' = ap suc e + + γ : is-odd (suc x) + γ = y , e' + +even-gives-odd-suc' : (x : ℕ) → is-even x → is-odd (suc x) +even-gives-odd-suc' x (y , e) = y , ap suc e + +odd-gives-even-suc : (x : ℕ) → is-odd x → is-even (suc x) +odd-gives-even-suc x (y , e) = γ + where + y' : ℕ + y' = 1 + y + + e' : suc x ≡ 2 * y' + e' = suc x ≡⟨ ap suc e ⟩ + suc (1 + 2 * y) ≡⟨ refl _ ⟩ + 2 + 2 * y ≡⟨ sym (*-+-distrib 2 1 y) ⟩ + 2 * (1 + y) ≡⟨ refl _ ⟩ + 2 * y' ∎ + + γ : is-even (suc x) + γ = y' , e' + +even-or-odd : (x : ℕ) → is-even x ∔ is-odd x +even-or-odd 0 = inl (0 , refl 0) +even-or-odd (suc x) = γ + where + IH : is-even x ∔ is-odd x + IH = even-or-odd x + + f : is-even x ∔ is-odd x → is-even (suc x) ∔ is-odd (suc x) + f (inl e) = inr (even-gives-odd-suc x e) + f (inr o) = inl (odd-gives-even-suc x o) + + γ : is-even (suc x) ∔ is-odd (suc x) + γ = f IH +``` + +```agda +even : ℕ → Bool +even 0 = true +even (suc x) = not (even x) + +even-true : (y : ℕ) → even (2 * y) ≡ true +even-true 0 = refl _ +even-true (suc y) = even (2 * suc y) ≡⟨ refl _ ⟩ + even (suc y + suc y) ≡⟨ refl _ ⟩ + even (suc (y + suc y)) ≡⟨ refl _ ⟩ + not (even (y + suc y)) ≡⟨ ap (not ∘ even) (+-step y y) ⟩ + not (even (suc (y + y))) ≡⟨ refl _ ⟩ + not (not (even (y + y))) ≡⟨ not-is-involution (even (y + y)) ⟩ + even (y + y) ≡⟨ refl _ ⟩ + even (2 * y) ≡⟨ even-true y ⟩ + true ∎ + +even-false : (y : ℕ) → even (1 + 2 * y) ≡ false +even-false 0 = refl _ +even-false (suc y) = even (1 + 2 * suc y) ≡⟨ refl _ ⟩ + even (suc (2 * suc y)) ≡⟨ refl _ ⟩ + not (even (2 * suc y)) ≡⟨ ap not (even-true (suc y)) ⟩ + not true ≡⟨ refl _ ⟩ + false ∎ + +div-by-2 : ℕ → ℕ +div-by-2 x = f (even-or-odd x) + where + f : is-even x ∔ is-odd x → ℕ + f (inl (y , _)) = y + f (inr (y , _)) = y + +even-odd-lemma : (y z : ℕ) → 1 + 2 * y ≡ 2 * z → 𝟘 +even-odd-lemma y z e = false-is-not-true impossible + where + impossible = false ≡⟨ sym (even-false y) ⟩ + even (1 + 2 * y) ≡⟨ ap even e ⟩ + even (2 * z) ≡⟨ even-true z ⟩ + true ∎ + +not-both-even-and-odd : (x : ℕ) → ¬ (is-even x × is-odd x) +not-both-even-and-odd x ((y , e) , (y' , o)) = even-odd-lemma y' y d + where + d = 1 + 2 * y' ≡⟨ sym o ⟩ + x ≡⟨ e ⟩ + 2 * y ∎ + +double : ℕ → ℕ +double 0 = 0 +double (suc x) = suc (suc (double x)) + +double-correct : (x : ℕ) → double x ≡ x + x +double-correct 0 = double 0 ≡⟨ refl _ ⟩ + 0 ≡⟨ refl _ ⟩ + 0 + 0 ∎ +double-correct (suc x) = γ + where + IH : double x ≡ x + x + IH = double-correct x + + γ : double (suc x) ≡ suc x + suc x + γ = double (suc x) ≡⟨ refl _ ⟩ + suc (suc (double x)) ≡⟨ ap (suc ∘ suc) IH ⟩ + suc (suc (x + x)) ≡⟨ ap suc (sym (+-step x x)) ⟩ + suc (x + suc x) ≡⟨ refl _ ⟩ + suc x + suc x ∎ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-type.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-type.lagda.md new file mode 100644 index 0000000..504be53 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/natural-numbers-type.lagda.md @@ -0,0 +1,83 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The type `ℕ` of natural numbers + +[Recall](introduction.lagda.md) that we defined the type of natural numbers inductively as follows: + +```notagda +data ℕ : Type where + zero : ℕ + suc : ℕ → ℕ + +{-# BUILTIN NATURAL ℕ #-} +``` + +## Elimination principle + +The elimination principle for all type formers follow the same pattern: they tell us how to define dependent functions *out* of the given type. In the case of natural numbers, the eliminator gives [primitive recursion](https://encyclopediaofmath.org/wiki/Primitive_recursion). Given a base case `a : A 0` and a step function `f : (k : ℕ) → A k → A (suc k)`, we get a function `h : (n : ℕ) → A n` defined by primitive recursion as follows: +```agda +ℕ-elim : {A : ℕ → Type} + → A 0 + → ((k : ℕ) → A k → A (suc k)) + → (n : ℕ) → A n +ℕ-elim {A} a f = h + where + h : (n : ℕ) → A n + h 0 = a + h (suc n) = f n (h n) +``` +In usual accounts of primitive recursion outside type theory, one encounters the following particular case of primitive recursion, which is the non-dependent version of the above. +```agda + +ℕ-nondep-elim : {A : Type} + → A + → (ℕ → A → A) + → ℕ → A +ℕ-nondep-elim {A} = ℕ-elim {λ _ → A} +``` +Notice that this is like `fold` for lists, if you know about this. +There is a further restricted version, which is usually called iteration: +```agda +ℕ-iteration : {A : Type} + → A + → (A → A) + → ℕ → A +ℕ-iteration a f = ℕ-nondep-elim a (λ k x → f x) +``` +Intuitively, `ℕ-iteration a f n = f (f (f (⋯ f a)))` where we apply `n` times the function `f` to the element `a`, which sometimes is written `fⁿ(a)` in the literature. + +## The induction principle for ℕ + +In logical terms, one can see immediately what the type of `ℕ-elim` is: it is simply the [principle of induction on natural numbers](https://en.wikipedia.org/wiki/Mathematical_induction), which says that in order to show that a property `A` holds for all natural numbers, it is enough to show that `A 0` holds (this is called the base case), and that if `A k` holds then so does `A (suc k)` (this is called the induction step). In Agda, in practice, we don't explicitly use this induction principle, but instead write definition recursively, just as we defined the above function `h` recursively. + +## Addition and multiplication + +As an **exercise**, you may try to define the following functions using some version of the above eliminators instead of using pattern matching and recursion: + +```agda +_+_ : ℕ → ℕ → ℕ +0 + y = y +suc x + y = suc (x + y) + +_*_ : ℕ → ℕ → ℕ +0 * y = 0 +suc x * y = x * y + y + +infixr 20 _+_ +infixr 30 _*_ +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/negation.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/negation.lagda.md new file mode 100644 index 0000000..dfd1bc6 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/negation.lagda.md @@ -0,0 +1,188 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Reasoning with negation + +[Recall that](empty-type.lagda.md) we defined the negation `¬ A` of a type `A` to be the function type `A → 𝟘`, +and that we also wrote `is-empty A` as a synonym of `¬ A`. + +## Emptiness of the empty type + +We have the following two proofs of "not false" or "the empty type is empty": +```agda +not-false : ¬ 𝟘 +not-false = 𝟘-elim + +not-false' : ¬ 𝟘 +not-false' = id +``` +A lot of things about negation don't depend on the fact that the target type of the function type is `𝟘`. We will begin by proving some things about negation by generalizing `𝟘` to any type `R` of "results". + +## Implication from disjunction and negation + +If `¬ A` or `B`, then `A implies B`: +```agda +implication-from-disjunction-and-negation : {A B : Type} → ¬ A ∔ B → (A → B) +implication-from-disjunction-and-negation (inl f) a = 𝟘-elim (f a) +implication-from-disjunction-and-negation (inr b) a = b + +implication-from-disjunction-and-negation-converse : {A B : Type} + → A ∔ ¬ A + → (A → B) → ¬ A ∔ B +implication-from-disjunction-and-negation-converse (inl x) f = inr (f x) +implication-from-disjunction-and-negation-converse (inr ν) f = inl ν +``` +In the handout on [decidability](decidability.lagda.md) we call *decidable* a type `A` such that `A ∔ ¬ A`. So the above two theorems together say that if `A` is decidable then `A → B` is logically equivalent to `¬ A ∔ B`. In classical mathematics, propositions are decidable, as this is what the principle of excluded middle says. +In MLTT the principle of excluded middle `(A : Type) → A ∔ ¬ A` is not provable or disprovable. In HoTT/UF, the principle of excluded middle is taken to be `(A : Type) → is-prop A → A ∔ ¬ A` and is also not provable or disprovable, but is validated in Voedvodsky's model of simplicial sets (assuming classical logic in the definition of this model) and hence can be consistently assumed for the purposes of doing classical mathematics in HoTT. And so is the axiom of choice - we refer the interested reader to [Introduction to Homotopy Type Theory and Univalent Foundations (HoTT/UF) with Agda](https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#contents). + +## Contrapositives + +If `A` implies `B`, then `B → R` implies `A → R`: +```agda +arrow-contravariance : {A B R : Type} + → (A → B) + → (B → R) → (A → R) +arrow-contravariance f g = g ∘ f +``` +A particular case of interest is the following. The [contrapositive](https://en.wikipedia.org/wiki/Contraposition) of an implication `A → B` is the implication `¬ B → ¬ A`: +```agda +contrapositive : {A B : Type} → (A → B) → (¬ B → ¬ A) +contrapositive {A} {B} = arrow-contravariance {A} {B} {𝟘} +``` +This can also be read as "if we have a function A → B and B is empty, then also A must be empty". + +## Double and triple negations + +We now introduce notation for double and triple negation, to reduce the number of needed brackets: + +```agda +¬¬_ ¬¬¬_ : Type → Type +¬¬ A = ¬(¬ A) +¬¬¬ A = ¬(¬¬ A) +``` +We have that `A` implies `¬¬ A`. This is called double negation introduction. More generally, we have the following: +```agda +dni : (A R : Type) → A → ((A → R) → R) +dni A R a u = u a + +¬¬-intro : {A : Type} → A → ¬¬ A +¬¬-intro {A} = dni A 𝟘 +``` +We don't always have `¬¬ A → A` in proofs-as-programs. This has to do with *computability theory*. But sometimes we do. For example, if we know that `A ∔ ¬ A` then `¬¬A → A` follows: + +```agda + ¬¬-elim : {A : Type} → A ∔ ¬ A → ¬¬ A → A + ¬¬-elim (inl x) f = x + ¬¬-elim (inr g) f = 𝟘-elim (f g) +``` +For more details, see the lecture notes on [decidability](decidability.lagda.md), where we discuss `¬¬-elim` again. +But three negations always imply one, and conversely: +```agda +three-negations-imply-one : {A : Type} → ¬¬¬ A → ¬ A +three-negations-imply-one = contrapositive ¬¬-intro + +one-negation-implies-three : {A : Type} → ¬ A → ¬¬¬ A +one-negation-implies-three = ¬¬-intro +``` + +## Negation of the identity type + +It is useful to introduce a notation for the negation of the [identity type](identity-type.lagda.md): +```agda +_≢_ : {X : Type} → X → X → Type +x ≢ y = ¬ (x ≡ y) + +≢-sym : {X : Type} {x y : X} → x ≢ y → y ≢ x +≢-sym = contrapositive sym + +false-is-not-true : false ≢ true +false-is-not-true () + +true-is-not-false : true ≢ false +true-is-not-false () +``` +The following is more interesting: +```agda +not-false-is-true : (x : Bool) → x ≢ false → x ≡ true +not-false-is-true true f = refl true +not-false-is-true false f = 𝟘-elim (f (refl false)) + +not-true-is-false : (x : Bool) → x ≢ true → x ≡ false +not-true-is-false true f = 𝟘-elim (f (refl true)) +not-true-is-false false f = refl false +``` + +## Disjointness of binary sums + +We now show something that is intuitively the case: +```agda +inl-is-not-inr : {X Y : Type} {x : X} {y : Y} → inl x ≢ inr y +inl-is-not-inr () +``` +Agda just knows it. + +## Disjunctions and negation + +If `A or B` holds and `B` is false, then `A` must hold: +```agda +right-fails-gives-left-holds : {A B : Type} → A ∔ B → ¬ B → A +right-fails-gives-left-holds (inl a) f = a +right-fails-gives-left-holds (inr b) f = 𝟘-elim (f b) + +left-fails-gives-right-holds : {A B : Type} → A ∔ B → ¬ A → B +left-fails-gives-right-holds (inl a) f = 𝟘-elim (f a) +left-fails-gives-right-holds (inr b) f = b +``` + +## Negation of the existential quantifier: + +If there is no `x : X` with `A x`, then for all `x : X` we have that not `A x`: +```agda +not-exists-implies-forall-not : {X : Type} {A : X → Type} + → ¬ (Σ x ꞉ X , A x) + → (x : X) → ¬ A x +not-exists-implies-forall-not f x a = f (x , a) +``` +The converse also holds: +```agda +forall-not-implies-not-exists : {X : Type} {A : X → Type} + → ((x : X) → ¬ A x) + → ¬ (Σ x ꞉ X , A x) +forall-not-implies-not-exists g (x , a) = g x a +``` +Notice how these are particular cases of [`curry` and `uncurry`](https://en.wikipedia.org/wiki/Currying). + +## Implication truth table + +Here is a proof of the implication truth-table: +```agda +open import empty-type +open import unit-type + +implication-truth-table : ((𝟘 → 𝟘) ⇔ 𝟙) + × ((𝟘 → 𝟙) ⇔ 𝟙) + × ((𝟙 → 𝟘) ⇔ 𝟘) + × ((𝟙 → 𝟙) ⇔ 𝟙) +implication-truth-table = ((λ _ → ⋆) , (λ _ → id)) , + ((λ _ → ⋆) , (λ _ _ → ⋆)) , + ((λ f → f ⋆) , (λ ⋆ _ → ⋆)) , + ((λ _ → ⋆) , (λ _ _ → ⋆)) +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/prelude.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/prelude.lagda.md new file mode 100644 index 0000000..1204fdf --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/prelude.lagda.md @@ -0,0 +1,25 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + +```agda +{-# OPTIONS --without-K --safe #-} + +module prelude where + +open import general-notation public +open import empty-type public +open import unit-type public +open import binary-sums public +open import binary-products public +open import sums public +open import products public +open import identity-type public +open import natural-numbers-type public +open import Bool public +open import List public +open import Vector public +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/products.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/products.lagda.md new file mode 100644 index 0000000..5d43f72 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/products.lagda.md @@ -0,0 +1,104 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + + +# Products + +We discuss function types of the form `A → B` (called non-dependent function types) and of the form `(x : A) → B x` (called dependent function types). The latter are also called *products* in type theory and this terminology comes from mathematics. + +## The identity function + +To *introduce* functions, we use `λ`-abstractions. For example, the identity function can be defined as follows: + +```agda + id : {A : Type} → A → A + id = λ x → x +``` + +But of course this function can be equivalently written as follows, which we take as our official definition: +```agda +id : {A : Type} → A → A +id x = x +``` + +## Logical implication + +A logical implication `A → B` is represented by the function type `A → B`, and it is a happy coincidence that both use the same notation. + +In terms of logic, the identity function implements the tautology "`A` implies `A`" when we understand the type `A` as a logical proposition. In general, most things we define in Agda have a dual interpretation type/proposition or program/proof, like this example. + +The *elimination* principle for function types is given by function application, which is built-in in Agda, but we can explicitly (re)define it. We do with the arguments swapped here: +```agda +modus-ponens : {A B : Type} → A → (A → B) → B +modus-ponens a f = f a +``` +[Modus ponens](https://en.wikipedia.org/wiki/Modus_ponens) is the rule logic that says that if the proposition `A` holds and `A` implies `B`, then `B` also holds. The above definition gives a computational realization of this. + +## Dependent functions + +As already discussed, a dependent function type `(x : A) → B x`, with `A : Type` and `B : A → Type`, is that of functions with input `x` in the type `A` and output in the type `B x`. It is called "dependent" because the output *type* `B x` depends on the input *element* `x : A`. + +## Universal quantification + +The logical interpretation of `(x : A) → B x` is "for all x : A, B x". +This is because in order to show that this universal quantification holds, we have to show that `B x` holds for each given `x:A`. The input is the given `x : A`, and the output is a justification of `B x`. We will see some concrete examples shortly. + +## Dependent function composition + +Composition can be defined for "non-dependent functions", as in usual mathematics, and, more generally, with dependent functions. With non-dependent functions, it has the following type and definition: + +```agda + _∘_ : {A B C : Type} → (B → C) → (A → B) → (A → C) + g ∘ f = λ x → g (f x) +``` + +In terms of computation, this means "first do f and then do g". For this reason the function composition `g ∘ f` is often read "`g` after `f`". In terms of logic, this implements "If B implies C and also A implies B, then A implies C". + +With dependent types, it has the following more general type but the same definition, which is the one we adopt: + +```agda +_∘_ : {A B : Type} {C : B → Type} + → ((y : B) → C y) + → (f : A → B) + → (x : A) → C (f x) +g ∘ f = λ x → g (f x) +``` + +Its computational interpretation is the same, "first do f and then do g", but its logical understanding changes: "If it is the case that for all y : B we have that C y holds, and we have a function f : A → B, then it is also the case that for all x : A, we have that C (f x) holds". + +## `Π` notation + +We have mentioned in the [propositions as types table](curry-howard.lagda) that the official notation in MLTT for the dependent function type uses `Π`, the Greek letter Pi, for *product*. We can, if we want, introduce the same notation in Agda, using a `syntax` declaration: +```agda +Pi : (A : Type) (B : A → Type) → Type +Pi A B = (x : A) → B x + +syntax Pi A (λ x → b) = Π x ꞉ A , b +``` +**Important.** We are using the alternative symbol `꞉` (typed "`\:4`" in the emacs mode for Agda), which is different from the reserved symbol "`:`". We will use the same alternative symbol when we define syntax for the sum type `Σ`. + +Notice that, for some reason, Agda has this kind of definition backwards. The end result of this is that now `(x : A) → B x` can be written as `Π x ꞉ A , B x` in Agda as in type theory. (If you happen to know a bit of maths, you may be familiar with the [cartesian product of a family of sets](https://en.wikipedia.org/wiki/Cartesian_product#Infinite_Cartesian_products), and this is the reason the letter `Π` is used.) + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/searchability.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/searchability.lagda.md new file mode 100644 index 0000000..d572ba4 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/searchability.lagda.md @@ -0,0 +1,151 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Examples of searchable sets + +Recall that we defined exhaustibly searchable types in the module [decidability](decidability.lagda.md). Let's adopt a shorter name for the purposes of this module. +```agda +is-searchable = is-exhaustively-searchable +``` + +## `𝟘` is searchable + +```agda +𝟘-is-searchable : is-searchable 𝟘 +𝟘-is-searchable A δ = inr (λ (x , a) → x) +``` + +## `𝟙` is searchable + +```agda +𝟙-is-searchable : is-searchable 𝟙 +𝟙-is-searchable A d = I (d ⋆) + where + I : A ⋆ ∔ ¬ A ⋆ → is-decidable (Σ x ꞉ 𝟙 , A x) + I (inl a) = inl (⋆ , a) + I (inr u) = inr (λ (⋆ , a) → u a) +``` + +## Searchable types are closed under `_∔_` + +```agda +∔-is-searchable : {X Y : Type} + → is-searchable X + → is-searchable Y + → is-searchable (X ∔ Y) +∔-is-searchable {X} {Y} c d A δ = II + where + I : is-decidable (Σ x ꞉ X , A (inl x)) + → is-decidable (Σ y ꞉ Y , A (inr y)) + → is-decidable (Σ z ꞉ X ∔ Y , A z) + I (inl (x , a)) _ = inl (inl x , a) + I (inr f) (inl (y , a)) = inl (inr y , a) + I (inr f) (inr g) = inr h + where + h : (Σ z ꞉ X ∔ Y , A z) → 𝟘 + h (inl x , a) = f (x , a) + h (inr y , a) = g (y , a) + + II : is-decidable (Σ z ꞉ X ∔ Y , A z) + II = I (c (λ x → A (inl x)) + (λ x → δ (inl x))) + (d (λ y → A (inr y)) + (λ y → δ (inr y))) +``` + +## `Fin' n` is searchable + +```agda +open import Fin-functions + +Fin'-is-searchable : (n : ℕ) → is-searchable (Fin' n) +Fin'-is-searchable 0 = 𝟘-is-searchable +Fin'-is-searchable (suc n) = ∔-is-searchable + 𝟙-is-searchable + (Fin'-is-searchable n) +``` + +## Searchable types are closed under isomorphism + +```agda +open import isomorphisms + +≅-is-searchable : {X Y : Type} + → is-searchable X + → X ≅ Y + → is-searchable Y +≅-is-searchable {X} {Y} s (Isomorphism f (Inverse g _ ε)) A d = III + where + B : X → Type + B x = A (f x) + + I : is-decidable-predicate B + I x = d (f x) + + II : is-decidable (Σ B) → is-decidable (Σ A) + II (inl (x , a)) = inl (f x , a) + II (inr u) = inr (λ (y , a) → u (g y , transport A (sym (ε y)) a)) + + III : is-decidable (Σ A) + III = II (s B I) +``` + +## `𝟚` is searchable + +```agda +open import binary-type + +𝟙∔𝟙-is-searchable : is-searchable (𝟙 ∔ 𝟙) +𝟙∔𝟙-is-searchable = ∔-is-searchable 𝟙-is-searchable 𝟙-is-searchable + +𝟙∔𝟙-is-𝟚 : 𝟙 ∔ 𝟙 ≅ 𝟚 +𝟙∔𝟙-is-𝟚 = Isomorphism f (Inverse g gf fg) + where + f : 𝟙 ∔ 𝟙 → 𝟚 + f (inl ⋆) = 𝟎 + f (inr ⋆) = 𝟏 + + g : 𝟚 → 𝟙 ∔ 𝟙 + g 𝟎 = inl ⋆ + g 𝟏 = inr ⋆ + + gf : g ∘ f ∼ id + gf (inl ⋆) = refl (inl ⋆) + gf (inr ⋆) = refl (inr ⋆) + + fg : f ∘ g ∼ id + fg 𝟎 = refl 𝟎 + fg 𝟏 = refl 𝟏 + +𝟚-is-searchable : is-searchable 𝟚 +𝟚-is-searchable = ≅-is-searchable + 𝟙∔𝟙-is-searchable + 𝟙∔𝟙-is-𝟚 +``` + +## `Fin n` is searchable + +```agda +open import Fin +open import isomorphism-functions + +Fin-is-searchable : (n : ℕ) → is-searchable (Fin n) +Fin-is-searchable n = ≅-is-searchable + (Fin'-is-searchable n) + (≅-sym (Fin-isomorphism n)) +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/sums-equality.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/sums-equality.lagda.md new file mode 100644 index 0000000..2405aa8 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/sums-equality.lagda.md @@ -0,0 +1,134 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Equality in `Σ` types + +The material here is rather subtle. + +## Equality in binary products + +Equality in binary product types behaves as expected: + +```agda +module _ {X A : Type} {(x , a) (y , b) : X × A} where + + from-×-≡ : (x , a) ≡ (y , b) + → (x ≡ y) × (a ≡ b) + from-×-≡ (refl (x , a)) = (refl x , refl a) + + to-×-≡ : (x ≡ y) × (a ≡ b) + → (x , a) ≡ (y , b) + to-×-≡ (refl x , refl a) = refl (x , a) + + ×-≡ : (x , a) ≡ (y , b) ⇔ (x ≡ y) × (a ≡ b) + ×-≡ = from-×-≡ , to-×-≡ +``` + +## Equality in `Σ` types + +Equality in `Σ` types is much less intuitive. + +Firstly, if we have a type `X`, a family `A : X → Type`, and pairs `(x , a)` and `(y , b)` in the type `Σ +A`, it is the case that if `(x , a) ≡ (y , b)` then `x ≡ y`, just as in the case of binary products: + +```agda +Σ-≡-fst : {X : Type} {A : X → Type} {(x , a) (y , b) : Σ A} + → (x , a) ≡ (y , b) → x ≡ y +Σ-≡-fst (refl (x , a)) = refl x +``` + +However, we don't have that `(x , a) ≡ (y , b)` implies `a ≡ b`. It is not that this is false, but instead that it doesn't even type-check. In fact we have + + * `a : A x` + * `b : A y` + +and the types `A x` and `A y` are not necessarily the *same*. But +equality is only defined for elements of the same type. Notice that `x +≡ y` doesn't say that `x` and `y` are the *same*. It only says that +there is an *identification* between them. And the distinction between +sameness and identification is important. + +We can use the function `transport` (defined in the module +[identity-type](identity-type.lagda)) + +```agda-repetition + transport : {X : Type} (A : X → Type) + → {x y : X} → x ≡ y → A x → A y + transport A (refl x) a = a +``` + +to solve this problem: + +```agda +Σ-≡-snd : {X : Type} {A : X → Type} {(x , a) (y , b) : Σ A} + → (e : (x , a) ≡ (y , b)) + → transport A (Σ-≡-fst e) a ≡ b +Σ-≡-snd (refl (x , a)) = refl a +``` + +We can pack these two functions into a single one as follows: + +```agda +from-Σ-≡ : {X : Type} {A : X → Type} {(x , a) (y , b) : Σ A} + → (x , a) ≡ (y , b) + → Σ e ꞉ x ≡ y , transport A e a ≡ b +from-Σ-≡ (refl (x , a)) = (refl x , refl a) +``` +This says that from `(x , a) ≡ (y , b)` we can get some `e : x ≡ y` such that `transport A e a ≡ b`. + + +The converse also holds: +```agda +to-Σ-≡ : {X : Type} {A : X → Type} {(x , a) (y , b) : Σ A} + → (Σ e ꞉ x ≡ y , transport A e a ≡ b) + → (x , a) ≡ (y , b) +to-Σ-≡ (refl x , refl a) = refl (x , a) +``` + +So the functions `from-Σ-≡` and `to-Σ-≡` give a complete characterization of equality in `Σ` types. + +```agda +Σ-≡ : {X : Type} {A : X → Type} {(x , a) (y , b) : Σ A} + → (x , a) ≡ (y , b) ⇔ (Σ e ꞉ x ≡ y , transport A e a ≡ b) +Σ-≡ = from-Σ-≡ , to-Σ-≡ +``` + +*Exercise.* The function `from-Σ-≡` is actually an isomorphism with + inverse `to-Σ-≡`, so that we actually get the stronger result + + > `((x , a) ≡ (y , b)) ≅ (Σ e ꞉ x ≡ y , transport A e a ≡ b)`. + +## Dependent version of `ap` + +Recall the function `ap` defined in the module +[identity-type](identity-type.lagda): + +```agda-repetition + ap : {A B : Type} (f : A → B) {x y : A} → x ≡ y → f x ≡ f y + ap f (refl x) = refl (f x) +``` + +A similar phenomenon as for `Σ` happens with dependent functions `f : (x : X) → A x`. We can't conclude that if `x ≡ y` then `f x ≡ f y`, because, again, the equality in the conclusion doesn't type check, as `f x` and `f y` belong to the types `A x` and `A y`, which are not the same in general, even if `x` and `y` are identified. + +But again we can use `transport` to solve this problem: +```agda +apd : {X : Type} {A : X → Type} (f : (x : X) → A x) {x y : X} + (e : x ≡ y) → transport A e (f x) ≡ f y +apd f (refl x) = refl (f x) +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/sums.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/sums.lagda.md new file mode 100644 index 0000000..78ade3d --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/sums.lagda.md @@ -0,0 +1,104 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The sum type former `Σ` + +Very often in computation we work with the type of pairs `(x , y)` with `x : A` and `y : B` where `A` and `B` are types. We will write `A × B` for the type of such pairs, and call it the [*cartesian product*](https://en.wikipedia.org/wiki/Cartesian_product). We will define this type as a particular case of a more general type, whose elements are again of the form `(x , y)` but with the difference that `x : A` and `y : B x` where `A : Type` and `B : A → Type`. The difference amounts to the fact that the type of the second component `y` depends on the first element `x`. The default notation for this type will be `Σ {A} B`, or simply `Σ B` when `A` can be inferred from the context, but we will also introduce the more common sum notation `Σ x ꞉ A , B x`. This type is also called the [disjoint union](https://en.wikipedia.org/wiki/Disjoint_union) of the type family `B : A → Type`. + +## Examples + +A simple example is the type `Σ xs ꞉ List X , Vector X (length xs)` with `X : Type`. An element of this type is a pair `(xs , ys)` where `xs` is a list and `ys` is a vector of the same length as `xs`. + +Another example, which iterates the `Σ` type construction, is `Σ x : ℕ , Σ y : ℕ , Σ z : ℕ , x ≡ y * z`. An element of this type is now a quadruple `(x , (y , (z , p)))` where `x y z : ℕ` and `p : x ≡ y * z`. That is, the fourth component ensures that `x y z : ℕ` are allowed in the tuple if, and only if, `x ≡ y * z`. We will see more interesting examples shortly. + +## Definition + +The `Σ` type can be defined in Agda using a `data` declaration as follows: + +```agda + data Σ {A : Type } (B : A → Type) : Type where + _,_ : (x : A) (y : B x) → Σ {A} B + + pr₁ : {A : Type} {B : A → Type} → Σ B → A + pr₁ (x , y) = x + + pr₂ : {A : Type} {B : A → Type} → (z : Σ B) → B (pr₁ z) + pr₂ (x , y) = y + ``` +Notice that the type of `pr₂` is dependent and uses `pr₁` to express the dependency. + +However, for a number of reasons to be explained later, we prefer to define it using a [record](https://agda.readthedocs.io/en/latest/language/record-types.html) definition: + +```agda +record Σ {A : Type } (B : A → Type) : Type where + constructor + _,_ + field + pr₁ : A + pr₂ : B pr₁ +``` + +Here we automatically get the projections with the same types and definitions as above and hence we don't need to provide them. In order for the projections `pr₁` and `pr₂` to be visible outside the scope of the record, we `open` the record. Moreover, we open it `public` so that when other files import this one, these two projections will be visible in the other files. The "constructor" allows to form an element of this type. Because "," is not a reserved symbol in Agda, we can use it as a binary operator to write `x , y`. However, following mathematical tradition, we will write brackets around that, to get `(x , y)`, even if this is not necessary. We also declare a fixity and precedence for this operator. + +```agda +open Σ public +infixr 0 _,_ +``` +Because we make `_,_` right associative, we can write `(x , y , z , p)` rather than `(x , (y , (z , p)))` as we did above. + +We also use a syntax declaration, [as we did](products.lagda.md) for dependent function types using Π, to get the more traditional type-theoretical notation. +```agda +Sigma : (A : Type) (B : A → Type) → Type +Sigma A B = Σ {A} B + +syntax Sigma A (λ x → b) = Σ x ꞉ A , b + +infix -1 Sigma +``` + +## Elimination principle + +We now define and discuss the elimination principle. +```agda +Σ-elim : {A : Type } {B : A → Type} {C : (Σ x ꞉ A , B x) → Type} + → ((x : A) (y : B x) → C (x , y)) + → (z : Σ x ꞉ A , B x) → C z +Σ-elim f (x , y) = f x y +``` +So the elimination principle for `Σ` is what was called `curry` in functional programming. The logical interpretation for this principle is that in order to show that "for all z : Σ x ꞉ A , B x) we have that C z holds", it is enough to show that "for all x : A and y : B x we have that C (x , y) holds". This condition is not only sufficient but also [necessary](https://en.wikipedia.org/wiki/Necessity_and_sufficiency): +```agda +Σ-uncurry : {A : Type } {B : A → Type} {C : (Σ x ꞉ A , B x) → Type} + → ((z : Σ x ꞉ A , B x) → C z) + → (x : A) (y : B x) → C (x , y) +Σ-uncurry g x y = g (x , y) +``` + +## Existential quantification + +Regarding logic, the `Σ` type is used to interpret the existential quantifier `∃`. The logical proposition `∃ x : X, A x`, that is, "there exists x : X such that A x", is interpreted as the type `Σ x ꞉ X , A x`. The reason is that to show that `∃ x : X, A x` we have to exhibit an example `x : X` and show that `x` satisfies the condition `A x` with some `y : A x`, in a pair `(x , y)`. + +For example, the type `Σ x : ℕ , Σ y : ℕ , Σ z : ℕ , x ≡ y * z` can be interpreted as saying that "there are natural numbers x, y, and z such that x = y * z", which is true as witnessed by the element `(6,2,3,refl 6)` of that type. But there are many other witnesses of this type, of course, such as `(10,5,2,refl 10)`. + +It is important to notice that it is possible to write types that correspond to false logical statements, and hence are empty. For example, consider `Σ x : ℕ , x ≡ x + 1`. There is no natural number that is its own successor, of course, and so this type is empty. While this type is empty, the type `¬ (Σ x : ℕ , x ≡ x + 1)` has an element, as we will see, which witnesses the fact that "there doesn't exist a natural number `x` such that `x = x + 1`". + +## Existential quantification in HoTT/UF + +In HoTT/UF it useful to have an alternative existential quantifier `∃ x : X , A x` defined to be `∥ Σ x : X , A x ∥` where `∥_∥` is a certain *propositional truncation* operation. + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/unit-type.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/unit-type.lagda.md new file mode 100644 index 0000000..415b231 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/unit-type.lagda.md @@ -0,0 +1,41 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# The unit type 𝟙 + +We now redefine the unit type as a record: +```agda +record 𝟙 : Type where + constructor + ⋆ + +open 𝟙 public +``` +In logical terms, we can interpret `𝟙` as "true", because we can always exhibit an element of it, namely `⋆`. +Its elimination principle is as follows: +```agda +𝟙-elim : {A : 𝟙 → Type} + → A ⋆ + → (x : 𝟙) → A x +𝟙-elim a ⋆ = a +``` +In logical terms, this says that it order to prove that a property `A` of elements of the unit type `𝟙` holds for all elements of the type `𝟙`, it is enough to prove that it holds for the element `⋆`. The non-dependent version says that if A holds, then "true implies A". +```agda +𝟙-nondep-elim : {A : Type} + → A + → 𝟙 → A +𝟙-nondep-elim {A} = 𝟙-elim {λ _ → A} +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Lecture-Notes/files/vector-and-list-isomorphisms.lagda.md b/src/HoTTEST/Agda/Lecture-Notes/files/vector-and-list-isomorphisms.lagda.md new file mode 100644 index 0000000..0adeac7 --- /dev/null +++ b/src/HoTTEST/Agda/Lecture-Notes/files/vector-and-list-isomorphisms.lagda.md @@ -0,0 +1,125 @@ + +[Martin Escardo](Https://www.Cs.Bham.Ac.Uk/~mhe/). +Notes originally written for the module *Advanced Functional Programming* of the [University of Birmingham](https://www.birmingham.ac.uk/index.aspx), UK. + + + +# Vector and list isomorphisms + +There are deliberate gaps in this file for you to fill. + +## The type of lists can be defined from that of vectors + +```agda +open import isomorphisms + +lists-from-vectors : {A : Type} → List A ≅ (Σ n ꞉ ℕ , Vector A n) +lists-from-vectors {A} = record { bijection = f ; bijectivity = f-is-bijection } + where + f : List A → Σ n ꞉ ℕ , Vector A n + f = {!!} + + g : (Σ n ꞉ ℕ , Vector A n) → List A + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +## The type of vectors can be defined from that of lists + +```agda +open import List-functions + +vectors-from-lists : {A : Type} (n : ℕ) → Vector A n ≅ (Σ xs ꞉ List A , (length xs ≡ n)) +vectors-from-lists {A} n = record { bijection = f ; bijectivity = f-is-bijection } + where + f : {!!} → {!!} + f = {!!} + + g : {!!} → {!!} + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +## The types of lists and vectors can be defined in basic MLTT + +```agda +Vector' : (A : Type) → ℕ → Type +Vector' A 0 = 𝟙 +Vector' A (suc n) = A × Vector' A n + +[]' : {A : Type} → Vector' A 0 +[]' = ⋆ + +_::'_ : {A : Type} {n : ℕ} → A → Vector' A n → Vector' A (suc n) +x ::' xs = x , xs + +List' : Type → Type +List' X = Σ n ꞉ ℕ , Vector' X n + +``` + +```agda +vectors-in-basic-MLTT : {A : Type} (n : ℕ) → Vector A n ≅ Vector' A n +vectors-in-basic-MLTT {A} n = record { bijection = f ; bijectivity = f-is-bijection } + where + f : {!!} → {!!} + f = {!!} + + g : {!!} → {!!} + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +``` +lists-in-basic-MLTT : {A : Type} → List A ≅ List' A +lists-in-basic-MLTT {A} = record { bijection = f ; bijectivity = f-is-bijection } + where + f : {!!} → {!!} + f = {!!} + + g : {!!} → {!!} + g = {!!} + + gf : g ∘ f ∼ id + gf = {!!} + + fg : f ∘ g ∼ id + fg = {!!} + + f-is-bijection : is-bijection f + f-is-bijection = record { inverse = g ; η = gf ; ε = fg } +``` + +[Go back to the table of contents](https://martinescardo.github.io/HoTTEST-Summer-School/) diff --git a/src/HoTTEST/Agda/Q-and-A/README.md b/src/HoTTEST/Agda/Q-and-A/README.md new file mode 100644 index 0000000..a2e7411 --- /dev/null +++ b/src/HoTTEST/Agda/Q-and-A/README.md @@ -0,0 +1,5 @@ +# Q & A + +This folder will contain PDFs of the Lecture Q&A +for the Agda track of the HoTTEST Summer School +2022. \ No newline at end of file diff --git a/src/HoTTEST/Agda/Q-and-A/qa1.pdf b/src/HoTTEST/Agda/Q-and-A/qa1.pdf new file mode 100644 index 0000000..bee9245 Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa1.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa3.pdf b/src/HoTTEST/Agda/Q-and-A/qa3.pdf new file mode 100644 index 0000000..78036f4 Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa3.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa4.pdf b/src/HoTTEST/Agda/Q-and-A/qa4.pdf new file mode 100644 index 0000000..0a57bd8 Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa4.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa5.pdf b/src/HoTTEST/Agda/Q-and-A/qa5.pdf new file mode 100644 index 0000000..118f310 Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa5.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa6.pdf b/src/HoTTEST/Agda/Q-and-A/qa6.pdf new file mode 100644 index 0000000..a75d0df Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa6.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa7.pdf b/src/HoTTEST/Agda/Q-and-A/qa7.pdf new file mode 100644 index 0000000..5b9ac5d Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa7.pdf differ diff --git a/src/HoTTEST/Agda/Q-and-A/qa9.pdf b/src/HoTTEST/Agda/Q-and-A/qa9.pdf new file mode 100644 index 0000000..deda23b Binary files /dev/null and b/src/HoTTEST/Agda/Q-and-A/qa9.pdf differ diff --git a/src/HoTTEST/Agda/README.md b/src/HoTTEST/Agda/README.md new file mode 100644 index 0000000..e8cef37 --- /dev/null +++ b/src/HoTTEST/Agda/README.md @@ -0,0 +1,41 @@ +## [Preferred link to see these lecture notes](https://martinescardo.github.io/HoTTEST-Summer-School/) + +## [Download the Agda files as a zip file](https://github.com/martinescardo/HoTTEST-Summer-School/archive/refs/heads/main.zip) + +## [How to install Agda 2.6.2.2](./Agda-installation.md) + +## [Agda manual](https://agda.readthedocs.io/en/latest/) + +## Lecture Notes + - [Lectures 1─3](./Lecture-Notes/README.md) delivered by Martín Escardó + 1. [Introduction to Agda, dependent types and dependent + functions](./Lecture-Log/lecture1.lagda.md) covers + (part of) the [introduction to + Agda](./Lecture-Notes/files/introduction.lagda.md) notes. + 1. [Basic MLTT Types](./Lecture-Log/lecture2.lagda.md) covers the + [empty type](./Lecture-Notes/files/empty-type.lagda.md), + [unit type](./Lecture-Notes/files/unit-type.lagda.md), + [binary type](./Lecture-Notes/files/binary-type.lagda.md), + [sum types](./Lecture-Notes/files/sums.lagda.md), + [binary sums](./Lecture-Notes/files/binary-sums.lagda.md) and + [propositions-as-types](./Lecture-Notes/files/curry-howard.lagda.md) notes. + 1. [Identity types, Σ-types and universes](./Lecture-Log/lecture3.lagda.md) + covers the + [binary sums](./Lecture-Notes/files/binary-sums.lagda.md), + [the identity type](./Lecture-Notes/files/identity-type.lagda.md) and + [equality in sums](./Lecture-Notes/files/sums-equality.lagda.md) notes. + - Lectures 4─6 delivered by Dan Licata + + 4. [Intro to higher inductive types](./HITs/Lecture4-notes.lagda.md) + 4. [Dependent elimination principles for HITs](./HITs/Lecture5-notes.lagda.md) + 4. [Calculating the fundamental group of the circle](./HITs/Lecture6-notes.lagda.md) + +## [Exercise Sheets](Exercises/README.md) + +## [Auxiliary files](Auxiliary/README.md) + +## Advanced further reading + + 1. [Introduction to Homotopy Type Theory and Univalent Foundations (HoTT/UF) with Agda](https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#contents). + +## [HoTTEST School github repository](https://github.com/martinescardo/HoTTEST-Summer-School) diff --git a/src/HoTTEST/Agda/agda-lectures.agda-lib b/src/HoTTEST/Agda/agda-lectures.agda-lib new file mode 100644 index 0000000..e23d982 --- /dev/null +++ b/src/HoTTEST/Agda/agda-lectures.agda-lib @@ -0,0 +1,9 @@ +name: agda-lectures +include: Lecture-Notes/files + Lecture-Log + Exercises + Auxiliary + Cubical +flags: + -WnoUnsupportedIndexedMatch + --cubical-compatible diff --git a/src/HoTTEST/Agda/support/web/template.html b/src/HoTTEST/Agda/support/web/template.html new file mode 100644 index 0000000..05a24f0 --- /dev/null +++ b/src/HoTTEST/Agda/support/web/template.html @@ -0,0 +1,218 @@ + + + + + + + HoTTEST Summer School 2022 Lecture Notes + + + + + + + +
+
+ +
+ +$body$ + +
+
+
+ + + diff --git a/src/MayConcise/Chapter1.pdf b/src/MayConcise/Chapter1.pdf new file mode 100644 index 0000000..db6052a Binary files /dev/null and b/src/MayConcise/Chapter1.pdf differ diff --git a/src/MayConcise/Chapter1.typ b/src/MayConcise/Chapter1.typ new file mode 100644 index 0000000..b6bbfb5 --- /dev/null +++ b/src/MayConcise/Chapter1.typ @@ -0,0 +1,24 @@ +#set page(width: 4.8in, height: 7.9in, margin: 0.4in) + +#let problem(content) = text(fill: maroon, content) +#let answer(content) = rect( + width: 100%, + fill: color.linear-rgb(200, 200, 255, 255), + content, +) +#let S1 = $bb(S)^1$ + + +== Problems + +1. #[ + #problem[ + Let $p$ be a polynomial function on $bb(C)$ which has no root on $S1$. Show that the number of roots of $p(z) = 0$ with $|z| < 1$ is the degree of the map $hat(p) : S1 arrow.r S1$ specified by $hat(p)(z) = p(z) \/ |p(z)|$. + ] + + #answer[ + TODO: + + - What is $bb(C)$? Complex numbers. + ] +] \ No newline at end of file