├── .gitignore ├── code-examples ├── agda │ ├── prelude │ │ ├── int.agda │ │ ├── Stream.agda │ │ ├── list.agda │ │ └── io.agda │ ├── dyn │ │ └── fib.agda │ ├── dyn.agda │ ├── chu │ │ └── lens.agda │ ├── poly │ │ └── core.agda │ ├── chu.agda │ ├── prelude.agda │ ├── functors.agda │ └── poly0.agda ├── poly_up_to_fibonacci.idr └── dyn.idr ├── README.md ├── .github └── workflows │ └── main.yml ├── katie_discussion.txt ├── LICENSE ├── P3-Vistas.tex ├── LMSN ├── P3-Vistas-depricated.tex ├── tikz_stuff.tex ├── P0-Preface-depricated.tex └── Book-Poly-depricated.tex ├── tikz_stuff.tex ├── P0-Preface.tex ├── Book-Poly.tex └── snippets └── temp-intro.tex /.gitignore: -------------------------------------------------------------------------------- 1 | .DS_Store 2 | .texpadtmp/* 3 | code-examples/build 4 | *.agdai 5 | *.idr~ 6 | *.agda~ 7 | \.\#* 8 | \#* 9 | MAlonzo 10 | *.pdf 11 | *.log 12 | *.aux 13 | *.bbl 14 | *.bcf 15 | *.blg 16 | *.fdb_latexmk 17 | *.fls 18 | *.idx 19 | *.out 20 | *.run.xml 21 | *.synctex.gz 22 | *.synctex(busy) 23 | *.toc 24 | solution-file*.tex 25 | *.ilg 26 | *.ind -------------------------------------------------------------------------------- /code-examples/agda/prelude/int.agda: -------------------------------------------------------------------------------- 1 | module prelude.int where 2 | open import Data.Nat as Nat renaming (suc to s; zero to z) 3 | open import Agda.Builtin.Int public 4 | 5 | ⁺1 : Int → Int 6 | ⁺1 (pos n) = pos (s n) 7 | ⁺1 (negsuc z) = pos z 8 | ⁺1 (negsuc (s n)) = negsuc n 9 | 10 | ⁻1 : Int → Int 11 | ⁻1 (pos Nat.zero) = negsuc Nat.zero 12 | ⁻1 (pos (Nat.suc n)) = pos n 13 | ⁻1 (negsuc n) = negsuc (Nat.suc n) 14 | -------------------------------------------------------------------------------- /code-examples/agda/prelude/Stream.agda: -------------------------------------------------------------------------------- 1 | module prelude.Stream where 2 | open import prelude 3 | open import Data.List as L using (List) 4 | 5 | record Stream (a : Set) : Set where 6 | constructor _∷_ 7 | coinductive 8 | field 9 | hd : a 10 | tl : Stream a 11 | open Stream 12 | 13 | 14 | take : ∀ {a} → ℕ → Stream a → List a 15 | take ℕz xs = L.[] 16 | take (ℕs n) xs = hd xs L.∷ take n (tl xs) 17 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # poly 2 | 3 | See the [`pdf` branch](https://github.com/ToposInstitute/poly/blob/pdf/poly-book.pdf) or visit [topos.site/poly-book.pdf](https://topos.site/poly-book.pdf) for the latest build. 4 | 5 | In the summer of 2021, we offered a course based on this book: visit https://topos.site/poly-course/ for more details. 6 | 7 | Feel free to share feedback at https://docs.google.com/document/d/1qY5hLglgSW_p2AascL7XM3c4dByGZrECcnh6M1d3by8/edit. 8 | -------------------------------------------------------------------------------- /code-examples/agda/dyn/fib.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module dyn.fib where 3 | open import prelude 4 | open import functors 5 | open import prelude.Stream using (take) 6 | open import prelude.io 7 | open import dyn 8 | 9 | 10 | fib : Dyn 11 | fib = (plus ⊠⊠⊠ delay ℕ) ⟫ 12 | λ ( f , f₋ ) → f₋ , λ (.tt) 13 | → ((f₋ , f) , f ) 14 | where plus = fun→dyn (uncurry _ℕ+_) 15 | 16 | main : IO ⊤ 17 | main = printStream (run fib auto (1 , 1)) 18 | -------------------------------------------------------------------------------- /.github/workflows/main.yml: -------------------------------------------------------------------------------- 1 | on: push 2 | name: Build PDF on push 3 | jobs: 4 | run-latex: 5 | runs-on: ubuntu-latest 6 | steps: 7 | - name: Get latest code 8 | uses: actions/checkout@v2.3.2 9 | - name: Compile LaTeX document 10 | uses: xu-cheng/latex-action@v2 11 | with: 12 | root_file: Book-Poly.tex 13 | latexmk_shell_escape: true 14 | post_compile: "mv Book-Poly.pdf poly-book.pdf && mkdir output/ && cp poly-book.pdf output/poly-book.pdf" 15 | - name: Deploy to build branch 16 | uses: JamesIves/github-pages-deploy-action@3.6.2 17 | with: 18 | GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} 19 | BRANCH: pdf 20 | FOLDER: output 21 | -------------------------------------------------------------------------------- /code-examples/agda/prelude/list.agda: -------------------------------------------------------------------------------- 1 | module prelude.list where 2 | open import Data.Nat renaming (zero to z; suc to s) 3 | open import Data.Fin as Fin using (Fin) 4 | import Data.Fin as Fin renaming (zero to z; suc to s) 5 | open import Agda.Builtin.Sigma 6 | open import Data.Product 7 | open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; refl; cong) 8 | open import Data.Empty 9 | open import Function 10 | 11 | Vect : ∀ {ℓ} → ℕ → Set ℓ → Set ℓ 12 | Vect n t = Fin n → t 13 | List : ∀ {ℓ} → Set ℓ → Set ℓ 14 | List t = Σ[ n ∈ ℕ ] Vect n t 15 | len : ∀ {ℓ} {a : Set ℓ} → List a → ℕ 16 | len = fst 17 | head : ∀ {ℓ} {a : Set ℓ} → (as : List a) → (len as ≢ z) → a 18 | head (z , as) n≢0 = ⊥-elim (n≢0 refl) 19 | head (s n , as) _ = as Fin.z 20 | tail : ∀ {ℓ} {a : Set ℓ} → (as : List a) → (len as ≢ z) → List a 21 | tail (z , as) n≢0 = ⊥-elim (n≢0 refl) 22 | tail (s n , as) _ = n , (as ∘ Fin.inject₁) 23 | -------------------------------------------------------------------------------- /code-examples/agda/dyn.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module dyn where 3 | open import prelude 4 | open import functors 5 | open import poly0 public 6 | open import prelude.Stream 7 | open Stream 8 | open import Data.List as L using (List) 9 | 10 | 11 | record Dyn : Set where 12 | constructor dyn 13 | field 14 | {state} : Set 15 | {body} : ∫ 16 | pheno : ∫[ (state , λ _ → state) , body ] 17 | open Dyn public 18 | 19 | run : (d : Dyn) → ∫[ body d , 𝒴 ] → state d → Stream (π₁ (body d)) 20 | hd (run d@(dyn l) e s₀) = s₀ ★ l 21 | tl (run d@(dyn l) e s₀) = run d e (s₀ # l ← hd (run d e s₀) # e ← tt) 22 | 23 | module _ (d₁ d₂ : Dyn) where 24 | _⊠⊠⊠_ : Dyn 25 | _⊠⊠⊠_ = dyn (pheno d₁ ⟦⊠⟧ pheno d₂) 26 | 27 | _⟫_ : (d : Dyn) → ∫[ body d , A ] → Dyn 28 | d ⟫ l = dyn (pheno d ▸ l) 29 | 30 | fun→dyn : ∀ {a b} → (a → b) → Dyn 31 | fun→dyn f = dyn (λ a⁺ → a⁺ , f) 32 | 33 | delay : Set → Dyn 34 | delay s = fun→dyn (id {A = s}) 35 | -------------------------------------------------------------------------------- /code-examples/agda/chu/lens.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module chu.lens where 3 | open import prelude 4 | open import functors 5 | open import chu 6 | open import poly.core 7 | 8 | →∫ : Chu → ∫ 9 | →∫ (A⁺ , A⁻ ! Ω) = A⁺ , λ a⁺ → ∃ (Ω a⁺) 10 | 11 | →Lens : {A B : Chu} → Chu[ A , B ] → ∫[ →∫ A , →∫ B ] 12 | →Lens (f ↔ fᵗ ! †) a⁺ = f a⁺ , λ (b⁻ , fa⁺Ωb⁻) → fᵗ b⁻ , subst id († a⁺ b⁻) fa⁺Ωb⁻ 13 | 14 | module _ {A B C : Chu} 15 | (F@(f ↔ fᵗ ! _†₁_) : Chu[ A , B ]) 16 | (G@(g ↔ gᵗ ! _†₂_) : Chu[ B , C ]) where 17 | 18 | comp₂ : ∀ a⁺ → π₂ (→Lens (F ▸ G) a⁺) 19 | ≡ π₂ ((→Lens F ▸ →Lens G) a⁺) 20 | comp₂ a⁺ = extensionality λ ( c⁻ , gfaΩc ) → (λ x → (fᵗ ∘ gᵗ) c⁻ , x) ⟨$⟩ 21 | subst⋯ id (f a⁺ †₂ c⁻) (a⁺ †₁ gᵗ c⁻) gfaΩc 22 | 23 | comp∀ : ∀ a⁺ → →Lens (F ▸ G) a⁺ ≡ (→Lens F ▸ →Lens G) a⁺ 24 | comp∀ a⁺ rewrite comp₂ a⁺ = refl 25 | 26 | instance 27 | open Chu[_,_] 28 | chu-lens-functor : Functor →∫ 29 | chu-lens-functor = φ: →Lens 30 | 𝒾: refl 31 | ▸: λ F G → extensionality (comp∀ F G) 32 | -------------------------------------------------------------------------------- /katie_discussion.txt: -------------------------------------------------------------------------------- 1 | Discussion with Katie Leach 2 | 2024 / 03 / 04 3 | 4 | 5 | Agenda: 6 | 1. Color 7 | 2. Length 8 | 3. Price 9 | 10 | Minutes: 11 | 12 | Length: 13 | LMS: $90 US 14 | It is niche 15 | It is more specialized, smaller market 16 | Need to price it higher to make up for lower sales 17 | Smaller than monograph pricing 18 | Need cost-saving, e.g. with color 19 | 15% Royalties 20 | Cut to 8% Royalties for color 21 | Timeline 22 | We can propose the timeline 23 | Draft to assistant 3 months ahead of deadline date 24 | Then to production 25 | Copy editor sends us marked up .pdf 26 | We implement those edits 27 | Ebook 28 | .pdf on Cambridge Core 29 | People or libraries pay to get it 30 | Sold on Kindle ($90), etc. 31 | No published version on website. 32 | 33 | 34 | 35 | 36 | Action items: 37 | 0. Decide on royalties issue and no published .pdf on website. 38 | 1. Try out new format; propose maximum length. 39 | 2. Environments should be outlined. 40 | 3. Color should accord to colorblindness guidelines. 41 | 4. Tell Katie the deadline date. 42 | 5. Get book (after 7A and major changes) to assistant. -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2020 David Spivak 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /code-examples/agda/poly/core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module poly.core where 3 | open import functors 4 | open import prelude 5 | 6 | _^ : Set → Set -- Presheaf 7 | I ^ = I → Set 8 | 9 | ∫ : Set -- Arena 10 | ∫ = ∃ _^ 11 | 12 | _⦅_⦆ : ∫ → Set → Set -- Interpret as a polynomial functor 13 | (A⁺ , A⁻) ⦅ y ⦆ = Σ[ a⁺ ∈ A⁺ ] (A⁻ a⁺ → y) 14 | 15 | module _ (A@(A⁺ , A⁻) B@(B⁺ , B⁻) : ∫) where 16 | ∫[_,_] : Set 17 | ∫[_,_] = (a⁺ : A⁺) → Σ[ b⁺ ∈ B⁺ ] (B⁻ b⁺ → A⁻ a⁺) 18 | module _ {A@(A⁺ , A⁻) B@(B⁺ , B⁻) : ∫} where 19 | lens : (get : A⁺ → B⁺) 20 | (set : (a⁺ : A⁺) → B⁻ (get a⁺) → A⁻ a⁺) 21 | → ∫[ A , B ] 22 | lens g s = λ a⁺ → g a⁺ , s a⁺ 23 | get : (l : ∫[ A , B ]) → A⁺ → B⁺ 24 | get l = π₁ ∘ l 25 | set : (A↝B : ∫[ A , B ]) → (a⁺ : A⁺) → B⁻ (get A↝B a⁺) → A⁻ a⁺ 26 | set l = π₂ ∘ l 27 | 28 | instance 29 | lens-cat : Category ∫[_,_] 30 | lens-cat = 𝒾: (λ a⁺ → a⁺ , id) 31 | ▸: (λ l1 l2 a⁺ → let b⁺ , setb = l1 a⁺ 32 | c⁺ , setc = l2 b⁺ 33 | in c⁺ , setb ∘ setc) 34 | 𝒾▸: (λ _ → refl) 35 | ▸𝒾: (λ _ → refl) 36 | -------------------------------------------------------------------------------- /code-examples/agda/chu.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module chu where 3 | open import functors 4 | open import prelude 5 | 6 | record Chu : Set where 7 | constructor _,_!_ 8 | field 9 | _⁺ _⁻ : Set 10 | _Ω_ : _⁺ → _⁻ → Set 11 | 12 | module _ (A@(A⁺ , A⁻ ! _Ω₁_) B@(B⁺ , B⁻ ! _Ω₂_) : Chu) where 13 | record Chu[_,_] : Set where -- Morphisms of chu spaces 14 | constructor _↔_!_ 15 | field 16 | to : A⁺ → B⁺ 17 | from : B⁻ → A⁻ 18 | adj : ∀ a⁺ b⁻ → to a⁺ Ω₂ b⁻ ≡ a⁺ Ω₁ from b⁻ 19 | module _ {A@(A⁺ , A⁻ ! _Ω₁_) 20 | B@(B⁺ , B⁻ ! _Ω₂_) 21 | C@(C⁺ , C⁻ ! _Ω₃_) : Chu} 22 | (F@(f ↔ fᵗ ! _†₁_) : Chu[ A , B ]) 23 | (G@(g ↔ gᵗ ! _†₂_) : Chu[ B , C ]) where 24 | adj-comp : ∀ a⁺ c⁻ → (g ∘ f) a⁺ Ω₃ c⁻ ≡ a⁺ Ω₁ (fᵗ ∘ gᵗ) c⁻ 25 | adj-comp a⁺ c⁻ = trans (f a⁺ †₂ c⁻) -- g (f a⁺) Ω₃ c⁻ 26 | ( a⁺ †₁ gᵗ c⁻) -- f a⁺ Ω₂ gᵗ c⁻ 27 | -- a⁺ Ω₁ fᵗ (gᵗ c⁻) 28 | chu-comp : Chu[ A , C ] 29 | chu-comp = (g ∘ f) ↔ (fᵗ ∘ gᵗ) ! adj-comp 30 | 31 | instance 32 | chu-cat : Category Chu[_,_] 33 | chu-cat = 𝒾: (id ↔ id ! λ _ _ → refl) 34 | ▸: chu-comp 35 | 𝒾▸: (λ (_ ↔ _ ! _†_) → (λ x → _ ↔ _ ! x) ⟨$⟩ 36 | extensionality2 λ a⁺ b⁻ → trans-refl (a⁺ † b⁻)) 37 | ▸𝒾: (λ _ → refl) 38 | -------------------------------------------------------------------------------- /code-examples/agda/prelude.agda: -------------------------------------------------------------------------------- 1 | module prelude where 2 | open import Function using (id; _∘_) public 3 | open import Data.Sum renaming (inj₁ to Σ₁; inj₂ to Σ₂) using (_⊎_) public 4 | open import Data.Product renaming (proj₁ to π₁; proj₂ to π₂) using (Σ; _×_; _,_; ∃; Σ-syntax) public 5 | open import Agda.Builtin.Unit using (⊤; tt) public 6 | open import Data.Empty using (⊥) public 7 | open import Data.Nat as Nat renaming (suc to ℕs; zero to ℕz; _+_ to _ℕ+_) using (ℕ) public 8 | open import Relation.Binary.PropositionalEquality.Core as Eq using (_≡_; _≢_; refl; sym; trans; subst) renaming (cong to _⟨$⟩_) public 9 | open Eq.≡-Reasoning using (_≡⟨⟩_; step-≡; _∎) public 10 | 11 | postulate 12 | extensionality : {A : Set} {B : A → Set} {f g : (x : A) → B x} 13 | → (∀ x → f x ≡ g x) → f ≡ g 14 | extensionality2 : {A : Set} {B : A → Set} {C : (x : A) → B x → Set} {f g : (x : A) (y : B x) → C x y } 15 | → (∀ x y → f x y ≡ g x y) → f ≡ g 16 | extensionality2 λλf≡g = extensionality λ x → extensionality λ y → λλf≡g x y 17 | 18 | trans-refl : {A : Set} {x y : A} (p : x ≡ y) → trans p refl ≡ p 19 | trans-refl p rewrite p = refl 20 | 21 | subst⋯ : {A : Set} {x y z : A} (P : A → Set) (p : x ≡ y) (q : y ≡ z) (Px : P x) 22 | → subst P (trans p q) Px ≡ subst P q (subst P p Px) 23 | subst⋯ _ p _ _ rewrite p = refl 24 | 25 | _×⁼_ : {A X : Set} {a b : A} {x y : X} → a ≡ b → x ≡ y → (a , x) ≡ (b , y) 26 | _×⁼_ refl refl = refl 27 | 28 | _⨄_ : {A B : Set} → (A → Set) → (B → Set) → A ⊎ B → Set 29 | (F ⨄ G) (Σ₁ x) = F x 30 | (F ⨄ G) (Σ₂ y) = G y 31 | _⨃_ _⨉_ : {A B : Set} → (A → Set) → (B → Set) → A × B → Set 32 | F ⨃ G = λ (a , b) → F a ⊎ G b 33 | F ⨉ G = λ (a , b) → F a × G b 34 | 35 | _$₁_ : ∀ {A B X : Set} (f : A → B) → A × X → B × X 36 | f $₁ (a , x) = f a , x 37 | 38 | _⁻¹ : {A B : Set} → (A → B) → B → Set 39 | _⁻¹ {A} {B} f b = Σ[ a ∈ A ] (f a ≡ b) 40 | 41 | uncurry : {a b c : Set} → (a → b → c) → (a × b) → c 42 | uncurry f (a , b) = f a b 43 | -------------------------------------------------------------------------------- /code-examples/agda/prelude/io.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module prelude.io where 3 | open import prelude 4 | open import functors 5 | open import Agda.Builtin.IO public using (IO) 6 | open import Data.String 7 | open import prelude.Stream as Stream using (Stream) 8 | open import Data.List as L using (List) 9 | 10 | IO[_,_] : Set → Set → Set 11 | IO[ A , B ] = A → IO B 12 | 13 | infixl 1 _>>=_ 14 | 15 | postulate 16 | return : ∀ {a} {A : Set a} → A → IO A 17 | _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → IO A → (A → IO B) → IO B 18 | return>>= : ∀ {a b} (f : a → IO b) → (λ x → return x >>= f) ≡ f 19 | >>=return : ∀ {a b} (f : a → IO b) → (λ x → f x >>= return) ≡ f 20 | 21 | 22 | {-# COMPILE GHC return = \_ _ -> return #-} 23 | {-# COMPILE GHC _>>=_ = \_ _ _ _ -> (>>=) #-} 24 | 25 | -- record Category {O : Set} (𝒞[_,_] : O → O → Set) : Set where 26 | -- constructor 𝒾:_▸:_𝒾▸:_▸𝒾: 27 | -- infixl 8 _▸_ 28 | -- field 29 | -- 𝒾 : ∀ {x} → 𝒞[ x , x ] 30 | -- _▸_ : ∀ {x y z} → 𝒞[ x , y ] → 𝒞[ y , z ] → 𝒞[ x , z ] 31 | -- 𝒾▸ : ∀ {x y} (f : 𝒞[ x , y ]) → (𝒾 ▸ f) ≡ f 32 | -- ▸𝒾 : ∀ {x y} (f : 𝒞[ x , y ]) → (f ▸ 𝒾) ≡ f 33 | -- open Category ⦃...⦄ public 34 | 35 | 36 | instance 37 | io-cat : Category IO[_,_] 38 | io-cat = 𝒾: return ▸: (λ f g x → f x >>= g) 𝒾▸: return>>= ▸𝒾: >>=return 39 | 40 | 41 | {-# FOREIGN GHC import qualified Data.Text as T #-} 42 | {-# FOREIGN GHC import qualified Data.Text.IO as T #-} 43 | 44 | postulate 45 | getContents : IO String 46 | putStrLn : String → IO ⊤ 47 | 48 | {-# COMPILE GHC getContents = T.getContents #-} 49 | {-# COMPILE GHC putStrLn = T.putStrLn #-} 50 | 51 | postulate 52 | showℕ : ℕ → String 53 | {-# COMPILE GHC showℕ = T.pack . show #-} 54 | 55 | {-# NON_TERMINATING #-} 56 | printStream : Stream ℕ → IO ⊤ 57 | printStream s = putStrLn (showℕ (Stream.hd s)) >>= λ _ → printStream (Stream.tl s) 58 | 59 | printList : List ℕ → IO ⊤ 60 | printList L.[] = return tt 61 | printList (x L.∷ xs) = putStrLn (showℕ x) >>= λ _ → printList xs 62 | -------------------------------------------------------------------------------- /P3-Vistas.tex: -------------------------------------------------------------------------------- 1 | % !TeX root = P3-Vistas.tex 2 | \documentclass[Book-Poly]{subfiles} 3 | \begin{document} 4 | 5 | \setcounter{chapter}{8}%Just finished 8. 6 | 7 | %------------ Chapter ------------% 8 | \chapter{New horizons} \label{ch.comon.vistas} 9 | 10 | In this brief chapter, we lay out some questions that whose answers may or may not be known, but which were not known to us at the time of writing. They vary from concrete to open-ended, they are not organized in any particular way, and are in no sense complete. Still we hope they may be useful to some readers. 11 | 12 | \begin{enumerate} 13 | \item What can you say about comonoids in the category of all functors $\smset\to\smset$, e.g. ones that aren't polynomial. 14 | \item What can you say about the internal logic for the topos $[\cofree{p},\smset]$ of dynamical systems with interface $p$, in terms of $p$? 15 | \item How does the logic of the topos $\cofree{p}$ help us talk about issues that might be useful in studying dynamical systems? 16 | \item Morphisms $p\to q$ in $\poly$ give rise to left adjoints $\cofree{p}\to\cofree{q}$ that preserve connected limits. These are not geometric morphisms in general; in some sense they are worse and in some sense they are better. They are worse in that they do not preserve the terminal object, but they are better in that they preserve every connected limit not just finite ones. How do these left adjoints translate statements from the internal language of $p$ to that of $q$? 17 | \item Consider the $\times$-monoids and $\otimes$-monoids in three categories: $\poly$, $\smcat^\sharp$, and $\bimod{}{}$. Find examples of these comonoids, and perhaps characterize them or create a theory of them. 18 | \item The category $\poly$ has pullbacks, so one can consider the bicategory of spans in $\poly$. Is there a functor from that to $\bimod{}{}$ that sends $p\mapsto\cofree{p}$? 19 | \item Databases are static things, whereas dynamical systems are dynamic; yet we see them both in terms of $\poly$. How do they interact? Can a dynamical system read from or write to a database in any sense? 20 | \item Can we do database aggregation in a nice dynamic way? 21 | \item In the theory of polynomial functors, sums of representable functors $\smset\to\smset$, what happens if we replace sets with homotopy types: how much goes through? Is anything improved? 22 | \item Are there any functors $\smset\to\smset$ that aren't polynomial, but which admit a comonoid structure with respect to composition $(\yon,\tri)$? 23 | \item Characterize the monads in poly. They're generalizations of one-object operads (which are the Cartesian ones), but how can we think about them? 24 | \item Describe the limits that exist in $\smcat^\sharp$ combinatorially. 25 | \item Since the forgetful functor $U\colon\smcat^\sharp\to\poly$ is faithful, it reflects monomorphisms: if $f\colon\cat{C}\cof\cat{D}$ is a retrofunctor whose underlying map on carriers is monic, then it is monic. Are all monomorphisms in $\smcat^\sharp$ of this form? 26 | \item Are there polynomials $p$ such that one use something like G\"odel numbers to encode logical propositions from the topos $[\tr_p,\smset]$ into a ``language'' that $p$-dynamical systems can ``work with''? 27 | \end{enumerate} 28 | 29 | \end{document} -------------------------------------------------------------------------------- /LMSN/P3-Vistas-depricated.tex: -------------------------------------------------------------------------------- 1 | % !TeX root = P3-Vistas.tex 2 | \documentclass[Book-Poly]{subfiles} 3 | \begin{document} 4 | 5 | \setcounter{chapter}{7}%Just finished 7. 6 | 7 | %------------ Chapter ------------% 8 | \chapter{New horizons}\label{sec.discussion_open_qs} 9 | 10 | In this brief chapter, we lay out some questions that whose answers may or may not be known, but which were not known to us at the time of writing. They vary from concrete to open-ended, they are not organized in any particular way, and are in no sense complete. Still we hope they may be useful to some readers. 11 | 12 | \begin{enumerate} 13 | \item What can you say about comonoids in the category of all functors $\smset\to\smset$, e.g. ones that aren't polynomial. 14 | \item What can you say about the internal logic for the topos $[\cofree{p},\smset]$ of dynamical systems with interface $p$, in terms of $p$? 15 | \item How does the logic of the topos $\cofree{p}$ help us talk about issues that might be useful in studying dynamical systems? 16 | \item Morphisms $p\to q$ in $\poly$ give rise to left adjoints $\cofree{p}\to\cofree{q}$ that preserve connected limits. These are not geometric morphisms in general; in some sense they are worse and in some sense they are better. They are worse in that they do not preserve the terminal object, but they are better in that they preserve every connected limit not just finite ones. How do these left adjoints translate statements from the internal language of $p$ to that of $q$? 17 | \item Consider the $\times$-monoids and $\otimes$-monoids in three categories: $\poly$, $\smcat^\sharp$, and $\bimod{}{}$. Find examples of these comonoids, and perhaps characterize them or create a theory of them. 18 | \item The category $\poly$ has pullbacks, so one can consider the bicategory of spans in $\poly$. Is there a functor from that to $\bimod{}{}$ that sends $p\mapsto\cofree{p}$? 19 | \item Databases are static things, whereas dynamical systems are dynamic; yet we see them both in terms of $\poly$. How do they interact? Can a dynamical system read from or write to a database in any sense? 20 | \item Can we do database aggregation in a nice dynamic way? 21 | \item In the theory of polynomial functors, sums of representable functors $\smset\to\smset$, what happens if we replace sets with homotopy types: how much goes through? Is anything improved? 22 | \item Are there any functors $\smset\to\smset$ that aren't polynomial, but which admit a comonoid structure with respect to composition $(\yon,\tri)$? 23 | \item Characterize the monads in poly. They're generalizations of one-object operads (which are the Cartesian ones), but how can we think about them? 24 | \item Describe the limits that exist in $\smcat^\sharp$ combinatorially. 25 | \item Since the forgetful functor $U\colon\smcat^\sharp\to\poly$ is faithful, it reflects monomorphisms: if $f\colon\cat{C}\cof\cat{D}$ is a cofunctor whose underlying map on carriers is monic, then it is monic. Are all monomorphisms in $\smcat^\sharp$ of this form? 26 | \item Are there polynomials $p$ such that one use something like G\"odel numbers to encode logical propositions from the topos $[\tr_p,\smset]$ into a ``language'' that $p$-dynamical systems can ``work with''? 27 | \end{enumerate} 28 | 29 | \end{document} -------------------------------------------------------------------------------- /code-examples/agda/functors.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module functors where 3 | open import prelude 4 | 5 | record Category {O : Set} (𝒞[_,_] : O → O → Set) : Set where 6 | constructor 𝒾:_▸:_𝒾▸:_▸𝒾: 7 | infixl 8 _▸_ 8 | field 9 | 𝒾 : ∀ {x} → 𝒞[ x , x ] 10 | _▸_ : ∀ {x y z} → 𝒞[ x , y ] → 𝒞[ y , z ] → 𝒞[ x , z ] 11 | 𝒾▸ : ∀ {x y} (f : 𝒞[ x , y ]) → (𝒾 ▸ f) ≡ f 12 | ▸𝒾 : ∀ {x y} (f : 𝒞[ x , y ]) → (f ▸ 𝒾) ≡ f 13 | open Category ⦃...⦄ public 14 | 15 | record Functor {A : Set} {𝒜[_,_] : A → A → Set} 16 | {B : Set} {ℬ[_,_] : B → B → Set} 17 | (f : A → B) : Set where 18 | constructor φ:_𝒾:_▸: 19 | field 20 | ⦃ cat₁ ⦄ : Category {A} 𝒜[_,_] 21 | ⦃ cat₂ ⦄ : Category {B} ℬ[_,_] 22 | φ : ∀ {a b} → 𝒜[ a , b ] → ℬ[ f a , f b ] 23 | functor-id : ∀ {a} → φ {a} 𝒾 ≡ 𝒾 24 | functor-comp : ∀ {A B C} → (f : 𝒜[ A , B ]) (g : 𝒜[ B , C ]) 25 | → φ (f ▸ g) ≡ φ f ▸ φ g 26 | open Functor ⦃...⦄ public 27 | 28 | record Nat {A : Set} {𝒜[_,_] : A → A → Set} 29 | {B : Set} {ℬ[_,_] : B → B → Set} 30 | {F : A → B} {G : A → B} 31 | (α : (x : A) → ℬ[ F x , G x ]) : Set where 32 | field 33 | overlap ⦃ fun₁ ⦄ : Functor {A} {𝒜[_,_]} {B} {ℬ[_,_]} F 34 | overlap ⦃ fun₂ ⦄ : Functor {A} {𝒜[_,_]} {B} {ℬ[_,_]} G 35 | nat : ∀ {x y : A} (f : 𝒜[ x , y ]) → let _▹_ = _▸_ ⦃ cat₂ ⦄ 36 | in φ f ▹ α y ≡ α x ▹ φ f 37 | 38 | module _ {A : Set} {𝒜[_,_] : A → A → Set} where 39 | instance 40 | id-functor : ⦃ Category 𝒜[_,_] ⦄ → Functor id 41 | id-functor = φ: id 42 | 𝒾: refl 43 | ▸: λ f g → refl 44 | record Monad (f : A → A) : Set where 45 | field 46 | ⦃ functor ⦄ : Functor {_} {𝒜[_,_]} {_} {𝒜[_,_]} f 47 | η : ∀ {a} → 𝒜[ a , f a ] 48 | μ : ∀ {a} → 𝒜[ f (f a) , f a ] 49 | open Monad ⦃...⦄ public 50 | 51 | {- 52 | record Lift[_⇒_] (pre post : (Set → Set) → Set) (t : (Set → Set) → Set → Set) : Set where 53 | field 54 | ⦃ lifted ⦄ : ∀ {f} → post (t f) 55 | lift : {f : Set → Set} {a : Set} → ⦃ pre f ⦄ → f a → t f a 56 | open Lift[_⇒_] ⦃...⦄ public 57 | -} 58 | 59 | Cat[_,_] : Set → Set → Set 60 | Cat[ a , b ] = a → b 61 | 62 | instance 63 | Cat-cat : Category Cat[_,_] 64 | Cat-cat = record { 𝒾 = id ; _▸_ = λ f g → g ∘ f ; 𝒾▸ = λ _ → refl; ▸𝒾 = λ _ → refl} 65 | instance 66 | ≡-cat : ∀ {A} → Category {A} (λ a b → a ≡ b) 67 | 𝒾 {{≡-cat}} = refl 68 | _▸_ {{≡-cat}} = trans 69 | 𝒾▸ {{≡-cat}} _ = refl 70 | ▸𝒾 {{≡-cat}} x≡y rewrite x≡y = refl 71 | instance 72 | ×-cat : ∀ {A 𝒜[_,_]} ⦃ 𝒜 : Category {A} 𝒜[_,_] ⦄ 73 | {B ℬ[_,_]} ⦃ ℬ : Category {B} ℬ[_,_] ⦄ 74 | → Category {A × B} (λ (a , b) (a′ , b′) → 𝒜[ a , a′ ] × ℬ[ b , b′ ]) 75 | 𝒾 {{×-cat}} = 𝒾 , 𝒾 76 | _▸_ {{×-cat}} (f , f′ ) ( g , g′ ) = f ▸ g , f′ ▸ g′ 77 | 𝒾▸ {{×-cat}} (f , g) = 𝒾▸ f ×⁼ 𝒾▸ g 78 | ▸𝒾 {{×-cat}} (f , g) = ▸𝒾 f ×⁼ ▸𝒾 g 79 | 80 | data _ᵒᵖ {A : Set} (𝒜 : A → A → Set) : A → A → Set where 81 | _ᵗ : ∀ {y x} → 𝒜(y )( x) → (𝒜 ᵒᵖ)(x )( y) 82 | 83 | instance 84 | op-cat : {A : Set} {𝒜[_,_] : A → A → Set} ⦃ 𝒜 : Category {A} 𝒜[_,_] ⦄ 85 | → Category (𝒜[_,_] ᵒᵖ) 86 | 𝒾 {{op-cat}} = 𝒾 ᵗ 87 | _▸_ {{op-cat}} (f ᵗ) (g ᵗ) = (g ▸ f) ᵗ 88 | 𝒾▸ {{op-cat}} (f ᵗ) = _ᵗ ⟨$⟩ ▸𝒾 f 89 | ▸𝒾 {{op-cat}} (f ᵗ) = _ᵗ ⟨$⟩ 𝒾▸ f 90 | -------------------------------------------------------------------------------- /code-examples/agda/poly0.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | module poly0 where 3 | open import prelude 4 | open import functors 5 | open import poly.core public 6 | 7 | variable 8 | A B C X Y : ∫ 9 | I A⁺ B⁺ C⁺ X⁺ Y⁺ : Set 10 | A⁻ : A⁺ → Set 11 | B⁻ : B⁺ → Set 12 | C⁻ : C⁺ → Set 13 | X⁻ : X⁺ → Set 14 | Y⁻ : Y⁺ → Set 15 | 16 | 17 | ∃⊤ ∃⊥ ⊤∫ ∫∫ : Set → ∫ 18 | ∃⊤ a = a , λ _ → ⊤ 19 | ∃⊥ a = a , λ _ → ⊥ 20 | ⊤∫ a = ⊤ , λ _ → a 21 | ∫∫ a = a , λ _ → a 22 | 𝒴 𝟘 : ∫ 23 | 𝒴 = ⊤ , λ _ → ⊤ 24 | 𝟘 = ⊥ , λ _ → ⊥ 25 | 26 | module _ {A@(A⁺ , A⁻) B@(B⁺ , B⁻) : ∫} where 27 | infixl 5 _★_ 28 | -- infix notatjon for get 29 | _★_ : A⁺ → (l : ∫[ A , B ]) → B⁺ 30 | σa ★ l = get l σa 31 | infixr 4 _#_←_ 32 | -- Infix notatjon for set 33 | _#_←_ : (a⁺ : A⁺) → (A↝B : ∫[ A , B ]) → B⁻ (a⁺ ★ A↝B) → A⁻ a⁺ 34 | a⁺ # l ← b⁻ = π₂ (l a⁺) b⁻ 35 | _↕_ : (get : A⁺ → B⁺) → (set : (a⁺ : A⁺) → B⁻ (get a⁺) → A⁻ a⁺) → ∫[ A , B ] 36 | g ↕ s = λ a⁺ → (g a⁺) , (s a⁺) 37 | 38 | module _ {A@(A⁺ , A⁻) C@(C⁺ , C⁻) : ∫} (l : ∫[ A , C ]) where 39 | -- vertical and cartesian factorization 40 | Factor : Σ[ B ∈ ∫ ] (∫[ A , B ]) × (∫[ B , C ]) 41 | Factor = (A⁺ , C⁻ ∘ get l) 42 | , id ↕ set l 43 | , get l ↕ λ _ → id 44 | 45 | module lenses (f : A⁺ → B⁺) where 46 | constant : ∫[ ∃⊥ A⁺ , ∃⊥ B⁺ ] 47 | emitter : ∫[ ∃⊤ A⁺ , ∃⊤ B⁺ ] 48 | sensor : ∫[ ⊤∫ B⁺ , ⊤∫ A⁺ ] 49 | constant a⁺ = f a⁺ , id 50 | emitter a⁺ = f a⁺ , λ _ → tt 51 | sensor _ = tt , f 52 | open lenses public 53 | enclose : ((a⁺ : A⁺) → B⁻ a⁺) → ∫[ (A⁺ , B⁻) , 𝒴 ] 54 | enclose f a⁺ = tt , λ _ → f a⁺ 55 | auto : ∫[ ∃⊤ A⁺ , 𝒴 ] 56 | auto = enclose λ _ → tt 57 | 58 | 59 | {- 60 | lift∫ : (f : Set → Set) → ∫ → ∫ 61 | lift∫ f (A⁺ , A⁻) = A⁺ , f ∘ A⁻ 62 | 63 | liftLens : ∀ {A B} (f : Set → Set) → ∫[ A , B ] → ∫[ lift∫ f A , lift∫ f B ] 64 | liftLens f l a⁺ with l a⁺ 65 | ... | b⁺ , setb = b⁺ , φ setb 66 | -} 67 | 68 | {- 69 | module lift_comonad (f : Set → Set) {A : ∫} ⦃ f_monad : Monad f ⦄ where 70 | extract : lift∫ f A ↝ A 71 | extract a⁺ = a⁺ , η f_monad 72 | -- id ↕ λ _ → η 73 | duplicate : lift∫ f A ↝ lift∫ f (lift∫ f A) 74 | duplicate a⁺ = a⁺ , μ 75 | -} 76 | 77 | module poly-ops where 78 | infixl 6 _⊗_ 79 | infixl 5 _⊕_ 80 | _⊕_ _⊗_ _⊠_ _⊚_ : ∫ → ∫ → ∫ 81 | (A⁺ , A⁻) ⊕ (B⁺ , B⁻) = (A⁺ ⊎ B⁺) , (A⁻ ⨄ B⁻) -- coproduct 82 | (A⁺ , A⁻) ⊗ (B⁺ , B⁻) = (A⁺ × B⁺) , (A⁻ ⨃ B⁻) -- product 83 | (A⁺ , A⁻) ⊠ (B⁺ , B⁻) = (A⁺ × B⁺) , (A⁻ ⨉ B⁻) -- juxtapose 84 | (A⁺ , A⁻) ⊚ (B⁺ , B⁻) = (Σ[ a⁺ ∈ A⁺ ](A⁻ a⁺ → B⁺)) -- compose 85 | , λ (_ , bs) → ∃ (B⁻ ∘ bs) 86 | 87 | -- N-ary 88 | Σ⊕ : (I → ∫) → ∫ 89 | Σ⊕ {I = I} A = Σ I (π₁ ∘ A) , λ (i , a⁺) → π₂ (A i) a⁺ 90 | Π⊗ : (I → ∫) → ∫ 91 | Π⊗ {I = I} a = ((i : I) → π₁ (a i)) 92 | , λ a⁺ → Σ[ i ∈ I ](π₂ (a i) (a⁺ i)) 93 | Π⊠ : (I → ∫) → ∫ 94 | Π⊠ {I = I} a = ((i : I) → π₁ (a i)) 95 | , (λ a⁺ → (i : I) -> π₂ (a i) (a⁺ i)) 96 | _ᵒ_ : ∫ → ℕ → ∫ 97 | _ ᵒ ℕz = 𝒴 98 | a ᵒ ℕs n = a ⊚ (a ᵒ n) 99 | open poly-ops public 100 | 101 | module lens-ops where 102 | _⟦+⟧_ : ∀ {a b x y} → ∫[ a , b ] → ∫[ x , y ] → ∫[ a ⊕ x , b ⊕ y ] -- coproduct 103 | _⟦|⟧_ : ∀ {a b x } → ∫[ a , x ] → ∫[ b , x ] → ∫[ a ⊕ b , x ] -- copair 104 | _⟦⊗⟧_ : ∀ {a b x y} → ∫[ a , b ] → ∫[ x , y ] → ∫[ a ⊗ x , b ⊗ y ] -- product 105 | _⟦×⟧_ : ∀ {x a b } → ∫[ x , a ] → ∫[ x , b ] → ∫[ x , a ⊗ b ] -- pair 106 | _⟦⊠⟧_ : ∀ {a b x y} → ∫[ a , b ] → ∫[ x , y ] → ∫[ a ⊠ x , b ⊠ y ] -- juxtaposition 107 | _⟦⊚⟧_ : ∀ {a b x y} → ∫[ a , b ] → ∫[ x , y ] → ∫[ a ⊚ x , b ⊚ y ] -- composition 108 | 109 | (a↝b ⟦+⟧ x↝y) = λ{(Σ₁ a⁺) → let b⁺ , setb = a↝b a⁺ in Σ₁ b⁺ , setb 110 | ;(Σ₂ x⁺) → let y⁺ , sety = x↝y x⁺ in Σ₂ y⁺ , sety} 111 | 112 | (a↝x ⟦|⟧ b↝x) = λ{(Σ₁ a⁺) → a↝x a⁺ 113 | ;(Σ₂ b⁺) → b↝x b⁺} 114 | 115 | (a↝b ⟦⊗⟧ x↝y) (a⁺ , x⁺) with a↝b a⁺ | x↝y x⁺ 116 | ... | b⁺ , setb | y⁺ , sety = (b⁺ , y⁺) 117 | , λ{(Σ₁ b⁻) → Σ₁ (setb b⁻) 118 | ;(Σ₂ y⁻) → Σ₂ (sety y⁻)} 119 | 120 | _⟦×⟧_ x↝a x↝b x⁺ with x↝a x⁺ | x↝b x⁺ 121 | ... | a⁺ , seta | b⁺ , setb = (a⁺ , b⁺) , λ{(Σ₁ a⁻) → seta a⁻ 122 | ;(Σ₂ b⁻) → setb b⁻} 123 | 124 | _⟦⊠⟧_ a↝b x↝y (a⁺ , x⁺) = ((a⁺ ★ a↝b) , (x⁺ ★ x↝y)) 125 | , λ (b⁻ , y⁻) → (a⁺ # a↝b ← b⁻) , (x⁺ # x↝y ← y⁻) 126 | 127 | (a↝b ⟦⊚⟧ x↝y) (a⁺ , a⁻→x⁺) with a↝b a⁺ 128 | ... | b⁺ , setb = (b⁺ , get x↝y ∘ a⁻→x⁺ ∘ setb) 129 | , λ (b⁻ , y⁻) → let a⁻ = setb b⁻ 130 | in a⁻ , (a⁻→x⁺ a⁻ # x↝y ← y⁻) 131 | 132 | -- N-ary 133 | Π⟦⊠⟧ : {as bs : I → ∫} 134 | → ((i : I) → (∫[ as i , bs i ])) 135 | → ∫[ Π⊠ as , Π⊠ bs ] 136 | Π⟦⊠⟧ ls as⁺ = (λ i → as⁺ i ★ ls i) 137 | , (λ dbs i → as⁺ i # ls i ← dbs i) 138 | _⟦ᵒ⟧_ : ∫[ A , B ] → (n : ℕ) → ∫[ A ᵒ n , B ᵒ n ] 139 | _⟦ᵒ⟧_ {a} {b} l = go where 140 | go : (n : ℕ) → ∫[ a ᵒ n , b ᵒ n ] 141 | go ℕz = 𝒾 {x = 𝒴} 142 | go (ℕs n) = l ⟦⊚⟧ (go n) 143 | open lens-ops public 144 | -------------------------------------------------------------------------------- /tikz_stuff.tex: -------------------------------------------------------------------------------- 1 | \pgfdeclarelayer{edgelayer} 2 | \pgfdeclarelayer{nodelayer} 3 | \pgfsetlayers{edgelayer,nodelayer,main} 4 | 5 | 6 | 7 | \tikzset{ 8 | tick/.style={postaction={ 9 | decorate, 10 | decoration={markings, mark=at position 0.5 with 11 | {\draw[-] (0,.4ex) -- (0,-.4ex);}}} 12 | } 13 | } 14 | 15 | \tikzset{ 16 | bottom base/.style={baseline=(current bounding box.south)} 17 | } 18 | 19 | 20 | 21 | \tikzset{trees/.style={ 22 | inner sep=0, 23 | minimum width=0, 24 | minimum height=0, 25 | level distance=.5cm, 26 | sibling distance=.5cm, 27 | % every child/.style={fill}, 28 | edge from parent/.style={shorten <= -2pt, draw, ->}, 29 | grow'=up, 30 | decoration={markings, mark=at position 0.75 with \arrow{stealth}} 31 | } 32 | } 33 | 34 | \newcommand{\idchild}{edge from parent[double, -]} 35 | 36 | \tikzset{ 37 | oriented WD/.style={%everything after equals replaces "oriented WD" in key. 38 | every to/.style={out=0,in=180,draw}, 39 | label/.style={ 40 | font=\everymath\expandafter{\the\everymath\scriptstyle}, 41 | inner sep=0pt, 42 | node distance=2pt and -2pt}, 43 | semithick, 44 | node distance=1 and 1, 45 | decoration={markings, mark=at position \stringdecpos with \stringdec}, 46 | ar/.style={postaction={decorate}}, 47 | execute at begin picture={\tikzset{ 48 | x=\bbx, y=\bby, 49 | }} 50 | }, 51 | string decoration/.store in=\stringdec, 52 | string decoration={\arrow{stealth};}, 53 | string decoration pos/.store in=\stringdecpos, 54 | string decoration pos=.7, 55 | dot size/.store in=\dotsize, 56 | dot size=3pt, 57 | dot/.style={ 58 | circle, draw, thick, inner sep=0, fill=black, minimum width=\dotsize 59 | }, 60 | bbx/.store in=\bbx, 61 | bbx = 1.5cm, 62 | bby/.store in=\bby, 63 | bby = 1.5ex, 64 | bb port sep/.store in=\bbportsep, 65 | bb port sep=1.5, 66 | % bb wire sep/.store in=\bbwiresep, 67 | % bb wire sep=1.75ex, 68 | bb port length/.store in=\bbportlen, 69 | bb port length=4pt, 70 | bb penetrate/.store in=\bbpenetrate, 71 | bb penetrate=0, 72 | bb min width/.store in=\bbminwidth, 73 | bb min width=1cm, 74 | bb rounded corners/.store in=\bbcorners, 75 | bb rounded corners=2pt, 76 | bb small/.style={bb port sep=1, bb port length=2.5pt, bbx=.4cm, bb min width=.4cm, 77 | bby=.7ex}, 78 | bb medium/.style={bb port sep=1, bb port length=2.5pt, bbx=.4cm, bb min width=.4cm, 79 | bby=.9ex}, 80 | bb/.code 2 args={%When you see this key, run the code below: 81 | \pgfmathsetlengthmacro{\bbheight}{\bbportsep * (max(#1,#2)+1) * \bby} 82 | \pgfkeysalso{draw,minimum height=\bbheight,minimum width=\bbminwidth,outer 83 | sep=0pt, 84 | rounded corners=\bbcorners,thick, 85 | prefix after command={\pgfextra{\let\fixname\tikzlastnode}}, 86 | append after command={\pgfextra{\draw 87 | \ifnum #1=0{} \else foreach \i in {1,...,#1} { 88 | ($(\fixname.north west)!{\i/(#1+1)}!(\fixname.south west)$) +(- 89 | \bbportlen,0) 90 | coordinate (\fixname_in\i) -- +(\bbpenetrate,0) coordinate (\fixname_in\i')}\fi 91 | %Define the endpoints of tickmarks 92 | \ifnum #2=0{} \else foreach \i in {1,...,#2} { 93 | ($(\fixname.north east)!{\i/(#2+1)}!(\fixname.south east)$) +(- 94 | \bbpenetrate,0) 95 | coordinate (\fixname_out\i') -- +(\bbportlen,0) coordinate (\fixname_out\i)}\fi; 96 | }}} 97 | }, 98 | bb name/.style={append after command={\pgfextra{\node[anchor=north] at 99 | (\fixname.north) {#1};}}} 100 | } 101 | 102 | 103 | \newcommand{\boxofbullets}[6][1]{% [y-spacing] {number of bullets} {left endpoint} {label} {text} {position} 104 | \foreach \i [evaluate=\i as \y using {#1/2*((#2-1)/2-\i)}] in {1,...,#2}{ 105 | \node (pt_#4_\i) at ($(0,\y)+#3$) {$\bullet$}; 106 | \node[#6, font=\tiny] (lab_#4_\i) at ($(0,\y)+#3$) {#5$_\i$}; 107 | } 108 | \node[draw, rounded corners, inner xsep=2pt, inner ysep=0pt, fit={(pt_#4_1) (pt_#4_#2) (lab_#4_1) (lab_#4_#2)}] (box_#4) {}; 109 | } 110 | 111 | \tikzset{polybox/.style={ 112 | poly/.style ={ 113 | rectangle split, 114 | rectangle split parts=2, 115 | rectangle split part align={bottom}, 116 | draw, 117 | anchor=south, 118 | inner ysep=4.5, 119 | prefix after command={\pgfextra{\let\fixname=\tikzlastnode}}, 120 | append after command={\pgfextra{ 121 | \node[inner xsep=0, inner ysep=0, 122 | fit=(\fixname.north west) (\fixname.two split east)] 123 | (\fixname_dir) {}; 124 | \node[inner xsep=0, inner ysep=0, 125 | fit=(\fixname.south west) (\fixname.two split east)] 126 | (\fixname_pos) {}; 127 | }} 128 | }, 129 | dom/.style={ 130 | rectangle split part fill={none, blue!10} 131 | }, 132 | cod/.style={ 133 | rectangle split part fill={blue!10, none} 134 | }, 135 | constant/.style={ 136 | rectangle split part fill={red!50, none} 137 | }, 138 | constant dom/.style={ 139 | rectangle split part fill={red!50, blue!10} 140 | }, 141 | terminal/.style={ 142 | rectangle split part fill={red!50, gray} 143 | }, 144 | identity/.style={ 145 | rectangle split part fill={gray, gray} 146 | }, 147 | linear/.style={ 148 | rectangle split part fill={gray, none} 149 | }, 150 | linear dom/.style={ 151 | rectangle split part fill={gray, blue!10} 152 | }, 153 | linear cod/.style={ 154 | rectangle split part fill={blue!10!gray, none} 155 | }, 156 | pure/.style={ 157 | rectangle split part fill={none, gray} 158 | }, 159 | pure dom/.style={ 160 | rectangle split part fill={none, blue!10!gray} 161 | }, 162 | pure cod/.style={ 163 | rectangle split part fill={blue!10, gray} 164 | }, 165 | shorten <=3pt, shorten >=3pt, 166 | first/.style={out=0, in=180}, 167 | climb/.style={out=180, in=180, looseness=2}, 168 | climb'/.style={out=0, in=0, looseness=2}, 169 | last/.style={out=180, in=0}, 170 | mapstos/.style={arrows={|->}}, 171 | tos/.style={arrows={->}}, 172 | font=\footnotesize, 173 | node distance=2ex and 1.5cm 174 | } 175 | } 176 | 177 | 178 | \tikzset{ 179 | biml/.tip={Glyph[glyph math command=triangleleft, glyph length=.95ex]}, 180 | bimr/.tip={Glyph[glyph math command=triangleright, glyph length=.95ex]}, 181 | } 182 | 183 | 184 | \newcommand{\bimodto}[1][]{ 185 | \begin{tikzcd}[ampersand replacement=\&, cramped]\ar[r, biml-bimr, "#1"]\&~\end{tikzcd} 186 | } 187 | \newcommand{\bimodfrom}[1][]{ 188 | \begin{tikzcd}[ampersand replacement=\&, cramped]\ar[r, bimr-biml, "#1"]\&~\end{tikzcd} 189 | } 190 | 191 | 192 | 193 | 194 | \newcommand{\earpic}{ 195 | \begin{tikzpicture}[polybox, tos] 196 | \node[poly, dom, "$m$" left] (m) {}; 197 | \node[poly, cod, right=of m, "$m$" right] (mm) {}; 198 | \node[poly, cod, above=of mm, "$C$" right] (C) {}; 199 | \node[poly, cod, below=of mm, "$D$" right] (D) {}; 200 | % 201 | \draw (m_pos) to[out=0, in=180] (D_pos); 202 | \draw (D_dir) to[climb] (mm_pos); 203 | \draw (mm_dir) to[climb] (C_pos); 204 | \draw (C_dir) to[last] (m_dir); 205 | \end{tikzpicture} 206 | } 207 | 208 | \newcommand{\treepic}{ 209 | \begin{tikzpicture}[trees, scale=1.5, 210 | level 1/.style={sibling distance=20mm}, 211 | level 2/.style={sibling distance=10mm}, 212 | level 3/.style={sibling distance=5mm}, 213 | level 4/.style={sibling distance=2.5mm}, 214 | level 5/.style={sibling distance=1.25mm}] 215 | \node[dgreen] (a) {$\bullet$} 216 | child {node[dgreen] {$\bullet$} 217 | child {node[dgreen] {$\bullet$} 218 | child {node[dgreen] {$\bullet$} 219 | child {node[dgreen] {$\bullet$} 220 | child {} 221 | child {} 222 | } 223 | child {node[dyellow] {$\bullet$} 224 | child {} 225 | child {} 226 | } 227 | } 228 | child {node[dyellow] {$\bullet$} 229 | child {node[dgreen] {$\bullet$} 230 | child {} 231 | child {} 232 | } 233 | child {node[red] {$\bullet$}} 234 | } 235 | } 236 | child {node[dyellow] {$\bullet$} 237 | child {node[dgreen] {$\bullet$} 238 | child {node[dgreen] {$\bullet$} 239 | child {} 240 | child {} 241 | } 242 | child {node[dyellow] {$\bullet$} 243 | child {} 244 | child {} 245 | } 246 | } 247 | child {node[red] {$\bullet$}} 248 | } 249 | } 250 | child {node[dyellow] {$\bullet$} 251 | child {node[dgreen] {$\bullet$} 252 | child {node[dgreen] {$\bullet$} 253 | child {node[dgreen] {$\bullet$} 254 | child {} 255 | child {} 256 | } 257 | child {node[dyellow] {$\bullet$} 258 | child {} 259 | child {} 260 | } 261 | } 262 | child {node[dyellow] {$\bullet$} 263 | child {node[dgreen] {$\bullet$} 264 | child {} 265 | child {} 266 | } 267 | child {node[red] {$\bullet$}} 268 | } 269 | } 270 | child {node[red] {$\bullet$} 271 | } 272 | } 273 | ; 274 | \end{tikzpicture} 275 | } 276 | 277 | \newcommand{\coverpic}{ 278 | \begin{tikzpicture} 279 | \node {\treepic}; 280 | \draw[blue, densely dotted] (current bounding box.south west) rectangle 281 | (current bounding box.north east); 282 | \end{tikzpicture} 283 | }%{\earpic} 284 | 285 | 286 | \tikzset{ 287 | mapsto/.style={->, densely dashed, blue, shorten <=\short, shorten 288 | >=\short, >=stealth, thick}, 289 | short/.store in=\short, 290 | short=0pt 291 | } 292 | 293 | 294 | 295 | 296 | 297 | 298 | 299 | 300 | 301 | 302 | 303 | 304 | 305 | 306 | 307 | 308 | -------------------------------------------------------------------------------- /LMSN/tikz_stuff.tex: -------------------------------------------------------------------------------- 1 | \pgfdeclarelayer{edgelayer} 2 | \pgfdeclarelayer{nodelayer} 3 | \pgfsetlayers{edgelayer,nodelayer,main} 4 | 5 | 6 | 7 | \tikzset{ 8 | tick/.style={postaction={ 9 | decorate, 10 | decoration={markings, mark=at position 0.5 with 11 | {\draw[-] (0,.4ex) -- (0,-.4ex);}}} 12 | } 13 | } 14 | 15 | \tikzset{ 16 | bottom base/.style={baseline=(current bounding box.south)} 17 | } 18 | 19 | 20 | 21 | \tikzset{trees/.style={ 22 | inner sep=0, 23 | minimum width=0, 24 | minimum height=0, 25 | level distance=.5cm, 26 | sibling distance=.5cm, 27 | % every child/.style={fill}, 28 | edge from parent/.style={shorten <= -2pt, draw, ->}, 29 | grow'=up, 30 | decoration={markings, mark=at position 0.75 with \arrow{stealth}} 31 | } 32 | } 33 | 34 | \newcommand{\idchild}{edge from parent[double, -]} 35 | 36 | \tikzset{ 37 | oriented WD/.style={%everything after equals replaces "oriented WD" in key. 38 | every to/.style={out=0,in=180,draw}, 39 | label/.style={ 40 | font=\everymath\expandafter{\the\everymath\scriptstyle}, 41 | inner sep=0pt, 42 | node distance=2pt and -2pt}, 43 | semithick, 44 | node distance=1 and 1, 45 | decoration={markings, mark=at position \stringdecpos with \stringdec}, 46 | ar/.style={postaction={decorate}}, 47 | execute at begin picture={\tikzset{ 48 | x=\bbx, y=\bby, 49 | }} 50 | }, 51 | string decoration/.store in=\stringdec, 52 | string decoration={\arrow{stealth};}, 53 | string decoration pos/.store in=\stringdecpos, 54 | string decoration pos=.7, 55 | dot size/.store in=\dotsize, 56 | dot size=3pt, 57 | dot/.style={ 58 | circle, draw, thick, inner sep=0, fill=black, minimum width=\dotsize 59 | }, 60 | bbx/.store in=\bbx, 61 | bbx = 1.5cm, 62 | bby/.store in=\bby, 63 | bby = 1.5ex, 64 | bb port sep/.store in=\bbportsep, 65 | bb port sep=1.5, 66 | % bb wire sep/.store in=\bbwiresep, 67 | % bb wire sep=1.75ex, 68 | bb port length/.store in=\bbportlen, 69 | bb port length=4pt, 70 | bb penetrate/.store in=\bbpenetrate, 71 | bb penetrate=0, 72 | bb min width/.store in=\bbminwidth, 73 | bb min width=1cm, 74 | bb rounded corners/.store in=\bbcorners, 75 | bb rounded corners=2pt, 76 | bb small/.style={bb port sep=1, bb port length=2.5pt, bbx=.4cm, bb min width=.4cm, 77 | bby=.7ex}, 78 | bb medium/.style={bb port sep=1, bb port length=2.5pt, bbx=.4cm, bb min width=.4cm, 79 | bby=.9ex}, 80 | bb/.code 2 args={%When you see this key, run the code below: 81 | \pgfmathsetlengthmacro{\bbheight}{\bbportsep * (max(#1,#2)+1) * \bby} 82 | \pgfkeysalso{draw,minimum height=\bbheight,minimum width=\bbminwidth,outer 83 | sep=0pt, 84 | rounded corners=\bbcorners,thick, 85 | prefix after command={\pgfextra{\let\fixname\tikzlastnode}}, 86 | append after command={\pgfextra{\draw 87 | \ifnum #1=0{} \else foreach \i in {1,...,#1} { 88 | ($(\fixname.north west)!{\i/(#1+1)}!(\fixname.south west)$) +(- 89 | \bbportlen,0) 90 | coordinate (\fixname_in\i) -- +(\bbpenetrate,0) coordinate (\fixname_in\i')}\fi 91 | %Define the endpoints of tickmarks 92 | \ifnum #2=0{} \else foreach \i in {1,...,#2} { 93 | ($(\fixname.north east)!{\i/(#2+1)}!(\fixname.south east)$) +(- 94 | \bbpenetrate,0) 95 | coordinate (\fixname_out\i') -- +(\bbportlen,0) coordinate (\fixname_out\i)}\fi; 96 | }}} 97 | }, 98 | bb name/.style={append after command={\pgfextra{\node[anchor=north] at 99 | (\fixname.north) {#1};}}} 100 | } 101 | 102 | 103 | \newcommand{\boxofbullets}[6][1]{% [y-spacing] {number of bullets} {left endpoint} {label} {text} {position} 104 | \foreach \i [evaluate=\i as \y using {#1/2*((#2-1)/2-\i)}] in {1,...,#2}{ 105 | \node (pt_#4_\i) at ($(0,\y)+#3$) {$\bullet$}; 106 | \node[#6, font=\tiny] (lab_#4_\i) at ($(0,\y)+#3$) {#5$_\i$}; 107 | } 108 | \node[draw, rounded corners, inner xsep=2pt, inner ysep=0pt, fit={(pt_#4_1) (pt_#4_#2) (lab_#4_1) (lab_#4_#2)}] (box_#4) {}; 109 | } 110 | 111 | \tikzset{polybox/.style={ 112 | poly/.style ={ 113 | rectangle split, 114 | rectangle split parts=2, 115 | rectangle split part align={bottom}, 116 | draw, 117 | anchor=south, 118 | inner ysep=4.5, 119 | prefix after command={\pgfextra{\let\fixname=\tikzlastnode}}, 120 | append after command={\pgfextra{ 121 | \node[inner xsep=0, inner ysep=0, 122 | fit=(\fixname.north west) (\fixname.two split east)] 123 | (\fixname_dir) {}; 124 | \node[inner xsep=0, inner ysep=0, 125 | fit=(\fixname.south west) (\fixname.two split east)] 126 | (\fixname_pos) {}; 127 | }} 128 | }, 129 | dom/.style={ 130 | rectangle split part fill={none, blue!10} 131 | }, 132 | cod/.style={ 133 | rectangle split part fill={blue!10, none} 134 | }, 135 | constant/.style={ 136 | rectangle split part fill={red!50, none} 137 | }, 138 | constant dom/.style={ 139 | rectangle split part fill={red!50, blue!10} 140 | }, 141 | terminal/.style={ 142 | rectangle split part fill={red!50, gray} 143 | }, 144 | identity/.style={ 145 | rectangle split part fill={gray, gray} 146 | }, 147 | linear/.style={ 148 | rectangle split part fill={gray, none} 149 | }, 150 | linear dom/.style={ 151 | rectangle split part fill={gray, blue!10} 152 | }, 153 | linear cod/.style={ 154 | rectangle split part fill={blue!10!gray, none} 155 | }, 156 | pure/.style={ 157 | rectangle split part fill={none, gray} 158 | }, 159 | pure dom/.style={ 160 | rectangle split part fill={none, blue!10!gray} 161 | }, 162 | pure cod/.style={ 163 | rectangle split part fill={blue!10, gray} 164 | }, 165 | shorten <=3pt, shorten >=3pt, 166 | first/.style={out=0, in=180}, 167 | climb/.style={out=180, in=180, looseness=2}, 168 | climb'/.style={out=0, in=0, looseness=2}, 169 | last/.style={out=180, in=0}, 170 | mapstos/.style={arrows={|->}}, 171 | tos/.style={arrows={->}}, 172 | font=\footnotesize, 173 | node distance=2ex and 1.5cm 174 | } 175 | } 176 | 177 | 178 | \tikzset{ 179 | biml/.tip={Glyph[glyph math command=triangleleft, glyph length=.95ex]}, 180 | bimr/.tip={Glyph[glyph math command=triangleright, glyph length=.95ex]}, 181 | } 182 | 183 | 184 | \newcommand{\bimodto}[1][]{ 185 | \begin{tikzcd}[ampersand replacement=\&, cramped]\ar[r, biml-bimr, "#1"]\&~\end{tikzcd} 186 | } 187 | \newcommand{\bimodfrom}[1][]{ 188 | \begin{tikzcd}[ampersand replacement=\&, cramped]\ar[r, bimr-biml, "#1"]\&~\end{tikzcd} 189 | } 190 | 191 | 192 | 193 | 194 | \newcommand{\earpic}{ 195 | \begin{tikzpicture}[polybox, tos] 196 | \node[poly, dom, "$m$" left] (m) {}; 197 | \node[poly, cod, right=of m, "$m$" right] (mm) {}; 198 | \node[poly, cod, above=of mm, "$C$" right] (C) {}; 199 | \node[poly, cod, below=of mm, "$D$" right] (D) {}; 200 | % 201 | \draw (m_pos) to[out=0, in=180] (D_pos); 202 | \draw (D_dir) to[climb] (mm_pos); 203 | \draw (mm_dir) to[climb] (C_pos); 204 | \draw (C_dir) to[last] (m_dir); 205 | \end{tikzpicture} 206 | } 207 | 208 | \newcommand{\treepic}{ 209 | \begin{tikzpicture}[trees, scale=1.5, 210 | level 1/.style={sibling distance=20mm}, 211 | level 2/.style={sibling distance=10mm}, 212 | level 3/.style={sibling distance=5mm}, 213 | level 4/.style={sibling distance=2.5mm}, 214 | level 5/.style={sibling distance=1.25mm}] 215 | \node[dgreen] (a) {$\bullet$} 216 | child {node[dgreen] {$\bullet$} 217 | child {node[dgreen] {$\bullet$} 218 | child {node[dgreen] {$\bullet$} 219 | child {node[dgreen] {$\bullet$} 220 | child {} 221 | child {} 222 | } 223 | child {node[dyellow] {$\bullet$} 224 | child {} 225 | child {} 226 | } 227 | } 228 | child {node[dyellow] {$\bullet$} 229 | child {node[dgreen] {$\bullet$} 230 | child {} 231 | child {} 232 | } 233 | child {node[red] {$\bullet$}} 234 | } 235 | } 236 | child {node[dyellow] {$\bullet$} 237 | child {node[dgreen] {$\bullet$} 238 | child {node[dgreen] {$\bullet$} 239 | child {} 240 | child {} 241 | } 242 | child {node[dyellow] {$\bullet$} 243 | child {} 244 | child {} 245 | } 246 | } 247 | child {node[red] {$\bullet$}} 248 | } 249 | } 250 | child {node[dyellow] {$\bullet$} 251 | child {node[dgreen] {$\bullet$} 252 | child {node[dgreen] {$\bullet$} 253 | child {node[dgreen] {$\bullet$} 254 | child {} 255 | child {} 256 | } 257 | child {node[dyellow] {$\bullet$} 258 | child {} 259 | child {} 260 | } 261 | } 262 | child {node[dyellow] {$\bullet$} 263 | child {node[dgreen] {$\bullet$} 264 | child {} 265 | child {} 266 | } 267 | child {node[red] {$\bullet$}} 268 | } 269 | } 270 | child {node[red] {$\bullet$} 271 | } 272 | } 273 | ; 274 | \end{tikzpicture} 275 | } 276 | 277 | \newcommand{\coverpic}{ 278 | \begin{tikzpicture} 279 | \node {\treepic}; 280 | \draw[blue, densely dotted] (current bounding box.south west) rectangle 281 | (current bounding box.north east); 282 | \end{tikzpicture} 283 | }%{\earpic} 284 | 285 | 286 | \tikzset{ 287 | mapsto/.style={->, densely dashed, blue, shorten <=\short, shorten 288 | >=\short, >=stealth, thick}, 289 | short/.store in=\short, 290 | short=0pt 291 | } 292 | 293 | 294 | 295 | 296 | 297 | 298 | 299 | 300 | 301 | 302 | 303 | 304 | 305 | 306 | 307 | 308 | -------------------------------------------------------------------------------- /LMSN/P0-Preface-depricated.tex: -------------------------------------------------------------------------------- 1 | % !TeX root = P0-Intro.tex 2 | \documentclass[Book-Poly]{subfiles} 3 | \begin{document} 4 | 5 | 6 | %------------ Chapter ------------% 7 | \chapter*{Preface}\label{chapter.0} 8 | 9 | \begin{quote} 10 | The proposal is also intended to [serve] equally as a foundation for the academic, intellectual, and technological, on the one hand, and for the curious, the moral, the erotic, the political, the artistic, and the sheerly obstreperous, on the other.\\ 11 | \mbox{}\hfill ---Brian Cantwell Smith\\ 12 | \mbox{}\hfill \emph{On the Origin of Objects} 13 | \end{quote} 14 | 15 | %\begin{quote} 16 | %For me, though, it is difficult to resist the idea that space-time is not essentially different from matter, which we understand more deeply. If so, it will consist of vast numbers of identical units---``particles of space''---each in contact with a few neighbors, exchanging messages, joining and breaking apart, giving birth and passing away.\\ 17 | %\mbox{}\hfill ---Frank Wilczek\\ 18 | %\mbox{}\hfill \emph{Fundamentals} 19 | %\end{quote} 20 | 21 | \begin{quote} 22 | And that is the way to get the greatest possible variety, but with all the order there could be; i.e. it is the way to get as much perfection as there could be.\\ 23 | \mbox{}\hfill ---Gottfried Wilhelm Leibniz\\ 24 | \mbox{}\hfill \emph{Monadology}, 58. 25 | \end{quote} 26 | 27 | \noindent During the Fifth International Conference on Applied Category Theory in 2022, at least twelve of the fifty-nine presentations and two of the ten posters referenced the category of polynomial functors and dependent lenses or its close cousins (categories of optics and Dialectica categories) and the way they model diverse forms of interactive behavior. 28 | At the same time, all that is needed to grasp the construction of this category---called $\poly$ for short---is an understanding of mathematical sets and functions. 29 | There is no need for the theory and applications of polynomial functors to remain the stuff of technical papers; $\poly$ is far too versatile, too full of potential, to be kept out of reach. 30 | 31 | \section*{Purpose and prerequisites} 32 | 33 | A categorical theory of general interaction must be interdisciplinary by its very nature. 34 | Already, drafts of this text have been read by everyone from algebraic geometers to neuroscientists and AI developers. 35 | We hope to extend our reach ever further, to bring together thinkers and tinkerers from a diverse array of backgrounds under a common language by which to study interactive systems categorically. 36 | In short---we know about $\poly$; you know about other things; but only our collective knowledge can reveal how $\poly$ could be applied to those other things. 37 | 38 | As such, we have strived to write a friendly and accessible expository text that can serve as a stepping stone toward further investigations into polynomial functors. 39 | We include exercises and, crucially, solutions to guide the learning process; we draw extensive analogies to provide motivation and develop intuition; we pose examples whenever necessary. 40 | Proofs may bear far more detail than you would find in a research paper, but not so much detail that it would clutter the key ideas. 41 | A few critical proofs are even argued through pictures, yet we contend that they are no less rigorous than the clouds of notation whose places they take. 42 | 43 | On the other hand, there is some deep mathematical substance to the work we will discuss, drawing from the well-established theory of categories. 44 | Although you will find, for example, a complete proof of the Yoneda lemma within these pages, we don't intend to build up everything from scratch. 45 | There are plenty of excellent resources for learning category theory out there, catering to a variety of needs, without adding our own to the mix when our primary goal is to introduce $\poly$. 46 | So for the sake of contributing only what is genuinely helpful, we assume a certain level of mathematical background. 47 | You are ready to read this book if you can define the following fundamental concepts from category theory, and give examples of each: 48 | \begin{itemize} 49 | \item categories, 50 | \item functors, 51 | \item natural transformations, 52 | \item (co)limits, 53 | \item adjunctions, and 54 | \item (symmetric) monoidal categories. 55 | \end{itemize} 56 | We will additionally assume a passing familiarity with the language of graph-theoretic trees (e.g.\ vertices, roots, leaves, paths). 57 | 58 | That said, with a little investment on your part, you could very well use this book as a way to teach yourself some category theory. 59 | If you have ever tried to learn category theory, only to become lost in abstraction or otherwise overwhelmed by seemingly endless lists of examples from foreign fields, perhaps you will benefit from a focused case study of one particularly fruitful category. 60 | If you encounter terms or ideas that you would like to learn more about, we encourage you to look them up elsewhere, and you may find yourself spending a pleasant afternoon doing a deep dive into a new definition or theorem. 61 | Then come back when you're ready---we'll be here. 62 | 63 | \section*{Choices and conventions} 64 | 65 | Throughout this book, we have chosen to focus on polynomial functors of a single variable on the category of sets. 66 | The motivation for this seemingly narrow scope is twofold: to keep matters as concrete and intuitive as possible, with easy access to elements that we can work with directly; and to demonstrate the immense versatility of even this small corner of the theory. 67 | 68 | Below is a list of conventions we adopt; while it is not necessarily comprehensive, most unusual choices are justified within the text, often as a footnote. 69 | 70 | The natural numbers include $0$, so $\nn\coloneqq\{0,1,2,\ldots\}$. 71 | 72 | The names of categories will be capitalized. 73 | We will not focus so much on size issues, but roughly speaking small categories will be written in script (e.g.$\ \cat{C}, \cat{D}$), while large categories (usually, but not always, named) will be written in bold (e.g.\ $\poly, \Cat{C}$). 74 | We use $\smset$ to denote the category of (small) sets and functions and $\smcat$ to denote the category of (small) categories and functors. 75 | We use exponential notation $\cat{D}^{\cat{C}}$ to denote the category of functors $\cat{C}\to\cat{D}$ and natural transformations. 76 | 77 | We write either $c\in\Ob\cat{C}$ or $c\in\cat{C}$ to denote an object $c$ of a category $\cat{C}$. 78 | We use $\sum$ rather than $\coprod$ to denote coproducts. 79 | We denote the collection of morphisms $f\colon c\to d$ in a category $\cat{C}$ by using the name of the category itself, followed by the ordered pair of objects: $\cat{C}(c,d)$. 80 | We denote the domain of $f$ by $\dom f$ and the codomain of $f$ by $\cod f$. 81 | We use $\coloneqq$ for definitions and temporary assignments, as opposed to $=$ for identifications that can be observed and proven. 82 | We use $\iso$ to indicate an isomorphism of objects and $=$ to indicate an equality of objects, although the choice of the former does not preclude the possibility of the latter, nor does the latter generally indicate anything special beyond an arbitrary selection that has been made. 83 | We will freely use the definite article ``the'' to refer to objects that are unique only up to unique isomorphism. 84 | 85 | We list nullary operations before binary ones: for example, we denote a monoidal category $\cat{C}$ with monoidal unit $I$ and monoidal product $\odot$ by $(\cat{C},I,\odot)$, or say that $(I,\odot)$ is a monoidal structure on $\cat{C}$. 86 | 87 | \section*{Past, present, and future} 88 | 89 | The idea for this book began in 2020, originally as part of a joint work with David Jaz Myers on using categories to model dynamical systems. 90 | It soon became clear, however, that our writing and his---while intimately related---would be better off as separate volumes. 91 | His book is nonetheless an excellent companion to ours: see \cite{jaz}. 92 | 93 | In the summer of 2021, we taught a course on a draft of this book that was livestreamed from the Topos Institute. 94 | Lecture recordings are freely available at \url{https://topos.site/poly-course/}. 95 | 96 | The theory and application of polynomial functors comprise an active area of research. An excellent reference is Joachim Kock's notes 97 | \begin{center} 98 | \url{https://mat.uab.cat/~kock/cat/polynomial.pdf}, 99 | \end{center} 100 | which includes many ideas we cannot discuss here, such as polynomial functors in many variables. 101 | We have laid our own foundations here, and the work is still ongoing. 102 | Even as we were writing this, we were discovering new results and uses for polynomial functors, which only goes to show how bountiful $\poly$ can be in its rewards---but of course, we had to cut things off somewhere. 103 | We say this in the hope that you will keep the following in mind: where this book ends, the story will have just begun. 104 | 105 | \section*{Acknowledgments} 106 | 107 | Special thanks to David Jaz Myers: a brilliant colleague, a wonderful conversation partner, a congenial housemate, a superb chef, and an all-around good guy. 108 | 109 | Thanks go to John Baez, Eric Bond, Spencer Breiner, Kris Brown, Matteo Capucci, Valeria de Paiva, Joseph Dorta, Brendan Fong, Richard Garner, Bruno Gavranovi\'c, Neil Ghani, Ben Goertzel, Tim Hosgood, Samantha Jarvis, Max Lieblich, Shaowei Lin, Owen Lynch, Joachim Kock, J\'er\'emie Koenig, Sophie Libkind, Joshua Meyers, Dominic Orchard, Nathaniel Osgood, Evan Patterson, Brandon Shapiro, Juliet Szatko, Tish Tanski, Todd Trimble, Adam Vandervorst, Jonathan Weinberger, and Christian Williams. 110 | 111 | We also thank our anonymous reviewers, who's advice made this book much more coherent. 112 | 113 | Finally, we thank our sponsors at the Air Force Office of Scientific Research, who has supported this work (and category theory itself) since the very beginning. In particular, this material is based upon work supported by the AFOSR under award numbers FA9550-20-1-0348 and FA9550-23-1-0376. 114 | 115 | \end{document} 116 | -------------------------------------------------------------------------------- /P0-Preface.tex: -------------------------------------------------------------------------------- 1 | % !TeX root = P0-Intro.tex 2 | \documentclass[Book-Poly]{subfiles} 3 | \begin{document} 4 | 5 | 6 | %------------ Chapter ------------% 7 | \chapter*{Preface}\label{chapter.0} 8 | 9 | \begin{quote} 10 | The proposal is also intended to [serve] equally as a foundation for the academic, intellectual, and technological, on the one hand, and for the curious, the moral, the erotic, the political, the artistic, and the sheerly obstreperous, on the other.\\ 11 | \mbox{}\hfill ---Brian Cantwell Smith\\ 12 | \mbox{}\hfill \emph{On the Origin of Objects} 13 | \end{quote} 14 | 15 | %\begin{quote} 16 | %For me, though, it is difficult to resist the idea that space-time is not essentially different from matter, which we understand more deeply. If so, it will consist of vast numbers of identical units---``particles of space''---each in contact with a few neighbors, exchanging messages, joining and breaking apart, giving birth and passing away.\\ 17 | %\mbox{}\hfill ---Frank Wilczek\\ 18 | %\mbox{}\hfill \emph{Fundamentals} 19 | %\end{quote} 20 | 21 | \begin{quote} 22 | And that is the way to get the greatest possible variety, but with all the order there could be; i.e. it is the way to get as much perfection as there could be.\\ 23 | \mbox{}\hfill ---Gottfried Wilhelm Leibniz\\ 24 | \mbox{}\hfill \emph{Monadology}, 58. 25 | \end{quote} 26 | 27 | \noindent During the Fifth International Conference on Applied Category Theory in 2022, at least twelve of the fifty-nine presentations and two of the ten posters referenced the category of polynomial functors and dependent lenses or its close cousins (categories of optics and Dialectica categories) and the way they model diverse forms of interactive behavior. 28 | At the same time, all that is needed to grasp the construction of this category---called $\poly$ for short---is an understanding of mathematical sets and functions. 29 | There is no need for the theory and applications of polynomial functors to remain the stuff of technical papers; $\poly$ is far too versatile, too full of potential, to be kept out of reach. 30 | 31 | \index{applied category theory}\index{lens}\index{dialectica category} 32 | 33 | Informally, a polynomial functor is a collection of elements we call \emph{positions} and, for each position, a collection of elements we call \emph{directions}. 34 | There is then a natural notion of a morphism between polynomial functors that sends positions forward and directions backward, modeling two-way communication. 35 | From these basic components, category theory allows us to construct an immense array of mathematical gadgets that model a diverse range of interactive processes. 36 | In this book, we will establish the theory of polynomial functors and categorical constructions on them while exploring how they model interaction. 37 | 38 | \section{Purpose and prerequisites} 39 | 40 | A categorical theory of general interaction must be interdisciplinary by its very nature. 41 | Already, drafts of this text have been read by everyone from algebraic geometers to neuroscientists and AI developers. 42 | We hope to extend our reach ever further, to bring together thinkers and tinkerers from a diverse array of backgrounds under a common language by which to study interactive systems categorically. 43 | In short---we know about $\poly$; you know about other things; but only our collective knowledge can reveal how $\poly$ could be applied to those other things. 44 | 45 | \index{interdisciplinary} 46 | 47 | As such, we have strived to write a friendly and accessible expository text that can serve as a stepping stone toward further investigations into polynomial functors. 48 | We include exercises and, crucially, solutions to guide the learning process; we draw extensive analogies to provide motivation and develop intuition; we pose examples whenever necessary. 49 | Proofs may bear far more detail than you would find in a research paper, but not so much detail that it would clutter the key ideas. 50 | A few critical proofs are even argued through pictures, yet we contend that they are no less rigorous than the clouds of notation whose places they take. 51 | 52 | On the other hand, there is some deep mathematical substance to the work we will discuss, drawing from the well-established theory of categories. 53 | Although you will find, for example, a complete proof of the Yoneda lemma within these pages, we don't intend to build up everything from scratch. 54 | There are plenty of excellent resources for learning category theory out there, catering to a variety of needs, without adding our own to the mix when our primary goal is to introduce $\poly$. 55 | So for the sake of contributing only what is genuinely helpful, we assume a certain level of mathematical background. 56 | You are ready to read this book if you can define the following fundamental concepts from category theory, and give examples of each: 57 | \begin{itemize} 58 | \item categories, 59 | \item functors, 60 | \item natural transformations, 61 | \item (co)limits, 62 | \item adjunctions, and 63 | \item (symmetric) monoidal categories. 64 | \end{itemize} 65 | We will additionally assume a passing familiarity with the language of graph-theoretic trees (e.g.\ vertices, roots, leaves, paths). 66 | 67 | That said, with a little investment on your part, you could very well use this book as a way to teach yourself some category theory. 68 | If you have ever tried to learn category theory, only to become lost in abstraction or otherwise overwhelmed by seemingly endless lists of examples from foreign fields, perhaps you will benefit from a focused case study of one particularly fruitful category. 69 | If you encounter terms or ideas that you would like to learn more about, we encourage you to look them up elsewhere, and you may find yourself spending a pleasant afternoon doing a deep dive into a new definition or theorem. 70 | Then come back when you're ready---we'll be here. 71 | 72 | \section*{Outline} 73 | 74 | This book is designed to be read linearly. 75 | \cref{part.poly} introduces the category $\poly$ and illustrates how it models interaction protocols; while \cref{part.comon} highlights a crucial operation on $\poly$, the composition product, which upgrades the theory so that it properly captures the time evolution of a dynamical system. 76 | 77 | \bigskip 78 | 79 | \cref{part.poly} (The category of polynomial functors) consists of: 80 | \begin{itemize} 81 | \item \cref{ch.poly.rep-sets} (Representable functors from the category of sets), in which we review category-theoretic constructions on sets and the Yoneda lemma; 82 | \item \cref{ch.poly.obj} (Polynomial functors), in which we introduce our objects of study and present several perspectives from which to view them; 83 | \item \cref{ch.poly.cat} (The category of polynomial functors), in which we define $\poly$ by introducing its morphisms and demonstrate many ways to work with them, including how they model interaction protocols; 84 | \item \cref{ch.poly.dyn_sys} (Dynamical systems as dependent lenses), in which we use a specific class of morphisms in $\poly$ to model discrete-time dynamical systems; and 85 | \item \cref{ch.poly.bonus} (More categorical properties of polynomials), in which we describe a smorgasbord of additional category-theoretic structures on $\poly$. 86 | \end{itemize} 87 | \bigskip 88 | 89 | \cref{part.comon} (A different category of categories) consists of: 90 | \begin{itemize} 91 | \item \cref{ch.comon.comp} (The composition product), in which we examine a monoidal structure on $\poly$ given by substituting one polynomial into another; 92 | \item \cref{ch.comon.sharp} (Polynomial comonoids and retrofunctors), in which we show that the category of comonoids in $\poly$ with respect to the composition product is equivalent to a category of small categories we call $\smcat^\sharp$ whose morphisms are not functors; 93 | \item \cref{ch.comon.cofree} (Categorical properties of polynomial comonoids), in which we study the structure and utility of $\smcat^\sharp$; and 94 | \item \cref{ch.comon.vistas} (New horizons), in which we list open questions. 95 | \end{itemize} 96 | 97 | \section{Choices and conventions}\label{sec.choices} 98 | 99 | Throughout this book, we have chosen to focus on polynomial functors of a single variable on the category of sets. 100 | The motivation for this seemingly narrow scope is twofold: to keep matters as concrete and intuitive as possible, with easy access to elements that we can work with directly; and to demonstrate the immense versatility of even this small corner of the theory. Furthermore, the subject of multi-variate polynomials arises by considering what are called comonoids and comodules in $\poly$ \cite{spivak2023functorial}. 101 | 102 | \index{polynomial functor!multivariate}\index{multivariate polynomial|see{polynomial functor, multivariate}} 103 | 104 | Below is a list of conventions we adopt; while it is not comprehensive, any unusual choices are justified within the text, often as a footnote. 105 | 106 | The natural numbers include $0$, so $\nn\coloneqq\{0,1,2,\ldots\}$. Throughout this book, when referring to finite sets, we will adopt the following convention: $\0\coloneqq\{\}=\varnothing,$ $\1\coloneqq\{1\},$ $\2\coloneqq\{1,2\},$ $\3\coloneqq\{1,2,3\}$, and so on, with $\ord{n}\coloneqq\{1,\ldots,n\}$, an $n$-element set, for each natural number $n$. 107 | For example, in standard font, $5$ represents the usual natural number, while in sans serif font, $\5$ represents the $5$-element set $\{1,2,3,4,5\}$. 108 | When the same variable name appears in both italicized and sans serif fonts, the italicized variable denotes a natural number and the sans serif variable denotes the corresponding set; for example, if we state that $m\in\nn$, then we also understand $\ord{m}$ to mean the set $\{1,\ldots,m\}$. 109 | \index{set!finite} 110 | 111 | \index{natural numbers, $\nn$}\index{set!ordinal} 112 | 113 | The names of categories will be capitalized. 114 | We will mostly ignore size issues, but roughly speaking small categories will be written in script (e.g.$\ \cat{C}, \cat{D}$), while large categories (usually, but not always, named) will be written in bold (e.g.\ $\poly, \Cat{C}$). 115 | We use $\smset$ to denote the category of (small) sets and functions and $\smcat$ to denote the category of (small) categories and functors. 116 | We use exponential notation $\cat{D}^{\cat{C}}$ to denote the category of functors $\cat{C}\to\cat{D}$ and natural transformations. 117 | \index{set!category of sets}\index{functor!category of functors} 118 | 119 | We write either $c\in\Ob\cat{C}$ or $c\in\cat{C}$ to denote an object $c$ of a category $\cat{C}$. 120 | We use $\sum$ rather than $\coprod$ to denote coproducts. 121 | We denote the collection of morphisms $f\colon c\to d$ in a category $\cat{C}$ by using the name of the category itself, followed by the ordered pair of objects: $\cat{C}(c,d)$. 122 | We denote the domain of $f$ by $\dom f$ and the codomain of $f$ by $\cod f$. 123 | We use $\coloneqq$ for definitions and temporary assignments, as opposed to $=$ for identifications that can be observed and proven. 124 | We use $\iso$ to indicate an isomorphism of objects and $=$ to indicate an equality of objects, although the choice of the former does not preclude the possibility of the latter, nor does the latter necessarily imply any significance beyond an arbitrary selection that has been made. 125 | We will freely use the definite article ``the'' to refer to objects that are unique only up to unique isomorphism. 126 | 127 | \index{coproduct}\index{product}\index{isomorphism}\index{equality!use of} 128 | 129 | We list nullary operations before binary ones: for example, we denote a monoidal category $\cat{C}$ with monoidal unit $I$ and monoidal product $\odot$ by $(\cat{C},I,\odot)$, or say that $(I,\odot)$ is a monoidal structure on $\cat{C}$. 130 | 131 | \section{Past, present, and future} 132 | 133 | The idea for this book began in 2020, originally as part of a joint work with David Jaz Myers on using categories to model dynamical systems. 134 | It soon became clear, however, that our writing and his---while intimately related---would be better off as separate volumes. 135 | His book is nonetheless an excellent companion to ours: see \cite{jaz}. 136 | 137 | In the summer of 2021, we taught a course on a draft of this book that was livestreamed from the Topos Institute. 138 | Lecture recordings are freely available at \url{https://topos.site/poly-course/}. 139 | A follow-up workshop with additional recorded lectures and write-ups was held in early 2024; see \url{https://topos.site/events/poly-at-work/}. 140 | 141 | The theory and application of polynomial functors comprise an active area of research. 142 | We have laid the foundations here, but work is still ongoing. 143 | Even while writing this, we discovered new results and uses for polynomial functors, which only goes to show how bountiful $\poly$ can be in its rewards---but of course, we had to cut things off somewhere. 144 | We say this in the hope that you will keep the following in mind: where this book ends, the story will have just begun. 145 | 146 | \section*{Acknowledgments} 147 | 148 | Special thanks to David Jaz Myers: a brilliant colleague, a wonderful conversation partner, a congenial housemate, a superb chef, and an all-around good guy. 149 | 150 | Thanks go to John Baez, Eric Bond, Spencer Breiner, Kris Brown, Matteo Capucci, Valeria de Paiva, Joseph Dorta, Brendan Fong, Richard Garner, Bruno Gavranovi\'c, Neil Ghani, Ben Goertzel, Tim Hosgood, Samantha Jarvis, Max Lieblich, Shaowei Lin, Owen Lynch, Joachim Kock, J\'er\'emie Koenig, Sophie Libkind, Joshua Meyers, Dominic Orchard, Nathaniel Osgood, Evan Patterson, Brandon Shapiro, Juliet Szatko, Tish Tanski, Todd Trimble, Adam Vandervorst, Jonathan Weinberger, and Christian Williams. 151 | 152 | This material is based upon work supported by the AFOSR under award numbers FA9550-20-1-0348 and FA9550-23-1-0376. 153 | 154 | \end{document} 155 | -------------------------------------------------------------------------------- /Book-Poly.tex: -------------------------------------------------------------------------------- 1 | \documentclass[11pt]{memoir} 2 | 3 | \settrims{0pt}{0pt} % page and stock same size 4 | \settypeblocksize{*}{34.5pc}{*} % {height}{width}{ratio} 5 | \setlrmargins{*}{*}{1} % {spine}{edge}{ratio} 6 | \setulmarginsandblock{1in}{1in}{*} % height of typeblock computed 7 | \setheadfoot{\onelineskip}{2\onelineskip} % {headheight}{footskip} 8 | \setheaderspaces{*}{1.5\onelineskip}{*} % {headdrop}{headsep}{ratio} 9 | \checkandfixthelayout 10 | 11 | \chapterstyle{bianchi} 12 | \newcommand{\titlefont}{\normalfont\Huge\bfseries} 13 | \renewcommand{\chaptitlefont}{\titlefont} 14 | 15 | %-------- Packages --------% 16 | 17 | \usepackage{amsthm} 18 | \usepackage{mathtools} 19 | 20 | \usepackage{imakeidx} 21 | \usepackage[framemethod=tikz]{mdframed} 22 | \usepackage{footnote} 23 | \usepackage{tablefootnote} 24 | 25 | \usepackage[inline]{enumitem} 26 | \setlistdepth{6} 27 | \usepackage{ifthen} 28 | \usepackage[utf8]{inputenc} %allows non-ascii in bib file 29 | \usepackage{xcolor} 30 | 31 | \usepackage{subfiles} 32 | \usepackage[backend=biber, backref=true, maxbibnames = 10, style = alphabetic]{biblatex} 33 | \usepackage{xr-hyper} 34 | \usepackage[bookmarks=true, colorlinks=true, linkcolor=blue!50!black, 35 | citecolor=orange!50!black, urlcolor=orange!50!black, pdfencoding=unicode]{hyperref} 36 | \usepackage[capitalize]{cleveref} 37 | 38 | % Multirow tables 39 | \usepackage{multirow} 40 | 41 | % String manipulation 42 | \usepackage{xstring} 43 | 44 | % Exercise solutions 45 | \usepackage{answers} 46 | 47 | % Graphics and diagrams 48 | \usepackage{tikz} 49 | \usepackage{varwidth} 50 | \usepackage[prefix=tikzsym]{tikzsymbols} 51 | \usepackage{makecell}%database table thickness 52 | %\usepackage{graphicx} 53 | %\usepackage[all]{xy} 54 | \usepackage{multirow} 55 | 56 | 57 | % Fonts 58 | \usepackage{amssymb} 59 | \usepackage{newpxtext} 60 | \usepackage[varg,bigdelims]{newpxmath} 61 | \usepackage{mathrsfs} 62 | %\usepackage{minted} 63 | \usepackage{dutchcal} 64 | \usepackage{fontawesome} 65 | 66 | % For matrices with labeled columns and rows 67 | % \usepackage{kbordermatrix}% http://www.hss.caltech.edu/~kcb/TeX/kbordermatrix.sty 68 | %\usepackage{accents} 69 | %\usepackage{eucal} 70 | 71 | % MathVersion kludge 72 | \DeclareMathVersion{normal2} 73 | 74 | % For external graphics. 75 | \graphicspath{graphics/} 76 | 77 | %\usepackage{changepage} %indented paragraphs in solutions 78 | %\usepackage{showkeys} %for drafting; prints labels in margin 79 | 80 | %-------- Package setup --------% 81 | 82 | % cleveref % 83 | \newcommand{\creflastconjunction}{, and\nobreakspace} % serial comma 84 | \crefformat{enumi}{\##2#1#3} 85 | \crefmultiformat{enumi}{\##2#1#3}{ and~\##2#1#3}{, \##2#1#3}{, and~\##2#1#3} 86 | 87 | % biblatex % 88 | \addbibresource{refs.bib} 89 | 90 | % makeidx % 91 | \makeindex 92 | 93 | % hyperref % 94 | \hypersetup{final} 95 | 96 | % enumitem % 97 | \setlist{nosep} 98 | 99 | % footnote 100 | \makesavenoteenv{tabular} 101 | 102 | % minted 103 | %\newenvironment{agda} 104 | % {\VerbatimEnvironment 105 | % \begin{minted}[escapeinside=??, mathescape=true, frame=single, framesep=5pt, tabsize=1 106 | % ]{Agda}} 107 | % {\end{minted}} 108 | 109 | 110 | 111 | % tikz % 112 | \usetikzlibrary{ 113 | cd, 114 | math, 115 | decorations.markings, 116 | decorations.pathreplacing, 117 | positioning, 118 | arrows.meta, 119 | shapes, 120 | shadows, 121 | shadings, 122 | calc, 123 | fit, 124 | quotes, 125 | intersections, 126 | circuits, 127 | circuits.ee.IEC 128 | } 129 | 130 | \tikzcdset{arrow style=tikz, diagrams={>=To}} 131 | 132 | \input{tikz_stuff} 133 | 134 | \newcommand{\unalignedDo}[1]{\framebox{\begin{tikzpicture}[baseline=(current bounding box.center),every node/.style={inner sep=0pt, minimum size=0pt,outer sep=0pt}] 135 | \node(do) {\textbf{do}}; 136 | \coordinate (11) at ($(do.east)+(0.25cm,0)$); 137 | \foreach \line [count=\k] in {#1} 138 | { 139 | \pgfmathtruncatemacro{\kn}{\k+1} 140 | \node(\kn1)[anchor=north west] at ($(\k1.south west)-(0,0.25cm)$) {\line}; 141 | } 142 | \end{tikzpicture}}} 143 | 144 | %\usepackage{circuitikz} 145 | 146 | 147 | % mdframed/tablefootnote% 148 | % This makes \tablefootnote allow construction of footnotes that appear at bottom of page instead of inside frame 149 | 150 | \makeatletter 151 | \AfterEndEnvironment{mdframed}{% 152 | \tfn@tablefootnoteprintout% 153 | \gdef\tfn@fnt{0}% 154 | } 155 | \makeatother 156 | 157 | % Longer nested lists 158 | \newlist{longenum}{enumerate}{6} 159 | \setlist[longenum,1]{label=\arabic*.} 160 | \setlist[longenum,2]{label=\arabic*.} 161 | \setlist[longenum,3]{label=\arabic*.} 162 | \setlist[longenum,4]{label=\arabic*.} 163 | \setlist[longenum,5]{label=\arabic*.} 164 | \setlist[longenum,6]{label=\arabic*.} 165 | 166 | % Theorem environments 167 | 168 | % Colored box background colours 169 | 170 | \colorlet{theoremcolor}{white!92!blue} 171 | \colorlet{definitioncolor}{white!92!purple} 172 | \colorlet{examplecolor}{white!91!green} 173 | \colorlet{exercisecolor}{white!96!green} 174 | 175 | 176 | \mdfdefinestyle{theoremframe}{ 177 | linewidth=0pt, 178 | %linecolor=blue, 179 | backgroundcolor=theoremcolor, 180 | roundcorner=6pt, 181 | nobreak=true, 182 | leftmargin=0, 183 | innerleftmargin=0, 184 | rightmargin=0, 185 | innerrightmargin=0, 186 | } 187 | 188 | \mdfdefinestyle{definitionframe}{ 189 | linewidth=0pt, 190 | %linecolor=blue, 191 | backgroundcolor=definitioncolor, 192 | roundcorner=6pt, 193 | leftmargin=0, 194 | innerleftmargin=0, 195 | rightmargin=0, 196 | innerrightmargin=0, 197 | } 198 | 199 | \mdfdefinestyle{exampleframe}{ 200 | linewidth=0pt, 201 | backgroundcolor=examplecolor, 202 | leftmargin=0, 203 | innerleftmargin=0, 204 | rightmargin=0, 205 | innerrightmargin=0, 206 | } 207 | 208 | \mdfdefinestyle{exerciseframe}{ 209 | linewidth=0pt, 210 | backgroundcolor=exercisecolor, 211 | leftmargin=0, 212 | innerleftmargin=0, 213 | rightmargin=0, 214 | innerrightmargin=0, 215 | } 216 | 217 | \newtheoremstyle{plain} 218 | {-\topsep} %space above 219 | {} %space below 220 | {\normalfont} %body font 221 | {} %indent amount 222 | {\bfseries} %theorem head font 223 | {.} %punctuation after theorem head 224 | {.5em} %space after theorem head 225 | {} %theorem head spec 226 | 227 | % amsthm % 228 | \theoremstyle{plain} 229 | \newmdtheoremenv[style=theoremframe]{theorem}[equation]{Theorem} 230 | \newmdtheoremenv[style=theoremframe]{proposition}[equation]{Proposition} 231 | \newmdtheoremenv[style=theoremframe]{corollary}[equation]{Corollary} 232 | \newmdtheoremenv[style=theoremframe]{lemma}[equation]{Lemma} 233 | 234 | \theoremstyle{plain} 235 | \newmdtheoremenv[style=definitionframe]{definition}[equation]{Definition} 236 | \newmdtheoremenv[style=definitionframe]{informal}[equation]{Informal Definition} 237 | \newmdtheoremenv[style=definitionframe]{roughDef}[equation]{Rough Definition} 238 | \crefname{roughDef}{Definition}{Definitions} 239 | \newtheorem{construction}[equation]{Construction} 240 | \newtheorem{notation}[equation]{Notation} 241 | \newtheorem{axiom}{Axiom} 242 | \newtheorem*{axiom*}{Axiom} 243 | 244 | \theoremstyle{remark} 245 | \newtheorem{remark}[equation]{Remark} 246 | \newtheorem{warning}[equation]{Warning} 247 | 248 | % Solution environment 249 | 250 | \newcommand{\finishSolutionChapter}{ 251 | %\vfill\hrulefill\\\noindent 252 | %\arabic{solcounterlocal} exercises in Chapter \arabic{section}, and \arabic{solcounterglobal} total exercises up to this point. 253 | % 254 | \clearpage 255 | } 256 | 257 | \makeatletter 258 | \newcommand{\nolisttopbreak}{\nobreak\@afterheading} 259 | \makeatother 260 | 261 | % \newcounter{solcounterlocal}[section] 262 | % \newcounter{solcounterglobal} 263 | 264 | % \newcommand{\sol}[4][noprint]{ 265 | 266 | % \stepcounter{solcounterlocal}\stepcounter{solcounterglobal} 267 | 268 | % \noindent\ignorespacesafterend\emph{Solution to} \cref{#2}.% 269 | % \nopagebreak% 270 | % \ifthenelse{\equal{#1}{print}}{ 271 | % \nopagebreak% 272 | % \begin{mdframed}[backgroundcolor=examplecolor,linewidth=0pt]% 273 | % #3% 274 | % \end{mdframed}% 275 | % \nopagebreak 276 | % }{}% 277 | % \nolisttopbreak 278 | % \begin{description}[leftmargin=2.5ex,itemindent=0pt,topsep=0ex,nosep] 279 | % \item\nopagebreak 280 | % #4 281 | % \end{description} 282 | % \bigskip 283 | % } 284 | 285 | \Newassociation{solution}{Sol}{solutions} 286 | 287 | \renewenvironment{Sol}[1] 288 | { 289 | \hypertarget{sol:#1}{} 290 | \noindent\ignorespacesafterend{\emph{Solution to} Exercise \hyperlink{ex:#1}{#1}.} 291 | \nopagebreak 292 | \nolisttopbreak 293 | \begin{description}[leftmargin=2.5ex,itemindent=0pt,topsep=0ex,nosep] 294 | \item\nopagebreak 295 | } 296 | { 297 | \end{description} \bigskip 298 | } 299 | 300 | 301 | % Aligned tikz environment 302 | \newenvironment{altikz}{ 303 | \begin{aligned} 304 | \begin{tikzpicture} 305 | } 306 | { 307 | \end{tikzpicture} 308 | \end{aligned} 309 | } 310 | 311 | %Getting ending symbols in example and exercise environments: 312 | 313 | \newmdtheoremenv[style=exampleframe]{example}[equation]{Example} 314 | 315 | \newmdtheoremenv[style=exerciseframe]{exc-inner}[equation]{Exercise} 316 | \newenvironment{exercise}[1][]{ 317 | \def\qedsymbol{$\lozenge$}% Set the QED symbol. 318 | \pushQED{\qed} 319 | \begin{exc-inner}[#1\ifthenelse{\equal{#1}{}}{\hyperlink{sol:\theequation}{Solution here}}{; \hyperlink{sol:\theequation}{solution here}}]\hypertarget{ex:\theequation}{}~ 320 | }{ 321 | \popQED 322 | \end{exc-inner} 323 | } 324 | \crefname{exercise}{Exercise}{Exercises} 325 | 326 | 327 | % Adjunctions 328 | \newcommand{\adj}[5][30pt]{%[size] Cat L, Left, Right, Cat R. 329 | \begin{tikzcd}[ampersand replacement=\&, column sep=#1] 330 | #2\ar[r, shift left=7pt, "#3"] 331 | \ar[r, phantom, "\scriptstyle\Rightarrow"]\& 332 | #5\ar[l, shift left=7pt, "#4"] 333 | \end{tikzcd} 334 | } 335 | 336 | \newcommand{\adjr}[5][30pt]{%[size] Cat R, Right, Left, Cat L. 337 | \begin{tikzcd}[ampersand replacement=\&, column sep=#1] 338 | #2\ar[r, shift left=7pt, "#3"]\& 339 | #5\ar[l, shift left=7pt, "#4"] 340 | \ar[l, phantom, "\scriptstyle\Leftarrow"] 341 | \end{tikzcd} 342 | } 343 | 344 | \newcommand{\slogan}[1]{\begin{center}\textit{#1}\end{center}} 345 | 346 | %-------- Single symbols --------% 347 | 348 | \DeclareSymbolFont{stmry}{U}{stmry}{m}{n} 349 | \DeclareMathSymbol\fatsemi\mathop{stmry}{"23} 350 | 351 | \DeclareFontFamily{U}{mathx}{\hyphenchar\font45} 352 | \DeclareFontShape{U}{mathx}{m}{n}{ 353 | <5> <6> <7> <8> <9> <10> 354 | <10.95> <12> <14.4> <17.28> <20.74> <24.88> 355 | mathx10 356 | }{} 357 | \DeclareSymbolFont{mathx}{U}{mathx}{m}{n} 358 | \DeclareFontSubstitution{U}{mathx}{m}{n} 359 | \DeclareMathAccent{\widecheck}{0}{mathx}{"71} 360 | 361 | 362 | %-------- Renewed commands --------% 363 | 364 | \renewcommand{\ss}{\subseteq} 365 | 366 | %-------- Other Macros --------% 367 | 368 | 369 | \DeclarePairedDelimiter{\present}{\langle}{\rangle} 370 | \DeclarePairedDelimiter{\copair}{[}{]} 371 | \DeclarePairedDelimiter{\floor}{\lfloor}{\rfloor} 372 | \DeclarePairedDelimiter{\ceil}{\lceil}{\rceil} 373 | \DeclarePairedDelimiter{\corners}{\ulcorner}{\urcorner} 374 | \DeclarePairedDelimiter{\ihom}{[}{]} 375 | 376 | \newcommand{\lchom}[2]{\genfrac{[}{]}{0pt}{}{#1}{#2}} 377 | \newcommand{\rchom}[1]{\overset{#1}{\frown}} 378 | 379 | \DeclareMathOperator{\Hom}{Hom} 380 | \DeclareMathOperator{\Mor}{Mor} 381 | \DeclareMathOperator{\dom}{dom} 382 | \DeclareMathOperator{\cod}{cod} 383 | \DeclareMathOperator{\idy}{idy} 384 | \DeclareMathOperator{\comp}{com} 385 | \DeclareMathOperator*{\colim}{colim} 386 | \DeclareMathOperator{\im}{im} 387 | \DeclareMathOperator{\Ob}{Ob} 388 | \DeclareMathOperator{\Tr}{Tr} 389 | \DeclareMathOperator{\dju}{\sqcup} 390 | 391 | 392 | 393 | \newcommand{\const}[1]{\texttt{#1}}%a constant, or named element of a set 394 | \newcommand{\Set}[1]{\mathsf{#1}}%a named set 395 | \newcommand{\ord}[1]{\mathsf{#1}}%an ordinal 396 | \newcommand{\cat}[1]{\mathcal{#1}}%a generic category 397 | \newcommand{\Cat}[1]{\mathbf{#1}}%a named category 398 | \newcommand{\fun}[1]{\mathrm{#1}}%a function 399 | \newcommand{\Fun}[1]{\mathrm{#1}}%a named functor 400 | \newcommand{\Sys}[1]{\mathsf{#1}}%dynamical system 401 | \newcommand{\com}[1]{\mathscr{#1}}%polynomial comonoid 402 | \newcommand{\oper}[1]{\operatorname{#1}} 403 | 404 | % System shorthands: 405 | \newcommand{\State}[1]{\Set{State}_{\Sys{#1}}} 406 | \newcommand{\Out}[1]{\Set{Out}_{\Sys{#1}}} 407 | \newcommand{\In}[1]{\Set{In}_{\Sys{#1}}} 408 | \newcommand{\return}[1]{\fun{return}_{\Sys{#1}}} 409 | \newcommand{\update}[1]{\fun{update}_{\Sys{#1}}} 410 | 411 | \newcommand{\id}{\mathrm{id}} 412 | \newcommand{\then}{\mathbin{\fatsemi}} 413 | 414 | \newcommand{\cocolon}{:\!} 415 | 416 | 417 | \newcommand{\iso}{\cong} 418 | \newcommand{\too}{\longrightarrow} 419 | \newcommand{\tto}{\rightrightarrows} 420 | \newcommand{\To}[1]{\xrightarrow{#1}} 421 | \renewcommand{\Mapsto}[1]{\xmapsto{#1}} 422 | \newcommand{\Tto}[3][13pt]{\begin{tikzcd}[sep=#1, cramped, ampersand replacement=\&, text height=1ex, text depth=.3ex]\ar[r, shift left=2pt, "#2"]\ar[r, shift right=2pt, "#3"']\&{}\end{tikzcd}} 423 | \newcommand{\Too}[1]{\xrightarrow{\;\;#1\;\;}} 424 | \newcommand{\from}{\leftarrow} 425 | \newcommand{\From}[1]{\xleftarrow{#1}} 426 | \newcommand{\Fromm}[1]{\xleftarrow{\;\;#1\;\;}} 427 | \newcommand{\surj}{\twoheadrightarrow} 428 | \newcommand{\inj}{\hookrightarrow} 429 | \newcommand{\wavyto}{\rightsquigarrow} 430 | \newcommand{\lollipop}{\multimap} 431 | \newcommand{\tickar}{\xtickar{}} 432 | \newcommand{\imp}{\Rightarrow} 433 | \renewcommand{\iff}{\Leftrightarrow} 434 | \newcommand{\down}{\mathbin{\downarrow}} 435 | \newcommand{\fromto}{\leftrightarrows} 436 | \newcommand{\xtickar}[1]{\begin{tikzcd}[baseline=-0.5ex,cramped,sep=small,ampersand 437 | replacement=\&]{}\ar[r,tick, "{#1}"]\&{}\end{tikzcd}} 438 | 439 | 440 | 441 | \newcommand{\inv}{^{-1}} 442 | \newcommand{\op}{^\tn{op}} 443 | 444 | \newcommand{\tn}[1]{\textnormal{#1}} 445 | \newcommand{\ol}[1]{\overline{#1}} 446 | \newcommand{\ul}[1]{\underline{#1}} 447 | \newcommand{\wt}[1]{\widetilde{#1}} 448 | \newcommand{\wh}[1]{\widehat{#1}} 449 | \newcommand{\wc}[1]{\widecheck{#1}} 450 | \newcommand{\ubar}[1]{\underaccent{\bar}{#1}} 451 | 452 | 453 | % OLD LMO: \newcommand{\LMO}[2][over]{\ifthenelse{\equal{#1}{over}}{\overset{#2}{\bullet}}{\underset{#2}{\bullet}}} 454 | \newcommand{\LMO}[2][over]{\LMOO[#1]{\phantom{\cdot}}{#2}} 455 | %% Convention for LMOO is 456 | %% {optional: over/under}{state name}{output name} 457 | \newcommand{\LMOO}[3][over]{\ifthenelse{\equal{#1}{over}} 458 | { 459 | \begin{tikzpicture}[inner sep=0pt, solid] \node[draw, circle, label={[label distance 460 | = 2pt]90:{$#3$}}] {$#2$}; \end{tikzpicture} 461 | }{ 462 | \begin{tikzpicture}[inner sep=0pt, solid] \node[draw, circle, label={[label distance 463 | = 2pt]270:{$#3$}}] {$#2$}; \end{tikzpicture} 464 | } 465 | } 466 | \newcommand{\LTO}[2][\bullet]{\overset{\tn{#2}}{#1}} 467 | 468 | 469 | \newcommand{\bb}{\mathbb{B}} 470 | \newcommand{\nn}{\mathbb{N}} 471 | \newcommand{\pp}{\mathbb{P}} 472 | \newcommand{\qq}{\mathbb{Q}} 473 | \newcommand{\zz}{\mathbb{Z}} 474 | \newcommand{\rr}{\mathbb{R}} 475 | 476 | \newcommand{\singleton}{\{1\}} 477 | \newcommand{\powset}{\Fun{P}} 478 | \newcommand{\probset}{\Fun{D}} 479 | 480 | 481 | \newcommand{\smset}{\Cat{Set}} 482 | \newcommand{\smcat}{\Cat{Cat}} 483 | \newcommand{\catsharp}{\smcat^\sharp} 484 | \newcommand{\comon}{\Cat{Comon}} 485 | \newcommand{\nat}{\Set{Nat}} 486 | \newcommand{\finset}{\Cat{FinSet}} 487 | \newcommand{\smctx}[1]{\Cat{Ctx}_{#1}} 488 | 489 | \newcommand{\List}{\Fun{List}} 490 | \newcommand{\set}{\tn{-}\Cat{Set}} 491 | 492 | 493 | 494 | \newcommand{\boxCD}[3][black] % [border color] {interior color} {math text} 495 | {\fcolorbox{#1}{#2}{\begin{varwidth}{\textwidth}\centering #3\end{varwidth}}} 496 | 497 | \newcommand{\erase}[2][]{{\color{red}#1}} 498 | \newcommand{\showhide}[2]{#1} 499 | 500 | 501 | % Spivak macros % 502 | 503 | \newenvironment{dedication} 504 | {\begin{flushright}\itshape} 505 | {\end{flushright}} 506 | 507 | \newcommand{\yon}{\mathcal{y}} 508 | \newcommand{\poly}[1][]{#1\Cat{Poly}} 509 | \newcommand{\0}{\textsf{0}} 510 | \newcommand{\1}{\textsf{1}} 511 | \newcommand{\2}{\textsf{2}} 512 | \newcommand{\3}{\textsf{3}} 513 | \newcommand{\4}{\textsf{4}} 514 | \newcommand{\5}{\textsf{5}} 515 | \newcommand{\6}{\textsf{6}} 516 | \newcommand{\7}{\textsf{7}} 517 | \newcommand{\8}{\textsf{8}} 518 | \newcommand{\9}{\textsf{9}} 519 | 520 | \definecolor{dgreen}{rgb}{0.0, 0.5, 0.3} 521 | \definecolor{dyellow}{rgb}{8.0, 0.74, 0} 522 | 523 | \newcommand{\pphantom}[2]{$\phantom{#1}$\nodepart{two}$\vphantom{#2}$} 524 | \newcommand{\at}{\nodepart{two}} 525 | 526 | \newcommand{\qand}{\quad\text{and}\quad} 527 | \newcommand{\qqand}{\qquad\text{and}\qquad} 528 | \newcommand{\qqor}{\qquad\text{or}\qquad} 529 | 530 | \newcommand{\bhline}{\Xhline{2\arrayrulewidth}} 531 | \newcommand{\bbhline}{\Xhline{2.5\arrayrulewidth}} 532 | \newcommand{\bbbhline}{\Xhline{3\arrayrulewidth}} 533 | 534 | \newcommand{\lst}{\Set{List}} 535 | 536 | \newcommand{\tr}{\Set{tree}} 537 | \newcommand{\vtx}{\Set{vtx}} 538 | \newcommand{\rt}{\fun{root}} 539 | \newcommand{\fs}{\fun{focus}} 540 | \newcommand{\lensput}{\fun{put}} 541 | \newcommand{\lensget}{\fun{get}} 542 | \newcommand{\deck}{\Fun{Deck}} 543 | 544 | \newcommand{\src}{\fun{src}} 545 | \newcommand{\tgt}{\fun{tgt}} 546 | 547 | \newcommand{\elts}{\textstyle{\int}} 548 | 549 | \newcommand{\mdot}{{\cdot}} 550 | \newcommand{\cofree}[1]{\com{T}_{#1}} 551 | \newcommand{\coto}{\nrightarrow} 552 | \newcommand{\cofun}{{\raisebox{2pt}{\resizebox{2.5pt}{2.5pt}{$\setminus$}}}} 553 | 554 | \newcommand{\lcmod}[1]{~_{#1}\Cat{Mod}} 555 | \newcommand{\rcmod}[1]{~\Cat{Mod}_{#1}} 556 | \newcommand{\bimod}[2]{~_{#1}\Cat{Mod}_{#2}} 557 | 558 | \newcommand{\car}[1]{\mathfrak{#1}} 559 | 560 | \newcommand{\cof}{\nrightarrow} 561 | 562 | \newcommand{\fl}[1]{\acute{#1}} 563 | 564 | \newcommand{\tri}{\mathbin{\triangleleft}} 565 | \newcommand{\tripow}[1]{^{\tri\,#1}} 566 | 567 | \newcommand{\bul}[1][black]{{\color{#1}\ensuremath{\bullet}}} 568 | 569 | \makeatletter 570 | \newcommand*{\bdiv}{% 571 | \nonscript\mskip-\medmuskip\mkern5mu% 572 | \mathbin{\operator@font div}\penalty900\mkern5mu% 573 | \nonscript\mskip-\medmuskip 574 | } 575 | \makeatother 576 | 577 | \newcommand{\centergraphics}[2][width=\textwidth]{ 578 | \begin{center} 579 | \includegraphics[#1]{#2} 580 | \end{center} 581 | } 582 | 583 | % for making notes in text: 584 | 585 | \newcounter{notecounter} 586 | 587 | \newcommand{\nnote}[1]{ 588 | \stepcounter{notecounter} 589 | \textcolor{red}{\;$\lozenge$\;\textbf{Nelson says}: {#1}\;$\lozenge$\;} 590 | } 591 | \newcommand{\dnote}[1]{ 592 | \stepcounter{notecounter} 593 | \textcolor{blue}{\;$\lozenge$\;\textbf{David says}: {#1}\;$\lozenge$\;} 594 | } 595 | 596 | \newcommand{\numnotes}{\clearpage\huge Number of todo notes in text, at last pass: \thenotecounter\normalsize} 597 | 598 | \newcommand{\sys}{\Cat{Sys}} 599 | 600 | % Booleans 601 | \newcommand{\true}{\const{true}} 602 | \newcommand{\false}{\const{false}} 603 | 604 | % Colors 605 | \newcommand{\Red}{{\color{red} \const{red}}} 606 | \newcommand{\Blue}{{\color{blue} \const{blue}}} 607 | \newcommand{\Green}{{\color{dgreen}\const{green}} } 608 | \newcommand{\Orange}{{\color{orange} \const{orange}}} 609 | 610 | % Grothendieck double construction: 611 | \newcommand{\rectangle}{{% 612 | \Large 613 | \ooalign{$\sqsubset\mkern3mu$\cr$\mkern3mu\sqsupset$\cr}% 614 | }} 615 | \newcommand{\sqiint}{{\rectangle}\kern-1.4em{\iint}} 616 | 617 | %% vdots without space at the top. 618 | \makeatletter 619 | \DeclareRobustCommand{\rvdots}{% 620 | \vbox{ 621 | \baselineskip4\p@\lineskiplimit\z@ 622 | \kern-\p@ 623 | \hbox{.}\hbox{.}\hbox{.} 624 | }} 625 | \makeatother 626 | 627 | % ---- Changeable document parameters ---- % 628 | 629 | \linespread{1.15} 630 | \allowdisplaybreaks 631 | \setsecnumdepth{subsection} 632 | \settocdepth{subsection} 633 | \setlength{\parindent}{15pt} 634 | \setcounter{tocdepth}{1} 635 | 636 | 637 | 638 | %--------------- Document ---------------% 639 | \begin{document} 640 | 641 | \frontmatter 642 | 643 | \title{\titlefont Polynomial Functors:\\\medskip 644 | A Mathematical Theory of Interaction} 645 | 646 | \author{ 647 | \LARGE Nelson Niu 648 | \and 649 | \LARGE David I. Spivak 650 | \normalsize} 651 | 652 | \posttitle{ 653 | \vspace{.8in} 654 | \normalsize 655 | \[ 656 | \coverpic 657 | \] 658 | \vspace{.5in} 659 | \endgroup 660 | } 661 | \date{\vfill Last updated: \today \\ Source: \url{https://github.com/ToposInstitute/poly}} 662 | 663 | \maketitle 664 | 665 | \thispagestyle{empty} 666 | \clearpage 667 | This page intentionally left blank. 668 | \clearpage 669 | \clearpage 670 | 671 | \mbox{} 672 | \vspace{2in} 673 | \begin{center} 674 | \begin{tabular}{lll} 675 | \LARGE Nelson Niu&~\hspace{.75in}~&\LARGE David I. Spivak\\ 676 | \large University of Washington&&\large Topos Institute\\ 677 | \large Seattle, WA&&\large Berkeley, CA 678 | \end{tabular} 679 | \end{center} 680 | \clearpage 681 | 682 | \begin{dedication} 683 | To Andr\'e Joyal 684 | \smallskip 685 | 686 | ---D.S. 687 | \bigskip 688 | 689 | To my graduate cohort at UW 690 | \smallskip 691 | 692 | ---N.N. 693 | \end{dedication} 694 | 695 | \subfile{P0-Preface} 696 | 697 | \clearpage 698 | \tableofcontents* 699 | \clearpage 700 | 701 | \mainmatter 702 | \subfile{P1-Polynomials} 703 | \subfile{P2-Comonoids} 704 | \subfile{P3-Vistas} 705 | 706 | \appendix 707 | \begingroup 708 | \footnotesize 709 | 710 | \backmatter 711 | 712 | \printbibliography 713 | \printindex 714 | 715 | \numnotes 716 | 717 | \end{document} -------------------------------------------------------------------------------- /LMSN/Book-Poly-depricated.tex: -------------------------------------------------------------------------------- 1 | \documentclass[10pt]{memoir} 2 | 3 | \settrims{0pt}{0pt} % page and stock same size 4 | \settypeblocksize{190mm}{110mm}{*} % {height}{width}{ratio} 5 | \setlrmargins{*}{*}{1} % {spine}{edge}{ratio} 6 | \setulmarginsandblock{1in}{1in}{*} % height of typeblock computed 7 | \setheadfoot{\onelineskip}{2\onelineskip} % {headheight}{footskip} 8 | \setheaderspaces{*}{1.5\onelineskip}{*} % {headdrop}{headsep}{ratio} 9 | \checkandfixthelayout 10 | 11 | \chapterstyle{bianchi} 12 | \newcommand{\titlefont}{\normalfont\Huge\bfseries} 13 | \renewcommand{\chaptitlefont}{\titlefont} 14 | 15 | %-------- Packages --------% 16 | 17 | \usepackage{amsthm} 18 | \usepackage{mathtools} 19 | 20 | \usepackage{makeidx} 21 | \usepackage[framemethod=tikz]{mdframed} 22 | \usepackage{footnote} 23 | \usepackage{tablefootnote} 24 | 25 | \usepackage[inline]{enumitem} 26 | \setlistdepth{6} 27 | \usepackage{ifthen} 28 | \usepackage[utf8]{inputenc} %allows non-ascii in bib file 29 | \usepackage{xcolor} 30 | 31 | \usepackage{subfiles} 32 | \usepackage[backend=biber, backref=true, maxbibnames = 10, style = alphabetic]{biblatex} 33 | \usepackage{xr-hyper} 34 | \usepackage[bookmarks=true, colorlinks=true, linkcolor=blue!50!black, 35 | citecolor=orange!50!black, urlcolor=orange!50!black, pdfencoding=unicode]{hyperref} 36 | \usepackage[capitalize]{cleveref} 37 | 38 | % Multirow tables 39 | \usepackage{multirow} 40 | 41 | % String manipulation 42 | \usepackage{xstring} 43 | 44 | % Exercise solutions 45 | \usepackage{answers} 46 | 47 | % Graphics and diagrams 48 | \usepackage{tikz} 49 | \usepackage{varwidth} 50 | \usepackage[prefix=tikzsym]{tikzsymbols} 51 | \usepackage{makecell}%database table thickness 52 | %\usepackage{graphicx} 53 | %\usepackage[all]{xy} 54 | \usepackage{multirow} 55 | 56 | 57 | % Fonts 58 | \usepackage{amssymb} 59 | \usepackage{newpxtext} 60 | \usepackage[varg,bigdelims]{newpxmath} 61 | \usepackage{mathrsfs} 62 | %\usepackage{minted} 63 | \usepackage{dutchcal} 64 | \usepackage{fontawesome} 65 | 66 | % For matrices with labeled columns and rows 67 | % \usepackage{kbordermatrix}% http://www.hss.caltech.edu/~kcb/TeX/kbordermatrix.sty 68 | %\usepackage{accents} 69 | %\usepackage{eucal} 70 | 71 | % MathVersion kludge 72 | \DeclareMathVersion{normal2} 73 | 74 | % For external graphics. 75 | \graphicspath{graphics/} 76 | 77 | %\usepackage{changepage} %indented paragraphs in solutions 78 | %\usepackage{showkeys} %for drafting; prints labels in margin 79 | 80 | %-------- Package setup --------% 81 | 82 | % cleveref % 83 | \newcommand{\creflastconjunction}{, and\nobreakspace} % serial comma 84 | \crefformat{enumi}{\##2#1#3} 85 | \crefmultiformat{enumi}{\##2#1#3}{ and~\##2#1#3}{, \##2#1#3}{, and~\##2#1#3} 86 | 87 | % biblatex % 88 | \addbibresource{refs.bib} 89 | 90 | % makeidx % 91 | \makeindex 92 | 93 | % hyperref % 94 | \hypersetup{final} 95 | 96 | % enumitem % 97 | \setlist{nosep} 98 | 99 | % footnote 100 | \makesavenoteenv{tabular} 101 | 102 | % minted 103 | %\newenvironment{agda} 104 | % {\VerbatimEnvironment 105 | % \begin{minted}[escapeinside=??, mathescape=true, frame=single, framesep=5pt, tabsize=1 106 | % ]{Agda}} 107 | % {\end{minted}} 108 | 109 | 110 | 111 | % tikz % 112 | \usetikzlibrary{ 113 | cd, 114 | math, 115 | decorations.markings, 116 | decorations.pathreplacing, 117 | positioning, 118 | arrows.meta, 119 | shapes, 120 | shadows, 121 | shadings, 122 | calc, 123 | fit, 124 | quotes, 125 | intersections, 126 | circuits, 127 | circuits.ee.IEC 128 | } 129 | 130 | \tikzcdset{arrow style=tikz, diagrams={>=To}} 131 | 132 | \input{tikz_stuff} 133 | 134 | \newcommand{\unalignedDo}[1]{\framebox{\begin{tikzpicture}[baseline=(current bounding box.center),every node/.style={inner sep=0pt, minimum size=0pt,outer sep=0pt}] 135 | \node(do) {\textbf{do}}; 136 | \coordinate (11) at ($(do.east)+(0.25cm,0)$); 137 | \foreach \line [count=\k] in {#1} 138 | { 139 | \pgfmathtruncatemacro{\kn}{\k+1} 140 | \node(\kn1)[anchor=north west] at ($(\k1.south west)-(0,0.25cm)$) {\line}; 141 | } 142 | \end{tikzpicture}}} 143 | 144 | %\usepackage{circuitikz} 145 | 146 | 147 | % mdframed/tablefootnote% 148 | % This makes \tablefootnote allow construction of footnotes that appear at bottom of page instead of inside frame 149 | 150 | \makeatletter 151 | \AfterEndEnvironment{mdframed}{% 152 | \tfn@tablefootnoteprintout% 153 | \gdef\tfn@fnt{0}% 154 | } 155 | \makeatother 156 | 157 | % Longer nested lists 158 | \newlist{longenum}{enumerate}{6} 159 | \setlist[longenum,1]{label=\arabic*.} 160 | \setlist[longenum,2]{label=\arabic*.} 161 | \setlist[longenum,3]{label=\arabic*.} 162 | \setlist[longenum,4]{label=\arabic*.} 163 | \setlist[longenum,5]{label=\arabic*.} 164 | \setlist[longenum,6]{label=\arabic*.} 165 | 166 | % Theorem environments 167 | 168 | % Colored box background colours 169 | 170 | \colorlet{theoremcolor}{white!52!blue} 171 | \colorlet{definitioncolor}{white!52!purple} 172 | \colorlet{examplecolor}{white!77!green} 173 | \colorlet{exercisecolor}{white!87!green} 174 | 175 | 176 | \mdfdefinestyle{theoremframe}{ 177 | linewidth=2pt, 178 | linecolor=theoremcolor, 179 | backgroundcolor=white, 180 | roundcorner=6pt, 181 | nobreak=true, 182 | leftmargin=0, 183 | innerleftmargin=0, 184 | rightmargin=0, 185 | innerrightmargin=0, 186 | } 187 | 188 | \mdfdefinestyle{definitionframe}{ 189 | linewidth=2pt, 190 | linecolor=definitioncolor, 191 | backgroundcolor=white, 192 | roundcorner=6pt, 193 | leftmargin=0, 194 | innerleftmargin=0, 195 | rightmargin=0, 196 | innerrightmargin=0, 197 | } 198 | 199 | \mdfdefinestyle{exampleframe}{ 200 | linewidth=2pt, 201 | linecolor=examplecolor, 202 | backgroundcolor=white, 203 | leftmargin=0, 204 | innerleftmargin=0, 205 | rightmargin=0, 206 | innerrightmargin=0, 207 | } 208 | 209 | \mdfdefinestyle{exerciseframe}{ 210 | linewidth=2pt, 211 | linecolor=exercisecolor, 212 | backgroundcolor=white, 213 | leftmargin=0, 214 | innerleftmargin=0, 215 | rightmargin=0, 216 | innerrightmargin=0, 217 | } 218 | 219 | \newtheoremstyle{plain} 220 | {-\topsep} %space above 221 | {} %space below 222 | {\normalfont} %body font 223 | {} %indent amount 224 | {\bfseries} %theorem head font 225 | {.} %punctuation after theorem head 226 | {.5em} %space after theorem head 227 | {} %theorem head spec 228 | 229 | % amsthm % 230 | \theoremstyle{plain} 231 | \newmdtheoremenv[style=theoremframe]{theorem}[equation]{Theorem} 232 | \newmdtheoremenv[style=theoremframe]{proposition}[equation]{Proposition} 233 | \newmdtheoremenv[style=theoremframe]{corollary}[equation]{Corollary} 234 | \newmdtheoremenv[style=theoremframe]{lemma}[equation]{Lemma} 235 | 236 | \theoremstyle{plain} 237 | \newmdtheoremenv[style=definitionframe]{definition}[equation]{Definition} 238 | \newmdtheoremenv[style=definitionframe]{informal}[equation]{Informal Definition} 239 | \newmdtheoremenv[style=definitionframe]{roughDef}[equation]{Rough Definition} 240 | \crefname{roughDef}{Definition}{Definitions} 241 | \newtheorem{construction}[equation]{Construction} 242 | \newtheorem{notation}[equation]{Notation} 243 | \newtheorem{axiom}{Axiom} 244 | \newtheorem*{axiom*}{Axiom} 245 | 246 | \theoremstyle{remark} 247 | \newtheorem{remark}[equation]{Remark} 248 | \newtheorem{warning}[equation]{Warning} 249 | 250 | % Solution environment 251 | 252 | \newcommand{\finishSolutionChapter}{ 253 | %\vfill\hrulefill\\\noindent 254 | %\arabic{solcounterlocal} exercises in Chapter \arabic{section}, and \arabic{solcounterglobal} total exercises up to this point. 255 | % 256 | \clearpage 257 | } 258 | 259 | \makeatletter 260 | \newcommand{\nolisttopbreak}{\nobreak\@afterheading} 261 | \makeatother 262 | 263 | % \newcounter{solcounterlocal}[section] 264 | % \newcounter{solcounterglobal} 265 | 266 | % \newcommand{\sol}[4][noprint]{ 267 | 268 | % \stepcounter{solcounterlocal}\stepcounter{solcounterglobal} 269 | 270 | % \noindent\ignorespacesafterend\emph{Solution to} \cref{#2}.% 271 | % \nopagebreak% 272 | % \ifthenelse{\equal{#1}{print}}{ 273 | % \nopagebreak% 274 | % \begin{mdframed}[backgroundcolor=examplecolor,linewidth=0pt]% 275 | % #3% 276 | % \end{mdframed}% 277 | % \nopagebreak 278 | % }{}% 279 | % \nolisttopbreak 280 | % \begin{description}[leftmargin=2.5ex,itemindent=0pt,topsep=0ex,nosep] 281 | % \item\nopagebreak 282 | % #4 283 | % \end{description} 284 | % \bigskip 285 | % } 286 | 287 | \Newassociation{solution}{Sol}{solutions} 288 | 289 | \renewenvironment{Sol}[1] 290 | { 291 | \hypertarget{sol:#1}{} 292 | \noindent\ignorespacesafterend{\emph{Solution to} Exercise \hyperlink{ex:#1}{#1}.} 293 | \nopagebreak 294 | \nolisttopbreak 295 | \begin{description}[leftmargin=2.5ex,itemindent=0pt,topsep=0ex,nosep] 296 | \item\nopagebreak 297 | } 298 | { 299 | \end{description} \bigskip 300 | } 301 | 302 | 303 | % Aligned tikz environment 304 | \newenvironment{altikz}{ 305 | \begin{aligned} 306 | \begin{tikzpicture} 307 | } 308 | { 309 | \end{tikzpicture} 310 | \end{aligned} 311 | } 312 | 313 | %Getting ending symbols in example and exercise environments: 314 | 315 | \newmdtheoremenv[style=exampleframe]{example}[equation]{Example} 316 | 317 | \newmdtheoremenv[style=exerciseframe]{exc-inner}[equation]{Exercise} 318 | \newenvironment{exercise}[1][]{ 319 | \def\qedsymbol{$\lozenge$}% Set the QED symbol. 320 | \pushQED{\qed} 321 | \begin{exc-inner}[#1\ifthenelse{\equal{#1}{}}{\hyperlink{sol:\theequation}{Solution here}}{; \hyperlink{sol:\theequation}{solution here}}]\hypertarget{ex:\theequation}{}~ 322 | }{ 323 | \popQED 324 | \end{exc-inner} 325 | } 326 | \crefname{exercise}{Exercise}{Exercises} 327 | 328 | 329 | % Adjunctions 330 | \newcommand{\adj}[5][30pt]{%[size] Cat L, Left, Right, Cat R. 331 | \begin{tikzcd}[ampersand replacement=\&, column sep=#1] 332 | #2\ar[r, shift left=7pt, "#3"] 333 | \ar[r, phantom, "\scriptstyle\Rightarrow"]\& 334 | #5\ar[l, shift left=7pt, "#4"] 335 | \end{tikzcd} 336 | } 337 | 338 | \newcommand{\adjr}[5][30pt]{%[size] Cat R, Right, Left, Cat L. 339 | \begin{tikzcd}[ampersand replacement=\&, column sep=#1] 340 | #2\ar[r, shift left=7pt, "#3"]\& 341 | #5\ar[l, shift left=7pt, "#4"] 342 | \ar[l, phantom, "\scriptstyle\Leftarrow"] 343 | \end{tikzcd} 344 | } 345 | 346 | \newcommand{\slogan}[1]{\begin{center}\textit{#1}\end{center}} 347 | 348 | %-------- Single symbols --------% 349 | 350 | \DeclareSymbolFont{stmry}{U}{stmry}{m}{n} 351 | \DeclareMathSymbol\fatsemi\mathop{stmry}{"23} 352 | 353 | \DeclareFontFamily{U}{mathx}{\hyphenchar\font45} 354 | \DeclareFontShape{U}{mathx}{m}{n}{ 355 | <5> <6> <7> <8> <9> <10> 356 | <10.95> <12> <14.4> <17.28> <20.74> <24.88> 357 | mathx10 358 | }{} 359 | \DeclareSymbolFont{mathx}{U}{mathx}{m}{n} 360 | \DeclareFontSubstitution{U}{mathx}{m}{n} 361 | \DeclareMathAccent{\widecheck}{0}{mathx}{"71} 362 | 363 | 364 | %-------- Renewed commands --------% 365 | 366 | \renewcommand{\ss}{\subseteq} 367 | 368 | %-------- Other Macros --------% 369 | 370 | 371 | \DeclarePairedDelimiter{\present}{\langle}{\rangle} 372 | \DeclarePairedDelimiter{\copair}{[}{]} 373 | \DeclarePairedDelimiter{\floor}{\lfloor}{\rfloor} 374 | \DeclarePairedDelimiter{\ceil}{\lceil}{\rceil} 375 | \DeclarePairedDelimiter{\corners}{\ulcorner}{\urcorner} 376 | \DeclarePairedDelimiter{\ihom}{[}{]} 377 | 378 | \newcommand{\lchom}[2]{\genfrac{[}{]}{0pt}{}{#1}{#2}} 379 | \newcommand{\rchom}[1]{\overset{#1}{\frown}} 380 | 381 | \DeclareMathOperator{\Hom}{Hom} 382 | \DeclareMathOperator{\Mor}{Mor} 383 | \DeclareMathOperator{\dom}{dom} 384 | \DeclareMathOperator{\cod}{cod} 385 | \DeclareMathOperator{\idy}{idy} 386 | \DeclareMathOperator{\comp}{com} 387 | \DeclareMathOperator*{\colim}{colim} 388 | \DeclareMathOperator{\im}{im} 389 | \DeclareMathOperator{\Ob}{Ob} 390 | \DeclareMathOperator{\Tr}{Tr} 391 | \DeclareMathOperator{\dju}{\sqcup} 392 | 393 | 394 | 395 | \newcommand{\const}[1]{\texttt{#1}}%a constant, or named element of a set 396 | \newcommand{\Set}[1]{\mathsf{#1}}%a named set 397 | \newcommand{\ord}[1]{\mathsf{#1}}%an ordinal 398 | \newcommand{\cat}[1]{\mathcal{#1}}%a generic category 399 | \newcommand{\Cat}[1]{\mathbf{#1}}%a named category 400 | \newcommand{\fun}[1]{\mathrm{#1}}%a function 401 | \newcommand{\Fun}[1]{\mathrm{#1}}%a named functor 402 | \newcommand{\Sys}[1]{\mathsf{#1}}%dynamical system 403 | \newcommand{\com}[1]{\mathscr{#1}}%polynomial comonoid 404 | 405 | % System shorthands: 406 | \newcommand{\State}[1]{\Set{State}_{\Sys{#1}}} 407 | \newcommand{\Out}[1]{\Set{Out}_{\Sys{#1}}} 408 | \newcommand{\In}[1]{\Set{In}_{\Sys{#1}}} 409 | \newcommand{\return}[1]{\fun{return}_{\Sys{#1}}} 410 | \newcommand{\update}[1]{\fun{update}_{\Sys{#1}}} 411 | 412 | \newcommand{\id}{\mathrm{id}} 413 | \newcommand{\then}{\mathbin{\fatsemi}} 414 | 415 | \newcommand{\cocolon}{:\!} 416 | 417 | 418 | \newcommand{\iso}{\cong} 419 | \newcommand{\too}{\longrightarrow} 420 | \newcommand{\tto}{\rightrightarrows} 421 | \newcommand{\To}[1]{\xrightarrow{#1}} 422 | \renewcommand{\Mapsto}[1]{\xmapsto{#1}} 423 | \newcommand{\Tto}[3][13pt]{\begin{tikzcd}[sep=#1, cramped, ampersand replacement=\&, text height=1ex, text depth=.3ex]\ar[r, shift left=2pt, "#2"]\ar[r, shift right=2pt, "#3"']\&{}\end{tikzcd}} 424 | \newcommand{\Too}[1]{\xrightarrow{\;\;#1\;\;}} 425 | \newcommand{\from}{\leftarrow} 426 | \newcommand{\From}[1]{\xleftarrow{#1}} 427 | \newcommand{\Fromm}[1]{\xleftarrow{\;\;#1\;\;}} 428 | \newcommand{\surj}{\twoheadrightarrow} 429 | \newcommand{\inj}{\hookrightarrow} 430 | \newcommand{\wavyto}{\rightsquigarrow} 431 | \newcommand{\lollipop}{\multimap} 432 | \newcommand{\tickar}{\xtickar{}} 433 | \newcommand{\imp}{\Rightarrow} 434 | \renewcommand{\iff}{\Leftrightarrow} 435 | \newcommand{\down}{\mathbin{\downarrow}} 436 | \newcommand{\fromto}{\leftrightarrows} 437 | \newcommand{\xtickar}[1]{\begin{tikzcd}[baseline=-0.5ex,cramped,sep=small,ampersand 438 | replacement=\&]{}\ar[r,tick, "{#1}"]\&{}\end{tikzcd}} 439 | 440 | 441 | 442 | \newcommand{\inv}{^{-1}} 443 | \newcommand{\op}{^\tn{op}} 444 | 445 | \newcommand{\tn}[1]{\textnormal{#1}} 446 | \newcommand{\ol}[1]{\overline{#1}} 447 | \newcommand{\ul}[1]{\underline{#1}} 448 | \newcommand{\wt}[1]{\widetilde{#1}} 449 | \newcommand{\wh}[1]{\widehat{#1}} 450 | \newcommand{\wc}[1]{\widecheck{#1}} 451 | \newcommand{\ubar}[1]{\underaccent{\bar}{#1}} 452 | 453 | 454 | % OLD LMO: \newcommand{\LMO}[2][over]{\ifthenelse{\equal{#1}{over}}{\overset{#2}{\bullet}}{\underset{#2}{\bullet}}} 455 | \newcommand{\LMO}[2][over]{\LMOO[#1]{\phantom{\cdot}}{#2}} 456 | %% Convention for LMOO is 457 | %% {optional: over/under}{state name}{output name} 458 | \newcommand{\LMOO}[3][over]{\ifthenelse{\equal{#1}{over}} 459 | { 460 | \begin{tikzpicture}[inner sep=0pt, solid] \node[draw, circle, label={[label distance 461 | = 2pt]90:{$#3$}}] {$#2$}; \end{tikzpicture} 462 | }{ 463 | \begin{tikzpicture}[inner sep=0pt, solid] \node[draw, circle, label={[label distance 464 | = 2pt]270:{$#3$}}] {$#2$}; \end{tikzpicture} 465 | } 466 | } 467 | \newcommand{\LTO}[2][\bullet]{\overset{\tn{#2}}{#1}} 468 | 469 | 470 | \newcommand{\bb}{\mathbb{B}} 471 | \newcommand{\nn}{\mathbb{N}} 472 | \newcommand{\pp}{\mathbb{P}} 473 | \newcommand{\qq}{\mathbb{Q}} 474 | \newcommand{\zz}{\mathbb{Z}} 475 | \newcommand{\rr}{\mathbb{R}} 476 | 477 | \newcommand{\singleton}{\{1\}} 478 | \newcommand{\powset}{\Fun{P}} 479 | \newcommand{\probset}{\Fun{D}} 480 | 481 | 482 | \newcommand{\smset}{\Cat{Set}} 483 | \newcommand{\smcat}{\Cat{Cat}} 484 | \newcommand{\comon}{\Cat{Comon}} 485 | \newcommand{\nat}{\Set{Nat}} 486 | \newcommand{\finset}{\Cat{FinSet}} 487 | \newcommand{\smctx}[1]{\Cat{Ctx}_{#1}} 488 | 489 | \newcommand{\List}{\Fun{List}} 490 | \newcommand{\set}{\tn{-}\Cat{Set}} 491 | 492 | 493 | 494 | \newcommand{\boxCD}[3][black] % [border color] {interior color} {math text} 495 | {\fcolorbox{#1}{#2}{\begin{varwidth}{\textwidth}\centering #3\end{varwidth}}} 496 | 497 | \newcommand{\erase}[2][]{{\color{red}#1}} 498 | \newcommand{\showhide}[2]{#1} 499 | 500 | 501 | % Spivak macros % 502 | 503 | \newenvironment{dedication} 504 | {\begin{flushright}\itshape} 505 | {\end{flushright}} 506 | 507 | \newcommand{\yon}{\mathcal{y}} 508 | \newcommand{\poly}[1][]{#1\Cat{Poly}} 509 | \newcommand{\0}{\textsf{0}} 510 | \newcommand{\1}{\textsf{1}} 511 | \newcommand{\2}{\textsf{2}} 512 | \newcommand{\3}{\textsf{3}} 513 | \newcommand{\4}{\textsf{4}} 514 | \newcommand{\5}{\textsf{5}} 515 | \newcommand{\6}{\textsf{6}} 516 | \newcommand{\7}{\textsf{7}} 517 | \newcommand{\8}{\textsf{8}} 518 | \newcommand{\9}{\textsf{9}} 519 | 520 | \definecolor{dgreen}{rgb}{0.0, 0.5, 0.3} 521 | \definecolor{dyellow}{rgb}{8.0, 0.74, 0} 522 | 523 | \newcommand{\pphantom}[2]{$\phantom{#1}$\nodepart{two}$\vphantom{#2}$} 524 | \newcommand{\at}{\nodepart{two}} 525 | 526 | \newcommand{\qand}{\quad\text{and}\quad} 527 | \newcommand{\qqand}{\qquad\text{and}\qquad} 528 | \newcommand{\qqor}{\qquad\text{or}\qquad} 529 | 530 | \newcommand{\bhline}{\Xhline{2\arrayrulewidth}} 531 | \newcommand{\bbhline}{\Xhline{2.5\arrayrulewidth}} 532 | \newcommand{\bbbhline}{\Xhline{3\arrayrulewidth}} 533 | 534 | \newcommand{\lst}{\Set{List}} 535 | 536 | \newcommand{\tr}{\Set{tree}} 537 | \newcommand{\vtx}{\Set{vtx}} 538 | \newcommand{\rt}{\fun{root}} 539 | \newcommand{\fs}{\fun{focus}} 540 | \newcommand{\lensput}{\fun{put}} 541 | \newcommand{\lensget}{\fun{get}} 542 | \newcommand{\deck}{\Fun{Deck}} 543 | 544 | \newcommand{\src}{\fun{src}} 545 | \newcommand{\tgt}{\fun{tgt}} 546 | 547 | \newcommand{\elts}{\textstyle{\int}} 548 | 549 | \newcommand{\mdot}{{\cdot}} 550 | \newcommand{\cofree}[1]{\com{T}_{#1}} 551 | \newcommand{\coto}{\nrightarrow} 552 | \newcommand{\cofun}{{\raisebox{2pt}{\resizebox{2.5pt}{2.5pt}{$\setminus$}}}} 553 | 554 | \newcommand{\lcmod}[1]{~_{#1}\Cat{Mod}} 555 | \newcommand{\rcmod}[1]{~\Cat{Mod}_{#1}} 556 | \newcommand{\bimod}[2]{~_{#1}\Cat{Mod}_{#2}} 557 | 558 | \newcommand{\car}[1]{\mathfrak{#1}} 559 | 560 | \newcommand{\cof}{\nrightarrow} 561 | 562 | \newcommand{\fl}[1]{\acute{#1}} 563 | 564 | \newcommand{\tri}{\mathbin{\triangleleft}} 565 | \newcommand{\tripow}[1]{^{\tri\,#1}} 566 | 567 | \newcommand{\bul}[1][black]{{\color{#1}\ensuremath{\bullet}}} 568 | 569 | \makeatletter 570 | \newcommand*{\bdiv}{% 571 | \nonscript\mskip-\medmuskip\mkern5mu% 572 | \mathbin{\operator@font div}\penalty900\mkern5mu% 573 | \nonscript\mskip-\medmuskip 574 | } 575 | \makeatother 576 | 577 | \newcommand{\centergraphics}[2][width=\textwidth]{ 578 | \begin{center} 579 | \includegraphics[#1]{#2} 580 | \end{center} 581 | } 582 | 583 | % for making notes in text: 584 | \newcommand{\niu}[1]{{\color{red} Niu: {#1}}} 585 | \newcommand{\spiz}[1]{{\color{blue} Spiz: {#1}}} 586 | 587 | % lenses 588 | \newcommand{\biglens}[2]{ 589 | \begin{pmatrix}{\vphantom{f_f^f}#1} \\ {\vphantom{f_f^f}#2} \end{pmatrix} 590 | } 591 | \newcommand{\littlelens}[2]{ 592 | \begin{psmallmatrix}{\vphantom{f}#1} \\ {\vphantom{f}#2} \end{psmallmatrix} 593 | } 594 | \newcommand{\lens}[2]{ 595 | \relax\if@display 596 | \biglens{#1}{#2} 597 | \else 598 | \littlelens{#1}{#2} 599 | \fi 600 | } 601 | 602 | \newcommand{\sys}{\Cat{Sys}} 603 | 604 | % category of arities 605 | \newcommand{\arity}{\Cat{Arity}} 606 | \newcommand{\XX}{\mathrm{X}} 607 | 608 | % Booleans 609 | \newcommand{\true}{\const{true}} 610 | \newcommand{\false}{\const{false}} 611 | 612 | % Colors 613 | \newcommand{\Red}{{\color{red} \const{red}}} 614 | \newcommand{\Blue}{{\color{blue} \const{blue}}} 615 | \newcommand{\Green}{{\color{dgreen}\const{green}} } 616 | \newcommand{\Orange}{{\color{orange} \const{orange}}} 617 | 618 | % Grothendieck double construction: 619 | \newcommand{\rectangle}{{% 620 | \Large 621 | \ooalign{$\sqsubset\mkern3mu$\cr$\mkern3mu\sqsupset$\cr}% 622 | }} 623 | \newcommand{\sqiint}{{\rectangle}\kern-1.4em{\iint}} 624 | 625 | %% vdots without space at the top. 626 | \makeatletter 627 | \DeclareRobustCommand{\rvdots}{% 628 | \vbox{ 629 | \baselineskip4\p@\lineskiplimit\z@ 630 | \kern-\p@ 631 | \hbox{.}\hbox{.}\hbox{.} 632 | }} 633 | \makeatother 634 | 635 | % ---- Changeable document parameters ---- % 636 | 637 | \linespread{1.15} 638 | \allowdisplaybreaks 639 | \setsecnumdepth{subsection} 640 | \settocdepth{subsection} 641 | \setlength{\parindent}{15pt} 642 | \setcounter{tocdepth}{1} 643 | 644 | 645 | 646 | %--------------- Document ---------------% 647 | \begin{document} 648 | 649 | \frontmatter 650 | 651 | % \title{\titlefont Categorical dynamics} 652 | % \title{\titlefont Positions and directions} 653 | % \title{\titlefont Systems Behaving Together\\in an Arena} 654 | % \title{\titlefont Ears, brains, and mouths} 655 | % \title{\titlefont Categorical Systems Theory} 656 | % \title{\titlefont P-Func:\\\bigskip\medskip\huge Math for the Mothership} 657 | % \title{\titlefont $\poly$ for the Sciences} 658 | % \title{\titlefont Operational Semantics for\\\medskip Distributed Systems} 659 | 660 | % \title{\titlefont The Dynamics of Polynomials} 661 | % \title{\titlefont Polynomial Functors: A Theory of Dynamic Interactions} 662 | % \title{\titlefont Polynomial Functors: A General Theory of Interaction} 663 | % \title{\titlefont Polynomial Functors: A Categorical Theory of Interaction} 664 | 665 | \title{\titlefont Polynomial Functors:\\\medskip 666 | A Mathematical Theory of Interaction} 667 | 668 | \author{ 669 | \LARGE Nelson Niu 670 | \and 671 | \LARGE David I. Spivak 672 | \normalsize} 673 | 674 | \posttitle{ 675 | \vspace{.8in} 676 | \normalsize 677 | \[ 678 | \coverpic 679 | \] 680 | \vspace{.5in} 681 | \endgroup 682 | } 683 | \date{\vfill Last updated: \today \\ Source: \url{https://github.com/ToposInstitute/poly}} 684 | 685 | \maketitle 686 | 687 | \thispagestyle{empty} 688 | \clearpage 689 | This page intentionally left blank. 690 | \clearpage 691 | \clearpage 692 | 693 | \mbox{} 694 | \vspace{2in} 695 | \begin{center} 696 | \begin{tabular}{lll} 697 | \LARGE Nelson Niu&~\hspace{.75in}~&\LARGE David I. Spivak\\ 698 | \large University of Washington&&\large Topos Institute\\ 699 | \large Seattle, WA&&\large Berkeley, CA 700 | \end{tabular} 701 | \end{center} 702 | \clearpage 703 | 704 | \begin{dedication} 705 | To Andr\'e Joyal 706 | \smallskip 707 | 708 | ---D.S. 709 | \bigskip 710 | 711 | To my graduate cohort at UW 712 | \smallskip 713 | 714 | ---N.N. 715 | \end{dedication} 716 | 717 | \subfile{P0-Preface} 718 | 719 | \clearpage 720 | \tableofcontents* 721 | \clearpage 722 | 723 | \mainmatter 724 | \subfile{P1-Polynomials} 725 | \subfile{P2-Comonoids} 726 | \subfile{P3-Vistas} 727 | 728 | \appendix 729 | \begingroup 730 | \footnotesize 731 | 732 | \backmatter 733 | 734 | \printbibliography 735 | \printindex 736 | 737 | \end{document} -------------------------------------------------------------------------------- /code-examples/poly_up_to_fibonacci.idr: -------------------------------------------------------------------------------- 1 | module Dynamical 2 | 3 | import Data.Vect 4 | 5 | %access public export 6 | 7 | --- Code by David I. Spivak and David Jaz Myers 8 | --- © 2020 9 | 10 | 11 | ------ The category of arenas ------ 12 | 13 | --Also called Cont, the category of containers. 14 | --Equivalent to Poly, the category of polynomial functors in one variable. 15 | 16 | 17 | --- Objects --- 18 | record Arena where 19 | constructor MkArena 20 | pos : Type 21 | dis : pos -> Type 22 | 23 | --- Morphisms--- 24 | 25 | record Lens (dom : Arena) (cod : Arena) where 26 | constructor MkLens 27 | observe : pos dom -> pos cod 28 | interpret : (p : pos dom) -> dis cod (observe p) -> dis dom p 29 | 30 | --- Identity --- 31 | 32 | idLens : (a : Arena) -> Lens a a 33 | idLens a = MkLens id (\_ => id) 34 | 35 | 36 | --- Composition --- 37 | 38 | infixr 4 <.> 39 | (<.>) : Lens a2 a3 -> Lens a1 a2 -> Lens a1 a3 40 | (<.>) lens23 lens12 = MkLens obs int 41 | where 42 | obs : pos a1 -> pos a3 43 | obs = (observe lens23) . (observe lens12) 44 | 45 | int : (p : pos a1) -> (dis a3 (obs p)) -> dis a1 p 46 | int p = (interpret lens12 p) . (interpret lens23 (observe lens12 p)) 47 | 48 | --- Manipulations on Arenas --- 49 | 50 | Display : Arena -> Type 51 | Display a = (p : pos a ** dis a p) 52 | 53 | AsFunctor : Arena -> Type -> Type 54 | AsFunctor a y = (p : pos a ** dis a p -> y) 55 | 56 | 57 | 58 | --- Special Arenas --- 59 | 60 | ArenaIO : (i : Type) -> (o : Type) -> Arena --positions as output and 61 | ArenaIO i o = MkArena o (\_ => i) --distinctions as input 62 | 63 | Self : Type -> Arena 64 | Self s = ArenaIO s s 65 | 66 | Closed : Arena 67 | Closed = ArenaIO () () 68 | 69 | 70 | --- Factorization of lenses --- 71 | 72 | Factor : {a, c : Arena} -> Lens a c -> (b : Arena ** (Lens b c, Lens a b)) 73 | Factor {a} {c} f = (b ** (cartf, vertf)) where 74 | b : Arena 75 | vertf : Lens a b 76 | cartf : Lens b c 77 | b = MkArena (pos a) $ dis c . observe f 78 | vertf = MkLens id (interpret f) 79 | cartf = MkLens (observe f) (\_ => id) 80 | 81 | 82 | --- Reflections to Type --- 83 | 84 | Exception : Type -> Arena 85 | Exception t = ArenaIO Void t 86 | 87 | Emitter : Type -> Arena 88 | Emitter t = ArenaIO () t 89 | 90 | Sensor : Type -> Arena 91 | Sensor t = ArenaIO t () 92 | 93 | ev0 : Arena -> Arena 94 | ev0 a = Exception $ AsFunctor a Void 95 | 96 | fromEv0 : (a : Arena) -> Lens (ev0 a) a 97 | fromEv0 a = MkLens o i 98 | where 99 | o : (p : pos a ** dis a p -> Void) -> pos a 100 | o = fst 101 | -- i : (x : pos (ev0 a)) -> dis a (o x) -> dis (ev0 a) x 102 | -- i : (x : AsFunctor a Void) -> dis a (o x) -> dis (ev0 a) x 103 | i : (x : (p : pos a ** dis a p -> Void)) -> dis a (fst x) -> dis (ev0 a) x 104 | i (p ** f) = f 105 | 106 | ev1 : Arena -> Arena 107 | ev1 a = Exception $ pos a -- = Exception $ AsFunctor a () 108 | 109 | toEv1 : (a : Arena) -> Lens a (ev1 a) 110 | toEv1 a = MkLens id (\_ => absurd) 111 | 112 | ev1y : Arena -> Arena 113 | ev1y a = Emitter $ pos a 114 | 115 | fromEv1y : (a : Arena) -> Lens (ev1y a) a 116 | fromEv1y a = MkLens id (\_, _ => ()) 117 | 118 | constantFunction : {t, u : Type} -> (t -> u) -> Lens (Exception t) (Exception u) 119 | constantFunction {t} {u} f = MkLens f (\_ => id) 120 | 121 | EmitterFunction : {t, u : Type} -> (t -> u) -> Lens (Emitter t) (Emitter u) 122 | EmitterFunction {t} {u} f = MkLens f (\_ => id) 123 | 124 | SensorFunction : {t, u: Type} -> (t -> u) -> Lens (Sensor u) (Sensor t) 125 | SensorFunction {t} {u} f = MkLens id (\_ => f) 126 | 127 | enclose : Arena -> Type 128 | enclose a = Lens a Closed 129 | 130 | encloseFunction : {t, u : Type} -> (t -> u) -> Lens (ArenaIO u t) Closed 131 | encloseFunction {t} {u} f = MkLens (\_ => ()) (\d => \_ => f d) 132 | 133 | auto : {m : Type} -> enclose (Emitter m) 134 | auto {m} = encloseFunction $ \_ => () 135 | 136 | --- functors and monads --- 137 | 138 | lift : (f : Type -> Type) -> Functor f => Arena -> Arena 139 | lift f ar = MkArena (pos ar) fdis 140 | where 141 | fdis : (p : pos ar) -> Type 142 | fdis p = f $ dis ar p 143 | 144 | LiftLens : {a, b : Arena} -> (f : Type -> Type) -> Functor f => 145 | Lens a b -> Lens (lift f a) (lift f b) 146 | LiftLens {a} {b} f lens = MkLens (observe lens) int 147 | where 148 | int : (p : pos a) -> f $ dis b (observe lens p) -> f $ dis a p 149 | int p = map $ interpret lens p 150 | 151 | extract : {a : Arena} -> (f : Type -> Type) -> Monad f => 152 | Lens (lift f a) a 153 | extract {a} f = MkLens id pur 154 | where 155 | pur : (p : pos a) -> dis a p -> dis (lift f a) p 156 | pur p = pure 157 | 158 | extend : {a : Arena} -> (f : Type -> Type) -> Monad f => 159 | Lens (lift f a) (lift f (lift f a)) 160 | extend {a} f = MkLens id joi 161 | where 162 | joi : (p : pos a) -> dis (lift f (lift f a)) p -> dis (lift f a) p 163 | joi p = join 164 | 165 | --- sum --- 166 | 167 | zero : Arena 168 | zero = ArenaIO Void Void 169 | 170 | infixr 4 <++> 171 | 172 | (<++>) : Arena -> Arena -> Arena 173 | (<++>) a b = MkArena posSum disSum 174 | where 175 | posSum : Type 176 | posSum = Either (pos a) (pos b) 177 | disSum : posSum -> Type 178 | disSum (Left p) = dis a p 179 | disSum (Right p) = dis b p 180 | 181 | sum : (ind : Type ** ind -> Arena) -> Arena 182 | sum (ind ** arena) = MkArena psum dsum 183 | where 184 | psum : Type 185 | psum = (i : ind ** pos (arena i)) 186 | dsum : psum -> Type 187 | dsum (i ** p) = dis (arena i) p 188 | 189 | 190 | 191 | sumLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 <++> a2) (b1 <++> b2) 192 | sumLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 193 | where 194 | o : pos (a1 <++> a2) -> pos (b1 <++> b2) 195 | o (Left p1) = Left (observe l1 p1) 196 | o (Right p2) = Right (observe l2 p2) 197 | i : (p : pos (a1 <++> a2)) -> dis (b1 <++> b2) (o p) -> dis (a1 <++> a2) p 198 | i (Left p1) d1 = interpret l1 p1 d1 199 | i (Right p2) d2 = interpret l2 p2 d2 200 | 201 | copair : {a1 : Arena} -> {a2 : Arena} -> {b : Arena} -> 202 | Lens a1 b -> Lens a2 b -> Lens (a1 <++> a2) b 203 | copair {a1} {a2} {b} l1 l2 = MkLens obs int 204 | where 205 | obs : pos (a1 <++> a2) -> pos b 206 | int : (p : pos (a1 <++> a2)) -> dis b (obs p) -> dis (a1 <++> a2) p 207 | obs (Left p1) = observe l1 p1 208 | obs (Right p2) = observe l2 p2 209 | int (Left p1) d1 = interpret l1 p1 d1 210 | int (Right p2) d2 = interpret l2 p2 d2 211 | 212 | --- product --- 213 | 214 | one : Arena 215 | one = ArenaIO Void () 216 | 217 | infixr 4 <**> 218 | 219 | (<**>) : Arena -> Arena -> Arena 220 | (<**>) a b = MkArena posProd disProd 221 | where 222 | posProd : Type 223 | posProd = (pos a, pos b) 224 | disProd : posProd -> Type 225 | disProd (pa, pb) = Either (dis a pa) (dis b pb) 226 | 227 | prodList : List Arena -> Arena 228 | prodList [] = one 229 | prodList [a] = a 230 | prodList (a :: as) = a <**> (prodList as) 231 | 232 | -- functoriality of prod 233 | prodLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 <**> a2) (b1 <**> b2) 234 | prodLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 235 | where 236 | o : pos (a1 <**> a2) -> pos (b1 <**> b2) 237 | o (p1, p2) = (observe l1 p1, observe l2 p2) 238 | i : (p : pos (a1 <**> a2)) -> dis (b1 <**> b2) (o p) -> dis (a1 <**> a2) p 239 | i (p1, p2) (Left d1) = Left (interpret l1 p1 d1) 240 | i (p1, p2) (Right d2) = Right (interpret l2 p2 d2) 241 | 242 | -- prod is the dependent product of polynomials; used in both inner homs 243 | prod : (ind : Type ** ind -> Arena) -> Arena 244 | prod (ind ** arena) = MkArena pprod dprod 245 | where 246 | pprod : Type 247 | pprod = (i : ind) -> pos (arena i) 248 | dprod : pprod -> Type 249 | dprod p = (i : ind ** dis (arena i) (p i)) 250 | 251 | pair : {a : Arena} -> {b1 : Arena} -> {b2 : Arena} -> 252 | Lens a b1 -> Lens a b2 -> Lens a (b1 <**> b2) 253 | pair {a} {b1} {b2} l1 l2 = MkLens obs int 254 | where 255 | obs : pos a -> (pos b1, pos b2) 256 | obs p = (observe l1 p, observe l2 p) 257 | int : (p : pos a) -> dis (b1 <**> b2) (obs p) -> dis a p 258 | int p (Left d1) = interpret l1 p d1 259 | int p (Right d2) = interpret l2 p d2 260 | 261 | --- Juxtaposition --- 262 | 263 | infixr 4 & 264 | 265 | (&) : Arena -> Arena -> Arena 266 | (&) a b = MkArena posJuxt disJuxt 267 | where 268 | posJuxt : Type 269 | posJuxt = (pos a, pos b) 270 | disJuxt : posJuxt -> Type 271 | disJuxt (pa, pb) = (dis a pa, dis b pb) 272 | 273 | juxtList : List Arena -> Arena 274 | juxtList [] = Closed 275 | juxtList [a] = a 276 | juxtList (a :: as) = a & (juxtList as) 277 | 278 | juxtLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 & a2) (b1 & b2) 279 | juxtLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 280 | where 281 | o : pos (a1 & a2) -> pos (b1 & b2) 282 | o (p1, p2) = (observe l1 p1, observe l2 p2) 283 | i : (p : pos (a1 & a2)) -> dis (b1 & b2) (o p) -> dis (a1 & a2) p 284 | i (p1, p2) (d1, d2) = (interpret l1 p1 d1, interpret l2 p2 d2) 285 | 286 | 287 | 288 | --- Composition product --- 289 | 290 | infixr 4 @@ 291 | (@@) : Arena -> Arena -> Arena 292 | (@@) a b = MkArena posCirc disCirc 293 | where 294 | posCirc : Type 295 | posCirc = (p : pos a ** dis a p -> pos b) 296 | disCirc : posCirc -> Type 297 | disCirc (p ** f) = (d : dis a p ** dis b (f d)) 298 | 299 | 300 | 301 | circLens : {a1 : Arena} -> {b1 : Arena} -> {a2 : Arena} -> {b2 : Arena} -> 302 | Lens a1 b1 -> Lens a2 b2 -> Lens (a1 @@ a2) (b1 @@ b2) 303 | circLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 304 | where 305 | o : pos (a1 @@ a2) -> pos (b1 @@ b2) 306 | o (p ** f) = (observe l1 p ** (observe l2) . f . (interpret l1 p)) 307 | i : (p : pos (a1 @@ a2)) -> dis (b1 @@ b2) (o p) -> dis (a1 @@ a2) p 308 | i (p ** f) (d1 ** d2) = (e1 ** interpret l2 (f e1) d2) 309 | where 310 | e1 : dis a1 p 311 | e1 = interpret l1 p d1 312 | 313 | CircPow : Arena -> Nat -> Arena 314 | CircPow _ Z = Closed 315 | CircPow a (S n) = a @@ CircPow a n 316 | 317 | CircPowLens : {a : Arena} -> {b : Arena} -> 318 | Lens a b -> (n : Nat) -> Lens (CircPow a n) (CircPow b n) 319 | CircPowLens {a} {b} _ Z = idLens Closed 320 | CircPowLens {a} {b} lens (S n) = circLens lens (CircPowLens lens n) 321 | 322 | EmitterPow : (t : Type) -> (n : Nat) -> Lens (CircPow (Emitter t) n) (Emitter (Vect n t)) 323 | EmitterPow t Z = MkLens (\_ => Nil) (\_, _ => ()) 324 | 325 | 326 | 327 | --- Comonoids --- 328 | 329 | record Comonoid where 330 | constructor MkComonoid 331 | domains : Arena 332 | identities : Lens domains Closed 333 | codsComps : Lens domains (domains @@ domains) 334 | 335 | codata Behavior : (ar : Arena) -> Type where 336 | (::) : (p : pos ar) -> (dis ar p -> Behavior ar) -> Behavior ar 337 | 338 | head : {ar : Arena} -> Behavior ar -> pos ar 339 | head {ar} (h :: _) = h 340 | 341 | tail : {ar : Arena} -> (b : Behavior ar) -> dis ar (head b) -> Behavior ar 342 | tail {ar} (_ :: t) = t 343 | 344 | 345 | 346 | 347 | MonSensor : (t : Type) -> t -> (t -> t -> t) -> Comonoid 348 | MonSensor t neut plus = MkComonoid sens idy cc 349 | where 350 | sens : Arena 351 | idy : Lens sens Closed 352 | cc : Lens sens (sens @@ sens) 353 | sens = (Sensor t) 354 | idy = SensorFunction $ \_ => neut 355 | cc = MkLens obs int 356 | where 357 | obs : pos sens -> (p : pos sens ** dis sens p -> pos sens) 358 | obs _ = (() ** \_ => ()) 359 | int : (p : pos sens) -> dis (sens @@ sens) (obs p) -> dis sens p 360 | int _ (d ** e) = d `plus` e 361 | 362 | TrajComon : Comonoid 363 | TrajComon = MonSensor Nat Z (+) 364 | 365 | 366 | 367 | --- Selves are comonoids --- 368 | 369 | 370 | 371 | counit : (s : Type) -> Lens (Self s) Closed 372 | counit s = MkLens o i 373 | where 374 | o : s -> () 375 | o _ = () 376 | i : s -> () -> s 377 | i s _ = s 378 | 379 | comult : (s : Type) -> Lens (Self s) ((Self s) @@ (Self s)) 380 | comult s = MkLens o i 381 | where 382 | o : s -> pos ((Self s) @@ (Self s)) 383 | o x = (x ** id) 384 | i : (x : s) -> (dis ((Self s) @@ (Self s)) (o x)) -> s 385 | i x (d1 ** d2) = d2 386 | 387 | ContrGrpd : Type -> Comonoid 388 | ContrGrpd t = MkComonoid (Self t) (counit t) (comult t) 389 | 390 | comultPow : (s : Type) -> (n : Nat) -> Lens (Self s) (CircPow (Self s) n) 391 | comultPow s Z = counit s 392 | comultPow s (S n) = (circLens (idLens (Self s)) (comultPow s n)) <.> (comult s) 393 | 394 | 395 | 396 | 397 | --- Duoidal --- 398 | 399 | duoidal : {a1, a2, b1, b2 : Arena} -> Lens ((a1 @@ a2) & (b1 @@ b2)) 400 | ((a1 & b1) @@ (a2 & b2)) 401 | duoidal {a1} {a2} {b1} {b2} = MkLens o i 402 | where 403 | x : Arena 404 | x = (a1 @@ a2) & (b1 @@ b2) 405 | y : Arena 406 | y = (a1 & b1) @@ (a2 & b2) 407 | o : pos x -> pos y 408 | o ((p1 ** p2), (q1 ** q2)) = pp 409 | where 410 | pp : (p : (pos a1, pos b1) ** dis (a1 & b1) p -> pos (a2 & b2)) 411 | pp = ((p1, q1) ** (\d : dis (a1 & b1) (p1, q1) => (p2 (fst d), q2 (snd d)))) 412 | i : (p : pos x) -> dis y (o p) -> dis x p 413 | i ((p1 ** p2), (q1 ** q2)) = ii 414 | where 415 | -- ii : dis y ((p1, q1) ** (\d : dis (a1 & b1) (p1, q1) => (p2 (fst d), q2 (snd d)))) 416 | -- -> dis x ((p1 ** p2), (q1 ** q2)) 417 | ii : (de1 : dis (a1 & b1) (p1, q1) ** dis (a2 & b2) (p2 (fst de1), q2 (snd de1))) 418 | -> dis x ((p1 ** p2), (q1 ** q2)) 419 | ii (de1 ** de2) = ((fst de1 ** fst de2), (snd de1 ** snd de2)) 420 | 421 | --- Exponentiation --- 422 | 423 | infixr 4 ^ 424 | 425 | --prod : (ind : Type ** ind -> Arena) -> Arena 426 | (^) : Arena -> Arena -> Arena 427 | (^) a b = prod (pos a ** arena) 428 | where 429 | arena : pos a -> Arena 430 | arena p = b @@ ((Exception $ dis a p) <++> Closed) 431 | 432 | --- Internal Hom --- 433 | 434 | infixr 4 ^^ 435 | (^^) : Arena -> Arena -> Arena 436 | (^^) a b = prod (pos a ** arena) 437 | where 438 | arena : pos a -> Arena 439 | arena p = b @@ (Emitter $ dis a p) 440 | 441 | {- 442 | eval : {a : Arena} -> {b : Arena} -> Lens (a & (b ^^ a)) b 443 | eval {a} {b} = MkLens obs int 444 | where 445 | obs : (pos a, pos (b ^^ a)) -> pos b 446 | int : (p : (pos a, pos (b ^^ a))) -> dis b (obs p) -> dis (a & (b ^^ a)) p 447 | obs (pa, pab) = ?evalo 448 | int p d = ?evali 449 | -} 450 | 451 | --- Dynamical systems --- 452 | 453 | 454 | record DynSystem where 455 | constructor MkDynSystem 456 | state : Type 457 | body : Arena 458 | pheno : Lens (Self state) body 459 | 460 | static : DynSystem 461 | static = MkDynSystem () Closed (EmitterFunction id) 462 | 463 | 464 | 465 | infixr 4 &&& 466 | (&&&) : DynSystem -> DynSystem -> DynSystem 467 | (&&&) dyn1 dyn2 = MkDynSystem state12 body12 pheno12 468 | where 469 | state12 : Type 470 | body12 : Arena 471 | pheno12 : Lens (Self state12) body12 472 | state12 = (state dyn1, state dyn2) 473 | body12 = (body dyn1) & (body dyn2) 474 | pheno12 = MkLens o i 475 | where 476 | o : (state dyn1, state dyn2) -> (pos (body dyn1), pos (body dyn2)) 477 | i : (s12 : (state dyn1, state dyn2)) -> dis body12 (o s12) -> state12 478 | o (s1, s2) = (observe (pheno dyn1) s1, observe (pheno dyn2) s2) 479 | i (s1, s2) (d1, d2) = 480 | (interpret (pheno dyn1) s1 d1, interpret (pheno dyn2) s2 d2) 481 | 482 | juxtapose : List DynSystem -> DynSystem 483 | juxtapose [] = static 484 | juxtapose [d] = d 485 | juxtapose (d :: ds) = d &&& (juxtapose ds) 486 | 487 | install : (d : DynSystem) -> (a : Arena) -> Lens (body d) a -> DynSystem 488 | install d a l = MkDynSystem (state d) a (l <.> (pheno d)) 489 | 490 | speedUp : DynSystem -> Nat -> DynSystem 491 | speedUp dyn n = MkDynSystem (state dyn) fastBody fastWork 492 | where 493 | fastBody : Arena 494 | fastWork : Lens (Self $ state dyn) fastBody 495 | fastBody = CircPow (body dyn) n 496 | fastWork = CircPowLens (pheno dyn) n <.> comultPow (state dyn) n 497 | 498 | 499 | 500 | run : (d : DynSystem) -> enclose (body d) -> (state d) -> Stream (pos $ body d) 501 | run d e s = outp :: (run d e next) 502 | where 503 | outp : pos $ body d 504 | next : state d 505 | outp = observe (pheno d) s 506 | next = interpret (pheno d) s $ interpret e outp () 507 | 508 | 509 | 510 | toStreamBehavior : {a : Arena} -> (b : Behavior a) -> (phys : enclose a) -> Stream (pos a) 511 | toStreamBehavior {a} b phys = currpos :: toStreamBehavior rest phys 512 | where 513 | currpos : pos a 514 | currpos = head b 515 | 516 | rest : Behavior a 517 | rest = tail b $ interpret phys currpos () 518 | 519 | dynBehavior : (d : DynSystem) -> (state d) -> Behavior (body d) 520 | dynBehavior dyn st = current :: choice 521 | where 522 | current : pos $ body dyn 523 | choice : dis (body dyn) current -> Behavior (body dyn) 524 | current = observe (pheno dyn) st 525 | choice d = dynBehavior dyn (interpret (pheno dyn) st d) 526 | 527 | runBehav : (d : DynSystem) -> enclose (body d) -> (state d) -> Stream (pos $ body d) 528 | runBehav dyn phys st = toStreamBehavior (dynBehavior dyn st) phys 529 | 530 | --- IO --- 531 | 532 | -- EmitterM : (m : Type -> Type) -> Monad m => (a : Type) -> DynSystem 533 | -- EmitterM m a = MkDynSystem (m a) (Emitter (m a)) passInput 534 | -- where 535 | -- passInput : Lens (Self (m a)) (Emitter (m a)) 536 | -- passInput = MkLens id 537 | 538 | IOEmitter : (a : Type) -> (String -> a) -> DynSystem 539 | IOEmitter a Cast = MkDynSystem (IO a) (Emitter $ IO a) passUserInput 540 | where 541 | passUserInput : Lens (Self $ IO a) (Emitter $ IO a) 542 | passUserInput = MkLens id listen 543 | where 544 | listen : (old : IO a) -> () -> IO a 545 | listen _ _ = map Cast getLine 546 | 547 | 548 | --- Debugging --- 549 | 550 | -- current : (d : DynSystem) -> state d -> pos (body d) 551 | -- current d s = observe (pheno d) s 552 | 553 | -- feed : (dyn : DynSystem) -> (s : state dyn) -> dis (body dyn) (observe (pheno dyn) s) -> state dyn 554 | -- feed dyn s d = interpret play s d 555 | -- where 556 | -- play : Lens (Self (state dyn)) (body dyn) 557 | -- play = pheno dyn 558 | 559 | --- Examples --- 560 | 561 | 562 | funcToDynSystem : {s : Type} -> {t : Type} -> (s -> t) -> DynSystem 563 | funcToDynSystem {s} {t} f = MkDynSystem t bodyf phenof 564 | where 565 | bodyf : Arena 566 | phenof : Lens (Self t) bodyf 567 | bodyf = ArenaIO s t 568 | phenof = MkLens id (\_ => f) 569 | 570 | 571 | 572 | 573 | delay : (s : Type) -> DynSystem 574 | delay s = funcToDynSystem (the (s -> s) id) 575 | 576 | plus : DynSystem 577 | plus = funcToDynSystem (uncurry (+)) 578 | 579 | 580 | Prefib : DynSystem 581 | Prefib = plus &&& (delay Integer) 582 | 583 | fibwd : Lens (body Prefib) (Emitter Integer) 584 | fibwd = MkLens observe interpret 585 | where 586 | observe : (Integer, Integer) -> Integer 587 | interpret : (p : (Integer, Integer)) -> () -> dis (body Prefib) p 588 | observe (pl, de) = de 589 | interpret (pl, de) = \_ => ((de, pl), pl) 590 | 591 | 592 | Fibonacci : DynSystem 593 | Fibonacci = install Prefib (Emitter Integer) fibwd 594 | 595 | 596 | FibSeq : Stream Integer 597 | FibSeq = run Fibonacci auto (1, 1) 598 | 599 | -------- Fibonacci in REPL -------- 600 | 601 | -- Run this in the REPL: 602 | -- take 10 FibSeq 603 | 604 | 605 | 606 | 607 | 608 | -- ========================== 609 | 610 | 611 | 612 | 613 | 614 | 615 | 616 | 617 | 618 | 619 | 620 | 621 | 622 | 623 | 624 | 625 | 626 | 627 | 628 | 629 | 630 | 631 | 632 | --- Distributivity --- 633 | 634 | prodSum : {a, b, c : Arena} -> Lens (a <**> (b <++> c)) ((a <**> b) <++> (a <**> c)) 635 | prodSum {a} {b} {c} = MkLens o i 636 | where 637 | a1 : Arena 638 | a1 = a <**> (b <++> c) 639 | a2 : Arena 640 | a2 = (a <**> b) <++> (a <**> c) 641 | o : pos a1 -> pos a2 642 | o (p, Left q) = Left (p, q) 643 | o (p, Right r) = Right (p, r) 644 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 645 | i (pa, Left pb) (Left da) = Left da 646 | i (pa, Left pb) (Right db) = Right db 647 | i (pa, Right pc) (Left da) = Left da 648 | i (pa, Right pc) (Right dc) = Right dc 649 | 650 | sumProd : {a, b, c : Arena} -> Lens ((a <**> b) <++> (a <**> c)) (a <**> (b <++> c)) 651 | sumProd {a} {b} {c} = MkLens o i 652 | where 653 | a1 : Arena 654 | a1 = (a <**> b) <++> (a <**> c) 655 | a2 : Arena 656 | a2 = a <**> (b <++> c) 657 | o : pos a1 -> pos a2 658 | o (Left (pa, pb)) = (pa, Left pb) 659 | o (Right (pa, pc)) = (pa, Right pc) 660 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 661 | i (Left (pa, pb)) (Left da) = Left da 662 | i (Left (pa, pb)) (Right db) = Right db 663 | i (Right (pa, pc)) (Left da) = Left da 664 | i (Right (pa, pc)) (Right dc) = Right dc 665 | 666 | juxtSum : {a, b, c : Arena} -> Lens (a & (b <++> c)) ((a & b) <++> (a & c)) 667 | juxtSum {a} {b} {c} = MkLens o i 668 | where 669 | a1 : Arena 670 | a1 = a & (b <++> c) 671 | a2 : Arena 672 | a2 = (a & b) <++> (a & c) 673 | o : pos a1 -> pos a2 674 | o (pa, Left pb) = Left (pa, pb) 675 | o (pa, Right pc) = Right (pa, pc) 676 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 677 | i (pa, Left pb) (da, db) = (da, db) 678 | i (pa, Right pc) (da, dc) = (da, dc) 679 | 680 | sumJuxt : {a, b, c : Arena} -> Lens ((a & b) <++> (a & c)) (a & (b <++> c)) 681 | sumJuxt {a} {b} {c} = MkLens o i 682 | where 683 | a1 : Arena 684 | a1 = (a & b) <++> (a & c) 685 | a2 : Arena 686 | a2 = a & (b <++> c) 687 | o : pos a1 -> pos a2 688 | o (Left (pa, pb)) = (pa, Left pb) 689 | o (Right (pa, pc)) = (pa, Right pc) 690 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 691 | i (Left (pa, pb)) (da, db) = (da, db) 692 | i (Right (pa, pc)) (da, dc) = (da, dc) 693 | 694 | 695 | -------------------------------------------------------------------------------- /code-examples/dyn.idr: -------------------------------------------------------------------------------- 1 | module dyn 2 | 3 | import Data.Vect 4 | import Data.List 5 | 6 | --- Code by David I. Spivak and David Jaz Myers and David Orion Girardo 7 | --- © 2020 8 | 9 | 10 | ------ The category of arenas ------ 11 | 12 | --Also called Cont, the category of containers. 13 | --Equivalent to Poly, the category of polynomial functors in one variable. 14 | 15 | 16 | --- Objects --- 17 | record Arena where 18 | constructor MkArena 19 | pos : Type 20 | dis : pos -> Type 21 | 22 | --- Morphisms--- 23 | 24 | record Lens (dom : Arena) (cod : Arena) where 25 | constructor MkLens 26 | observe : pos dom -> pos cod 27 | interpret : (p : pos dom) -> dis cod (observe p) -> dis dom p 28 | 29 | --- Identity --- 30 | 31 | idLens : (a : Arena) -> Lens a a 32 | idLens a = MkLens id (\_ => id) 33 | 34 | 35 | --- Composition --- 36 | 37 | infixr 4 <.> 38 | (<.>) : Lens a2 a3 -> Lens a1 a2 -> Lens a1 a3 39 | (<.>) lens23 lens12 = MkLens obs int 40 | where 41 | obs : pos a1 -> pos a3 42 | obs = (observe lens23) . (observe lens12) 43 | 44 | int : (p : pos a1) -> (dis a3 (obs p)) -> dis a1 p 45 | int p = (interpret lens12 p) . (interpret lens23 (observe lens12 p)) 46 | 47 | --- Manipulations on Arenas --- 48 | 49 | Display : Arena -> Type 50 | Display a = (p : pos a ** dis a p) 51 | 52 | AsFunctor : Arena -> Type -> Type 53 | AsFunctor a y = (p : pos a ** dis a p -> y) 54 | 55 | 56 | 57 | --- Special Arenas --- 58 | 59 | ArenaIO : (i : Type) -> (o : Type) -> Arena --positions as output and 60 | ArenaIO i o = MkArena o (\_ => i) --distinctions as input 61 | 62 | Self : Type -> Arena 63 | Self s = ArenaIO s s 64 | 65 | Closed : Arena 66 | Closed = ArenaIO () () 67 | 68 | 69 | --- Factorization of lenses --- 70 | 71 | Factor : {a, c : Arena} -> Lens a c -> (b : Arena ** (Lens b c, Lens a b)) 72 | Factor f 73 | = let vertf = MkLens id (interpret f) 74 | cartf = MkLens (observe f) (\_ => id) 75 | in ((MkArena (pos a) $ dis c . observe f) ** (cartf, vertf)) 76 | 77 | --- Reflections to Type --- 78 | 79 | Exception : Type -> Arena 80 | Exception t = ArenaIO Void t 81 | 82 | Emitter : Type -> Arena 83 | Emitter t = ArenaIO () t 84 | 85 | Sensor : Type -> Arena 86 | Sensor t = ArenaIO t () 87 | 88 | ev0 : Arena -> Arena 89 | ev0 a = Exception $ AsFunctor a Void 90 | 91 | fromEv0 : (a : Arena) -> Lens (ev0 a) a 92 | fromEv0 a = MkLens o i 93 | where 94 | o : (p : pos a ** dis a p -> Void) -> pos a 95 | o = fst 96 | -- i : (x : pos (ev0 a)) -> dis a (o x) -> dis (ev0 a) x 97 | -- i : (x : AsFunctor a Void) -> dis a (o x) -> dis (ev0 a) x 98 | i : (x : (p : pos a ** dis a p -> Void)) -> dis a (fst x) -> dis (ev0 a) x 99 | i (p ** f) = f 100 | 101 | ev1 : Arena -> Arena 102 | ev1 a = Exception $ pos a -- = Exception $ AsFunctor a () 103 | 104 | toEv1 : (a : Arena) -> Lens a (ev1 a) 105 | toEv1 a = MkLens id (\_ => absurd) 106 | 107 | ev1y : Arena -> Arena 108 | ev1y a = Emitter $ pos a 109 | 110 | fromEv1y : (a : Arena) -> Lens (ev1y a) a 111 | fromEv1y a = MkLens id (\_, _ => ()) 112 | 113 | constantFunction : {t, u : Type} -> (t -> u) -> Lens (Exception t) (Exception u) 114 | constantFunction {t} {u} f = MkLens f (\_ => id) 115 | 116 | EmitterFunction : {t, u : Type} -> (t -> u) -> Lens (Emitter t) (Emitter u) 117 | EmitterFunction {t} {u} f = MkLens f (\_ => id) 118 | 119 | SensorFunction : {t, u: Type} -> (t -> u) -> Lens (Sensor u) (Sensor t) 120 | SensorFunction {t} {u} f = MkLens id (\_ => f) 121 | 122 | enclose : Arena -> Type 123 | enclose a = Lens a Closed 124 | 125 | encloseFunction : {t, u : Type} -> (t -> u) -> Lens (ArenaIO u t) Closed 126 | encloseFunction {t} {u} f = MkLens (\_ => ()) (\d => \_ => f d) 127 | 128 | enclosed : {m : Type} -> enclose (Emitter m) 129 | enclosed {m} = encloseFunction $ \_ => () 130 | 131 | --- functors and monads --- 132 | 133 | lift : (f : Type -> Type) -> Functor f => Arena -> Arena 134 | lift f ar = MkArena (pos ar) fdis 135 | where 136 | fdis : (p : pos ar) -> Type 137 | fdis p = f $ dis ar p 138 | 139 | LiftLens : {a, b : Arena} -> (f : Type -> Type) -> Functor f => 140 | Lens a b -> Lens (lift f a) (lift f b) 141 | LiftLens {a} {b} f lens = MkLens (observe lens) int 142 | where 143 | int : (p : pos a) -> f $ dis b (observe lens p) -> f $ dis a p 144 | int p = map $ interpret lens p 145 | 146 | extract : {a : Arena} -> (f : Type -> Type) -> Monad f => 147 | Lens (lift f a) a 148 | extract {a} f = MkLens id pur 149 | where 150 | pur : (p : pos a) -> dis a p -> dis (lift f a) p 151 | pur p = pure 152 | 153 | extend : {a : Arena} -> (f : Type -> Type) -> Monad f => 154 | Lens (lift f a) (lift f (lift f a)) 155 | extend {a} f = MkLens id joi 156 | where 157 | joi : (p : pos a) -> dis (lift f (lift f a)) p -> dis (lift f a) p 158 | joi p = join 159 | 160 | --- sum --- 161 | 162 | zero : Arena 163 | zero = ArenaIO Void Void 164 | 165 | infixr 4 <++> 166 | 167 | (<++>) : Arena -> Arena -> Arena 168 | (<++>) a b = MkArena (Either (pos a) (pos b)) (\x => case x of {(Left p) => dis a p; (Right p) => dis b p}) 169 | 170 | sum : (ind : Type ** ind -> Arena) -> Arena 171 | sum (ind ** arena) = MkArena (i : ind ** pos (arena i)) (\ (i ** p) => dis (arena i) p) 172 | 173 | 174 | 175 | 176 | sumLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 <++> a2) (b1 <++> b2) 177 | sumLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 178 | where 179 | o : pos (a1 <++> a2) -> pos (b1 <++> b2) 180 | o (Left p1) = Left (observe l1 p1) 181 | o (Right p2) = Right (observe l2 p2) 182 | i : (p : pos (a1 <++> a2)) -> dis (b1 <++> b2) (o p) -> dis (a1 <++> a2) p 183 | i (Left p1) d1 = interpret l1 p1 d1 184 | i (Right p2) d2 = interpret l2 p2 d2 185 | 186 | copair : {a1 : Arena} -> {a2 : Arena} -> {b : Arena} -> 187 | Lens a1 b -> Lens a2 b -> Lens (a1 <++> a2) b 188 | copair {a1} {a2} {b} l1 l2 = MkLens obs int 189 | where 190 | obs : pos (a1 <++> a2) -> pos b 191 | int : (p : pos (a1 <++> a2)) -> dis b (obs p) -> dis (a1 <++> a2) p 192 | obs (Left p1) = observe l1 p1 193 | obs (Right p2) = observe l2 p2 194 | int (Left p1) d1 = interpret l1 p1 d1 195 | int (Right p2) d2 = interpret l2 p2 d2 196 | 197 | --- product --- 198 | 199 | one : Arena 200 | one = ArenaIO Void () 201 | 202 | infixr 4 <**> 203 | 204 | (<**>) : Arena -> Arena -> Arena 205 | (<**>) a b = MkArena (pos a, pos b) (\ (pa,pb) => Either (dis a pa) (dis b pb)) 206 | 207 | prodList : List Arena -> Arena 208 | prodList [] = one 209 | prodList [a] = a 210 | prodList (a :: as) = a <**> (prodList as) 211 | 212 | -- functoriality of prod 213 | prodLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 <**> a2) (b1 <**> b2) 214 | prodLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 215 | where 216 | o : pos (a1 <**> a2) -> pos (b1 <**> b2) 217 | o (p1, p2) = (observe l1 p1, observe l2 p2) 218 | i : (p : pos (a1 <**> a2)) -> dis (b1 <**> b2) (o p) -> dis (a1 <**> a2) p 219 | i (p1, p2) (Left d1) = Left (interpret l1 p1 d1) 220 | i (p1, p2) (Right d2) = Right (interpret l2 p2 d2) 221 | 222 | -- prod is the dependent product of polynomials; used in both inner homs 223 | prod : (ind : Type ** ind -> Arena) -> Arena 224 | prod (ind ** arena) = MkArena ((i : ind) -> pos (arena i)) (\p => (i : ind ** 225 | dis (arena i) (p i))) 226 | 227 | pair : {a : Arena} -> {b1 : Arena} -> {b2 : Arena} -> 228 | Lens a b1 -> Lens a b2 -> Lens a (b1 <**> b2) 229 | pair {a} {b1} {b2} l1 l2 = MkLens obs int 230 | where 231 | obs : pos a -> (pos b1, pos b2) 232 | obs p = (observe l1 p, observe l2 p) 233 | int : (p : pos a) -> dis (b1 <**> b2) (obs p) -> dis a p 234 | int p (Left d1) = interpret l1 p d1 235 | int p (Right d2) = interpret l2 p d2 236 | 237 | --- Juxtaposition --- 238 | 239 | (*) : Arena -> Arena -> Arena 240 | (*) a b = MkArena (pos a, pos b) \ (pa, pb) => (dis a pa, dis b pb) 241 | 242 | juxtList : List Arena -> Arena 243 | juxtList [] = Closed 244 | juxtList [a] = a 245 | juxtList (a :: as) = a * (juxtList as) 246 | 247 | juxt : (ind : Type) -> (ind -> Arena) -> Arena 248 | juxt ind arena = MkArena ((i : ind) -> pos (arena i)) 249 | (\p => (i : ind) -> dis (arena i) (p i)) 250 | 251 | juxtLens : Lens a1 b1 -> Lens a2 b2 -> Lens (a1 * a2) (b1 * b2) 252 | juxtLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 253 | where 254 | o : pos (a1 * a2) -> pos (b1 * b2) 255 | o (p1, p2) = (observe l1 p1, observe l2 p2) 256 | i : (p : pos (a1 * a2)) -> dis (b1 * b2) (o p) -> dis (a1 * a2) p 257 | i (p1, p2) (d1, d2) = (interpret l1 p1 d1, interpret l2 p2 d2) 258 | 259 | 260 | 261 | 262 | --- Composition product --- 263 | 264 | infixr 4 @@ 265 | (@@) : Arena -> Arena -> Arena 266 | (@@) a b = MkArena (p : pos a ** dis a p -> pos b) \ (p ** f) => (d : dis a p ** dis b (f d)) 267 | 268 | 269 | 270 | circLens : {a1 : Arena} -> {b1 : Arena} -> {a2 : Arena} -> {b2 : Arena} -> 271 | Lens a1 b1 -> Lens a2 b2 -> Lens (a1 @@ a2) (b1 @@ b2) 272 | circLens {a1} {b1} {a2} {b2} l1 l2 = MkLens o i 273 | where 274 | o : pos (a1 @@ a2) -> pos (b1 @@ b2) 275 | o (p ** f) = (observe l1 p ** (observe l2) . f . (interpret l1 p)) 276 | i : (p : pos (a1 @@ a2)) -> dis (b1 @@ b2) (o p) -> dis (a1 @@ a2) p 277 | i (p ** f) (d1 ** d2) = (e1 ** interpret l2 (f e1) d2) 278 | where 279 | e1 : dis a1 p 280 | e1 = interpret l1 p d1 281 | 282 | CircPow : Arena -> Nat -> Arena 283 | CircPow _ Z = Closed 284 | CircPow a (S n) = a @@ CircPow a n 285 | 286 | CircPowLens : {a : Arena} -> {b : Arena} -> 287 | Lens a b -> (n : Nat) -> Lens (CircPow a n) (CircPow b n) 288 | CircPowLens {a} {b} _ Z = idLens Closed 289 | CircPowLens {a} {b} lens (S n) = circLens lens (CircPowLens lens n) 290 | 291 | EmitterPow : (t : Type) -> (n : Nat) -> Lens (CircPow (Emitter t) n) (Emitter (Vect n t)) 292 | EmitterPow t Z = MkLens (\_ => Nil) (\_, _ => ()) 293 | EmitterPow t (S n) = ?pown 294 | 295 | 296 | 297 | --- Comonoids --- 298 | 299 | record Comonoid where 300 | constructor MkComonoid 301 | domains : Arena 302 | identities : Lens domains Closed 303 | codsComps : Lens domains (domains @@ domains) 304 | 305 | data Behavior : (ar : Arena) -> Type where 306 | (::) : (p : pos ar) -> (dis ar p -> Behavior ar) -> Behavior ar 307 | 308 | head : {ar : Arena} -> Behavior ar -> pos ar 309 | head {ar} (h :: _) = h 310 | 311 | tail : {ar : Arena} -> (b : Behavior ar) -> dis ar (head b) -> Behavior ar 312 | tail {ar} (_ :: t) = t 313 | 314 | 315 | 316 | 317 | MonSensor : (t : Type) -> t -> (t -> t -> t) -> Comonoid 318 | MonSensor t neut plus 319 | = let 320 | sens : Arena 321 | idy : Lens sens Closed 322 | cc : Lens sens (sens @@ sens) 323 | sens = (Sensor t) 324 | idy = SensorFunction $ \_ => neut 325 | cc = MkLens obs int 326 | where 327 | obs : pos sens -> (p : pos sens ** dis sens p -> pos sens) 328 | obs _ = (() ** \_ => ()) 329 | int : (p : pos sens) -> dis (sens @@ sens) (obs p) -> dis sens p 330 | int _ (d ** e) = d `plus` e 331 | in MkComonoid sens idy cc 332 | 333 | TrajComon : Comonoid 334 | TrajComon = MonSensor Nat Z (+) 335 | 336 | 337 | 338 | --- Selves are comonoids --- 339 | 340 | 341 | 342 | counit : (s : Type) -> Lens (Self s) Closed 343 | counit s = MkLens o i 344 | where 345 | o : s -> () 346 | o _ = () 347 | i : s -> () -> s 348 | i s _ = s 349 | 350 | comult : (s : Type) -> Lens (Self s) ((Self s) @@ (Self s)) 351 | comult s = MkLens o i 352 | where 353 | o : s -> pos ((Self s) @@ (Self s)) 354 | o x = (x ** id) 355 | i : (x : s) -> (dis ((Self s) @@ (Self s)) (o x)) -> s 356 | i x (d1 ** d2) = d2 357 | 358 | ContrGrpd : Type -> Comonoid 359 | ContrGrpd t = MkComonoid (Self t) (counit t) (comult t) 360 | 361 | comultPow : (s : Type) -> (n : Nat) -> Lens (Self s) (CircPow (Self s) n) 362 | comultPow s Z = counit s 363 | comultPow s (S n) = (circLens (idLens (Self s)) (comultPow s n)) <.> (comult s) 364 | 365 | 366 | 367 | 368 | --- Duoidal --- 369 | 370 | duoidal : {a1, a2, b1, b2 : Arena} -> Lens ((a1 @@ a2) * (b1 @@ b2)) 371 | ((a1 * b1) @@ (a2 * b2)) 372 | duoidal {a1} {a2} {b1} {b2} = 373 | let 374 | x : Arena 375 | x = (a1 @@ a2) * (b1 @@ b2) 376 | y : Arena 377 | y = (a1 * b1) @@ (a2 * b2) 378 | o : pos x -> pos y 379 | o ((p1 ** p2), (q1 ** q2)) = pp 380 | where 381 | pp : (p : (pos a1, pos b1) ** dis (a1 * b1) p -> pos (a2 * b2)) 382 | pp = ((p1, q1) ** (\d : dis (a1 * b1) (p1, q1) => (p2 (fst d), q2 (snd d)))) 383 | i : (p : pos x) -> dis y (o p) -> dis x p 384 | i ((p1 ** p2), (q1 ** q2)) = ii 385 | where 386 | -- ii : dis y ((p1, q1) ** (\d : dis (a1 * b1) (p1, q1) => (p2 (fst d), q2 (snd d)))) 387 | -- -> dis x ((p1 ** p2), (q1 ** q2)) 388 | ii : (de1 : dis (a1 * b1) (p1, q1) ** dis (a2 * b2) (p2 (fst de1), q2 (snd de1))) 389 | -> dis x ((p1 ** p2), (q1 ** q2)) 390 | ii (de1 ** de2) = ((fst de1 ** fst de2), (snd de1 ** snd de2)) 391 | in MkLens o i 392 | 393 | --- Exponentiation --- 394 | 395 | infixr 4 ^ 396 | 397 | --prod : (ind : Type ** ind -> Arena) -> Arena 398 | (^) : Arena -> Arena -> Arena 399 | (^) a b = prod (pos a ** arena) 400 | where 401 | arena : pos a -> Arena 402 | arena p = b @@ ((Exception $ dis a p) <++> Closed) 403 | 404 | --- Internal Hom --- 405 | 406 | infixr 4 ^^ 407 | (^^) : Arena -> Arena -> Arena 408 | (^^) b a = prod (pos a ** arena) 409 | where 410 | arena : pos a -> Arena 411 | arena p = b @@ (Emitter $ dis a p) 412 | 413 | --eval : {a : Arena} -> {b : Arena} -> Lens (a * (b ^^ a)) b 414 | --eval {a} {b} = MkLens (\(pa,pab) => let (pb ** _) = pab pa in pb) (\(pa, pab) , 415 | --db => let (paa ** da) = (pab pa) in (da db , (pa ** ?xx))) 416 | 417 | --- Dynamical systems --- 418 | 419 | 420 | record DynSystem where 421 | constructor MkDynSystem 422 | state : Type 423 | body : Arena 424 | pheno : Lens (Self state) body 425 | 426 | static : DynSystem 427 | static = MkDynSystem () Closed (EmitterFunction id) 428 | 429 | infixr 4 *** 430 | (***) : DynSystem -> DynSystem -> DynSystem 431 | (***) dyn1 dyn2 = 432 | let 433 | state12 : Type 434 | body12 : Arena 435 | pheno12 : Lens (Self state12) body12 436 | state12 = (state dyn1, state dyn2) 437 | body12 = (body dyn1) * (body dyn2) 438 | pheno12 439 | = MkLens (\(s1, s2) => (observe (pheno dyn1) s1, observe (pheno dyn2) s2)) 440 | (\(s1, s2),(d1, d2) => (interpret (pheno dyn1) s1 d1, interpret (pheno dyn2) s2 d2)) 441 | in MkDynSystem state12 body12 pheno12 442 | 443 | --juxtapose : List DynSystem -> DynSystem 444 | --juxtapose [] = static 445 | --juxtapose [d] = d 446 | --juxtapose (d :: ds) = d *** (juxtapose ds) 447 | 448 | install : (d : DynSystem) -> (a : Arena) -> Lens (body d) a -> DynSystem 449 | install d a l = MkDynSystem (state d) a (l <.> (pheno d)) 450 | 451 | speedUp : DynSystem -> Nat -> DynSystem 452 | speedUp dyn n = let 453 | fastBody : Arena 454 | fastWork : Lens (Self $ state dyn) fastBody 455 | fastBody = CircPow (body dyn) n 456 | fastWork = CircPowLens (pheno dyn) n <.> comultPow (state dyn) n 457 | in MkDynSystem (state dyn) fastBody fastWork 458 | 459 | 460 | run : (d : DynSystem) -> enclose (body d) -> (state d) -> Stream (pos $ body d) 461 | run d e s = outp :: (run d e next) 462 | where 463 | outp : pos $ body d 464 | next : state d 465 | outp = observe (pheno d) s 466 | next = interpret (pheno d) s $ interpret e outp () 467 | 468 | 469 | 470 | toStreamBehavior : {a : Arena} -> (b : Behavior a) -> (phys : enclose a) -> Stream (pos a) 471 | toStreamBehavior {a} b phys = currpos :: toStreamBehavior rest phys 472 | where 473 | currpos : pos a 474 | currpos = head b 475 | 476 | rest : Behavior a 477 | rest = tail b $ interpret phys currpos () 478 | 479 | dynBehavior : (d : DynSystem) -> (state d) -> Behavior (body d) 480 | dynBehavior dyn st = let 481 | current : pos $ body dyn 482 | choice : dis (body dyn) current -> Behavior (body dyn) 483 | current = observe (pheno dyn) st 484 | choice d = dynBehavior dyn (interpret (pheno dyn) st d) 485 | in current :: choice 486 | 487 | runBehav : (d : DynSystem) -> enclose (body d) -> (state d) -> Stream (pos $ body d) 488 | runBehav dyn phys st = toStreamBehavior (dynBehavior dyn st) phys 489 | 490 | --- IO --- 491 | 492 | -- EmitterM : (m : Type -> Type) -> Monad m => (a : Type) -> DynSystem 493 | -- EmitterM m a = MkDynSystem (m a) (Emitter (m a)) passInput 494 | -- where 495 | -- passInput : Lens (Self (m a)) (Emitter (m a)) 496 | -- passInput = MkLens id 497 | 498 | IOEmitter : (a : Type) -> (String -> a) -> DynSystem 499 | IOEmitter a cast = let 500 | passUserInput : Lens (Self $ IO a) (Emitter $ IO a) 501 | passUserInput = MkLens id listen 502 | where 503 | listen : (old : IO a) -> () -> IO a 504 | listen _ _ = map cast getLine 505 | in MkDynSystem (IO a) (Emitter $ IO a) passUserInput 506 | 507 | 508 | --- Debugging --- 509 | 510 | -- current : (d : DynSystem) -> state d -> pos (body d) 511 | -- current d s = observe (pheno d) s 512 | 513 | -- feed : (dyn : DynSystem) -> (s : state dyn) -> dis (body dyn) (observe (pheno dyn) s) -> state dyn 514 | -- feed dyn s d = interpret play s d 515 | -- where 516 | -- play : Lens (Self (state dyn)) (body dyn) 517 | -- play = pheno dyn 518 | 519 | --- Examples --- 520 | 521 | 522 | funcToDynSystem : {s : Type} -> {t : Type} -> (s -> t) -> DynSystem 523 | funcToDynSystem {s} {t} f = let 524 | bodyf : Arena 525 | phenof : Lens (Self t) bodyf 526 | bodyf = ArenaIO s t 527 | phenof = MkLens id (\_ => f) 528 | in MkDynSystem t bodyf phenof 529 | 530 | delay : (s : Type) -> DynSystem 531 | delay s = funcToDynSystem (the (s -> s) id) 532 | 533 | plus : DynSystem 534 | plus = funcToDynSystem (uncurry (+)) 535 | 536 | inv : {a,b : Type} -> (f : a -> b) -> b -> Type 537 | inv f y = DPair a (\ x => (f x = y)) 538 | 539 | record Graph (Node : Type) (Edge : Type) where 540 | constructor MkGraph 541 | source, target : Edge -> Node 542 | 543 | namespace graph_examples 544 | grph1 : Graph (Nat , Nat) (Nat, Nat, Nat, Nat) 545 | grph1 = MkGraph (\(i,j,m,n) => (m,n)) (\(i,j,m,n) => (m + i,n + j)) 546 | 547 | -- Exmaple 1.122 548 | --emanation : {n, e : Type} -> Graph n e -> Arena 549 | --emanation g = MkArena n $ inv (source g) 550 | 551 | --graphDyn : {node, edge : Type} -> Graph node edge -> DynSystem 552 | --graphDyn g = MkDynSystem node (emanation g) (MkLens id \ _ , (e ** p) => target g e) 553 | --- 554 | 555 | 556 | -- Exmaple 1.137 557 | emanation : {node,edge : Type} -> (g : Graph node edge) -> (t : node -> Type) -> node -> Arena 558 | emanation g t v = let dis = ((a : inv (source g) v) -> t (target g (fst a))) 559 | in ArenaIO dis (t v) 560 | 561 | --juxt : (ind : Type) -> (ind -> Arena) -> Arena 562 | --juxt ind arena = MkArena ((i : ind) -> pos (arena i)) 563 | --(\p => (i : ind) -> dis (arena i) (p i)) 564 | 565 | data Direction = N | NE | E | SE | S | SW | W | NW 566 | tgt_gol : ((Integer, Integer), Direction) -> (Integer, Integer) 567 | tgt_gol ((i,j),d) = case d of 568 | N => (i,j+1) 569 | S => (i,j-1) 570 | E => (i+1,j) 571 | W => (i-1,j) 572 | NE => (i+1,j+1) 573 | NW => (i-1,j+1) 574 | SE => (i+1,j-1) 575 | SW => (i-1,j-1) 576 | 577 | 578 | juxtLenses : (ind : Type) -> (as,bs : ind -> Arena) 579 | -> ((i : ind) -> Lens (as i) (bs i)) 580 | -> Lens (juxt ind as) (juxt ind bs) 581 | juxtLenses i as bs ls = 582 | MkLens (\pas, i => observe (ls i) (pas i)) 583 | (\pas, dbs, i => interpret (ls i) (pas i) (dbs i)) 584 | 585 | juxtapose : {ind : Type} -> (ind -> DynSystem) -> DynSystem 586 | juxtapose ds = 587 | MkDynSystem ((v : ind) -> state (ds v)) 588 | (juxt ind $ \v => body (ds v)) 589 | (juxtLenses ind (\i => Self (state (ds i))) 590 | (\i => body (ds i)) 591 | (\i => pheno (ds i))) 592 | 593 | cellArena : Arena 594 | cellArena = MkArena Bool (\_ => Direction -> Bool) 595 | 596 | bool2Nat : Bool -> Nat 597 | bool2Nat True = 1 598 | bool2Nat False = 0 599 | 600 | countNear : (Direction -> Bool) -> Nat 601 | countNear f = let f' = bool2Nat . f in f' N + f' S + f' E + f' W + f' NE + f' NW + f' SE + f' SW 602 | 603 | cellUpdt : (p : Bool) -> (Direction -> Bool) -> Bool 604 | cellUpdt True f = countNear f == 2 || countNear f == 3 605 | cellUpdt False f = countNear f == 3 606 | 607 | --install : (d : DynSystem) -> (a : Arena) -> Lens (body d) a -> DynSystem 608 | --install d a l = MkDynSystem (state d) a (l <.> (pheno d)) 609 | 610 | graphWD : {node, edge : Type} -> (g : Graph node edge) 611 | -> (t : node -> Type) -> Lens (juxt node $ emanation g t) 612 | (Emitter ((v : node) -> t v)) 613 | graphWD g t = MkLens id ?st -- (\f, (), _, a => f (target g (fst a))) 614 | 615 | 616 | cellDyn : DynSystem 617 | cellDyn = MkDynSystem Bool cellArena $ MkLens id (\p, updt => cellUpdt p updt) 618 | 619 | golGraph : Graph (Integer, Integer) ((Integer, Integer), Direction) 620 | golGraph = MkGraph (\(ij,_) => ij) tgt_gol 621 | 622 | golDyn : DynSystem 623 | golDyn = let cellDyns : (Integer,Integer) -> DynSystem 624 | cellDyns _ = cellDyn 625 | in juxtapose cellDyns 626 | 627 | 628 | GoL : DynSystem 629 | GoL = let ll = graphWD golGraph (\_ => Bool) 630 | in install golDyn (Emitter ((Integer,Integer) -> Bool)) 631 | (MkLens id (\f, (), ij, d => f (tgt_gol (ij, d)))) 632 | 633 | 634 | initGol : (Integer,Integer) -> Bool 635 | initGol (i,j) = j == 0 && i >= (-1) && i <= 1 636 | 637 | runGol : Stream ((Integer,Integer) -> Bool) 638 | runGol = run GoL enclosed initGol 639 | 640 | viewGol : Stream (Bool,Bool,Bool) 641 | viewGol = map (\f => (f (-1,0), f (0,0), f (1,0))) runGol 642 | 643 | -- (graphWD golGraph ?ss) 644 | 645 | Prefib : DynSystem 646 | Prefib = plus *** (delay Integer) 647 | 648 | fibwd : Lens (body Prefib) (Emitter Integer) 649 | fibwd = MkLens observe interpret 650 | where 651 | observe : (Integer, Integer) -> Integer 652 | interpret : (p : (Integer, Integer)) -> () -> dis (body Prefib) p 653 | observe (pl, de) = de 654 | interpret (pl, de) = \_ => ((de, pl), pl) 655 | 656 | Fibonacci : DynSystem 657 | Fibonacci = install Prefib (Emitter Integer) fibwd 658 | 659 | FibSeq : Stream Integer 660 | FibSeq = run Fibonacci enclosed (1, 1) 661 | 662 | FibSeq5 : Lens (MkArena (p : Integer ** () -> (p : Integer ** () -> ())) ?t) Closed 663 | -> Stream (pos (body (speedUp Fibonacci 2))) 664 | FibSeq5 = \x => run (speedUp Fibonacci 2) x (1,1) 665 | 666 | -------- Fibonacci in REPL -------- 667 | main : IO () 668 | main = print $ take 10000 FibSeq 669 | 670 | -- Run this in the REPL: 671 | -- take 10 FibSeq 672 | 673 | -- ========================== 674 | 675 | 676 | --- Distributivity --- 677 | 678 | prodSum : {a, b, c : Arena} -> Lens (a <**> (b <++> c)) ((a <**> b) <++> (a <**> c)) 679 | prodSum {a} {b} {c} = let 680 | a1 : Arena 681 | a1 = a <**> (b <++> c) 682 | a2 : Arena 683 | a2 = (a <**> b) <++> (a <**> c) 684 | o : pos a1 -> pos a2 685 | o (p, Left q) = Left (p, q) 686 | o (p, Right r) = Right (p, r) 687 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 688 | i (pa, Left pb) (Left da) = Left da 689 | i (pa, Left pb) (Right db) = Right db 690 | i (pa, Right pc) (Left da) = Left da 691 | i (pa, Right pc) (Right dc) = Right dc 692 | in MkLens o i 693 | sumProd : {a, b, c : Arena} -> Lens ((a <**> b) <++> (a <**> c)) (a <**> (b <++> c)) 694 | sumProd {a} {b} {c} = let 695 | a1 : Arena 696 | a1 = (a <**> b) <++> (a <**> c) 697 | a2 : Arena 698 | a2 = a <**> (b <++> c) 699 | o : pos a1 -> pos a2 700 | o (Left (pa, pb)) = (pa, Left pb) 701 | o (Right (pa, pc)) = (pa, Right pc) 702 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 703 | i (Left (pa, pb)) (Left da) = Left da 704 | i (Left (pa, pb)) (Right db) = Right db 705 | i (Right (pa, pc)) (Left da) = Left da 706 | i (Right (pa, pc)) (Right dc) = Right dc 707 | in MkLens o i 708 | 709 | juxtSum : {a, b, c : Arena} -> Lens (a * (b <++> c)) ((a * b) <++> (a * c)) 710 | juxtSum {a} {b} {c} = let 711 | a1 : Arena 712 | a1 = a * (b <++> c) 713 | a2 : Arena 714 | a2 = (a * b) <++> (a * c) 715 | o : pos a1 -> pos a2 716 | o (pa, Left pb) = Left (pa, pb) 717 | o (pa, Right pc) = Right (pa, pc) 718 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 719 | i (pa, Left pb) (da, db) = (da, db) 720 | i (pa, Right pc) (da, dc) = (da, dc) 721 | in MkLens o i 722 | 723 | sumJuxt : {a, b, c : Arena} -> Lens ((a * b) <++> (a * c)) (a * (b <++> c)) 724 | sumJuxt {a} {b} {c} = let 725 | a1 : Arena 726 | a1 = (a * b) <++> (a * c) 727 | a2 : Arena 728 | a2 = a * (b <++> c) 729 | o : pos a1 -> pos a2 730 | o (Left (pa, pb)) = (pa, Left pb) 731 | o (Right (pa, pc)) = (pa, Right pc) 732 | i : (p : pos a1) -> dis a2 (o p) -> dis a1 p 733 | i (Left (pa, pb)) (da, db) = (da, db) 734 | i (Right (pa, pc)) (da, dc) = (da, dc) 735 | in MkLens o i 736 | -------------------------------------------------------------------------------- /snippets/temp-intro.tex: -------------------------------------------------------------------------------- 1 | \Opensolutionfile{solutions}[solution-file1] 2 | 3 | %------------ Chapter ------------% 4 | \chapter{Introduction}\label{ch.poly.intro} 5 | 6 | \begin{quote} 7 | It is a treasury box!\\ 8 | Full of unexpected connections!\\ 9 | It is fascinating!\\ 10 | I will think about it.\\ 11 | \mbox{}\hfill ---Andr\'e Joyal, Summer 2020,\\ 12 | \mbox{}\hfill personal communication. 13 | \end{quote} 14 | 15 | %-------- Section --------% 16 | \section{Perspectives on polynomials} 17 | 18 | In this book we will investigate a remarkable category called $\poly$ and its intimate relationships with interactive and dynamic processes. 19 | 20 | 21 | Throughout this book, we will see several other perspectives from which we can view polynomials. 22 | Here is a table of terminology, capturing five different perspectives from which we may view our objects of study. 23 | The first row shows the algebraic notation, as in \eqref{eqn.poly_example}; 24 | the second row shows the indexed family terminology, as in \eqref{eqn.arena_example}; 25 | the third shows the pictorial terminology of trees, as in \eqref{eqn.forest_example}; the fourth shows decision-making terminology, giving our polynomials semantics that we will introduce in \cref{sec.poly.intro.dec}; and the fifth row shows dynamical systems terminology, which we will explore in \cref{ch.poly.dyn_sys}. 26 | 27 | \begin{equation}%\label{eqn.table_terminology} 28 | \footnotesize 29 | \begin{tabular}{l|l|l|l} 30 | \multicolumn{4}{c}{\normalsize Polynomial Terminology}\\[3pt] 31 | \textbf{algebra} & $p \coloneqq \sum_{i \in p(\1)} \yon^{p[i]}$ & $i \in p(\1)$ & $a \in p[i]$ \\ 32 | \textbf{indexed families} & arena & position & direction \\ 33 | \textbf{tree pictures} & corolla forest & root $\bullet$ & leaf $\uparrow$ \\ 34 | \textbf{decisions} & scenario & menu & option \\ 35 | \textbf{dynamics} & interface & output & input 36 | \end{tabular} 37 | \end{equation} 38 | 39 | We will freely switch between these equivalent terms depending on context. 40 | For instance, though a polynomial will turn out to be a \emph{functor}, while an arena is an \emph{indexed family}, they are so closely related that we often do not make a distinction between a polynomial $p$ and its arena $(p[i])_{i \in I}$; they are two different syntaxes for the same object. 41 | In particular, we will directly refer to the positions and directions of a polynomial when we mean the positions and directions of its associated arena. 42 | 43 | 44 | 45 | 46 | 47 | Before we can really get into this story, let's summarize where we're going: polynomials are going to have really surprising applications to modeling general interaction, and $\poly$ as a category has an abundance of structure that we can draw on. 48 | We speak superlatively of $\poly$: 49 | \slogan{ 50 | The category of polynomials is a jackpot. Its beauty flows without bound} 51 | but we have not yet begun to deliver. So let's introduce some of the applications and mathematics to come. 52 | 53 | %-------- Section --------% 54 | \section{Decisions and interaction in series} \label{sec.poly.intro.dec} 55 | We return to the example polynomial $\yon^\2 + \2\yon + \1$ from \eqref{eqn.poly_example} and its corresponding corolla forest, in which positions are expressed as roots and directions are represented as leaves (this time, we've colored the roots to distinguish them): 56 | \begin{equation} \label{eqn.forest2110} 57 | \begin{tikzpicture}[trees] 58 | \node[dgreen] (1) {$\bullet$} 59 | child {} 60 | child {}; 61 | \node[right=.5 of 1,blue] (2) {$\bullet$} 62 | child {}; 63 | \node[right=.5 of 2,dyellow] (3) {$\bullet$} 64 | child {}; 65 | \node[right=.5 of 3,red] (4) {$\bullet$}; 66 | \end{tikzpicture} 67 | \end{equation} 68 | As in our arena language, we will call the corollas in the forest for $p$ the $p$-corollas, whose roots are $p$-roots and whose leaves at each root $i$ are the $p[i]$-leaves. 69 | Since each corolla can be identified with its root, we will often use the same name $i$ for both the $p$-corolla and its $p$-root. 70 | 71 | Concretely, we might think of each position as representing a \emph{menu}. 72 | Associated to every menu is a set of \emph{options} (directions). 73 | The four menus we exhibit in \eqref{eqn.forest2110} are particularly interesting: they respectively have two options, one option, one option, and no options. 74 | Having two options is familiar from life---it's the classic yes/no decision---as well as from Claude Shannon's information theory. 75 | Having one option is also familiar both theoretically and in life: ``sorry, ya just gotta go through it.'' 76 | Having no options is when you actually don't get through it: an impossible decision, a sort of ``dead end.'' 77 | We refer to the menus and options represented by the arena as a whole as the \emph{scenario}. 78 | 79 | This decision-making perspective will prove insightful when we start to discuss the \emph{morphisms} between two polynomials. 80 | A morphism will model how two different agents, each in their own scenario, may interact. 81 | If we think of our lives as series of decisions we make, we interact with others by letting our own decisions depend on the decisions of others. 82 | In some sense, any interaction between two agents is a way of \emph{delegating} decisions from one scenario to another. 83 | We'll cover this in more detail in \cref{sec.poly.func_nat.morph}. 84 | 85 | For now, let's hone in on that ``in series'' part a bit. 86 | Life is not a single decision---it is a series of decisions, each dependent on the last. 87 | A game theorist might model such a series with a \emph{decision tree}, like the ones below (the first two are infinite, but that’s hard to draw, so we’ve included just the first five levels): 88 | \[ 89 | \begin{tikzpicture}[trees] 90 | \begin{scope}[ 91 | level 1/.style={sibling distance=20mm}, 92 | level 2/.style={sibling distance=10mm}, 93 | level 3/.style={sibling distance=5mm}, 94 | level 4/.style={sibling distance=2.5mm}, 95 | level 5/.style={sibling distance=1.25mm}] 96 | \node[dgreen] (a) {$\bullet$} 97 | child {node[dgreen] {$\bullet$} 98 | child {node[dgreen] {$\bullet$} 99 | child {node[dgreen] {$\bullet$} 100 | child {node[dgreen] {$\bullet$} 101 | child {} 102 | child {} 103 | } 104 | child {node[dyellow] {$\bullet$} 105 | child {} 106 | } 107 | } 108 | child {node[blue] {$\bullet$} 109 | child {node[dgreen] {$\bullet$} 110 | child {} 111 | child {} 112 | } 113 | } 114 | } 115 | child {node[dgreen] {$\bullet$} 116 | child {node[dgreen] {$\bullet$} 117 | child {node[dgreen] {$\bullet$} 118 | child {} 119 | child {} 120 | } 121 | child {node[blue] {$\bullet$} 122 | child {} 123 | } 124 | } 125 | child {node[red] {$\bullet$}} 126 | } 127 | } 128 | child {node[dyellow] {$\bullet$} 129 | child {node[dgreen] {$\bullet$} 130 | child {node[dgreen] {$\bullet$} 131 | child {node[dgreen] {$\bullet$} 132 | child {} 133 | child {} 134 | } 135 | child {node[blue] {$\bullet$} 136 | child {} 137 | } 138 | } 139 | child {node[dyellow] {$\bullet$} 140 | child {node[red] {$\bullet$}} 141 | } 142 | } 143 | } 144 | ; 145 | \end{scope} 146 | \begin{scope}[ 147 | level 1/.style={sibling distance=13mm}, 148 | level 2/.style={sibling distance=10mm}, 149 | level 3/.style={sibling distance=5mm}, 150 | level 4/.style={sibling distance=2.5mm}, 151 | level 5/.style={sibling distance=1.25mm}] 152 | \node (b) [right=4 of a, blue] {$\bullet$} 153 | child {node[dgreen] {$\bullet$} 154 | child {node[dgreen] {$\bullet$} 155 | child {node[dgreen] {$\bullet$} 156 | child {node[dgreen] {$\bullet$} 157 | child {} 158 | child {} 159 | } 160 | child {node[blue] {$\bullet$} 161 | child {} 162 | } 163 | } 164 | child {node[dyellow] {$\bullet$} 165 | child {node[dgreen] {$\bullet$} 166 | child {} 167 | child {} 168 | } 169 | } 170 | } 171 | child {node[dgreen] {$\bullet$} 172 | child {node[red] {$\bullet$}} 173 | child {node[dgreen] {$\bullet$} 174 | child {node[blue] {$\bullet$} 175 | child {} 176 | } 177 | child {node[dyellow] {$\bullet$} 178 | child {} 179 | } 180 | } 181 | } 182 | } 183 | ; 184 | \end{scope} 185 | \node (c) [red, right=2 of b] {$\bullet$}; 186 | \end{tikzpicture} 187 | \] 188 | Each tree models a series of decisions that may be presented to us. 189 | We start at the bottom---the root---and climb up the tree, making a choice of which arrow to follow at each step. 190 | Sometimes we are presented with two options; sometimes just one; and sometimes no option at all. 191 | Eventually, our choices trace out a path up the tree. 192 | 193 | But squint at these trees, and you'll see that they are made up of something quite familiar---every tree is built out of several smaller corollas, of exactly the form we drew in \eqref{eqn.forest2110}! 194 | It makes sense now why we interpret our polynomials both as roots and leaves of corollas and as menus and options of decisions---when we arrange these corollas and decisions in series, we get a decision tree. 195 | We call each tree that we can build out of the corollas of $\yon^\2+\2\yon+\1$, like the three trees drawn above, a \emph{$(\yon^\2+\2\yon+\1)$-tree}. 196 | Similarly, we could define and construct a $p$-tree for any polynomial $p$. 197 | 198 | What's amazing is that this is not just an arbitrarily imposed reading of our polynomials---it is directly from the categorical structure of $\poly$ that such trees arise. 199 | But we will have to wait before we have all the machinery we need to discuss what's happening here---we'll revisit these trees in \cref{sec.comon.cofree.cons}. 200 | 201 | \begin{exercise}\label{exc.decision_streams} 202 | \begin{enumerate} 203 | \item Draw the first three levels of a $(\yon^\2+\yon^\0)$-tree. 204 | \item Draw the first four levels of a $\yon$-tree. 205 | \item Draw a the first three levels of a $\nn\yon^\2$-tree by labeling every node with a natural number. 206 | \qedhere 207 | \end{enumerate} 208 | 209 | \begin{solution} 210 | \begin{enumerate} 211 | \item Here are the first three levels of a $(\yon^\2+\yon^\0)$-tree. 212 | \[ 213 | \begin{tikzpicture}[trees, 214 | level 1/.style={sibling distance=20mm}, 215 | level 2/.style={sibling distance=10mm}, 216 | level 3/.style={sibling distance=5mm}, 217 | level 4/.style={sibling distance=2.5mm}] 218 | \node (a) {$\bullet$} 219 | child {node {$\bullet$} 220 | child {node {$\bullet$} 221 | child 222 | child 223 | } 224 | child {node {$\bullet$} 225 | } 226 | } 227 | child {node {$\bullet$} 228 | child {node {$\bullet$}} 229 | child {node {$\bullet$} 230 | } 231 | } 232 | ; 233 | \end{tikzpicture} 234 | \] 235 | \item Here are the first four levels of a $\yon$-tree. 236 | \[ 237 | \begin{tikzpicture}[trees] 238 | \node (a) {$\bullet$} 239 | child {node {$\bullet$} 240 | child {node {$\bullet$} 241 | child {node {$\bullet$} 242 | child 243 | }}}; 244 | \end{tikzpicture} 245 | \] 246 | \item Here are the first three levels of a $\nn\yon^\2$-tree, where we indicate the position of each node by labeling it with a natural number. 247 | \[ 248 | \begin{tikzpicture}[trees, 249 | level 1/.style={sibling distance=20mm}, 250 | level 2/.style={sibling distance=10mm}, 251 | level 3/.style={sibling distance=5mm}, 252 | level 4/.style={sibling distance=2.5mm}] 253 | \node (a) {$27$} 254 | child {node {$5040$} 255 | child {node {$192$} 256 | child 257 | child 258 | } 259 | child {node {$0$} 260 | child 261 | child 262 | } 263 | } 264 | child {node {$314159$} 265 | child {node {$1000$} 266 | child 267 | child 268 | } 269 | child {node {$1296$} 270 | child 271 | child 272 | } 273 | } 274 | ; 275 | \end{tikzpicture} 276 | \] 277 | \end{enumerate} 278 | \end{solution} 279 | \end{exercise} 280 | 281 | 282 | 283 | %-------- Section --------% 284 | \section{Dynamical systems} \label{sec.poly.intro.dyn_sys} 285 | 286 | When we say ``dynamical system,'' we are referring to a concept that may be familiar: a machine that stores an internal state. 287 | The machine may return output according to its current state, while it may also receive input that updates this state. 288 | 289 | For example, the internal state of a digital clock may consist of the time and the display format (``12-hour'' vs. ``24-hour''). 290 | If the time is six minutes past noon and the display format is ``12-hour,'' the clock will display ``12:06pm'' as its output. 291 | If the time is twenty minutes till midnight and the display format is ``24-hour,'' the clock will instead display ``23:40.'' 292 | 293 | Meanwhile, the clock may receive input via one button that increments its current time by one minute and another button that toggles between the two display formats. 294 | At the press of a button, the internal state will change depending on which button was pressed and what the previous internal state was. 295 | 296 | So our digital clock is a very simple dynamical system. 297 | We can think of its input buttons and its output display as the way in which the clock interacts with the outside world---its \emph{interface}. 298 | 299 | Here's one way the clock might interact with the outside world: my finger is on the minute button, while your finger is on the format button, and the clock is facing you so that only you can read its display. 300 | In fact, we could think of ourselves as a couple of (particularly complex) dynamical systems as well, with each of our interfaces connected with the interface of the clock. 301 | We could model this situation with the following picture, called a \emph{wiring diagram}, showing how each system can receive input from and send output to other systems in a particular interaction pattern: 302 | \begin{equation*} 303 | \begin{tikzpicture}[oriented WD, every fit/.style={inner xsep=\bbx, inner ysep=\bby}, bb min width =.5cm, bbx=.5cm, bb port sep =1,bb port length=0, bby=.15cm] 304 | \node[bb={3}{3}, green!25!black] (X12) {\tiny me}; 305 | \node[bb={2}{2}, green!25!black, below right = -1 and 1.5 of X12] (X21) {\tiny clock}; 306 | \node[bb={1}{2}, green!25!black, above right=-1 and 1 of X21] (X22) {\tiny you}; 307 | \draw (X21_out1) to (X22_in1); 308 | \draw let \p1=(X22.north east), \p2=(X21.north west), \n1={\y1+\bby}, \n2=\bbportlen in 309 | (X22_out1) to[in=0] (\x1+\n2,\n1) -- (\x2-\n2,\n1) to[out=180] (X21_in1); 310 | \draw (X12_out2) to (X21_in2); 311 | \end{tikzpicture} 312 | \end{equation*} 313 | 314 | Of course, there is a lot going on in the world around us that we haven't drawn. 315 | We each have some input ports: our eyes, our ears, etc., and some output ports: our speech, our gestures, etc. We can connect with other systems: our family, our colleagues, etc. And we can think of all of these systems as subsystems of one larger system interacting with the world. 316 | \begin{equation} \label{eqn.wired_forever} 317 | \begin{tikzpicture}[oriented WD, every fit/.style={inner xsep=\bbx, inner ysep=\bby}, bb min width =.5cm, bbx=.5cm, bb port sep =1,bb port length=0, bby=.15cm] 318 | \node[bb={2}{2}, green!25!black] (X11) {\tiny Alice}; 319 | \node[bb={3}{3}, green!25!black, below right=of X11] (X12) {\tiny me}; 320 | \node[bb={2}{1}, green!25!black, above right=of X12] (X13) {\tiny Bob}; 321 | \node[bb={2}{2}, green!25!black, below right = -1 and 1.5 of X12] (X21) {\tiny clock}; 322 | \node[bb={1}{2}, green!25!black, above right=-1 and 1 of X21] (X22) {\tiny you}; 323 | \node[bb={2}{2}, fit = {($(X11.north east)+(-1,4)$) (X11) (X12) (X13) ($(X21.south)$) ($(X22.east)+(.5,0)$)}, bb name = {\small Wired together like this forever?}] (Z) {}; 324 | \draw (X21_out1) to (X22_in1); 325 | \draw let \p1=(X22.north east), \p2=(X21.north west), \n1={\y1+\bby}, \n2=\bbportlen in 326 | (X22_out1) to[in=0] (\x1+\n2,\n1) -- (\x2-\n2,\n1) to[out=180] (X21_in1); 327 | \draw (X11_out1) to (X13_in1); 328 | \draw (X11_out2) to (X12_in1); 329 | \draw (X12_out1) to (X13_in2); 330 | \draw (Z_in1'|-X11_in2) to (X11_in2); 331 | \draw (Z_in2'|-X12_in2) to (X12_in2); 332 | \draw (X12_out2) to (X21_in2); 333 | % \draw (X21_out2) to (Z_out2'|-X21_out2); 334 | \draw let \p1=(X12.south east), \p2=(X12.south west), \n1={\y1-\bby}, \n2=\bbportlen in 335 | (X12_out3) to[in=0] (\x1+\n2,\n1) -- (\x2-\n2,\n1) to[out=180] (X12_in3); 336 | \draw let \p1=(X22.north east), \p2=(X11.north west), \n1={\y2+\bby}, \n2=\bbportlen in 337 | (X22_out2) to[in=0] (\x1+\n2,\n1) -- (\x2-\n2,\n1) to[out=180] (X11_in1); 338 | \draw (X13_out1) to (Z_out1'|-X13_out1); 339 | \end{tikzpicture} 340 | \end{equation} 341 | 342 | We wrote a little question for you at the top of the diagram. 343 | Isn't there something a little funny about the way we've connected these systems? 344 | Maybe for very simple machines, you would wire things together once and they'd stay like that for the life of the machine. 345 | But you could turn your eyes away from the clock to look at Bob, Alice drops her connection to me for weeks at a time, and I would really like to be able to lift my finger off of the clock to do something else. 346 | So 347 | \slogan{the way systems connect can change over time.} 348 | In fact, $\poly$ will let us express this. 349 | 350 | % In this book, we'll be looking at dynamical systems with a lot of interesting new options: 351 | % \begin{enumerate} 352 | % \item The interface of the system---the way in which it can be interacted with---can change shape through time. 353 | % \item The wiring diagram connecting a bunch of systems can change through time. 354 | % % \item One can speed up the dynamics of a system. 355 | % % \item One can introduce ``effects,'' i.e.\ as defined by monads on $\smset$. 356 | % % \item The dynamical systems on any interface form a topos. 357 | % \end{enumerate} 358 | % To give some intuition for the first two, imagine yourself as a system, wired up to other systems. 359 | 360 | 361 | \begin{example}\label{ex.changing_wiring_bonds_supplier_assemble} 362 | Here are some familiar circumstances where we see interaction patterns changing over time. 363 | \begin{enumerate}[itemsep=0pt] 364 | % \item Airplanes only communicate when they get near enough; 365 | % \item A person can choose when to open (receive input through) their eyes and when to speak (produce output);\goodbreak 366 | \item When too much force is applied to a material, bonds can break: 367 | \end{enumerate} 368 | \[ 369 | \begin{tikzpicture}[oriented WD, bb small, bb port length=0] 370 | \foreach \i in {0,...,4} { 371 | \node[bb={1}{1}, fill=blue!10] at (1.7*\i,0) (X\i) {}; 372 | } 373 | % \node[bb={1}{1}, fit=(X0) (X4)] (X) {}; 374 | \foreach \i in {0,...,3} { 375 | \draw[thick] (X\i_out1) -- (X\the\numexpr\i+1\relax_in1); 376 | }; 377 | \draw[thick, ->] (X0_in1) -- node[above, font=\tiny] {Force} +(-2.5,0); 378 | \draw[thick, ->] (X4_out1) -- node[above, font=\tiny] {Force} +(2.5,0) node (R) {}; 379 | % 380 | \def\x{21}; 381 | \foreach \i in {0,...,2} { 382 | \node[bb={1}{1}, fill=blue!10] at (\x+1.7*\i,0) (Y\i) {}; 383 | } 384 | \foreach \i in {3,...,4} { 385 | \node[bb={1}{1}, fill=blue!10] at (\x+1.3+1.7*\i,0) (Y\i) {}; 386 | } 387 | % \node[bb={1}{1}, fit=(Y0) (Y4)] (Y) {}; 388 | \foreach \i in {0,1,3} { 389 | \draw[thick] (Y\i_out1) -- (Y\the\numexpr\i+1\relax_in1); 390 | }; 391 | \draw[thick, ->] (Y0_in1) -- node[above, font=\tiny] {Force} +(-2.5,0) node (L) {}; 392 | \draw[thick, ->] (Y4_out1) -- node[above, font=\tiny] {Force} +(2.5,0); 393 | \node[starburst, draw, minimum width=2cm, minimum height=1.5cm,red,fill=orange,line width=1.5pt] at ($(L)!.5!(R)$) 394 | {Snap!}; 395 | \end{tikzpicture} 396 | \] 397 | \begin{quote} 398 | In materials science, the Young's modulus accounts for how much force can be transferred across a material as its endpoints are pulled apart. When the material breaks, the two sides can no longer feel evidence of each other. Thinking of pulling as sending a signal (a signal of force), we might say that the ability of internal entities to send signals to each other---the connectivity of the wiring diagram---is being measured by the Young's modulus. It will also be visible within $\poly$. 399 | \end{quote} 400 | \begin{enumerate}[resume] 401 | \item A company may change its supplier at any time: 402 | \end{enumerate} 403 | \begin{equation*}%\label{eqn.supplier} 404 | \begin{tikzpicture}[oriented WD, font=\ttfamily, every node/.style={fill=blue!10}, baseline=(c)] 405 | \node[bb={0}{1}] (s1) {Supplier 1}; 406 | \node[bb={0}{1}, below=of s1] (s2) {Supplier 2}; 407 | \coordinate (helper) at ($(s1)!.5!(s2)$); 408 | \node[bb={1}{0}, right=1.5 of helper] (c) {Company}; 409 | \draw (s1_out1) to (c_in1); 410 | \draw (s2_out1) to +(5pt,0) node[fill=none] {$\bullet$}; 411 | \begin{scope}[xshift=3.5in] 412 | \node[bb={0}{1}] (s1') {Supplier 1}; 413 | \node[bb={0}{1}, below=of s1'] (s2') {Supplier 2}; 414 | \coordinate (helper') at ($(s1')!.5!(s2')$); 415 | \node[bb={1}{0}, right=1.5 of helper'] (c') {Company}; 416 | \draw (s2'_out1) to (c'_in1); 417 | \draw (s1'_out1) to +(5pt,0) node[fill=none] {$\bullet$}; 418 | \end{scope} 419 | \node[starburst, draw, minimum width=2cm, minimum height=2cm,align=center,fill=green!10, font=\small, fill=white, line width=1.5pt] at ($(c)!.5!(helper')$) 420 | {Change\\supplier!}; 421 | \end{tikzpicture} 422 | \end{equation*} 423 | \begin{quote} 424 | The company can get widgets either from supplier 1 or supplier 2; we could imagine this choice is completely up to the company. The company can decide based on the quality of widgets it has received in the past: when the company gets a bad widget, it updates an internal variable, and sometimes that variable passes a threshold making the company switch states. Whatever its strategy for deciding, we should be able to encode it in $\poly$. 425 | \end{quote} 426 | \begin{enumerate}[resume] 427 | \item When someone assembles a machine, their own outputs dictate the connection pattern of the machine's components. 428 | \end{enumerate} 429 | \begin{equation*}%\label{eqn.someone} 430 | \begin{tikzpicture}[oriented WD, font=\ttfamily, bb port length=0, every node/.style={fill=blue!10}, baseline=(someone.north)] 431 | \node[bb port sep=.5, bb={0}{1}] (A) {unit A}; 432 | \node[bb port sep=.5, bb={1}{0}, right=of A] (B) {unit B}; 433 | \coordinate (helper) at ($(A)!.5!(B)$); 434 | \node[bb={1}{1}, below=2 of helper] (someone) {\tikzsymStrichmaxerl[3]}; 435 | \draw[->, dashed, blue] (someone_in1) to[out=180, in=270] (A.270); 436 | \draw[->, dashed, blue] (someone_out1) to[out=0, in=270] (B.270); 437 | \draw (A_out1) -- +(10pt,0); 438 | \draw (B_in1) -- +(-10pt,0); 439 | % 440 | \begin{scope}[xshift=3.5in] 441 | \node[bb port sep=.5, bb={0}{1}] (A') {unit A}; 442 | \node[bb port sep=.5, bb={1}{0}, right=.5of A'] (B') {unit B}; 443 | \coordinate (helper') at ($(A')!.5!(B')$); 444 | \node[bb={1}{1}, below=2 of helper'] (someone') {\tikzsymStrichmaxerl[3]}; 445 | \draw[->, dashed, blue] (someone'_in1) to[out=180, in=270] (A'.270); 446 | \draw[->, dashed, blue] (someone'_out1) to[out=0, in=270] (B'.270); 447 | \draw (A'_out1) -- (B'_in1); 448 | \end{scope} 449 | % 450 | \node[starburst, draw, minimum width=2cm, minimum height=2cm,fill=blue!50,line width=1.5pt, align=center, font=\upshape] at ($(B)!.5!(A')-(0,.6cm)$) 451 | {Attach!}; 452 | \end{tikzpicture} 453 | \end{equation*} 454 | \begin{quote} 455 | Have you ever assembled something? Your internal states dictate the interaction pattern of some other things. We can say this in $\poly$. 456 | \end{quote} 457 | 458 | All of the examples discussed here will be presented in some detail once we have the requisite mathematical theory (\cref{ex.bonds_break,ex.supplier_change,ex.assemble_machine}). 459 | \end{example} 460 | 461 | \begin{exercise}%\label{exc.changing_types} 462 | Think of another example where systems are sending each other information, but where who the information is being sent to or received from can change based on the states of the systems involved. You might have more than two, say $\rr$-many, different interaction patterns in your setting. 463 | \begin{solution} 464 | When I am carrying my phone in my house, my phone will connect to my Wi-Fi router. But my phone may send me a reminder that tells me to leave my house. 465 | Once I carry my phone far enough away, it will disconnect from my router and connect to a cellular network instead. 466 | While I am outside, I might then press a button on my phone to disconnect it from the cellular network to reduce my data usage, so that it is no longer connectd to any network. 467 | \end{solution} 468 | \end{exercise} 469 | 470 | But there's more that's intuitively wrong or limiting about the picture in \eqref{eqn.wired_forever}. Ever notice how you can change how you interface with the world? Sometimes I close my eyes, which makes that particular way of sending me information inaccessible: that port vanishes, and you need to use your words. Sometimes I'm in a tunnel and my car can't receive a radio signal. Sometimes I extend my hand to give or receive an object from another person, but sometimes I don't. So 471 | \slogan{a system's interface itself can change over time.} % with the setting? 472 | We will be able to say all this using $\poly$ as well. 473 | 474 | And there's even more that's wrong with the above description. 475 | Namely, when I use my muscles or mouth to express things, my very position changes: my tongue moves, my body moves. 476 | The display of an analog clock is literally the positions of its hands. 477 | So 478 | \slogan{the output of a system is essentially the position it takes.} 479 | Moreover, when I move my eyes, that's something you can actually see---you can tell if I'm looking at you. 480 | When I turn around, I see different things, and \emph{you can notice I'm turned around}! 481 | In other words, 482 | \slogan{the range of inputs a system can receive depends on the position it currently outputs.} 483 | This is integral to our model of dynamical systems in $\poly$, and why we say that outputs correspond to positions and inputs to directions---with an entire interface represented by a polynomial. 484 | 485 | \begin{example}\label{ex.pond_eyeballs} 486 | Imagine a million little eyeballs, each of which has a tiny brain inside it, all together in a pond of eyeballs. All that an individual eyeball $e$ can do is open and close. When $e$ is open, it can make some distinction about all the rest of the eyeballs in view: maybe it can count how many are open, or maybe it can see whether just one certain eyeball $e'$ is open or closed. But when $e$ is closed, it can't see anything; whatever else is happening, it's all the same to $e$. All it can do in that state is process previous information. 487 | 488 | Each eyeball in this system will correspond to the polynomial $\yon^\ord{n}+\yon$, which consists of two positions: an ``open'' position with $n$-many possible inputs it may perceive, and a ``closed'' position with only one. For simplicity, we could assume $n=2$, so that each eyeball makes a single yes-no distinction whenever it's open. 489 | 490 | The point, however, is that any other eyeball may be capable of noticing if $e$ is open or closed. We can imagine some interesting dynamics in this system, e.g.\ waves of openings or closings sweeping over the group, a ripple of openings expanding through the pond. 491 | 492 | Talk about real-world applications! 493 | \end{example} 494 | 495 | \begin{exercise} 496 | Give another example of a system where the range of possible inputs the system can receive is dependent on what output the system is currently providing. 497 | \begin{solution} 498 | Consider a computer application whose output is an image displayed on the screen. 499 | This image consists of buttons I can click, and each button-click is a possible input that the application can receive. 500 | Clicking a button will alter the display, including which buttons are now available to be clicked. 501 | For example, when I click the ``File'' button, the ``File'' menu will appear, allowing me to click other buttons like the ``New Window'' button. 502 | So the set of inputs I can send to the system (i.e.\ buttons I can click) is directly dependent on the output the system sends to me (i.e.\ the image displayed on the screen). 503 | \end{solution} 504 | \end{exercise} 505 | 506 | Hopefully you now have an idea of what we call \emph{mode-dependence}: interfaces and interaction patterns changing over time, based on the states of all the systems involved. We'll see that $\poly$ speaks about mode-dependent systems and interaction patterns in this sense. 507 | 508 | \begin{remark} 509 | We ended \cref{ex.pond_eyeballs} by joking about ``real-world applications,'' because a pond of eyeballs is about the most bizarre thing one can imagine. But recall Nobel physicist Frank Wilczek's quote from the preface: 510 | \begin{quote} 511 | For me, though, it is difficult to resist the idea that space-time is not essentially different from matter, which we understand more deeply. If so, it will consist of vast numbers of identical units---``particles of space''---each in contact with a few neighbors, exchanging messages, joining and breaking apart, giving birth and passing away. 512 | \end{quote} 513 | Suppose the world was made out of a vast number of identical units, each with its own behavior, able to connect and disconnect with neighbors, and even disappear from the world of cause and effect. We may not even be interested in what our world is actually made of---just what these units are able to do. Is there such an elementary unit that could produce all other dynamical systems? The $\yon^\2+\yon$ eyeballs give a sense of a very simple interface---open and perceiving a single distinction about the world, or closed and making no distinctions---that we could imagine building an entire world from. 514 | \end{remark} 515 | 516 | %-------- Section --------% 517 | 518 | \section{A note on implementation} \label{sec.poly.intro.code} 519 | 520 | In theory, everything we will discuss could be rigorously implemented and verified on a computer---as long as you have access to a functional programming language that supports dependent types, such as Agda or Idris. 521 | One implementation of $\poly$ that uses Cubical Agda can be found at \cite{1lab-poly}. 522 | But even in languages that are not strictly functional, the theory of $\poly$ may still guide programming paradigms for the design and execution of interactive systems. 523 | In this book, we'll focus on the math, but we encourage you to explore the field of computational implementations if that is where your interests lie. 524 | 525 | If you do decide to embark on a computational journey, a word of warning: what we have been calling polynomials---things like $\yon^\2+\2\yon+\1$---are often called \emph{containers} in the computer science (and particularly functional programming) literature. A container consists of a type $S$, usually called the type of \emph{shapes}, and a type $P(s)$ for each term $s:S$, called the type of \emph{positions} in shape $s$. It's mildly unfortunate that the names clash with our own: for us a container-shape is a position and a container-position is a direction. 526 | 527 | %-------- Section --------% 528 | \section{Mathematical theory} \label{sec.poly.intro.math_theory} 529 | 530 | % TODO: change this to a chapter-by-chapter overview of the book, rather than statements of results 531 | 532 | The applications of $\poly$ are quite diverse and interesting, encapsulating decision-based interactions and dynamical systems. 533 | However it is how the mathematics of $\poly$ supports these applications that is so fantastic. 534 | For reference, we list some of the major results about $\poly$ that we will cover in this book; note that this list is by no means comprehensive. 535 | 536 | \begin{proposition} 537 | $\poly$ has all products and coproducts and is completely distributive. 538 | $\poly$ also has exponential objects, making it a bicartesian closed category. 539 | It therefore supports the simply typed lambda calculus. 540 | \end{proposition} 541 | \begin{proof} 542 | We will prove that $\poly$ has coproducts in \cref{prop.poly_coprods}, that it has products in \cref{prop.poly_prods}, that it is completely distributive in \cref{prop.poly_completely_distributive}, and that it has exponential objects in \cref{thm.poly_cart_closed}. 543 | \end{proof} 544 | 545 | \begin{proposition} 546 | Beyond the cocartesian and cartesian monoidal structures $(\0,+)$ and $(\1,\times)$, the category $\poly$ has two additional monoidal structures, denoted $(\yon,\otimes)$ and $(\yon,\circ)$, which are together duoidal.\footnote{We will follow the convention of writing the tensor unit before the tensor product when specifying a monoidal structure.} Moreover $\otimes$ is a closed monoidal structure that distributes over coproducts, while $\circ$ is a left coclosed monoidal structure that preserves connected limits. 547 | \end{proposition} 548 | \begin{proof} 549 | We will define $\otimes$ in \cref{def.parallel} and prove that $(\yon, \otimes)$ is a monoidal structure on $\poly$ in \cref{prop.parallel_monoidal}, and we will define $\circ$ in \cref{def.comp} and prove that $(\yon, \circ)$ is a monoidal structure on $\poly$ in \cref{cor.comp_monoidal}. 550 | Then we will show that $\circ$ is duoidal over $\otimes$ in \cref{prop.duoidal}. 551 | 552 | In \cref{prop.day}, we will show that $\otimes$ distributes over coproducts. 553 | Then in \cref{prop.parallel_closure}, we will prove that $\otimes$ is closed. 554 | In \cref{prop.comp_left_coclosed}, we will show that $\circ$ is left coclosed, and in \cref{thm.connected_limits}, we will show that $\circ$ preserves connected limits. 555 | \end{proof} 556 | 557 | \begin{proposition}\label{prop.adjoint_quadruple} 558 | $\poly$ has an adjoint quadruple with $\smset$ and an adjoint pair with $\smset\op$: 559 | \begin{equation*}%\label{eqn.adjoints_galore} 560 | \begin{tikzcd}[column sep=60pt,background color=definitioncolor] 561 | \smset 562 | \ar[r, shift left=7pt, "A" description] 563 | \ar[r, shift left=-21pt, "A\yon"']& 564 | \poly 565 | \ar[l, shift right=21pt, "p(\0)"'] 566 | \ar[l, shift right=-7pt, "p(\1)" description] 567 | \ar[l, phantom, "\scriptstyle\Leftarrow"] 568 | \ar[l, phantom, shift left=14pt, "\scriptstyle\Rightarrow"] 569 | \ar[l, phantom, shift right=14pt, "\scriptstyle\Rightarrow"] 570 | \end{tikzcd} 571 | \hspace{.6in} 572 | \adjr[50pt]{\smset\op}{\yon^A}{\Gamma(p)}{\poly}.\footnote{ 573 | We use the notation 574 | $\begin{tikzcd}[ampersand replacement=\&] 575 | \cat{C}\ar[r, shift left=4pt, "L"]\& 576 | \cat{D}\ar[l, shift left=4pt, "R"]\ar[l, phantom, "\scriptstyle\Rightarrow"] 577 | \end{tikzcd}$ 578 | to denote an adjunction $L \dashv R$. The double arrow, always pointing in the same direction as the left adjoint, indicates both the unit $\cat{C}\Rightarrow R \circ L$ and the counit $L \circ R\Rightarrow \cat{D}$ of the adjunction. 579 | } 580 | \end{equation*} 581 | Each functor is labeled by where it sends $p\in\poly$ or $A\in\smset$; in particular, $\Gamma(p) \coloneqq \poly(p, \yon)$. 582 | \end{proposition} 583 | \begin{proof} 584 | We will prove that $\poly$ has an adjoint quadruple with $\smset$ in \cref{thm.adjoint_quadruple}, and that it has an adjoint pair with $\smset\op$ in \cref{prop.yoneda_left_adjoint}. 585 | \end{proof} 586 | 587 | \begin{proposition} 588 | $\poly$ has all limits and colimits. 589 | \end{proposition} 590 | \begin{proof} 591 | See \cref{thm.poly_limits} for a proof that $\poly$ has all limits, and \cref{thm.poly_colimits} for a proof that $\poly$ has all colimits. 592 | \end{proof} 593 | 594 | There's a lot we're leaving out of this summary, just so we can hit the highlights. 595 | But here's where things get really interesting. 596 | 597 | \begin{proposition}[Ahman-Uustalu]\label{prop.ahman_uustalu1} 598 | There is a one-to-one isomorphism-preserving correspondence between comonoids in $(\poly,\yon,\circ)$ and small categories. 599 | \end{proposition} 600 | \begin{proof} 601 | See \cref{thm.ahman_uustalu}. 602 | \end{proof} 603 | 604 | \begin{proposition} 605 | The forgetful functor $U\colon\comon(\poly)\to\poly$ has a right adjoint 606 | \[ 607 | \adjr{\poly}{\cofree{-}}{U}{\comon(\poly)} 608 | \] 609 | called the \emph{cofree comonoid} construction. It is lax monoidal with respect to $\otimes$. 610 | \end{proposition} 611 | \begin{proof} 612 | See \cref{thm.cofree}; lax monoidality is proven in \cref{prop.cofree_lax_monoidal}. 613 | \end{proof} 614 | 615 | \begin{proposition} 616 | The category $\comon(\poly)$ of comonoids in $(\poly,\yon,\circ)$ has all limits and colimits. 617 | In particular, coproducts in $\comon(\poly)$ agree with those in $\smcat$. 618 | \end{proposition} 619 | \begin{proof} 620 | See \cref{cor.comon_compl} for a proof that $\comon(\poly)$ has all small limits and \cref{cor.comon_cocomp} for a proof that $\comon(\poly)$ has all small colimits. 621 | We explain why coproducts in $\comon(\poly)$ agree with those in $\smcat$ in \cref{ex.comon_coprod}. 622 | \end{proof} 623 | 624 | \begin{proposition} 625 | The category $\comon(\poly)$ has a symmetric monoidal structure $(\yon,\otimes)$ making the forgetful functor $U\colon(\comon(\poly),\yon,\otimes)\to(\poly,\yon,\otimes)$ strong monoidal. 626 | The monoidal product of comonoids in $\comon(\poly)$ with respect to $\otimes$ coincides with the product of categories in $\smcat$. 627 | \end{proposition} 628 | \begin{proof} 629 | See \cref{prop.parallel_on_catsharp}. 630 | \end{proof} 631 | 632 | For the following results, let $\com{C}$ be the comonoid in $(\poly,\yon,\circ)$ that corresponds to a category $\cat{C}$ under \cref{prop.ahman_uustalu1}, and similarly for $\com{D}$ and $\cat{D}$. 633 | 634 | % \begin{proposition} 635 | % There is an equivalence of categories 636 | % \[ 637 | % \Cat{Bimod}(\com{C},0)\cong[\cat{C},\smset] 638 | % \] 639 | % between $(\com{C},0)$-bicomodules and $\cat{C}$-copresheaves. 640 | % \end{proposition} 641 | % \begin{proof} 642 | % 643 | % \end{proof} 644 | 645 | \begin{proposition} 646 | The category of left $\cat{C}$-comodules is equivalent to the category of functors $\cat{C}\to\poly$. 647 | Meanwhile, a right $\cat{D}$-comodule with polynomial carrier $m$ can be identified (up to isomorphism) with a functor $\cat{D}\to\smset^{m(\1)}$. 648 | \end{proposition} 649 | \begin{proof} 650 | See \cref{prop.left_comod} for the result about left $\cat{C}$-comodules and \cref{prop.right_comod} for the result about right $\cat{D}$-comodules. 651 | \end{proof} 652 | 653 | \begin{proposition}[Garner] 654 | There is an equivalence 655 | \[ 656 | \Cat{Bimod}(\com{C},\com{D})\cong\Cat{pra}([\cat{D},\smset],[\cat{C},\smset]) 657 | \] 658 | between the category of bicomodules over comonoids in $\poly$ and parametric right adjoints between copresheaf categories. 659 | \end{proposition} 660 | \begin{proof} 661 | See \cref{prop.prafunctor}. 662 | \end{proof} 663 | 664 | If you skipped over any of that---or all of that---it'll be no problem whatsoever! We will cover each of the above results in detail over the course of this book. As you read, we encourage you to periodically check back to see how many more of these results you undesrtand. 665 | 666 | There are many avenues for study, but we need to push forward. 667 | We'll begin in the next chapter. 668 | 669 | %-------- Section --------% 670 | \section{Exercise solutions} 671 | \Closesolutionfile{solutions} 672 | {\footnotesize 673 | \input{solution-file1}} --------------------------------------------------------------------------------