├── .gitignore ├── Chapter1 └── Normal.agda ├── Chapter2 └── LambdaCalculus.agda ├── Chapter3 ├── Containers.agda └── W.agda ├── Chapter4 ├── Apocrypha │ └── Roman.agda ├── ITree.agda ├── IndexedContainers.agda └── Jacobian.agda ├── Chapter5 └── InductionRecursion.agda ├── Chapter6 └── OTT.agda ├── Meta ├── Basics.agda ├── Control │ ├── Applicative.agda │ ├── Monad.agda │ └── Traversable.agda ├── Data │ ├── EndoFunctor.agda │ ├── Fin.agda │ ├── Functor │ │ ├── Container.agda │ │ ├── Container │ │ │ └── Indexed.agda │ │ └── Normal.agda │ ├── Inductive │ │ ├── Apocrypha │ │ │ └── Roman.agda │ │ ├── ITree.agda │ │ ├── InductionRecursion.agda │ │ ├── Jacobian.agda │ │ ├── Recursion.agda │ │ └── W.agda │ ├── Monoid.agda │ ├── Nat.agda │ └── Vector.agda └── Language │ ├── LambdaCalculus.agda │ └── OTT.agda └── meta.pdf /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | -------------------------------------------------------------------------------- /Chapter1/Normal.agda: -------------------------------------------------------------------------------- 1 | module Chapter1.Normal where 2 | 3 | open import Meta.Basics 4 | open import Agda.Primitive 5 | 6 | open import Meta.Data.Fin 7 | open import Meta.Data.Nat 8 | open import Meta.Data.Vector 9 | open import Meta.Data.Monoid 10 | open import Meta.Control.Applicative 11 | open import Meta.Control.Traversable 12 | 13 | -- A Normal Functor is given, up to isomorphism, by: 14 | record Normal : Set₁ where 15 | constructor _/_ 16 | field 17 | Shape : Set -- a set of shapes 18 | size : Shape → ℕ -- and a function which assigns to each shape a size. 19 | ⟦_⟧ℕ : Set → Set -- Normal functors can be interpreted as dependent vectors. 20 | ⟦_⟧ℕ X = Σ Shape λ s → Vec X (size s) 21 | open Normal public 22 | infixr 0 _/_ 23 | 24 | -- Vectors are normal functors with a unique shape. 25 | VecN : ℕ → Normal 26 | VecN n = One / const n 27 | 28 | -- Lists are normal functors whose shape is their size. 29 | ListN : Normal 30 | ListN = ℕ / id 31 | 32 | -- "But let us not get ahead of ourselves. We can build a kit for normal 33 | -- functors cor- responding to the type constructors that we often define, then 34 | -- build up composite structures." 35 | 36 | -- The normal functor for any particular point in a type is given by that point 37 | -- and a singular shape: that of the point which has size 0. 38 | Kℕ : Set → Normal 39 | Kℕ A = A / const 0 40 | 41 | -- The identity can be lifted to vectors quite easily. 42 | IKℕ : Normal 43 | IKℕ = VecN 1 44 | 45 | -- The sum of normal functors is the sum of the shapes and a selector for sizes. 46 | _+N_ : Normal → Normal → Normal 47 | (ShF / szF) +N (ShG / szG) = (ShF ⊎ ShG) / vv szF ⟨?⟩ szG 48 | 49 | -- The product of normal functors is the product of shapes and the sum of sizes. 50 | _×N_ : Normal → Normal → Normal 51 | (ShF / szF) ×N (ShG / szG) = (ShF × ShG) / vv λ f g → szF f + szG g 52 | 53 | -- If you "lower" 2 normal functors to their vector representation, you 54 | -- should be able to interpret the sum of vectors as a vector of the sum 55 | -- of the normals. 56 | nInj : ∀ { X } (F G : Normal) → ⟦ F ⟧ℕ X ⊎ ⟦ G ⟧ℕ X → ⟦ F +N G ⟧ℕ X 57 | nInj F G (tt , (ShF , xs)) = (tt , ShF) , xs 58 | nInj F G (ff , (ShG , xs)) = (ff , ShG) , xs 59 | 60 | -- nInj is also surjective. 61 | nCase : ∀ { X } F G (s :  ⟦ F +N G ⟧ℕ X) → nInj F G ⁻¹ s 62 | nCase F G ((tt , ShF) , xs) = from (tt , (ShF , xs)) 63 | nCase F G ((ff , ShG) , xs) = from (ff , (ShG , xs)) 64 | 65 | -- 66 | nOut : ∀ { X }(F G : Normal) → ⟦ F +N G ⟧ℕ X → ⟦ F ⟧ℕ X ⊎ ⟦ G ⟧ℕ X 67 | nOut F G xs' with nCase F G xs' 68 | nOut F G . (nInj F G xs) | from xs = xs 69 | 70 | -- The product of vectors is the vector of inner products 71 | nPair : ∀ { X }(F G : Normal) → ⟦ F ⟧ℕ X × ⟦ G ⟧ℕ X → ⟦ F ×N G ⟧ℕ X 72 | nPair F G ((ShFx , fs) , (ShGx , gs)) = (ShFx , ShGx) , (fs ++ gs) 73 | 74 | ----------------------------------------------------------- 75 | -- Exercise 1.14 76 | -- 77 | -- "While you are in this general area, construct (from readily available 78 | -- components) the usual monoid structure for our normal presentation of lists." 79 | ----------------------------------------------------------- 80 | 81 | split : ∀ m n {X}(xs : Vec X (m + n)) → (vv (_++_ {m}{n}{X})) ⁻¹ xs 82 | split zero n xs = from (⟨⟩ , xs) 83 | split (suc m) n (x , xs) with split m n xs 84 | split (suc m) n (x , .(ys ++ zs)) | from (ys , zs) = from ((x , ys) , zs) 85 | 86 | nSplit : ∀ {X}(F G : Normal)(fgx : ⟦ F ×N G ⟧ℕ X) → nPair F G ⁻¹ fgx 87 | nSplit F G ((f , g) , xs) with split (size F f) (size G g) xs 88 | nSplit F G ((f , g) , .(ys ++ zs)) | from (ys , zs) = from ((f , ys) , (g , zs)) 89 | 90 | concatSurjectivity : forall {m n : ℕ} {X} → (x : Vec X (m + n)) → (vv λ (u : Vec X m) (v : Vec X n) → u ++ v) ⁻¹ x 91 | concatSurjectivity {zero} v = from (⟨⟩ , v) 92 | concatSurjectivity {suc m} (x , v) with concatSurjectivity {m} v 93 | concatSurjectivity {suc m} (x , .(u ++ w)) | from (u , w) = from ((x , u) , w) 94 | 95 | unAppend : ∀ m n {X} → (xsys : Vec X (m + n)) → (vv _++_ {m}{n}) ⁻¹ xsys 96 | unAppend zero n ys = from (⟨⟩ , ys) 97 | unAppend (suc m) n (x , xsys) with unAppend m n xsys 98 | unAppend (suc m) n (x , .(xs ++ ys)) | from (xs , ys) = from ((x , xs) , ys) 99 | 100 | instance 101 | listNMonoid : {X : Set} → Monoid (⟦ ListN ⟧ℕ X) 102 | listNMonoid = record 103 | { ε = (zero , ⟨⟩) 104 | ; _•_ = vv λ xn xs → vv λ yn ys → (xn + yn) , (xs ++ ys) 105 | } 106 | ----------------------------------------------------------- 107 | ----------------------------------------------------------- 108 | 109 | listNMonoidOK : {X : Set} → MonoidOK (⟦ ListN ⟧ℕ X) {{listNMonoid}} 110 | listNMonoidOK {X} = record 111 | { absorbL = λ _ → refl 112 | ; absorbR = vv ++<> 113 | ; assoc = vv <+> 114 | } where 115 | ++<> : ∀ n (xs : Vec X n) → ((n , xs) • ε) ≃ (n , xs) 116 | ++<> zero ⟨⟩ = refl 117 | ++<> (suc n) (x , xs) = cong (vv λ m ys → suc m , (x , ys)) (++<> n xs) 118 | 119 | <+> : ∀ n (t : Vec X (size ListN n)) 120 | (y z : ⟦ ListN ⟧ℕ X) → 121 | ((n , t) • y) • z ≃ (n , t) • (y • z) 122 | <+> zero ⟨⟩ _ _ = refl 123 | <+> (suc n) (x , xs) (i , ys) (j , zs) = subst (<+> n xs (i , ys) (j , zs)) 124 | (vv λ m ws → _≃_ {_}{⟦ ListN ⟧ℕ X} 125 | (suc ((n + i) + j) , (x , ((xs ++ ys) ++ zs))) (suc m , (x , ws))) 126 | refl 127 | 128 | -- Normal functors are traversable because Vectors are traversable. 129 | instance 130 | normalTraversable : (F : Normal) → Traversable ⟦ F ⟧ℕ 131 | normalTraversable F = record 132 | { traverse = λ { {{aG}} f (sh , xs) → pure {{aG}} (_,_ sh) ⍟ traverse f xs } } 133 | 134 | 135 | -- "We have already seen that the identity functor VecN 1 is Normal, but can we 136 | -- define composition?" 137 | -- 138 | -- Yeah. Yeah we can. 139 | _◦N_ : Normal → Normal → Normal 140 | F ◦N (ShG / szG) = ⟦ F ⟧ℕ ShG / crush {{normalTraversable F}} szG 141 | 142 | -- The size of a traversable thing can be obtained using the Monoidness of ℕ. 143 | sizeT : ∀ {F}{{TF : Traversable F}}{X} → F X → ℕ 144 | sizeT {F} = crush {F = F} (λ _ → 1) 145 | 146 | -- And hey, if you can get its size that way, you can make a Normal. 147 | normalT : ∀ F {{TF : Traversable F}} → Normal 148 | normalT F = F One / (sizeT {F = F}) 149 | 150 | -- 151 | shapeT : ∀ {F}{{TF : Traversable F}}{X} → F X → F One 152 | shapeT {F}{{TF}} = traverse (λ _ → <>) 153 | 154 | one : ∀ {X} → X → ⟦ ListN ⟧ℕ X 155 | one x = 1 , (x , ⟨⟩) 156 | 157 | contentsT : ∀ {F}{{TF : Traversable F}}{X} → F X → ⟦ ListN ⟧ℕ X 158 | contentsT {F} {{TF}} {X} = crush {{TF}}{{listNMonoid}} one 159 | 160 | _⟶N_ : Normal → Normal → Set 161 | F ⟶N G = (s : Shape F) → ⟦ G ⟧ℕ (Fin (size F s)) 162 | 163 | nMorph : ∀ {F G} → F ⟶N G → ∀ {X} → ⟦ F ⟧ℕ X → ⟦ G ⟧ℕ X 164 | nMorph f (s , xs) with f s 165 | nMorph f (x , xs) | (s' , is) = s' , vmap (proj xs) is 166 | 167 | morphN : ∀ {F G} → (∀ {X} → ⟦ F ⟧ℕ X → ⟦ G ⟧ℕ X) → F ⟶N G 168 | morphN f s = f (s , tabulate id) 169 | 170 | {- 171 | toNormal : ∀ {F}{{TF : Traversable F}} → TraversableOKP F → ∀ {X} → F X → ⟦ normalT F ⟧ℕ X 172 | toNormal tokf fx = shapeT fx , subst (lengthContentsSizeShape tokf fx) (Vec _) (snd (contentsT fx)) 173 | where 174 | lengthContentsSizeShape : ∀ {F} {{TF : Traversable F}} → TraversableOKP F → ∀ {X} (fx : F X) → fst (contentsT fx) ≃ sizeT (shapeT fx) 175 | lengthContentsSizeShape tokF fx = 176 | fst (contentsT fx) 177 | ⟨ TraversableOKP.lawPHom tokF {{monoidApplicative {{_}}}} {{monoidApplicative {{_}}}} fst one (monoidApplicativeHom {{_}}{{_}} fst {{_}}) fx ]= 178 | sizeT fx 179 | ⟨ TraversableOKP.lawPCo tokF {{monoidApplicative {{_}}}} {{applicativeId}} (λ _ → 1) (λ _ → <>) fx ]= 180 | sizeT {{_}} (shapeT fx) 181 | ∎ 182 | -} 183 | 184 | _⊗_ : Normal → Normal → Normal 185 | (ShF / szF) ⊗ (ShG / szG) = (ShF × ShG) / λ s → szF (fst s) * szG (snd s) 186 | 187 | swap : (F G : Normal) → (F ⊗ G) ⟶N (G ⊗ F) 188 | swap F G (ShF , ShG) rewrite *-comm (size F ShF) (size G ShG) = (ShG , ShF) , allFin _ 189 | 190 | drop : (F G : Normal) → (F ⊗ G) ⟶N (F ◦N G) 191 | drop F G (ShF , ShG) = (ShF , pure ShG) , subst (lemma (size F ShF)) (Vec _) (tabulate id) where 192 | lemma : ∀ n → (n * size G ShG) ≃ (crush (size G) (vec {n} ShG)) 193 | lemma zero = refl 194 | lemma (suc n) rewrite lemma n = refl 195 | 196 | fstHom : ∀ {X} → MonoidHom {⟦ ListN ⟧ℕ X}{{listNMonoid}}{ℕ}{{sumMonoid}} fst 197 | fstHom = record 198 | { respε = refl 199 | ; resp• = λ _ _ → refl 200 | } 201 | 202 | Batch : Set → Set → Set 203 | Batch X Y = Σ ℕ λ n → Vec X n → Y 204 | 205 | batchApplicative : {X : Set} → Applicative (Batch X) 206 | batchApplicative {X} = record 207 | { pure = λ y → (0 , λ _ → y) 208 | ; _⍟_ = app 209 | } where 210 | app : {S T : Set} → Batch X (S → T) → Batch X S → Batch X T 211 | app (n , f) (m , g) = (n + m) , (λ xs → let ss = chop n xs in f (fst ss) (g (snd ss))) 212 | 213 | batcher : ∀ {F} {{TF : Traversable F}} → F One → ∀ {X} → Batch X (F X) 214 | batcher {F} {{TF}} sF {X} = traverse {F} {{TF}} {Batch X} {{batchApplicative}} (λ _ → 1 , λ { (x , _) → x }) sF 215 | 216 | {- 217 | coherence : ∀ {F} {{TF : Traversable F}} {X} → TraversableOKP F 218 | → (sF : F One) → 219 | fst (batcher {F} {{TF}} sF {X}) ≃ 220 | traverse {F} {{TF}} {λ _ → ℕ} {One} {One} {{monoidApplicative {{sumMonoid}}}} (λ _ → 1) sF 221 | coherence {F} {{TF}} {X} tokF u = {! !} 222 | 223 | fst (traverse TF (λ _ → 1 , unpack) u) 224 | ⟨ TraversableOKP.lawPHom {F} {TF} tokF {Batch X} {{ABatch}} {λ _ → ℕ} {{monoidApplicative {{sumMonoid}}}} fst (λ _ → 1 , unpack) 225 | (record { respPure = λ {X₁} x → refl 226 | ; resp-⍟ = λ f s → refl 227 | }) u ]= 228 | (traverse TF {λ _ → ℕ} {One {lzero}} {X} 229 | {{record { pure = λ {_} _ → 0; _<*>_ = λ {_} {_} → _+_ }}} 230 | (λ a → 1) u) 231 | =[ TraversableOKP.lawPCo {F} {TF} tokF 232 | {G = id} {{AG = applicativeId}} 233 | {H = λ _ → ℕ} {{AH = monoidApplicative}} 234 | (λ _ → <>) (λ _ → 1) u ⟩ 235 | (traverse TF {λ _ → ℕ} {One {lzero}} {One} 236 | {{record { pure = λ {_} _ → 0; _<*>_ = λ {_} {_} → _+_ }}} 237 | (λ a → 1) u) 238 | ∎ 239 | 240 | fromNormal : ∀ {F}{{TF : Traversable F}} → TraversableOKP F → 241 | ∀ {X} → ⟦ normalT F ⟧ℕ X → F X 242 | fromNormal {F} {{TF}} tokf {X} (sF , cF) with (coherence {F} {{TF}} {X} tokf sF) 243 | fromNormal {F} {{TF}} tokf {X} (sF , cF) | q with batcher {F} {{TF}} sF {X} 244 | fromNormal {F} {{TF}} tokf {X} (sF , cF) | q | n , csF = csF (subst (sym q) (λ u → Vec X u) cF) 245 | -} 246 | 247 | data Tree (N : Normal) : Set where 248 | ⟨_⟩ : ⟦ N ⟧ℕ (Tree N) → Tree N 249 | 250 | NatT : Normal 251 | NatT = Two / 0 ⟨?⟩ 1 252 | 253 | zeroT : Tree NatT 254 | zeroT = ⟨ tt , ⟨⟩ ⟩ 255 | 256 | sucT : Tree NatT → Tree NatT 257 | sucT n = ⟨ ff , (n , ⟨⟩) ⟩ 258 | 259 | NatInd : ∀ {l} (P : Tree NatT → Set l) → 260 | P zeroT → 261 | ((n : Tree NatT) → P n → P (sucT n)) → 262 | (n : Tree NatT) → P n 263 | NatInd P z s ⟨ tt , ⟨⟩ ⟩ = z 264 | NatInd P z s ⟨ ff , (n , ⟨⟩) ⟩ = s n (NatInd P z s n) 265 | 266 | All : ∀ {l X}(P : X → Set l) {n} → Vec X n → Set l 267 | All P ⟨⟩ = One 268 | All P (x , xs) = P x × All P xs 269 | 270 | induction : ∀ (N : Normal) {l} (P : Tree N → Set l) → 271 | ((s : Shape N)(ts : Vec (Tree N) (size N s)) → All P ts → P  ⟨ s , ts ⟩) → 272 | (t : Tree N) → P t 273 | induction N P p ⟨ s , ts ⟩ = p s ts (hyps ts) where 274 | hyps : ∀ {n} (ts : Vec (Tree N) n) → All P ts 275 | hyps ⟨⟩ = <> 276 | hyps (t , ts) = induction N P p t , hyps ts 277 | 278 | eq? : (N : Normal)(sheq? : (s s` : Shape N) → Dec (s ≃ s`)) → 279 | (t t` : Tree N) → Dec (t ≃ t`) 280 | eq? N sheq? ⟨ vt , vts ⟩ ⟨ vt' , vts' ⟩ with sheq? vt vt' 281 | eq? N sheq? ⟨ vt , vts ⟩ ⟨ vt' , vts' ⟩ | tt , refl = {! !} 282 | eq? N sheq? ⟨ vt , vts ⟩ ⟨ vt' , vts' ⟩ | ff , falsum = ff , falsum ∘ lemma where 283 | lemma : ∀ {s s' : Shape N} → 284 | {vts : Vec (Tree N) (size N s)}{vts' : Vec (Tree N) (size N s')} → 285 | ⟨ s , vts ⟩ ≃ ⟨ s' , vts' ⟩ → s ≃ s' 286 | lemma refl = refl 287 | -------------------------------------------------------------------------------- /Chapter2/LambdaCalculus.agda: -------------------------------------------------------------------------------- 1 | module Chapter2.LambdaCalculus where 2 | 3 | open import Meta.Basics 4 | 5 | -- Start with base types closed under function spaces. 6 | data ⋆ : Set where 7 | ι : ⋆ 8 | _▹_ : ⋆ → ⋆ → ⋆ 9 | infixl 5 _▹_ 10 | 11 | -- Add a dash of context 12 | data Cx (X : Set) : Set where 13 | ε : Cx X 14 | _,_ : Cx X → X → Cx X 15 | infixl 4 _,_ 16 | 17 | -- And a skosh of de Bruijn indices. 18 | data _∈_ (τ : ⋆) : Cx ⋆ → Set where 19 | zero : ∀ {Γ} → τ ∈ Γ , τ 20 | suc : ∀ {Γ σ} → τ ∈ Γ → τ ∈ Γ , σ 21 | infix 3 _∈_ 22 | 23 | -- Mix well with syntax-directed typing judgements and serve immediately. 24 | data _⊢_ (Γ : Cx ⋆) : ⋆ → Set where 25 | var : ∀ {τ} 26 | → τ ∈ Γ 27 | ------- 28 | → Γ ⊢ τ 29 | lam : ∀ {σ τ} 30 | → Γ , σ ⊢ τ 31 | ------------ 32 | → Γ ⊢ σ ▹ τ 33 | app : ∀ {σ τ} 34 | → Γ ⊢ σ ▹ τ → Γ ⊢ σ 35 | ------------------- 36 | → Γ ⊢ τ 37 | infix 2 _⊢_ 38 | 39 | -- Decode into Agda-isms. This looks a hell of a lot like Augustsson and 40 | -- Carlsson's first go at a decoder in "An exercise in dependent types: a well 41 | -- typed interpreter." 42 | ⟦_⟧⋆ : ⋆ → Set 43 | ⟦ ι ⟧⋆ = ℕ -- "by way of being nontrivial" 44 | ⟦ σ ▹ τ ⟧⋆ = ⟦ σ ⟧⋆ → ⟦ τ ⟧⋆ 45 | 46 | -- Decode a context into an environment with a projection function. 47 | ⟦_⟧Cx : Cx ⋆ → (⋆ → Set) → Set 48 | ⟦ ε ⟧Cx V = One 49 | ⟦ (Γ , σ) ⟧Cx V = ⟦ Γ ⟧Cx V × V σ 50 | 51 | -- Decode and project a term in context. 52 | ⟦_⟧∈ : ∀ {Γ τ V} → τ ∈ Γ → ⟦ Γ ⟧Cx V → V τ 53 | ⟦ zero ⟧∈ (γ , t) = t 54 | ⟦ suc i ⟧∈ (γ , s) = ⟦ i ⟧∈ γ 55 | 56 | -- Finally, decode terms. 57 | ⟦_⟧⊢ : ∀ {Γ τ} → Γ ⊢ τ → ⟦ Γ ⟧Cx ⟦_⟧⋆ → ⟦ τ ⟧⋆ 58 | ⟦ var i ⟧⊢ γ = ⟦ i ⟧∈ γ 59 | ⟦ lam t ⟧⊢ γ = λ s → ⟦ t ⟧⊢ (γ , s) 60 | ⟦ app f s ⟧⊢ γ = ⟦ f ⟧⊢ γ (⟦ s ⟧⊢ γ) 61 | 62 | -- Ren: Simultaneous Renamings 63 | -- Sub: Substitutions 64 | Ren Sub : Cx ⋆ → Cx ⋆ → Set 65 | Ren Γ Δ = ∀ {τ} → τ ∈ Γ → τ ∈ Δ 66 | Sub Γ Δ = ∀ {τ} → τ ∈ Γ → Δ ⊢ τ 67 | 68 | -- Extend a context with a list of variables. 69 | _<><_ : ∀ { X } → Cx X → List X → Cx X 70 | xz <>< ⟨⟩ = xz 71 | xz <>< (x , xs) = xz , x <>< xs 72 | infixl 3 _<><_ 73 | 74 | -- Shiftable Simultaneous Substitutions extend both contexts with the new list 75 | -- of variables, then executes a Simultaneous substitution of the first into the 76 | -- second. 77 | Shub : Cx ⋆ → Cx ⋆ → Set 78 | Shub Γ Δ = ∀ Ξ → Sub (Γ <>< Ξ) (Δ <>< Ξ) 79 | 80 | -- With this kit, pushing substitions under binders is trivial 81 | _//_ : ∀ {Γ Δ}(θ : Shub Γ Δ) {τ} → Γ ⊢ τ → Δ ⊢ τ 82 | θ // var i = θ ⟨⟩ i -- Sub var is a context extension 83 | θ // lam t = lam ((θ ∘ _,_ _) // t) -- Sub lam is lam of sub applied in the binding. 84 | θ // app f s = app (θ // f) (θ // s) -- Sub app is sub down the binder and body. 85 | 86 | -- Weakens both contexts in a renaming with σ. 87 | wkr : ∀ {Γ Δ σ} → Ren Γ Δ → Ren (Γ , σ) (Δ , σ) 88 | wkr r zero = zero -- Weakening a renaming in the empty context does nothing. 89 | wkr r (suc n) = suc (r n) -- Else iteratively weaken the renaming with the variable. 90 | 91 | -- So a renaming can be made into a shiftable substitution simply by iteratively 92 | -- weakening the renaming. 93 | ren : ∀ {Γ Δ} → Ren Γ Δ → Shub Γ Δ 94 | ren r ⟨⟩ = var ∘ r -- Shub for the empty renaming is just the names we already know. 95 | ren r (_ , Ξ) = ren (wkr r) Ξ -- Else Shub by weakening the context. 96 | 97 | -- So a substitution can be made into a shiftable substitution simply by 98 | -- iteratively weakening the substitution. 99 | sub : ∀ {Γ Δ} → Sub Γ Δ → Shub Γ Δ 100 | sub s ⟨⟩ = s 101 | sub s (_ , Ξ) = sub (wks s) Ξ 102 | where 103 | -- Weaks both contexts in a substitution with σ. 104 | wks : ∀ {Γ Δ σ} → Sub Γ Δ → Sub (Γ , σ) (Δ , σ) 105 | wks s zero = var zero -- Weakening a substitution in the empty context does nothing. 106 | wks s (suc n) = ren suc // s n -- Else iteratively weaken the substitution. 107 | 108 | -- "A renaming that shifts past any context extension." 109 | weak : ∀ {Γ} Ξ → Ren Γ (Γ <>< Ξ) 110 | weak ⟨⟩ i = i -- Shift once in the empty context. 111 | weak (_ , Ξ) i = weak Ξ (suc i) -- Shift 1 + the context 112 | 113 | lambda' : ∀ {Γ σ τ} → ((∀ {Ξ} → Γ , σ <>< Ξ ⊢ σ) → Γ , σ ⊢ τ) → Γ ⊢ σ ▹ τ 114 | lambda' f = lam (f λ {Ξ} → var (weak Ξ zero)) 115 | 116 | -- "... Constructor-based unification is insufficient to solve for the prefix of 117 | -- a context, given a common suffix. 118 | -- 119 | -- By contrast, solving for a suffix is easy when the prefix is just a value: it 120 | -- requires only the stripping off of matching constructors. So, we can cajole 121 | -- Agda into solving the problem by working with its reversal, via the ‘chips’ 122 | -- operator:" 123 | _<>>_ : ∀ {X} → Cx X → List X → List X 124 | ε <>> ys = ys 125 | (xz , x) <>> ys = xz <>> (x , ys) 126 | 127 | open import Agda.Primitive 128 | 129 | lambda : ∀ {Γ σ τ} → ((∀ {Δ Ξ} {{_ : Δ <>> ⟨⟩ ≃ Γ <>> (σ , Ξ)}} → Δ ⊢ σ) → Γ , σ ⊢ τ) → Γ ⊢ σ ▹ τ 130 | lambda {Γ} f = lam ((f λ {Δ Ξ}{{q}} → subst (lem Δ Γ (_ , Ξ) q) (λ Γ → Γ ⊢ _) (var (weak Ξ zero)))) 131 | where 132 | {- This is Conor's. He wasn't kidding about the ugly. -} 133 | sucI : (a b : ℕ) → (_≃_ {lzero}{ℕ} (suc a) (suc b)) → a ≃ b 134 | sucI .b b refl = refl 135 | 136 | grr : (x y : ℕ) → suc x + y ≃ x + suc y 137 | grr zero y = refl 138 | grr (suc x) y rewrite grr x y = refl 139 | 140 | _+a_ : ℕ → ℕ → ℕ 141 | zero +a y = y 142 | suc x +a y = x +a suc y 143 | 144 | noc' : (x y : ℕ) → suc (x + y) ≃ y → {A : Set} → A 145 | noc' x zero () 146 | noc' x (suc y) q = noc' x y 147 | (suc (x + y) =[ grr x y ⟩ x + suc y =[ sucI _ _ q ⟩ y ∎) 148 | 149 | noc : (x k y : ℕ) → x +a (suc k + y) ≃ y → {A : Set} → A 150 | noc zero k y q = noc' k y q 151 | noc (suc x) k y q = noc x (suc k) y q 152 | 153 | len : ∀ {X} → Cx X → ℕ 154 | len ε = zero 155 | len (xz , x) = suc (len xz) 156 | 157 | lenlem : ∀ {X}(xz : Cx X)(xs : List X) → length (xz <>> xs) ≃ len xz +a length xs 158 | lenlem ε xs = refl 159 | lenlem (xz , x) xs = lenlem xz (x , xs) 160 | 161 | lem0 : ∀ {X}(xz yz : Cx X)(xs ys : List X) → length xs ≃ length ys → xz <>> xs ≃ yz <>> ys → (xz ≃ yz) × (xs ≃ ys) 162 | lem0 ε ε xs ys q q' = refl , q' 163 | lem0 ε (yz , x) .(yz <>> (x , ys)) ys q refl rewrite lenlem yz (x , ys) = noc (len yz) zero (length ys) q 164 | lem0 (xz , x) ε xs .(xz <>> (x , xs)) q refl rewrite lenlem xz (x , xs) = noc (len xz) zero (length xs) (_ ⟨ q ]= _ ∎) 165 | lem0 (xz , x) (yz , y) xs ys q q' with lem0 xz yz (x , xs) (y , ys) (cong suc q) q' 166 | lem0 (.yz , .y) (yz , y) .ys ys q q' | refl , refl = refl , refl 167 | 168 | lem : ∀ {X} (Δ Γ : Cx X) Ξ → Δ <>> ⟨⟩ ≃ Γ <>> Ξ → Γ <>< Ξ ≃ Δ 169 | lem Δ Γ ⟨⟩ q = Γ 170 | ⟨ fst (lem0 Δ Γ ⟨⟩ ⟨⟩ refl q) ]= Δ 171 | ∎ 172 | lem Δ Γ (x , Ξ) q = lem Δ (Γ , x) Ξ q 173 | 174 | {- 175 | myTest : ε ⊢ ι ▹ ι 176 | myTest = lambda λ x → x 177 | 178 | myTest2 : ε ⊢ (ι ▹ ι) ▹ (ι ▹ ι) 179 | myTest2 = lambda λ f → lambda λ x → app f (app f x) 180 | -} 181 | 182 | -- "The right-nested spine representation of β-normal η-long forms. 183 | -- In other words, Γ ⊨ τ is a normal form in τ and Γ ⊨* τ is a spine for τ 184 | -- that'll give you a term ι if you ask nicely. 185 | mutual 186 | -- Γ ⊨ τ is the type of normal forms in τ 187 | data _⊨_ (Γ : Cx ⋆) : ⋆ → Set where 188 | lam : ∀ {σ τ} → Γ , σ ⊨ τ → Γ ⊨ σ ▹ τ 189 | _$_ : ∀ {τ} → τ ∈ Γ → Γ ⊨* τ → Γ ⊨ ι 190 | 191 | -- Γ ⊨* τ is the type of spines for τ delivering ι. 192 | data _⊨*_ (Γ : Cx ⋆) : ⋆ → Set where 193 | -- Just deliver ι 194 | ⟨⟩ : Γ ⊨* ι 195 | -- ι comes with a sequence of destructors applied to it that have to 196 | -- be normalized away. 197 | _,_ : ∀ {σ τ} → Γ ⊨ σ → Γ ⊨* τ → Γ ⊨* σ ▹ τ 198 | infix 3 _⊨_ _⊨*_ 199 | infix 3 _$_ 200 | 201 | -- Removes the term x from the context. 202 | _-×_ : ∀ (Γ : Cx ⋆) {τ} (x : τ ∈ Γ) → Cx ⋆ 203 | ε -× () 204 | (Γ , _) -× zero = Γ 205 | (Γ , sg) -× suc x = (Γ -× x) , sg 206 | infixl 4 _-×_ 207 | 208 | -- Embeds a smaller context into a larger one through a 209 | _≠_ : ∀ {Γ σ} (x : σ ∈ Γ) → Ren (Γ -× x) Γ 210 | zero ≠ y = suc y 211 | suc x ≠ zero = zero 212 | suc x ≠ suc y = suc (x ≠ y) 213 | 214 | {- 215 | ⟨_|⟶_⟩_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ -× x ⊨ σ → Γ ⊨ τ → Γ -× x ⊨ τ 216 | ⟨ x |⟶ s ⟩ lam t = lam (⟨ suc x |⟶ {! !} ⟩ t) 217 | ⟨ x |⟶ s ⟩ (x₁ $ x₂) = {! !} 218 | -} 219 | 220 | -- Boolean equality is insufficient to identify whether a term is a suitable 221 | -- spine for the expression - we need to know its representation in Γ -× x. 222 | data Veq? {Γ σ}(x : σ ∈ Γ) : ∀ {τ} → τ ∈ Γ → Set where 223 | same : Veq? x x 224 | diff : ∀ {τ}(y : τ ∈ Γ -× x) → Veq? x (x ≠ y) 225 | 226 | -- Show that every |y| is discriminable with respect to a given |x|. 227 | 228 | veq? : ∀ {Γ σ τ}(x : σ ∈ Γ)(y : τ ∈ Γ) → Veq? x y 229 | veq? zero zero = same 230 | veq? zero (suc y) = diff y 231 | veq? (suc x) zero = diff zero 232 | veq? (suc x) (suc y) with veq? x y 233 | veq? (suc x) (suc .x) | same = same 234 | veq? (suc x) (suc .(x ≠ y)) | diff y = diff (suc y) 235 | 236 | -- Show how to propagate a renaming through a normal form. 237 | mutual 238 | -- Renaming for normal forms. 239 | renNm : ∀ {Γ Δ τ} → Ren Γ Δ → Γ ⊨ τ → Δ ⊨ τ 240 | -- Push the renaming under the binder and push the renaming through the body 241 | -- of the binder by iteratively weakening the context at each level with the 242 | -- renaming. 243 | renNm ρ (lam n) = lam (renNm (wkr ρ) n) 244 | -- Rename the bound variable and push the renaming through the spine of the 245 | -- right side of the application. 246 | renNm ρ (f $ x) = ρ f $ (renSp ρ x) 247 | 248 | -- Renaming with spines. Maps a normal form renaming over each destructor. 249 | renSp : ∀ {Γ Δ τ} → Ren Γ Δ → Γ ⊨* τ → Δ ⊨* τ 250 | renSp ρ ⟨⟩ = ⟨⟩ 251 | renSp ρ (x , ss) = renNm ρ x , renSp ρ ss 252 | 253 | -- Implement hereditary substitution for normal forms and spines, defined 254 | -- mutually with application of a normal form to a spine, performing β-reduction. 255 | mutual 256 | -- Substitution for normal forms. 257 | ⟨_↦_⟩_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ -× x ⊨ σ → Γ ⊨ τ → Γ -× x ⊨ τ 258 | -- Push the substitution under the binder and recurse. 259 | ⟨ x ↦ s ⟩ lam t = lam (⟨ suc x ↦ renNm (_≠_ zero) s ⟩ t) 260 | -- Are you ready for a substitution? 261 | ⟨ x ↦ s ⟩ x₁ $ x₂ with veq? x x₁ 262 | -- Yep! Sub in s in the binder and recurse into the spine to keep renaming. 263 | ⟨ x ↦ s ⟩ .x $ xs | same = s $$ (⟨ x ↦ s ⟩* xs) 264 | -- Nope! Don't sub, but recurse into the spine to try again. 265 | ⟨ x ↦ s ⟩ .(x ≠ y) $ xs | diff y = y $ (⟨ x ↦ s ⟩* xs) 266 | 267 | -- Substitution for spines. As before, map normal-form substitution down the 268 | -- spine. 269 | ⟨_↦_⟩*_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ -× x ⊨ σ → Γ ⊨* τ → Γ -× x ⊨* τ 270 | ⟨ x ↦ s ⟩* ⟨⟩ = ⟨⟩ 271 | ⟨ x ↦ s ⟩* (t , ts) = (⟨ x ↦ s ⟩ t) , (⟨ x ↦ s ⟩* ts) 272 | 273 | -- Evaluation of the destructors in a spine 274 | _$$_ : ∀ {Γ τ} → Γ ⊨ τ → Γ ⊨* τ → Γ ⊨ ι 275 | -- Yield ι 276 | f $$ ⟨⟩ = f 277 | -- Substitute the argument into the body of the function. 278 | lam f $$ (s , ss) = (⟨ zero ↦ s ⟩ f) $$ ss 279 | infix 3 _$$_ 280 | infix 2 ⟨_↦_⟩_ 281 | 282 | -- Delivers a variable x in η-long form. 283 | -- Naturally, this is Conor's. I can't really think of a better way to do this. 284 | η : ∀ {Γ σ}(x : σ ∈ Γ) τ → (∀ {Δ} → Ren Γ Δ → Δ ⊨* τ → Δ ⊨* σ) → Γ ⊨ τ 285 | η x ι f = x $ f id ⟨⟩ 286 | η x (σ ▹ τ) f = lam (η (suc x) τ λ ρ ss → f (ρ ∘ suc) ((η (ρ zero) σ (λ _ → id)) , ss)) 287 | 288 | -- The usual normalization, but η-long 289 | normalize : ∀ {Γ τ} → Γ ⊢ τ → Γ ⊨ τ 290 | normalize (var x) = η x _ λ _ → id -- Do nothing 291 | normalize (lam t) = lam (normalize t) -- Normalize under the binder 292 | normalize (app f s) with normalize f | normalize s -- Normalize both sides 293 | normalize (app f s) | lam t | s2 = ⟨ zero ↦ s2 ⟩ t -- Hereditary sub into 294 | -- nothing to apply, then normalize under the binder 295 | 296 | {- 297 | try₁ : ε ⊨ ((ι ▹ ι) ▹ (ι ▹ ι)) ▹ (ι ▹ ι) ▹ (ι ▹ ι) 298 | try₁ = normalize (lambda id) 299 | 300 | church₂ : ∀ {τ} → ε ⊢ (τ ▹ τ) ▹ τ ▹ τ 301 | church₂ = lambda λ f → lambda λ x → app f (app f x) 302 | 303 | try₂ : ε ⊨ (ι ▹ ι) ▹ (ι ▹ ι) 304 | try₂ = normalize (app (app church₂ church₂) church₂) 305 | -} 306 | 307 | -- Creative normalization involves: 308 | data Stop (Γ : Cx ⋆) (τ : ⋆) : Set where 309 | var : τ ∈ Γ → Stop Γ τ -- non-η-long forms that don't understand app 310 | _$_ : ∀ {σ} → Stop Γ (σ ▹ τ) → Γ ⊨ σ → Stop Γ τ -- and η-long forms that understand app 311 | 312 | -- Lift a renaming into some creative normalization. 313 | renSt : ∀ {Γ Δ τ} → Ren Γ Δ → Stop Γ τ → Stop Δ τ 314 | renSt r (var x) = var (r x) 315 | renSt r (u $ s) = renSt r u $ renNm r s 316 | 317 | -- Applying a Stop term to a spine yields a normal form. 318 | stopSp : ∀ {Γ τ} → Stop Γ τ → Γ ⊨* τ → Γ ⊨ ι 319 | stopSp (var x) ss = x $ ss 320 | stopSp (u $ x) ss = stopSp u (x , ss) 321 | 322 | -- Look, semantics in a context. Looks positively Kripkean. 323 | mutual 324 | -- Values either Go or Stop 325 | Val : Cx ⋆ → ⋆ → Set 326 | Val Γ τ = Go Γ τ ⊎ Stop Γ τ 327 | 328 | Go : Cx ⋆ → ⋆ → Set 329 | Go Γ ι = Zero 330 | Go Γ (σ ▹ τ) = ∀ {Δ} → Ren Γ Δ → Val Δ σ → Val Δ τ 331 | 332 | -- Show that values admit renaming. Extend renaming to environments storing 333 | -- values. Construct the identity environment mapping each variable to itself. 334 | renVal : ∀ {Γ Δ} τ → Ren Γ Δ → Val Γ τ → Val Δ τ 335 | renVal τ r (ff , u) = ff , renSt r u 336 | renVal ι r (tt , ()) 337 | renVal (σ ▹ τ) r (tt , f) = tt , (λ r' s → f (r' ∘ r) s) 338 | 339 | renVals : ∀ Θ {Γ Δ} → Ren Γ Δ → ⟦ Θ ⟧Cx (Val Γ) → ⟦ Θ ⟧Cx (Val Δ) 340 | renVals ε r _ = <> 341 | renVals (Θ , σ) r (θ , τ) = (renVals Θ r θ) , renVal σ r τ 342 | 343 | idEnv : ∀ Γ → ⟦ Γ ⟧Cx (Val Γ) 344 | idEnv ε = <> 345 | idEnv (Γ , σ) = renVals Γ suc (idEnv Γ) , (ff , var zero) 346 | 347 | mutual 348 | apply : ∀ {Γ σ τ} → Val Γ (σ ▹ τ) → Val Γ σ → Val Γ τ 349 | apply (tt , f) s = f id s 350 | apply (ff , u) s = ff , (u $ quo _ s) 351 | 352 | quo : ∀ {Γ} τ → Val Γ τ → Γ ⊨ τ 353 | quo ι (tt , ()) 354 | quo ι (ff , u) = stopSp u ⟨⟩ 355 | quo (σ ▹ τ) v = lam (quo τ (apply (renVal _ suc v) (ff , var zero))) 356 | 357 | -- Look, Kripe again. 358 | eval : ∀ {Γ Δ τ} → Γ ⊢ τ → ⟦ Γ ⟧Cx (Val Δ) → Val Δ τ 359 | eval (var x) γ = ⟦ x ⟧∈ γ 360 | eval {Γ}{_}{_} (lam t) γ = tt , λ r s → eval t (renVals Γ r γ , s) 361 | eval (app f s) γ = apply (eval f γ) (eval s γ) 362 | 363 | -- "With all the pieces in place, we get:" 364 | normByEval : ∀ {Γ τ} → Γ ⊢ τ → Γ ⊨ τ 365 | normByEval {Γ}{τ} t = quo τ (eval t (idEnv Γ)) 366 | 367 | {- 368 | zero : Γ ⊢ ι 369 | suc : Γ ⊢ ι → Γ ⊢ ι 370 | rec : ∀ {τ} → Γ ⊢ τ → Γ ⊢ (ι ▹ τ ▹ τ) → Γ ⊢ ι → Γ ⊢ τ 371 | -} 372 | -------------------------------------------------------------------------------- /Chapter3/Containers.agda: -------------------------------------------------------------------------------- 1 | module Chapter3.Containers where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | 6 | -- "Containers are the infinitary generalization of normal functors" by 7 | record Con : Set₁ where 8 | constructor _◃_ 9 | field 10 | Sh : Set -- a set of shapes 11 | Po : Sh → Set -- And a function taking shapes to "positions" 12 | ⟦_⟧◃ : Set → Set -- Containers can be interpreted in the dependent reader monad. 13 | ⟦_⟧◃ X = Σ Sh λ s → Po s → X 14 | open Con public 15 | infixr 1 _◃_ 16 | 17 | -- Conor says: "We may readily check that the polynomials are all containers." 18 | 19 | -- Types are containers with a unique shape. 20 | -- K◃ : X → (_ → 0) 21 | K◃ : Set → Con 22 | K◃ A = A ◃ λ _ → Zero 23 | 24 | -- The identity container 25 | -- I◃ : 1 → (_ → 1) 26 | I◃ : Con 27 | I◃ = One ◃ λ _ → One 28 | 29 | -- The sum of containers is a container of sums 30 | -- (X , XS) +◃ (Y : YS) = (X + Y) → (choose (XS + YS)) 31 | _+◃_ : Con → Con → Con 32 | (S ◃ P) +◃ (S₁ ◃ P₁) = (S ⊎ S₁) ◃ vv P ⟨?⟩ P₁ 33 | 34 | -- The product of containers is a container of products 35 | -- (X , XS) ×◃ (Y : YS) = (X × Y) → (selector → (XS + YS)) 36 | _×◃_ : Con → Con → Con 37 | (S ◃ P) ×◃ (S₁ ◃ P₁) = (S × S₁) ◃ vv λ s s₁ → P s ⊎ P₁ s₁ 38 | 39 | -- Containers are closed under dependent pairs. 40 | Σ◃ : (A : Set)(C : A → Con) → Con 41 | Σ◃ A C = (Σ A λ a → Sh (C a)) ◃ vv λ a s → Po (C a) s 42 | 43 | -- Containers are closed under dependent functions. 44 | Π◃ : (A : Set)(C : A → Con) → Con 45 | Π◃ A C = ((a : A) → Sh (C a)) ◃ (λ f → Σ A λ a → Po (C a) (f a)) 46 | 47 | -- Because of that, containers are closed under composition. 48 | _∘◃_ : Con → Con → Con 49 | (S ◃ P) ∘◃ C = Σ◃ S λ s → Π◃ (P s) λ p → C 50 | 51 | conEndoFunctor : {C : Con} → EndoFunctor ⟦ C ⟧◃ 52 | conEndoFunctor {S ◃ P} = record 53 | { map = λ f x → (fst x) , (λ z → f (snd x z)) } 54 | 55 | conEndoFunctorOKP : {C : Con} → EndoFunctorOKP ⟦ C ⟧◃ {{conEndoFunctor}} 56 | conEndoFunctorOKP {S ◃ P} = record 57 | { endoFunctorId = λ x → refl 58 | ; endoFunctorCo = λ f g x → refl 59 | } 60 | 61 | -- Also, containers model a natural transformation between the functors in the 62 | -- each container by hooking up the input shapes of one container to the output 63 | -- of the other then by hooking up the output positions to the input positions 64 | -- form which to fetch elements. 65 | _⟶◃_ : Con → Con → Set 66 | (S ◃ P) ⟶◃ (S₁ ◃ P₁) = Σ (S → S₁) λ f → (s : S) → P₁ (f s) → P s 67 | 68 | -- Action of the natural transformation. 69 | _/◃_ : ∀ {C C₁} → C ⟶◃ C₁ → ∀ {X} → ⟦ C ⟧◃ X → ⟦ C₁ ⟧◃ X 70 | (to , fro) /◃ (s , k) = to s , k ∘ fro s 71 | 72 | -- Check that you can represent any natural transformation between containers as 73 | -- a container morphism 74 | 75 | morph◃ : ∀ {C C₁} → (∀ {X} → ⟦ C ⟧◃ X → ⟦ C₁ ⟧◃ X) → C ⟶◃ C₁ 76 | morph◃ f = (λ x → fst (f (x , id))) , (λ s → snd (f (s , id))) 77 | 78 | -- The identity natural transformation for containers takes shapes back to 79 | -- shapes and output positions to input positions. 80 | id⟶◃ : ∀ {C} → C ⟶◃ C 81 | id⟶◃ = id , (λ _ x → x) 82 | 83 | -- Composition of container natural transformations. 84 | _∘⟶◃_ : ∀ {C D E} → (D ⟶◃ E) → (C ⟶◃ D) → (C ⟶◃ E) 85 | (toe , froe) ∘⟶◃ (tod , frod) = toe ∘ tod , (λ s z → frod s (froe (tod s) z)) 86 | -------------------------------------------------------------------------------- /Chapter3/W.agda: -------------------------------------------------------------------------------- 1 | module Chapter3.W where 2 | 3 | open import Meta.Basics 4 | open import Chapter3.Containers 5 | open import Meta.Control.Monad 6 | 7 | -- The least fixed point of a container is a Well-Founded Type. 8 | data W (C : Con) : Set where 9 | ⟨_⟩ : ⟦ C ⟧◃ (W C) → W C 10 | 11 | -- The well-founded type of natural numbers is the type with Two constructors 12 | -- given by Zero (⊥) or One (⊤) 13 | NatW : Set 14 | NatW = W (Two ◃ Zero ⟨?⟩ One) 15 | 16 | -- Now for the constructors: 17 | 18 | -- zero selects the Zero constructor that has a... er... "magic" shape. Such is 19 | -- life. 20 | zeroW : NatW 21 | zeroW = ⟨ tt , magic ⟩ 22 | 23 | -- suc(n) has a shape that selects the underlying n. 24 | sucW : NatW → NatW 25 | sucW n = ⟨ tt , (λ _ → n) ⟩ 26 | 27 | -- Primitive recursion over well-founded nats. 28 | precW : ∀ {l}{T : Set l} → T → (NatW → T → T) → NatW → T 29 | precW z _ ⟨ tt , _ ⟩ = z 30 | precW z s ⟨ ff , k ⟩ = s (k <>) (precW z s (k <>)) 31 | 32 | -- Addition is then defined using primitive recursion by standard application of 33 | -- the suc constructor. 34 | addW : NatW → NatW → NatW 35 | addW x y = precW y (λ _ → sucW) x 36 | 37 | -- 38 | _*◃_ : Con → Set → Set 39 | C *◃ X = W (K◃ X +◃ C) 40 | 41 | -- The free monad for a Container. 42 | freeMonad : (C : Con) → Monad (_*◃_ C) 43 | freeMonad C = record 44 | { return = λ x → ⟨ (tt , x) , magic ⟩ -- Like zeroW but carries an element 45 | ; _>>=_ = bind 46 | } where 47 | bind : ∀ {S T : Set} → C *◃ S → (S → C *◃ T) → C *◃ T 48 | bind ⟨ (tt , a) , _ ⟩ f = f a 49 | bind ⟨ (ff , s) , k ⟩ f = ⟨ (ff , s) , (λ x → bind (k x) f) ⟩ 50 | 51 | -- Free monad closure. Respects C * X ≅ ⟦ C*◃ ⟧◃ X 52 | _*◃ : Con → Con 53 | _*◃ C = C *◃ One ◃ Path where 54 | Path : C *◃ One → Set 55 | Path ⟨ (tt , _) , _ ⟩ = One 56 | Path ⟨ (ff , s) , k ⟩ = Σ (Po C s) λ p → Path (k p) 57 | 58 | -- Performs one command-response interaction. Look, it `bind`. 59 | call : ∀ {C} → (s : Sh C) → C *◃ Po C s 60 | call s = ⟨ (ff , s) , (λ x → ⟨ (tt , x) , magic ⟩) ⟩ 61 | 62 | -- "We can model the general recursive function space as the means to perform 63 | -- finite, on demand expansion of call trees." 64 | Π⊥ : (S : Set) (T : S → Set) → Set 65 | Π⊥ S T = (s : S) → (S ◃ T) *◃ (T s) 66 | 67 | -- Give the 'gasoline-drive' interpreter for this function space, delivering 68 | -- a result provided the call tree does not expand more times than a given 69 | -- number. 70 | gas : ∀ {S T} → ℕ → Π⊥ S T → (s : S) → T s ⊎ One 71 | gas zero f s = ff , <> 72 | gas {S}{T} (suc n) f s = combust (f s) where 73 | combust : ∀ {X} → (S ◃ T) *◃ X → X ⊎ One 74 | combust ⟨ (tt , s) , l ⟩ = tt , s 75 | combust ⟨ (ff , s) , l ⟩ with gas n f s 76 | combust ⟨ (ff , _) , l ⟩ | tt , t = combust (l t) 77 | combust ⟨ (ff , _) , l ⟩ | ff , _ = ff , <> 78 | 79 | -- "We have 80 | -- ⟦ S ◃ P ⟧◃ X = Σ S λ s → P s → X 81 | -- but we could translate the right-hand side into a more mathematical notation 82 | -- and observe that a container is something a bit like a power series 83 | -- ⟦ S ◃ P ⟧◃ X = Σ_(s : S) X^(P s) 84 | -- We might imagine computing a formal derivative of such a series, 'multiplying 85 | -- down each index, then subtracting one', but we are not merely counting data 86 | -- - they have individual existences. Let us define a kind of 'dependent 87 | -- decrement' subtracting a *particular* element from a type. 88 | -- 89 | -- That is, an element of X ─ x is some element for X which is known to be 90 | -- other than x." 91 | _─_ : (X : Set) (x : X) → Set 92 | X ─ x = Σ X λ x₁ → x₁ ≃ x → Zero 93 | 94 | -- The formal derivative of a container. 95 | -- 96 | -- The shape of a derivative is the pair of a shape with one position, which 97 | -- we call the 'hole', and the positions in the derivative are 'everywhere but 98 | -- the hole.' 99 | ∂ : Con → Con 100 | ∂ (S ◃ P) = Σ S P ◃ vv λ s p → P s ─ p 101 | 102 | -- A container morphism that witnesses the ability to fill the hole, provided 103 | -- equality on positions is decidable. 104 | plug : ∀ {C} → ((s : Sh C)(p p₁ : Po C s) → Dec (p ≃ p₁)) → (∂ C ×◃ I◃) ⟶◃ C 105 | plug {C} poeq? = (fst ∘ fst) , (λ { ((s , x) , _) x₁ → nplay s x x₁ }) where 106 | nplay : (s : Sh C)(x x₁ : Po C s) → (Po C s ─ x) ⊎ One 107 | nplay s x x₁ with poeq? s x₁ x 108 | nplay s x .x | tt , refl = ff , <> 109 | nplay s x x₁ | ff , pp = tt , x₁ , pp 110 | -------------------------------------------------------------------------------- /Chapter4/Apocrypha/Roman.agda: -------------------------------------------------------------------------------- 1 | module Chapter4.Apocrypha.Roman where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Functor.Container.Indexed 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container 7 | open import Meta.Data.Inductive.W 8 | 9 | -- A Roman Container is a Container decorated by functions that attach input 10 | -- indices to positions and an output index to the shape. 11 | -- 12 | -- "The reason Roman containers are so called is that they invoke equality and 13 | -- its mysterious capacity for transubstantiation." 14 | record Roman (I J : Set) : Set₁ where 15 | constructor SPqr 16 | field 17 | S : Set -- A set of shapes 18 | P : S → Set -- a function taking shapes to arities. 19 | q : S → J -- a function taking shapes to an index in the family. 20 | r : (s : S) → P s → I -- a function taking arities to indices. 21 | 22 | -- Lower a Roman Container back to a plain container by dropping the 23 | -- input/output distinction and the indexing description. 24 | Plain : Con 25 | Plain = S ◃ P 26 | 27 | -- 28 | ⟦_⟧R : (I → Set) → (J → Set) 29 | ⟦_⟧R X j = Σ (Σ S λ s → q s ≃ j) (vv λ s _ → (p : P s) → X (r s p)) 30 | Plain = Roman.Plain 31 | ⟦_⟧R = Roman.⟦_⟧R 32 | 33 | -- We can turn Roman Containers into indexed containers 34 | FromRoman : ∀ {I J} → Roman I J → I ▷ J 35 | FromRoman (SPqr S P q r) = (λ j → Σ S λ s → q s ≃ j) -- Cleverly encodes that the new Roman Container never actually dropped the old indexing structure. 36 | ◃ (λ j → P ∘ fst) -- Just drop the proof and grab arities. 37 | $ (λ f → r ∘ fst) -- As before, drop the proof and grab the indexing structure. 38 | -- Whose meanings match 'on the nose'. 39 | onTheNose : ∀ {I J} (C : Roman I J) → ⟦ C ⟧R ≃ ⟦ FromRoman C ⟧ᵢ 40 | onTheNose C = refl 41 | 42 | -- Similarly, we can convert indexed containers to Roman Containers. 43 | -- Exercise 4.14 (ToRoman): Show how to construct the Roman container isomorphic 44 | -- to a given indexed container and exhibit the isomorphism. 45 | -- 46 | -- To project into a Roman container, we re-arrange the signature as follows: 47 | -- S: A product giving us access to the family's index and the operation given at that index. 48 | -- P: Cleverly, this is precisely the data we need to project arities out of the container. 49 | -- q: Just get the family's index out of the product. 50 | -- r: Once again, we set the shapes up with the induction principle and knock 51 | -- them down with the container's indexing description. 52 | ToRoman : ∀ {I J} → I ▷ J → Roman I J 53 | ToRoman {I}{J} (Shlx ◃ Polx $ rilx) = SPqr (Σ J Shlx) (vv Polx) fst (vv rilx) 54 | 55 | toRoman : ∀ {I J} (C : I ▷ J) → 56 | ∀ {X j} → ⟦ C ⟧ᵢ X j → ⟦ ToRoman C ⟧R X j 57 | toRoman C (fst , snd) = ((_ , fst) , refl) , snd 58 | 59 | fromRoman : ∀ {I J} (C : I ▷ J) → 60 | ∀ {X j} → ⟦ ToRoman C ⟧R X j → ⟦ C ⟧ᵢ X j 61 | fromRoman C (((j , snd) , refl) , snd₂) = snd , snd₂ 62 | 63 | toAndFromRoman : ∀ {I J} (C : I ▷ J){X j} 64 | → (∀ xs → toRoman C {X}{j} (fromRoman C {X}{j} xs) ≃ xs) 65 | × (∀ xs → fromRoman C {X}{j} (toRoman C {X}{j} xs) ≃ xs) 66 | toAndFromRoman C = (λ { (((_ , _) , refl) , _) → refl }) , (λ xs → refl) 67 | 68 | -- The general purpose tree type for Roman containers looks a lot like the 69 | -- inductive families you find in Agda or the GADTs of Haskell. 70 | -- 71 | -- The RomanData type looks a lot like a W-type, albeit festooned with 72 | -- equations. 73 | data RomanData {I} (C : Roman I I) : I → Set where 74 | _,_ : (s : Roman.S C) → ((p : Roman.P C s) → RomanData C (Roman.r C s p)) → RomanData C (Roman.q C s) 75 | 76 | -- Exercise 4.15: Construct a function which takes plain W-type data for a 77 | -- Roman container and marks up each node with the index required of it, 78 | -- using Roman.r. 79 | 80 | ideology : ∀ {I} (C : Roman I I) → I → W (Plain C) → W (Plain C ×◃ K◃ I) 81 | ideology C i ⟨ s , k ⟩ = ⟨ ((s , i) , (λ { (tt , p) → ideology C (Roman.r C s p) (k p) ; (ff , ()) })) ⟩ 82 | 83 | -- Construct a function which takes plain W-type data for a Roman container and 84 | -- marks up each node with the index delivered by it, using Roman.q. 85 | 86 | phenomenology : ∀ {I} (C : Roman I I) → W (Plain C) → W (Plain C ×◃ K◃ I) 87 | phenomenology C ⟨ s , k ⟩ = ⟨ ((s , Roman.q C s) , (λ { (tt , snd) → phenomenology C (k snd) ; (ff , ()) })) ⟩ 88 | 89 | -- Take the W-type interpretation of a Roman container to be the plain data for 90 | -- which the required indices are delivered. 91 | 92 | RomanW : ∀ {I} → Roman I I → I → Set 93 | RomanW C i = Σ (W (Plain C)) λ t → phenomenology C t ≃ ideology C i t 94 | 95 | -- Now, check that you can extract RomanData from RomanW. 96 | 97 | fromRomanW : ∀ {I} (C : Roman I I) {i} → RomanW C i → RomanData C i 98 | fromRomanW C = vv lemma3 C where 99 | lemma1 : ∀ {I} (C : Roman I I) {s i i' k k'} → _≃_ {X = W (Plain C ×◃ K◃ I)} ⟨ (s , i) , k ⟩ ⟨ (s , i') , k' ⟩ → (i ≃ i') 100 | lemma1 C₁ refl = refl 101 | 102 | lemma2 : ∀ {I} (C : Roman I I) {s i i' k k'} → _≃_ {X = W (Plain C ×◃ K◃ I)} ⟨ (s , i) , k ⟩ ⟨ (s , i') , k' ⟩ → (k ≃ k') 103 | lemma2 C₁ refl = refl 104 | 105 | lemma3 : ∀ {I} (C : Roman I I) {i} → (t : W (Plain C)) → phenomenology C t ≃ ideology C i t → RomanData C i 106 | lemma3 C {i} ⟨ s , k ⟩ g rewrite (i ⟨ lemma1 C g ]= Roman.q C s ∎) = s , λ p → lemma3 C (k p) (cong (_$′ (tt , p)) (lemma2 C g)) 107 | 108 | -- To go the other way, it is easy to construct the plain tree, but to prove the 109 | -- constraint, you will need to establish equality of functions. Using 110 | 111 | postulate 112 | extensionality : ∀ {S : Set}{T : S → Set} (f g : (s : S) → T s) → ((s : S) → f s ≃ g s) → f ≃ g 113 | 114 | -- construct 115 | 116 | toRomanW : ∀ {I} (C : Roman I I) {i} → RomanData C i → RomanW C i 117 | toRomanW C (s , x) = {! !} 118 | 119 | -- 120 | -- Reflexive Transitive Closure 121 | -- 122 | 123 | -- Consider the reflexive transitive closure of a relation, also known as the 124 | -- ‘paths in a graph’. 125 | data _** {I : Set} (R : I × I → Set) : I × I → Set where 126 | ⟨⟩ : {i : I} → (R **) (i , i) 127 | _,_ : {i j k : I} → R (i , j) → (R **) (j , k) → (R **) (i , k) 128 | infix 1 _** 129 | 130 | -- You can construct the natural numbers as an instance. 131 | NAT : Set 132 | NAT = (Loop **) _ where 133 | Loop : One × One → Set 134 | Loop _ = One 135 | 136 | -- Exercise 4.17 (monadic operations) Implement these such that the 137 | -- monad laws hold. 138 | 139 | one : ∀ {I}{R : I × I → Set} → R -:> (R **) 140 | one (fst , snd) x = x , ⟨⟩ 141 | 142 | join** : ∀ {I}{R : I × I → Set} → ((R **) **) -:> (R **) 143 | join** _ ⟨⟩ = ⟨⟩ 144 | join** i (x , xs) = x ◅◅ join** _ xs where 145 | _◅◅_ : ∀ {I}{R : I × I → Set}{i j k : I} → (R **) (i , j) → (R **) (j , k) → (R **) (i , k) 146 | ⟨⟩ ◅◅ ys = ys 147 | (x , xs) ◅◅ ys = x , (xs ◅◅ ys) 148 | 149 | -- We have two ways to formulate a notion of ‘subset’ in type theory. We can 150 | -- define a subset of X as a predicate in 151 | -- X → Set 152 | -- giving a proof-relevant notion of evidence that a given X : X belongs, or we 153 | -- can pick out some elements of X as the image of a function 154 | -- Σ Set λ I → I → X 155 | -- so we have a family of X s indexed by some set. 156 | 157 | -- Are these notions the same? That turns out to be a subtle question. A lot 158 | -- turns on the size of X , so we had best be formal about it. 159 | 160 | -- In general, X is large. 161 | Pow : Set₁ → Set₁ 162 | Pow X = X → Set 163 | 164 | Fam : Set₁ → Set₁ 165 | Fam X = Σ Set λ I → I → X 166 | 167 | -- Exercise 4.18 (small Pow and Fam): Show that, given a suitable notion of 168 | -- propositional equality, Pow ∘ ⇑ and Fam ∘ ⇑ capture essentially the same 169 | -- notion of subset. 170 | 171 | p2f : (Pow ∘ ⇑) -:> (Fam ∘ ⇑) 172 | p2f X P = (Σ X (P ∘ ↑)) , (↑ ∘ fst) 173 | 174 | f2p : (Fam ∘ ⇑) -:> (Pow ∘ ⇑) 175 | f2p X (I , f) (↑ x) = Σ I λ i → ↓ (f i) ≃ x 176 | 177 | -- Exercise 4.19 (functoriality of Pow and Fam): Equip Pow with a contravariant 178 | -- functorial action and Fam with a covariant functorial action. 179 | 180 | _$p_ : ∀ {I J} → (J → I) → Pow I → Pow J 181 | f $p P = P ∘ f 182 | 183 | _$F_ : ∀ {I J} → (I → J) → Fam I → Fam J 184 | f $F (F , i) = F , (f ∘ i) 185 | 186 | -- "Fam Set is Martin-Löf’s notion of a universe, naming a bunch of sets by the 187 | -- elements of some indexing set. Meanwhile, the ‘representation type’ method of 188 | -- describing types concretely in Haskell is just using 189 | -- Pow Set in place of Fam Set. It is good to get used to recognizing when 190 | -- concepts are related just by exchanging Fam and Pow." 191 | 192 | -- "Modulo currying and λ-lifting of parameters, the distinction between 193 | -- Roman I J and our Hancock-style I ▹ J is just that the former represents 194 | -- indexed shapes by a Fam (so Roman.q reads off the shape) whilst the latter 195 | -- uses a Pow (so the shapes pertain to a given index). 196 | -- Both use Fams for positions. 197 | 198 | ROMAN : Set → Set → Set1 199 | ROMAN I J = Σ (Fam (⇑ J)) λ { (S , q) → S → Fam (⇑ I) } 200 | 201 | HANCOCK : Set → Set → Set1 202 | HANCOCK I J = Σ (Pow (⇑ J)) λ S → Σ J (S ∘ ↑) → Fam (⇑ I) 203 | 204 | -- A ‘Nottingham’ indexed container switches the positions to a Pow 205 | -- (see Altenkirch and Morris). 206 | 207 | NOTTINGHAM : Set → Set → Set₁ 208 | NOTTINGHAM I J = Σ (Pow (⇑ J)) λ S → Σ J (S ∘ ↑) → Pow (⇑ I) 209 | 210 | -- which amounts to a presentation of shapes and positions as predicates: 211 | -- NSh : J ! Set 212 | -- NPo : (j : J)!NShj !I !Set 213 | 214 | -- For HANCOCK and NOTTINGHAM, we can abstract the whole construction 215 | -- over J , obtaining: 216 | HANCOCK' : Set → Set → Set₁ 217 | HANCOCK' I J = J → Fam (Fam (⇑ I)) 218 | 219 | NOTTINGHAM' : Set → Set → Set₁ 220 | NOTTINGHAM' I J = J → Fam (Pow (⇑ I)) 221 | -------------------------------------------------------------------------------- /Chapter4/ITree.agda: -------------------------------------------------------------------------------- 1 | module Chapter4.ITree where 2 | 3 | open import Agda.Primitive 4 | 5 | open import Meta.Basics 6 | open import Meta.Language.LambdaCalculus 7 | 8 | open import Chapter4.IndexedContainers 9 | 10 | -- The universal inductive family or the fixpoint of an Indexed Container. 11 | -- 12 | -- To encode data type definitions, the non-recursive constructors are the 13 | -- shapes, the recursive constructors are the positions, and the indexing 14 | -- structure is the typing discipline. 15 | data ITree {J : Set} (C : J ▷ J) (j : J) : Set where 16 | ⟨_⟩ : ⟦ C ⟧ᵢ (ITree C) j → ITree C j 17 | 18 | -- Natural numbers are usually given as 19 | -- 20 | -- data 𝕟 : Set where 21 | -- zero : 𝕟 22 | -- suc : (n : 𝕟) → 𝕟 23 | -- 24 | -- Let's derive a signature: 25 | NatC : One ▷ One 26 | NatC = (λ _ → Two) -- We have two constructors, z, and s. 27 | ◃ (λ _ → Zero ⟨?⟩ One) -- At position 0 we have no arguments, at s we recur with the predecessor. 28 | $ _ -- A trivial indexing structure yields a trivial typing (⊤) 29 | 30 | -- We can inhabit this signature with a bit of magic. 31 | zeroC : ITree NatC <> 32 | zeroC = ⟨ tt , magic ⟩ 33 | 34 | sucC : ITree NatC <> → ITree NatC <> 35 | sucC n = ⟨ ff , (λ _ → n) ⟩ 36 | 37 | -- The indexing structure of a dependent vector is given below: 38 | VecC : Set → ℕ ▷ ℕ 39 | VecC X = VS ◃ VP $ Vr where -- depending on the length 40 | VS : ℕ → Set 41 | VS zero = One -- at sort 0, there is a single operation: nil 42 | VS (suc n) = X -- cons carries an element 43 | 44 | VP : (n : ℕ) → VS n → Set 45 | VP zero _ = Zero -- nil has no children 46 | VP (suc n) _ = One -- cons has one child 47 | 48 | Vr : (n : ℕ)(s : VS n)(p : VP n s) → ℕ 49 | Vr zero <> () -- nil has no children to index 50 | Vr (suc n) x <> = n -- the tail of a cons has the length one less 51 | 52 | -- Once again, we inhabit the signature with sorcery. 53 | vnil' : ∀ {X} → ITree (VecC X) 0 54 | vnil' = ⟨ <> , (λ ()) ⟩ 55 | 56 | vcons' : ∀ {X n} → X → ITree (VecC X) n → ITree (VecC X) (suc n) 57 | vcons' x xs = ⟨ (x , (λ _ → xs)) ⟩ 58 | 59 | -- 4.2: Define the simply typed λ-terms as Petersson-Synek trees. 60 | 61 | -- A little hack Conor uses to case on the term constructors 62 | IsArr : ⋆ → Set 63 | IsArr ι = Zero 64 | IsArr (_ ▹ _) = One 65 | 66 | -- The simply typed lambda calculus has 67 | -- Shapes: Given by an ambient context and a term [and our case for its constructor] 68 | -- Positions: We have 3 cases: 69 | -- ∘ Case 1 (an arrow to a thing) there are no open positions 70 | -- ∘ Case 2 (a thing and an arrow) there are potentially 2 holes to fill 71 | -- ∘ Case 3 (a thing and a thing) there is only one hole 72 | -- Items: We have 5 cases 73 | -- ∘ Case 1 (an arrow and a thing) there are no positions, so we can't construct anything 74 | -- ∘ Case 2 (a thing and proof of an arrow) construct an arrow in the context 75 | -- ∘ Case 3 (a thing and proof of a term) construct a term in the context 76 | -- ∘ Case 4 nonsense 77 | -- ∘ Case 5 (an arrow and two holes) Throw the terms into the context. 78 | STLC : (Cx ⋆ × ⋆) ▷ (Cx ⋆ × ⋆) 79 | STLC = (vv λ Γ T → (T ∈ Γ) ⊎ (⋆ ⊎ IsArr T)) ◃ (λ { (Γ , τ) (tt , _) → Zero 80 | ; (Γ , τ) (ff , (tt , l)) → Two 81 | ; (Γ , τ) (ff , (ff , l)) → One 82 | }) $ 83 | (λ { (Γ , τ) (tt , _) () 84 | ; (Γ , τ) (ff , (tt , σ)) tt → Γ , (σ ▹ τ) 85 | ; (Γ , τ) (ff , (tt , σ)) ff → Γ , σ 86 | ; (Γ , ι) (ff , (ff , ())) _ 87 | ; (Γ , (σ ▹ τ)) (ff , (ff , p)) _ → (Γ , σ) , τ 88 | }) 89 | 90 | -- The universal identity, satisfies 91 | -- ⟦ IdIx ⟧i X i ≅ X i 92 | -- Shapes: The only possible shape is ⊤ 93 | -- Positions: The only possible position is also ⊤ 94 | -- Items: Just send whatever comes our way back out again. 95 | IdIx : ∀ {I} → I ▷ I 96 | IdIx = (λ _ → One) ◃ (λ j _ → One) $ (λ j _ _ → j) 97 | 98 | -- The universal composition, satisfies 99 | -- ⟦ Colx C C′ ⟧i X k ≅ ⟦ C ⟧i (⟦ C′ ⟧i X) k 100 | -- 101 | -- From the law above we infer the following construction 102 | -- Shapes: Piped from I shapes to K shapes through J shapes along the 103 | -- second shape function 104 | -- Positions: Piped from I positions to K positions through J positions along the 105 | -- second position function. 106 | -- Items: Piped from I items to K items through J items along the second item 107 | -- function. 108 | Colx : ∀ {I J K} → J ▷ K → I ▷ J → I ▷ K 109 | Colx (S ◃ P $ r) (S₁ ◃ P₁ $ r₁) = (λ x → Σ (S x) λ s → (p : P x s) → S₁ (r x s p)) 110 | ◃ (λ x → vv λ s f → Σ (P x s) λ p → P₁ (r x s p) (f p)) 111 | $ (λ { j (s , f) (p , p₁) → r₁ (r j s p) (f p) p₁ }) 112 | 113 | -- A richer description of the class of indexed functors. 114 | -- It comes as no surprise that this looks a lot like λΠω 115 | data Desc {l} (I : Set l) : Set (lsuc l) where 116 | var : I → Desc I 117 | σ π : (A : Set l) (D : A → Desc I) → Desc I 118 | _×D_ : Desc I → Desc I → Desc I 119 | κ : Set l → Desc I 120 | infixr 4 _×D_ 121 | 122 | -- Interpret a description as Agda types with a little help from a function that 123 | -- maps the index to things Agda understands. 124 | ⟦_⟧D : ∀ {l}{I : Set l} → Desc I → (I → Set l) → Set l 125 | ⟦ (var i) ⟧D X = X i -- Just apply the index yielded by the term to the type former. 126 | ⟦ (σ A D) ⟧D X = Σ A λ a → ⟦ D a ⟧D X -- Lift σ terms to Σ terms 127 | ⟦ (π A D) ⟧D X = (a : A) → ⟦ D a ⟧D X -- Lift π terms to Π terms 128 | ⟦ (D ×D E) ⟧D X = ⟦ D ⟧D X × ⟦ E ⟧D X -- Interpretation of products is the product of interpretations. 129 | ⟦ (κ A) ⟧D X = A -- Extract the type from inside the description. 130 | 131 | -- Every indexed container has a description. 132 | ixConDesc : ∀ {I J} → I ▷ J → J → Desc I 133 | ixConDesc (S ◃ P $ r) j = σ (S j) λ s → π (P j s) λ p → var (r j s p) 134 | 135 | -- Likewise every description has an indexed container. 136 | -- 137 | -- Naturally, the indexed container is constructed pretty much indentically to 138 | -- the interpretation function above. 139 | descIxCon : ∀ {I J} → (J → Desc I) → I ▷ J 140 | descIxCon F = (DSh ∘ F) ◃ (DPo ∘ F) $ (Dri ∘ F) where 141 | DSh : {I : Set} → Desc I → Set 142 | DSh (var x) = One 143 | DSh (σ A D) = Σ A λ a → DSh (D a) 144 | DSh (π A D) = (a : A) → DSh (D a) 145 | DSh (D ×D D₁) = DSh D × DSh D₁ 146 | DSh (κ A) = A 147 | 148 | DPo : ∀ {I} (D : Desc I) → DSh D → Set 149 | DPo (var x) x₁ = One 150 | DPo (σ A D) (x , y) = DPo (D x) y 151 | DPo (π A D) f = Σ A λ a → DPo (D a) (f a) 152 | DPo (D ×D D₁) (x , y) = DPo D x ⊎ DPo D₁ y 153 | DPo (κ A) s = Zero 154 | 155 | Dri : ∀ {I}(D : Desc I)(s : DSh D) → DPo D s → I 156 | Dri (var x) s p = x 157 | Dri (σ A D) (x , y) p = Dri (D x) y p 158 | Dri (π A D) f (x , y) = Dri (D x) (f x) y 159 | Dri (D ×D D₁) (x , y) (tt , p) = Dri D x p 160 | Dri (D ×D D₁) (x , y) (ff , p) = Dri D₁ y p 161 | Dri (κ x) s () 162 | 163 | {- 164 | vecD : Set → ℕ → Desc ℕ 165 | vecD X n = 166 | σ Two ( κ (n ≃ 0) 167 | ⟨?⟩ σ ℕ λ k → κ X ×D var k ×D κ (n ≃ ℕ.suc k) 168 | ) 169 | -} 170 | 171 | -- "Descriptions are quite a lot like inductive family declarations." Only this 172 | -- time we have the full power of Agda at our disposal. 173 | vecD : Set → ℕ → Desc ℕ 174 | vecD X zero = κ One 175 | vecD X (suc n) = κ X ×D var n 176 | 177 | -- Datatypes from description. 178 | data Data {l}{J : Set l} (F : J → Desc J)(j : J) : Set l where 179 | ⟨_⟩ : ⟦ F j ⟧D (Data F) → Data F j 180 | 181 | -- "Let us once again construct vectors" 182 | vnil : ∀ {X} → Data (vecD X) 0 183 | vnil = ⟨ <> ⟩ 184 | 185 | vcons : ∀ {X n} → X → Data (vecD X) n → Data (vecD X) (suc n) 186 | vcons x xs = ⟨ x , xs ⟩ 187 | 188 | -- Construct a family of descriptions which describes a type like Desc. As Agda 189 | -- is not natively cumulative, you will need to shunt types up through the Set l 190 | -- hierarchy by hand, with this gadget. 191 | record ⇑ {l}(X : Set l) : Set (lsuc l) where 192 | constructor ↑ -- Shunt up a level 193 | field 194 | ↓ : X -- And shunt down a level 195 | open ⇑ public 196 | 197 | {- 198 | data Desc⋆ {l} : Set l where 199 | varD σD πD ×D κD : Desc⋆ 200 | 201 | DescD : ∀ {l}(I : Set l) → One {lsuc l} → Desc (One {lsuc l}) 202 | DescD {l} I _ = Σ Desc⋆ (λ 203 | { varD → ? -- κD (⇑ I) 204 | ; σD → Σ (Set l) (λ A → π (⇑ A) λ _ → var <>) 205 | ; πD → Σ (Set l) (λ A → π (⇑ A) λ _ → var <>) 206 | ; ×D → ? 207 | ; κD → κD (Set l) 208 | }) 209 | -} 210 | 211 | -- 212 | -- Predicate Transformers 213 | -- 214 | 215 | Everywhere : ∀ {I J} (C : I ▷ J)(X : I → Set) → Σ I X ▷ Σ J (⟦ C ⟧ᵢ X) 216 | Everywhere (S ◃ P $ r) X 217 | = (λ _ → One) 218 | ◃ (λ {(j , (s , k)) _ → P j s }) 219 | $ (λ { (j , (s , k)) _ p → r j s p , k p }) 220 | 221 | allTrivial : ∀ {I J} (C : I ▷ J)(X : I → Set) jc → ⟦ Everywhere C X ⟧ᵢ (λ _ → One) jc 222 | allTrivial C X _ = <> , λ _ → <> 223 | 224 | Somewhere : ∀ {I J}(C : I ▷ J)(X : I → Set) → Σ I X ▷ Σ J (⟦ C ⟧ᵢ X) 225 | Somewhere (S ◃ P $ r) X 226 | = (λ { (j , (s , k)) → P j s}) 227 | ◃ (λ { (j , (s , k)) _ → One}) 228 | $ (λ { (j , (s , k)) p _ → r j s p , k p }) 229 | 230 | noMagic : ∀ {I J}(C : I ▷ J)(X : I → Set) jc → ⟦ Somewhere C X ⟧ᵢ (λ _ → Zero) jc → Zero 231 | noMagic C X _ (p , m) = m <> 232 | 233 | lookup : ∀ {I J}(C : I ▷ J)(X : I → Set) jc {Q R} → ⟦ Everywhere C X ⟧ᵢ Q jc → ⟦ Somewhere C X ⟧ᵢ R jc → ⟦ Somewhere C X ⟧ᵢ (λ ix → Q ix × R ix) jc 234 | lookup C X jc (_ , q) (p , r) = p , (λ _ → (q p) , (r <>)) 235 | 236 | treeInd : ∀ {I}(C : I ▷ I)(P : Σ I (ITree C) → Set) → 237 | (⟦ Everywhere C (ITree C) ⟧ᵢ P -:> 238 | (vv λ i ts → P (i , ⟨ ts ⟩))) → 239 | (i : I)(t : ITree C i) → P (i , t) 240 | treeInd C P m i ⟨ s , k ⟩ = m (i , (s , k)) (<> , λ p → treeInd C P m _ (k p)) 241 | 242 | treeFold : ∀ {I}(C : I ▷ I)(P : I → Set) → 243 | (⟦ C ⟧ᵢ P -:> P) → 244 | (ITree C -:> P) 245 | treeFold C P m i ⟨ s , k ⟩ = m i (s , λ p → treeFold C P m _ (k p)) 246 | 247 | Children : ∀ {I} (C : I ▷ I) → Σ I (ITree C) ▷ Σ I (ITree C) 248 | Children C = Colx (descIxCon (var ∘ buds)) (Everywhere C (ITree C)) 249 | where 250 | buds : ∀ {I}{C : I ▷ I} → Σ I (ITree C) → Σ I λ i → ⟦ C ⟧ᵢ (ITree C) i 251 | buds (i , ⟨ xs ⟩) = i , xs 252 | 253 | treeFoldInd : ∀ {I} (C : I ▷ I) P → 254 | (⟦ Children C ⟧ᵢ P -:> P) → 255 | ∀ it → P it 256 | treeFoldInd C P m (i , t) = treeFold (Children C) P m (i , t) (children C i t) 257 | where 258 | children : ∀ {I}(C : I ▷ I) i t → ITree (Children C) (i , t) 259 | children C i ⟨ s , k ⟩ = ⟨ _ , (vv (λ _ p → children C _ (k p))) ⟩ 260 | 261 | EverywhereD SomewhereD : {I : Set}(D : Desc I)(X : I → Set) → ⟦ D ⟧D X → Desc (Σ I X) 262 | EverywhereD (var x) X xs = var (x , xs) 263 | EverywhereD (σ A D) X (x , xs) = EverywhereD (D x) X xs 264 | EverywhereD (π A D) X f = π A λ a → EverywhereD (D a) X (f a) 265 | EverywhereD (D ×D D₁) X (x , xs) = EverywhereD D X x ×D EverywhereD D₁ X xs 266 | EverywhereD (κ x) X xs = κ One 267 | SomewhereD (var x) X xs = var (x , xs) 268 | SomewhereD (σ A D) X (x , xs) = SomewhereD (D x) X xs 269 | SomewhereD (π A D) X f = σ A λ a → SomewhereD (D a) X (f a) 270 | SomewhereD (D ×D D₁) X (x , xs) = σ Two (SomewhereD D X x ⟨?⟩ SomewhereD D₁ X xs) 271 | SomewhereD (κ x) X xs = κ Zero 272 | 273 | dataInd : ∀ {I : Set} (F : I → Desc I)(P : Σ I (Data F) → Set) → 274 | ((i : I)(ds : ⟦ F i ⟧D (Data F)) → 275 | ⟦ EverywhereD (F i) (Data F) ds ⟧D P → P (i , ⟨ ds ⟩)) → 276 | ∀ i d → P (i , d) 277 | dataInd F P m i ⟨ ds ⟩ = m i ds (lem (F i) ds) where 278 | lem : (D : Desc _)(ds : ⟦ D ⟧D (Data F)) → ⟦ EverywhereD D (Data F) ds ⟧D P 279 | lem (var x) d = dataInd F P m x d 280 | lem (σ A D) (l , r) = lem (D l) r 281 | lem (π A D) f a = lem (D a) (f a) 282 | lem (D ×D E) (l , r) = lem D l , lem E r 283 | lem (κ x) y = <> 284 | 285 | vecNodeIx : (One ⊎ ℕ) ▷ ℕ 286 | vecNodeIx = descIxCon {J = ℕ} λ 287 | { zero → κ One 288 | ; (suc n) → var (tt , <>) ×D var (ff , n) 289 | } 290 | 291 | μlx : ∀ {I J} → (I ⊎ J) ▷ J → I ▷ J 292 | μlx {I}{J} F = (ITree F₁ ∘ _,_ ff) ◃ (P₁ ∘ _,_ ff) $ (r₁ ∘ _,_ ff) where 293 | F₁ : (I ⊎ J) ▷ (I ⊎ J) 294 | F₁ = (vv (λ i → One) ⟨?⟩ ShIx F) 295 | ◃ (vv (λ _ _ → Zero) ⟨?⟩ PoIx F) 296 | $ (vv (λ t s ()) ⟨?⟩ riIx F) 297 | P₁ : (x : I ⊎ J) → ITree F₁ x → Set 298 | P₁ (tt , i) _ = One 299 | P₁ (ff , j) ⟨ s , k ⟩ = Σ (PoIx F j s) λ p → P₁ (riIx F j s p) (k p) 300 | 301 | r₁ : (x : I ⊎ J) (t : ITree F₁ x) → P₁ x t → I 302 | r₁ (tt , i) _ _ = i 303 | r₁ (ff , j) ⟨ s , k ⟩ (p , ps) = r₁ _ (k p) ps 304 | 305 | vecIx : One ▷ ℕ 306 | vecIx = μlx vecNodeIx 307 | 308 | Vec'' : Set → ℕ → Set 309 | Vec'' X = ⟦ vecIx ⟧ᵢ (λ _ → X) 310 | 311 | vnil'' : ∀ {X} → Vec'' X 0 312 | vnil'' = ⟨ (<> , λ ()) ⟩ , (vv λ ()) 313 | 314 | vcons'' : ∀ {X n} → X → Vec'' X n → Vec'' X (suc n) 315 | vcons'' x (s , k) 316 | = ⟨ _ , (λ { (tt , _) → ⟨ (_ , λ ()) ⟩ 317 | ; (ff , _) → s 318 | }) ⟩ 319 | , (λ { ((tt , _) , _) → x 320 | ; ((ff , _) , p) → k p 321 | }) 322 | 323 | data Descμ (I : Set) : Set₁ where 324 | var : I → Descμ I 325 | σ π : (A : Set)(D : A → Descμ I) → Descμ I 326 | _×D_ : Descμ I → Descμ I → Descμ I 327 | κ : Set → Descμ I 328 | μ : (J : Set) → (J → Descμ (I ⊎ J)) → J → Descμ I 329 | 330 | {- 331 | mutual 332 | ⟦_⟧Dμ : ∀ {I} → Descμ I → (I → Set) → Set 333 | ⟦ x ⟧Dμ X = ? 334 | 335 | data Dataμ {I J} (F : J → Descμ (I ⊎ J))(X : I → Set) (j : J) : Set where 336 | ⟨_⟩ : ⟦ F j ⟧Dμ (vv X ⟨?⟩ Dataμ F X) → Dataμ F X j 337 | -} 338 | -------------------------------------------------------------------------------- /Chapter4/IndexedContainers.agda: -------------------------------------------------------------------------------- 1 | module Chapter4.IndexedContainers where 2 | 3 | open import Agda.Primitive 4 | open import Meta.Basics 5 | open import Meta.Language.LambdaCalculus 6 | 7 | -- An I ▷ J describes J sorts of structures in terms of I sorts of elements. 8 | -- 9 | -- "Hancock calls these indexed containers 'interaction structures.' Consider 10 | -- J to be the set of possible 'states of the world' before an interaction, 11 | -- and I to be the possible states afterward. The 'before' states will 12 | -- determine a choice of commands we can issue, each of which has a set of 13 | -- possible responses which will then determine the state 'after'. An 14 | -- interaction structures thus describes the predicate transformer which 15 | -- describes the precondition for achieving a postcondition by one step of 16 | -- interaction. We are just using proof-relevant Hoare logic as the type system!" 17 | record _▷_ (I J : Set) : Set₁ where 18 | constructor _◃_$_ 19 | field 20 | ShIx : J → Set -- A J-indexed family of operations 21 | PoIx : (j : J) → ShIx j → Set -- for each of these, a family of arities 22 | riIx : (j : J) (s : ShIx j) (p : PoIx j s) → I -- and a typing discipline, matching sorts to arities. 23 | ⟦_⟧ᵢ : (I → Set) → (J → Set) 24 | ⟦_⟧ᵢ X j = Σ (ShIx j) λ s → (p : PoIx j s) → X (riIx j s p) 25 | open _▷_ public 26 | 27 | -- The interpretation of indexed containers as operations on indexed families of 28 | -- sets. 29 | _-:>_ : ∀ {k l}{I : Set k} → (I → Set l) → (I → Set l) → Set (l ⊔ k) 30 | X -:> Y = ∀ i → X i → Y i 31 | 32 | -- The functorial action for indexed containers as operations on indexed families 33 | -- of sets. 34 | ixMap : ∀ {I J}{C : I ▷ J}{X Y} → (X -:> Y) → ⟦ C ⟧ᵢ X -:> ⟦ C ⟧ᵢ Y 35 | ixMap f j (s , k) = s , λ p → f _ (k p) 36 | -------------------------------------------------------------------------------- /Chapter4/Jacobian.agda: -------------------------------------------------------------------------------- 1 | module Chapter4.Jacobian where 2 | 3 | open import Meta.Basics 4 | open import Chapter4.IndexedContainers 5 | open import Chapter4.ITree 6 | open import Meta.Data.Inductive.W 7 | open import Meta.Data.Functor.Container 8 | 9 | -- The derivative of I ▷ J is I ▷ (J × I) where a (j, i) derivative is a 10 | -- structure of index j with a hole of index i. 11 | -- Shapes: (i, j) derivatives 12 | -- Positions: the hole at index i 13 | -- Elements: 14 | Jac : ∀ {I J} → I ▷ J → I ▷ (J × I) 15 | Jac (S ◃ P $ r) 16 | = (λ { (j , i) → Σ (S j) λ s → r j s ⁻¹ i }) 17 | ◃ (λ { (j , .(r j s p)) (s , from p) → P j s ─ p }) 18 | $ (λ { (j , .(r j s p)) (s , from p) (p' , _) → r j s p' }) 19 | 20 | plugIx : ∀ {I J} (C : I ▷ J) → ((j : J)(s : ShIx C j)(p p₁ : PoIx C j s) → Dec (p ≃ p₁)) → 21 | ∀ {i j X} → ⟦ Jac C ⟧ᵢ X (j , i) → X i → ⟦ C ⟧ᵢ X j 22 | plugIx C eq? {X = X} ((s , from p) , k) x = s , help where 23 | help : (p : PoIx C _ s) → X (riIx C _ s p) 24 | help p' with eq? _ s p' p 25 | help .p | tt , refl = x 26 | help p' | ff , np = k (p' , np) 27 | 28 | -- 4.12 (the Zipper) For a given C : I ▷ I, construct the indexed container 29 | -- Zipper C : (I × I) ▷ (I × I) such that ITree (Zipper C) (ir, sh) represents 30 | -- a one ih-hole context in a ITree C ir, represented as a sequence of 31 | -- hole-to-root layers. 32 | Zipper : ∀ {I} → I ▷ I → (I × I) ▷ (I × I) 33 | Zipper {I} C = (λ { (fst , snd) → (fst ≃ snd) ⊎ Σ I λ ip → ⟦ Jac C ⟧ᵢ (ITree C) (ip , snd) }) 34 | ◃ (λ { _ (tt , _) → Zero ; _ (ff , _) → One }) 35 | $ (λ { _ (tt , snd) () ; (fst , snd) (ff , (fst₁ , _)) _ → fst , fst₁ }) 36 | 37 | -- Check that you can zipper all the way out to the root. 38 | zipOut : ∀ {I} (C : I ▷ I) {ir ih} → 39 | ((i : I)(s : ShIx C i)(p p₁ : PoIx C i s) → Dec (p ≃ p₁)) → 40 | ITree (Zipper C) (ir , ih) → ITree C ih → ITree C ir 41 | zipOut C eq? ⟨ (tt , refl) , _ ⟩ t = t 42 | zipOut C eq? ⟨ (ff , (_ , c)) , cz ⟩ t = zipOut C eq? (cz <>) ⟨ plugIx C eq? c t ⟩ 43 | 44 | -- 4.13 (differentiating Desc) The notion corresponding to Jac for descriptions 45 | -- is ▽, computing a 'vector' of partial derivatives. Define it symbolically. 46 | ▽ : {I : Set} → Desc I → I → Desc I 47 | -- The derivative of a constant is 0 48 | ▽ (κ x) h = κ Zero 49 | -- ▽ X is a constant (⊤) 50 | ▽ (var x) h = κ (x ≃ h) 51 | -- ▽ (Σ (a : A) (D a)) is the partial derivative of D projected from A. 52 | ▽ (σ A D) h = σ A λ a → ▽ (D a) h 53 | -- ▽ ((a : A) → D) is the partial deriative of D projected from A and an arrow 54 | -- computing further derivatives of values distinct from a. 55 | ▽ (π A D) h = σ A λ a → ▽ (D a) h ×D π (A ─ a) λ { (a₁ , _) → D a₁ } 56 | -- The product rule 57 | ▽ (D ×D E) h = σ Two ((▽ D h ×D E) ⟨?⟩ (D ×D ▽ E h)) 58 | -------------------------------------------------------------------------------- /Chapter5/InductionRecursion.agda: -------------------------------------------------------------------------------- 1 | module Chapter5.InductionRecursion where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Fin 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container.Indexed 7 | 8 | -- Finite sums are finite, so we can compute their sizes. 9 | sum prod : (n : ℕ) → (Fin n → ℕ) → ℕ 10 | sum zero _ = 0 11 | sum (suc n) f = f zero + sum n (f ∘ suc) 12 | 13 | -- Finite products are finite, so we can compute their sizes. 14 | prod zero _ = 1 15 | prod (suc n) f = f zero * sum n (f ∘ suc) 16 | 17 | {- 18 | data FTy : Set where 19 | fin : ℕ → FTy 20 | σ π : (S : FTy) (T : Fin {!!} → FTy) → FTy 21 | -} 22 | 23 | {- 24 | mutual 25 | data FTy : Set where 26 | fin : ℕ → FTy 27 | σ π : (S : FTy) (T : Fin (# S) → FTy) → FTy 28 | 29 | # : FTy → ℕ 30 | # (fin x) = x 31 | # (σ S T) = sum (# S) λ s → # (T s) 32 | # (π S T) = prod (# S) λ s → # (T s) 33 | 34 | -- A forgetful map from (Fin n) back to ℕ to illustrate induction-recursion. 35 | fog : ∀ {n} → Fin n → ℕ 36 | fog zero = zero 37 | fog (suc i) = suc (fog i) 38 | -} 39 | 40 | -- "Where an inductive definition tells us how to perform construction of data 41 | -- incrementally, induction-recursion tells us how to perform 42 | -- construction-with-interpretation incrementally." 43 | mutual 44 | -- Thus, Construction 45 | data FTy : Set where 46 | fin : ℕ → FTy 47 | σ π : (S : FTy) (T : FEI S → FTy) → FTy 48 | 49 | -- And interpretation. 50 | FEI : FTy → Set 51 | FEI (fin x) = Fin x 52 | FEI (σ S T) = Σ (FEI S) λ s → FEI (T s) 53 | FEI (π S T) = (s : FEI S) → FEI (T s) 54 | 55 | -- The above definition is an important step up in power. Rather than before 56 | -- where `FTy ∘ #` yielded an unstructured interpretation of a family of sets 57 | -- indexed by ℕ, this time `FTy ∘ FEI` yields a subset Types that have names in 58 | -- FTy while still being small enough to be a set. 59 | -- 60 | -- "IR stands for Incredible Ray that shrinks large sets to small encodings of 61 | -- subsets of them. 62 | 63 | -- Exercise 5.1 By means of a suitable choice of recursive interpretation, fill 64 | -- the ? with a condition which ensures that FreshLists have distinct elements. 65 | -- Try to make sure that, for any concrete FreshList, ok can be inferred 66 | -- trivially. 67 | module FRESHLIST (X : Set) (Xeq? : (x x₁ : X) → Dec (x ≃ x₁)) where 68 | mutual 69 | data FreshList : Set where 70 | [] : FreshList 71 | _,_ : (x : X)(xs : FreshList) {ok : x ∉ xs} → FreshList 72 | 73 | -- The distinctness (freshness) guarantee says that. 74 | _∉_ : X → FreshList → Set 75 | x ∉ [] = One -- There's nothing to do with an x and the empty list. 76 | x ∉ (x₁ , xs) with Xeq? x x₁ -- Otherwise destruct a decidable eq on head. 77 | x ∉ (x₁ , xs) | tt , _ = Zero -- We have a match! It isn't fresh. 78 | x ∉ (x₁ , xs) | ff , _ = x ∉ xs -- Otherwise try again. 79 | 80 | -- "Randy Pollack identified the task of modelling record types as a key early 81 | -- use of induction-recursion , motivated to organise libraries for mathematical 82 | -- structure." 83 | -- In this case, we don't need the full power of the IR Ray because you can 84 | -- get there with Desc and right-nested Σ's. Essentially, the Σ's give us a way 85 | -- of creating a structure where later fields can depend on earlier fields. 86 | data RecR : Set₁ where 87 | ⟨⟩ : RecR 88 | _,_ : (A : Set) (R : A → RecR) → RecR 89 | 90 | -- Lower everything back down to Agda's level. 91 | ⟦_⟧RR : RecR → Set 92 | ⟦ ⟨⟩ ⟧RR = One 93 | ⟦ A , R ⟧RR = Σ A λ a → ⟦ R a ⟧RR 94 | 95 | -- Exercise 5.2: Show how to compute the size of a record, the define the 96 | -- projections, first of types, then of values. 97 | 98 | -- For the sizes of records: 99 | sizeRR : (R : RecR) → ⟦ R ⟧RR → ℕ 100 | sizeRR ⟨⟩ r = 0 -- If we have nothing we're done. 101 | sizeRR (A , R) (fld , rest) = suc (sizeRR (R fld) rest) -- Otherwise we add 1 field and try again. 102 | 103 | -- For types: 104 | TyRR : (R : RecR)(r : ⟦ R ⟧RR) → Fin (sizeRR R r) → Set 105 | TyRR ⟨⟩ _ () -- It makes no sense to project the type out of nothing. 106 | TyRR (A , R) (_ , _) zero = A -- If the structure has one field we just project the type of that thing. 107 | TyRR (A , R) (fld , rest) (suc i) = TyRR (R fld) rest i -- Else try again. 108 | 109 | -- For values: 110 | vaRR : (R : RecR)(r : ⟦ R ⟧RR)(i : Fin (sizeRR R r)) → TyRR R r i 111 | vaRR ⟨⟩ <> () -- It makes no sense to project a value out of nothing. 112 | vaRR (_ , _) (fld , rest) zero = fld -- If the structure has one field we just project the value out of that. 113 | vaRR (A , R) (fld , rest) (suc i) = vaRR (R fld) rest i -- Else try again. 114 | 115 | -- But if you want left-nesting of record types (i.e. Dependent Contexts), you 116 | -- do need the IR Ray. 117 | mutual 118 | data RecL : Set₁ where 119 | ε : RecL 120 | _,_ : (R : RecL)(A : ⟦ R ⟧RL → Set) → RecL 121 | 122 | ⟦_⟧RL : RecL → Set 123 | ⟦ ε ⟧RL = One 124 | ⟦ R , A ⟧RL = Σ ⟦ R ⟧RL A 125 | 126 | -- Exercise 5.3: Show how to compute the size of a RecL without knowing the 127 | -- individual record. Show how to interpret a projection as a function from a 128 | -- record, first for types, then values. 129 | sizeRL : RecL → ℕ 130 | sizeRL ε = 0 -- The size of the empty context is zero. 131 | sizeRL (R , A) = suc (sizeRL R) -- Else add one and try again. 132 | 133 | -- For Types: 134 | TyRL : (R : RecL) → Fin (sizeRL R) → ⟦ R ⟧RL → Set 135 | TyRL ε () <> -- It makes no sense to project a type out of the empty structure. 136 | TyRL (R , A) zero (idx , _) = A idx -- If the structure has one field we project it out by indexing with the field. 137 | TyRL (R , A) (suc i) (idx , _) = TyRL R i idx -- Else try again. 138 | 139 | -- For values: 140 | vaRL : (R : RecL)(i : Fin (sizeRL R))(r : ⟦ R ⟧RL) → TyRL R i r 141 | vaRL ε () <> -- It makes no sense to project a value out of the empty structure. 142 | vaRL (R , A) zero (_ , fld) = fld -- If the structure has one field just project the value out. 143 | vaRL (R , A) (suc i) (idx , snd) = vaRL R i idx 144 | 145 | -- Exercise 5.4: Show how to truncate a record signature from a given field and 146 | -- compute the corresponding projection on structures. 147 | 148 | -- To truncate a context: 149 | TruncRL : (R : RecL) → Fin (sizeRL R) → RecL 150 | TruncRL ε () -- It makes no sense to truncate the empty context. 151 | TruncRL (R , A) zero = R -- For a context with a single part, yield it. 152 | TruncRL (R , A) (suc i) = TruncRL R i -- Else try again. 153 | 154 | -- To project out of a truncation: 155 | truncRL : (R : RecL)(i : Fin (sizeRL R)) → ⟦ R ⟧RL → ⟦ TruncRL R i ⟧RL 156 | truncRL ε () <> -- How did we get here? 157 | truncRL (R , A) zero (fst , snd) = fst -- For a singleton context, give up the subpart. 158 | truncRL (R , A) (suc i) (fst , snd) = truncRL R i fst -- Else try again. 159 | 160 | -- A Manifest is a record whose values are computed from earlier fields. "It is 161 | -- rather like allowing definitions in contexts." 162 | data Manifest {A : Set} : A → Set where 163 | ⟨_⟩ : (a : A) → Manifest a 164 | 165 | mutual 166 | -- "I index by size, to save on measuring." 167 | data RecM : ℕ → Set₁ where 168 | ε : RecM 0 -- The empty manifest. 169 | _,_ : {n : ℕ} (R : RecM n)(A : ⟦ R ⟧RM → Set) → RecM (suc n) -- Old-school fields 170 | _,_∋_ : {n : ℕ} (R : RecM n) (A : ⟦ R ⟧RM → Set)(a : (r : ⟦ R ⟧RM) → A r) → RecM (suc n) -- A definition. 171 | 172 | ⟦_⟧RM : {n : ℕ} → RecM n → Set 173 | ⟦ ε ⟧RM = One 174 | ⟦ R , A ⟧RM = Σ ⟦ R ⟧RM A 175 | ⟦ R , A ∋ a ⟧RM = Σ ⟦ R ⟧RM (Manifest ∘ a) 176 | 177 | -- Exercise 5.5: Implement projection for RecM. 178 | 179 | -- For types: 180 | TyRM : {n : ℕ} (R : RecM n) → Fin n → ⟦ R ⟧RM → Set 181 | -- Makes no sense to project out of nothing. 182 | TyRM ε () <> 183 | -- If the structure has one field we project it out by indexing with the field 184 | TyRM (R , A) zero (idx , _) = A idx 185 | -- Else try again. 186 | TyRM (R , A) (suc i) (idx , _) = TyRM R i idx 187 | -- If the structure has one definition we project it out by indexing with the field. 188 | TyRM (R , A ∋ a) zero (idx , _) = A idx 189 | -- Else try again. 190 | TyRM (R , A ∋ a) (suc i) (idx , _) = TyRM R i idx 191 | 192 | -- For values: 193 | vaRM : {n : ℕ}(R : RecM n)(i : Fin n)(r : ⟦ R ⟧RM) → TyRM R i r 194 | -- Makes no sense to project out of nothing. 195 | vaRM ε () r 196 | -- If the structure has one field, project it out. 197 | vaRM (R , A) zero (_ , fld) = fld 198 | -- Else try again. 199 | vaRM (R , A) (suc i) (idx , _) = vaRM R i idx 200 | -- If the structure has one manifest field, project it out by feeding it what it depends on. 201 | vaRM (R , A ∋ a) zero (idx , _) = a idx 202 | -- Else try again. 203 | vaRM (R , A ∋ a) (suc i) (idx , _) = vaRM R i idx 204 | 205 | {- 206 | mutual 207 | data REx : {n m : ℕ} → RecM n → RecM m → Set₁ where 208 | ε : REx ε ε 209 | 210 | rfog : ∀ {n m}{R : RecM n}{R₁ : RecM m} (X : REx R R₁) → ⟦ R₁ ⟧RM → ⟦ R ⟧RM 211 | rfog ε <> = <> 212 | -} 213 | 214 | mutual 215 | -- Conor's Favorite Universe Featuring: 216 | data TU : Set where 217 | -- A scattering of base types. 218 | Zero' One' Two' : TU 219 | -- Dependent pairs (and functions). 220 | Σ' Π' : (S : TU)(R : ⟦ S ⟧TU → TU) → TU 221 | -- And Petersson-Synek Trees. 222 | Tree' : (I : TU)(F : ⟦ I ⟧TU → Σ TU λ S → 223 | ⟦ S ⟧TU → Σ TU λ P → 224 | ⟦ P ⟧TU → ⟦ I ⟧TU) 225 | (i : ⟦ I ⟧TU) → TU 226 | ⟦_⟧TU : TU → Set 227 | ⟦ Zero' ⟧TU = Zero 228 | ⟦ One' ⟧TU = One 229 | ⟦ Two' ⟧TU = Two 230 | ⟦ Σ' S R ⟧TU = Σ ⟦ S ⟧TU λ s → ⟦ R s ⟧TU 231 | ⟦ Π' S R ⟧TU = (s : ⟦ S ⟧TU) → ⟦ R s ⟧TU 232 | ⟦ Tree' I F i ⟧TU = ITree 233 | ( (λ i → ⟦ fst (F i) ⟧TU) 234 | ◃ (λ i s → ⟦ fst (snd (F i) s) ⟧TU) 235 | $ (λ i s p → snd (snd (F i) s) p) 236 | ) i 237 | 238 | -- Unfortunately, because this is an inductive-recursive structure in and of 239 | -- itself: you can't use the IR Ray on an IR Ray. 240 | 241 | Pow : Set₁ → Set₁ 242 | Pow X = X → Set 243 | 244 | Fam : Set₁ → Set₁ 245 | Fam X = Σ Set λ I → I → X 246 | 247 | mutual 248 | -- A predicative hierarchy of small universes built using induction-recursion. 249 | data NU (X : Fam Set) : Set where 250 | U′ : NU X 251 | El′ : fst X → NU X 252 | Nat′ : NU X 253 | Π′ : (S : NU X) (T : ⟦ S ⟧NU → NU X) → NU X 254 | 255 | ⟦_⟧NU : ∀ {X} → NU X → Set 256 | ⟦_⟧NU { U , El } U′ = U 257 | ⟦_⟧NU { U , El } (El′ T) = El T 258 | ⟦ Nat′ ⟧NU = ℕ 259 | ⟦ Π′ S T ⟧NU = (s : ⟦ S ⟧NU) → ⟦ T s ⟧NU 260 | 261 | -- Universes can be "jacked up" as far as we like. 262 | 263 | EMPTY : Fam Set 264 | EMPTY = Zero , λ () 265 | 266 | LEVEL : ℕ → Fam Set 267 | LEVEL zero = (NU EMPTY) , ⟦_⟧NU 268 | LEVEL (suc n) = (NU (LEVEL n)) , ⟦_⟧NU 269 | 270 | mutual 271 | -- To eliminate redudancy of lower universes, we parametrize the universe by a 272 | -- de Bruijin indexed collection of the previous universes. 273 | data HU {n} (U : Fin n → Set) : Set where 274 | U′ : Fin n → HU U 275 | Nat′ : HU U 276 | Π′ : (S : HU U) (T : ⟦ S ⟧HU → HU U) → HU U 277 | ⟦_⟧HU : ∀ {n} {U : Fin n → Set} → HU U → Set 278 | ⟦_⟧HU {U = U} (U′ i) = U i 279 | ⟦ Nat′ ⟧HU = ℕ 280 | ⟦ Π′ S T ⟧HU = (s : ⟦ S ⟧HU) → ⟦ T s ⟧HU 281 | 282 | -- "To finish the job, we must build the collections of levels to hand to HU. At 283 | -- each step, level zero is the new top level, built with a fresh appeal to HU, 284 | -- but lower levels can be projected from the previous collection." 285 | 286 | HPREDS : (n : ℕ) → Fin n → Set 287 | HPREDS zero () 288 | HPREDS (suc n) zero = HU (HPREDS n) 289 | HPREDS (suc n) (suc x) = HPREDS n x 290 | 291 | HSET : ℕ → Set 292 | HSET n = HU (HPREDS n) 293 | 294 | {- rejected. 295 | mutual 296 | data VV : Set where 297 | V′ : VV 298 | Π′ : (S : VV) (T : ⟦ S ⟧VV → VV) → VV 299 | ⟦_⟧VV : VV → Set 300 | ⟦ V′ ⟧VV = VV 301 | ⟦ Π′ S T ⟧VV = (s : ⟦ S ⟧VV) → ⟦ T s ⟧VV 302 | -} 303 | 304 | -- An encoding of allowed inductive-recursive things [Dybjer and Setzer 1999]. 305 | -- The encoding is as follows: 306 | -- ∘ Describe one node of inductive recursive types in the manner of RecR but 307 | -- where each J-value read gives us a way of reading off I values. 308 | data DS (I J : Set₁) : Set₁ where 309 | ι : J → DS I J -- no more fields 310 | σ : (S : Set) (T : S → DS I J) → DS I J -- ordinary field 311 | δ : (H : Set) (T : (H → I) → DS I J) → DS I J -- child field. 312 | 313 | -- A DS is a functor. 314 | -- 315 | -- "In each case, we must say which set is being encoded and how to read off a J 316 | -- from a value in that set. 317 | ⟦_⟧DS : ∀ {I J} → DS I J → Fam I → Fam J 318 | -- The ι constructor carries exactly the j required. The other two specify a 319 | -- field in the node structure, from which the computation of the J gains some 320 | -- information. 321 | ⟦_⟧DS (ι x) x₁ 322 | = One 323 | , (λ {<> → x}) 324 | -- The σ specifies a field of type S, and the rest of the structure may depend 325 | -- on a value of type S. 326 | ⟦_⟧DS (σ S T) x 327 | = (Σ S λ s → fst (⟦ T s ⟧DS x)) 328 | , (λ { (s , t) → snd (⟦ T s ⟧DS x) t }) 329 | -- The δ case is the clever bit. It specifies a place for an H-indexed bunch of 330 | -- children, and even through we do not fix what set represents the children, we 331 | -- do know that they allow us to read off an I. Correspondingly, the rest of 332 | -- the structure can at least depend on knowing a function in H → I which gives 333 | -- access to the interpretation of the children. Once we plugin a specific 334 | -- (X, xi) : Fam I, we represent the field by the small function space 335 | -- hx : H → X, then the composition xi ∘ hx tells us how to compute the large 336 | -- meaning of each child." 337 | ⟦_⟧DS (δ H T) (X , xi) 338 | = (Σ (H → X) λ hx → fst (⟦ T (xi ∘ hx) ⟧DS (X , xi))) 339 | , (λ { (hx , t) → snd (⟦ T (xi ∘ hx) ⟧DS (X , xi)) t }) 340 | 341 | -- Exercise 5.10: A morphism form (X, xi) to (Y, yi) in Fam I is a function 342 | -- f : X → Y such that xi = yi ∘ f. Construct a code for the identity functor 343 | -- on Fam I such that ⟦ idDS ⟧DS ≅ id. 344 | idDS : { I : Set₁ } → DS I I 345 | idDS = δ One λ f → ι (f <>) 346 | 347 | {- -- fails positivity check and termination check. 348 | mutual 349 | data DataDS {I} (D : DS I I) : Set where 350 | ⟨_⟩ : fst (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) → DataDS D 351 | ⟦_⟧ds : {I : Set₁}{D : DS I I} → DataDS D → I 352 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = snd (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) ds 353 | -} 354 | 355 | mutual 356 | -- We get out of this jam by inlining the interpretation. 357 | data DataDS {I} (D : DS I I) : Set where 358 | ⟨_⟩ : NoDS D D → DataDS D 359 | 360 | ⟦_⟧ds : {I : Set₁} {D : DS I I} → DataDS D → I 361 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = DeDS D D ds 362 | 363 | NoDS : {I : Set₁} (D D′ : DS I I) → Set 364 | NoDS D (ι i) = One 365 | NoDS D (σ S T) = Σ S λ s → NoDS D (T s) 366 | NoDS D (δ H T) = Σ (H → DataDS D) λ hd → NoDS D (T (λ h → ⟦ hd h ⟧ds)) 367 | 368 | DeDS : {I : Set₁} (D D′ : DS I I) → NoDS D D′ → I 369 | DeDS D (ι x) <> = x 370 | DeDS D (σ S T) (s , t) = DeDS D (T s) t 371 | DeDS D (δ H T) (hd , t) = DeDS D (T (λ h → ⟦ hd h ⟧ds)) t 372 | 373 | 374 | -- Exercise 5.12 (bindDS and its meaning): Implement the appropriate bindDS 375 | -- operator, corresponding to substitution at ◆. 376 | bindDS : ∀ {I J K} → DS I J → (J → DS I K) → DS I K 377 | bindDS (ι x) U = U x 378 | bindDS (σ S T) U = σ S (λ s → bindDS (T s) U) 379 | bindDS (δ H T) U = δ H (λ f → bindDS (T f) U) 380 | 381 | -- Show that bindDS corresponds to a kind of Σ by implementing pairing and 382 | -- projections: 383 | pairDS : ∀ {I J K} (T : DS I J) (U : J → DS I K) {X : Fam I} → 384 | (t : fst (⟦ T ⟧DS X)) (u : fst (⟦ U (snd (⟦ T ⟧DS X) t) ⟧DS X)) 385 | → fst (⟦ bindDS T U ⟧DS X) 386 | pairDS (ι x) U <> u = u 387 | pairDS (σ S T) U (s , t) u = s , pairDS (T s) U t u 388 | pairDS (δ H T) U {_ , d} (f , t) u = f , (pairDS (T (d ∘ f)) U t u) 389 | 390 | {- 391 | coDS : ∀ {I J K} → DS J K → DS I J → DS I K 392 | coDS E D = {!!} 393 | -} 394 | 395 | mutual 396 | -- An Indutive-Recursive definition of Indutive-Recursive things that have 397 | -- 'I'nformation in them. It is an open problem as to whether this actually 398 | -- encodes IR, so for now it is just IR-ish. 399 | data Irish (I : Set₁) : Set₁ where 400 | ι : Irish I 401 | κ : Set → Irish I 402 | π : (S : Set) (T : S → Irish I) → Irish I 403 | σ : (S : Irish I) (T : Info S → Irish I) → Irish I 404 | 405 | Info : ∀ {I} → Irish I → Set₁ 406 | Info {I} ι = I 407 | Info (κ A) = ⇑ A 408 | Info (π S T) = (s : S) → Info (T s) 409 | Info (σ S T) = Σ (Info S) λ s → Info (T s) 410 | 411 | -- To interpret π and σ, we shall need to equip Fam with pointwise lifting and 412 | -- dependent pairs, respectively. 413 | ΠF : (S : Set) {J : S → Set₁} (T : (s : S) → Fam (J s)) → Fam ((s : S) → J s) 414 | ΠF S T = ((s : S) → fst (T s)) , (λ f s → snd (T s) (f s)) 415 | 416 | ΣF : {I : Set₁} (S : Fam I) {J : I → Set₁} (T : (i : I) → Fam (J i)) → Fam (Σ I J) 417 | ΣF S T = (Σ (fst S) (fst ∘ (T ∘ snd S))) , (λ { (s , t) → snd S s , snd (T (snd S s)) t }) 418 | 419 | -- Now, for any T : Irish I, if someone gives us a Fam I to represent children, 420 | -- we can compute a Fam (Info T) — a small node structure from which the large 421 | -- Info T can be extracted. 422 | Node : ∀ {I} (T : Irish I) → Fam I → Fam (Info T) 423 | Node ι X = X 424 | Node (κ A) X = A , ↑ 425 | Node (π S T) X = ΠF S λ s → Node (T s) X 426 | Node (σ S T) X = ΣF (Node S X) λ iS → Node (T iS) X 427 | 428 | -- A functor from Fam I to Fam J is then given by a pair 429 | IF : Set₁ → Set₁ → Set₁ 430 | IF I J = Σ (Irish I) λ T → Info T → J 431 | 432 | _$F_ : ∀ {I J} → (I → J) → Fam I → Fam J 433 | f $F (F , i) = F , (f ∘ i) 434 | 435 | ⟦_⟧IF : ∀ {I J} → IF I J → Fam I → Fam J 436 | ⟦ (T , d) ⟧IF X = d $F Node T X 437 | 438 | {- 439 | -- With a certain tedious inevitability, we find that Agda rejects the obvious 440 | -- attempt to tie the knot. 441 | 442 | mutual 443 | data DataIF {I} (F : IF I I) : Set where 444 | ⟨_⟩ : fst (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) → DataIF F 445 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 446 | ⟦_⟧if {F = F} ⟨ ds ⟩ = snd (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) ds 447 | -} 448 | 449 | mutual 450 | -- Again, specialization of Node fixes the problem 451 | data DataIF {I} (F : IF I I) : Set where 452 | ⟨_⟩ : NoIF F (fst F) → DataIF F 453 | 454 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 455 | ⟦_⟧if {F = F} ⟨ rs ⟩ = snd F (DelF F (fst F) rs) 456 | 457 | NoIF : ∀ {I} (F : IF I I) (T : Irish I) → Set 458 | NoIF F ι = DataIF F 459 | NoIF F (κ A) = A 460 | NoIF F (π S T) = (s : S) → NoIF F (T s) 461 | NoIF F (σ S T) = Σ (NoIF F S) λ s → NoIF F (T (DelF F S s)) 462 | 463 | DelF : ∀ {I} (F : IF I I) (T : Irish I) → NoIF F T → Info T 464 | DelF F ι r = ⟦ r ⟧if 465 | DelF F (κ A) r = ↑ r 466 | DelF F (π S T) f = λ s → DelF F (T s) (f s) 467 | DelF F (σ S T) (s , t) = let s′ = DelF F S s in s′ , DelF F (T s′) t 468 | 469 | -- Exercise 5.15 (Irish-to-Swedish): Show how to define this, such that 470 | -- such that ⟦ DSIF T ⟧DS ≅ ⟦ T ⟧IF 471 | DSIF : ∀ {I J} → DS I J → IF I J 472 | DSIF (ι x) = κ One , (λ _ → x) 473 | DSIF (σ S T) = (σ (κ S) (λ s → fst (DSIF (T (↓ s))))) , (λ { (s , t) → snd (DSIF (T (↓ s))) t }) 474 | DSIF (δ H T) = (σ (π H λ _ → ι) (λ f → fst (DSIF (T f)))) , (λ { (f , t) → snd (DSIF (T f)) t }) 475 | 476 | -- We clearly have identity for the Irish IR. 477 | idIF : ∀ {I} → IF I I 478 | idIF = ι , id 479 | 480 | -- Exercise 5.16 (subIF): Construct a substitution operator for Irish J with a 481 | -- refinement of the following type. 482 | subIF : ∀ {I J} (T : Irish J) (F : IF I J) → IF I (Info T) 483 | subIF ι F = F 484 | subIF (κ A) F = κ A , (λ z → ↑ (↓ z)) 485 | subIF (π S T) F = (π S λ s → fst (subIF (T s) F)) , (λ f s → snd (subIF (T s) F) (f s)) 486 | subIF (σ S T) F with subIF S F 487 | ... | (SF , f) = (σ SF λ sf → fst (subIF (T (f sf)) F)) , (λ { (sf , tf) → f sf , snd (subIF (T (f sf)) F) tf }) 488 | 489 | -- Exercise 5.17 (coIF): Now define composition for Irish IR functors. 490 | colF : ∀ {I J K} → IF J K → IF I J → IF I K 491 | colF (S , x) F with subIF S F 492 | ... | TF , f = TF , (x ∘ f) 493 | 494 | -- "Some of us are inclined to suspect that IF does admit more functors than DS, 495 | -- but the exact status of Irish induction-recursion remains the stuff of future 496 | -- work." 497 | -------------------------------------------------------------------------------- /Chapter6/OTT.agda: -------------------------------------------------------------------------------- 1 | module Chapter6.OTT where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Inductive.Recursion 5 | open import Meta.Data.Functor.Container.Indexed 6 | open import Meta.Data.Inductive.ITree 7 | 8 | {- 9 | "We cannot have an equality which is both extensional and decidable." 10 | 11 | Or, that is to say, in order to give up either extensionality or decidability 12 | you wind up in one of two bad places: 13 | 1) Either you give up extensionality and are right back where you started 14 | (hence the definition below): trapped in a framework in which you can't 15 | prove anything more interesting than the ensemble of constructors you have 16 | and some equally terrible definition of substitution will allow. This is 17 | because `favorite` below is an "intensional Predicate" whose only canonical 18 | form is `λ x → zero + x`. 19 | 2) Or you give up decidability and, well, no. Don't do that. 20 | 21 | The trouble here is binders. If all the Favourite type chose to do with its 22 | parameter was apply it to arguments, extensionality would come naturally. But 23 | to decide equality under binders in general requires a framework for computing 24 | under them that we simply don't have (yet) with vanilla refl. 25 | 26 | data Favourite : (ℕ → ℕ) → Set where 27 | favourite : Favourite (λ x → zero + x) 28 | 29 | If we try again taking some fictitious extensional equality ≃ then lift things 30 | into Favorite, then Favorite becomes just as extensional as the chosen equality. 31 | 32 | "If q′ : f ≃ g, then we can transport favorite q from Favorite f to Favorite g, 33 | returning, not the original data but 34 | 35 | favourite ((λ x → zero +ℕ x) =[ q ⟩ f =[ q′ ⟩ q ∎) 36 | 37 | with a modified proof." 38 | 39 | data Favourite (f : ℕ → ℕ) : Set where 40 | favourite : (λ x → zero + x) ≃ f → Favourite f 41 | -} 42 | 43 | -- The mechanism that allows us a more general yet extensional equality of terms 44 | -- is Observational Equality, where essentially we use a set of local 45 | -- mini-propositional equalities on the result of inspecting the types and 46 | -- values of terms. 47 | 48 | -- "Inasmuch as types depend on values, we shall also need to say when values 49 | -- are equal. There is no reasn to presume that we shall be interested only to 50 | -- consider the equality of values in types which are judgmentally equal, for we 51 | -- know that judgmental equality is too weak to recognize the sameness of some 52 | -- types whose values are interchangeable. Correspondingly, let us weaken our 53 | -- requirement for the formation of value equalities and have a heterogeneous 54 | -- equality, Eq X x Y y. We have some options for how to do that: 55 | -- 56 | -- ∘ We could add the requirement X ↔ Y to the formation rule for Eq 57 | -- ∘ We could allow the formation of any Eq X x Y y, but ensure that it holds 58 | -- only if X ↔ Y. 59 | -- ∘ We could allow the formation of any Eq X x Y y, but ensure that proofs of 60 | -- such equations are useless information unless X ↔ Y." 61 | -- 62 | -- Spoiler: We choose the third. 63 | 64 | mutual -- By recursion on types in TU: 65 | EQ : (X Y : TU) → TU × (⟦ X ⟧TU → ⟦ Y ⟧TU → TU) 66 | -- "Base types equal only themselves, and we need no help to transport a value 67 | -- from a type to itself."" 68 | EQ Zero' Zero' = One' , λ _ _ → One' 69 | EQ One' One' = One' , λ _ _ → One' 70 | EQ Two' Two' = One' , λ 71 | { tt tt → One' 72 | ; ff ff → One' 73 | ; _ _ → Zero' 74 | } 75 | -- Σ-types+ 76 | EQ (Σ' S T) (Σ' S' T') = (Σ' (S ↔ S') λ _ → -- are interchangeable if 77 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 78 | T s ↔ T' s') 79 | , (λ { (s , t) (s' , t') → -- their components are. 80 | Σ' (Eq S s S' s') λ _ → Eq (T s) t (T' s') t' }) 81 | -- Equality of Π-types is straightforward, then. Because Π-types 82 | EQ (Π' S T) (Π' S' T') = (Σ' (S' ↔ S) λ _ → -- are interchangeable if 83 | Π' S' λ s' → Π' S λ s → Π' (Eq S' s' S s) λ _ → 84 | T s ↔ T' s') 85 | , (λ { f f' → -- their projected components are. 86 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 87 | Eq (T s) (f s) (T' s') (f' s') }) 88 | -- Equality of Tree-types procedes structurally. That is, Tree-types 89 | EQ (Tree' I F i) (Tree' I' F' i') 90 | = (Σ' (I ↔ I') λ _ → Σ' (Eq I i I' i') λ _ → -- are interchangeable if 91 | Π' I λ i → Π' I' λ i' → Π' (Eq I i I' i') λ _ → 92 | let (S , K) = F i ; S' , K' = F' i' 93 | in Σ' (S ↔ S') λ _ → -- we get equal shapes here 94 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 95 | let (P , r) = K s ; (P' , r') = K' s' 96 | in Σ' (P' ↔ P) λ _ → -- and equal shapes there 97 | Π' P' λ p' → Π' P λ p → Π' (Eq P' p' P p) λ _ → 98 | Eq I (r p) I' (r' p')) -- 99 | , teq i i' where 100 | -- Demands equal subtrees of the two given terms in TU 101 | teq : (i : ⟦ I ⟧TU) → (i' : ⟦ I' ⟧TU) → ⟦ Tree' I F i ⟧TU → ⟦ Tree' I' F' i' ⟧TU → TU 102 | teq i i' ⟨ s , k ⟩ ⟨ s' , k' ⟩ 103 | = let (S , K) = F i ; (S' , K') = F' i' 104 | (P , r) = K s ; (P' , r') = K' s' 105 | in Σ' (Eq S s S' s') λ _ → -- Subtrees are interchangeable when 106 | Π' P λ p → Π' P' λ p' → Π' (Eq P p P' p') λ _ → -- the shapes and points match and 107 | teq (r p) (r' p') (k p) (k' p') -- when further subtrees are interchangeable. 108 | -- Otherwise ex falso. After all, if you got this far, you deserve a medal. 109 | EQ _ _ = Zero' , λ _ _ → One' 110 | 111 | -- Extract the type realized by the observational equality above. 112 | _↔_ : TU → TU → TU 113 | X ↔ Y = fst (EQ X Y) 114 | 115 | -- Extract the equality realized by the observational equality above. 116 | Eq : (X : TU) (x : ⟦ X ⟧TU) → (Y : TU) (y : ⟦ Y ⟧TU) → TU 117 | Eq X x Y y = snd (EQ X Y) x y 118 | 119 | -- Coercion that realises equality as transportation. 120 | coe : (X Y : TU) → ⟦ X ↔ Y ⟧TU → ⟦ X ⟧TU → ⟦ Y ⟧TU 121 | postulate 122 | coh : (X Y : TU) (Q : ⟦ X ↔ Y ⟧TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x Y (coe X Y Q x) ⟧TU 123 | coe Zero' Zero' <> x = x 124 | coe Zero' One' () x 125 | coe Zero' Two' () x 126 | coe Zero' (Σ' Y R) () x 127 | coe Zero' (Π' Y R) () x s 128 | coe Zero' (Tree' Y F i) () x 129 | coe One' Zero' () x 130 | coe One' One' <> x = x 131 | coe One' Two' () x 132 | coe One' (Σ' Y R) () x 133 | coe One' (Π' Y R) () x s 134 | coe One' (Tree' Y F i) () x 135 | coe Two' Zero' () x 136 | coe Two' One' () x 137 | coe Two' Two' <> x = x 138 | coe Two' (Σ' Y R) () x 139 | coe Two' (Π' Y R) () x s 140 | coe Two' (Tree' Y F i) () x 141 | coe (Σ' X R) Zero' () x 142 | coe (Σ' X R) One' () x 143 | coe (Σ' X R) Two' () x 144 | coe (Σ' X XT) (Σ' Y YT) (TYEQ , TEQ) (x , y) 145 | = let x′ = coe X Y TYEQ x 146 | y′ = coe (XT x) (YT x′) (TEQ x x′ (coh X Y TYEQ x)) y 147 | in (x′ , y′) 148 | coe (Σ' X R) (Π' Y R₁) () x s 149 | coe (Σ' X R) (Tree' Y F i) () x 150 | coe (Π' X R) Zero' () x 151 | coe (Π' X R) One' () x 152 | coe (Π' X R) Two' () x 153 | coe (Π' X R) (Σ' Y R₁) () x 154 | coe (Π' XL XR) (Π' YL YR) (LEQ , REQ) f x′ 155 | = let x = coe YL XL LEQ x′ 156 | y = f x 157 | in coe (XR x) (YR x′) (REQ x′ x (coh YL XL LEQ x′)) y 158 | coe (Π' X R) (Tree' Y F i) () x 159 | coe (Tree' X F i) Zero' () x 160 | coe (Tree' X F i) One' () x 161 | coe (Tree' X F i) Two' () x 162 | coe (Tree' X F i) (Σ' Y R) () x 163 | coe (Tree' X F i) (Π' Y R) () x s 164 | coe (Tree' I F i) (Tree' I′ F′ i′) (_ , (TREQ , IEQ)) x = tcoe i i′ TREQ x 165 | where 166 | -- This is Conor's 167 | tcoe : (i : ⟦ I ⟧TU)(i′ : ⟦ I′ ⟧TU)(iq : ⟦ Eq I i I′ i′ ⟧TU) → ⟦ Tree' I F i ⟧TU → ⟦ Tree' I′ F′ i′ ⟧TU 168 | tcoe i i' ieq ⟨ s , k ⟩ = 169 | let (S , K) = F i ; (S' , K') = F′ i' 170 | (SQ , TQ) = IEQ i i' ieq 171 | s' = coe S S' SQ s ; sh = coh S S' SQ s 172 | (P , r) = K s ; (P' , r') = K' s' 173 | (STEQ , rq) = TQ s s' sh 174 | in ⟨ (s' , λ p' → 175 | let p = coe P' P STEQ p' 176 | peq = coh P' P STEQ p' 177 | in tcoe (r p) (r' p') (rq p' p peq) (k p)) ⟩ 178 | 179 | -- "If you look at the definition of EQ quite carefully, you will notice that we 180 | -- did not use all of the types in TU to express equations. There is never any 181 | -- choice about how to be equal, so we need never use Two; meanwhile, we can 182 | -- avoid expressing tree equality as itself a tree just by using structural 183 | -- recursion. As a result, the only constructor pattern matching coe need ever 184 | -- perform on proofs is on pairs, which is just sugar for the lazy use of 185 | -- projections. Correspondingly, the only way coercion of canonical values 186 | -- between canonical types can get stuck is if those types are conspicuously 187 | -- different. Although we postulated coherence, no computation which relies on 188 | -- it is strict in equality proofs, so it is no source of blockage. 189 | -- 190 | -- The only way a closed coercion can get stuck is if we can prove a false 191 | -- equation. The machinery works provided the theory is consistent, but we can 192 | -- prove no equations which do not also hold in extensional type theories which 193 | -- are known to be consistent. In general, we are free to assert consistent 194 | -- equations." 195 | postulate 196 | reflTU : (X : TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x X x ⟧TU 197 | 198 | {- Exercise 6.2: Try proving this; where do you get stuck? 199 | reflTU : (X : TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x X x ⟧TU 200 | reflTU X x = ? 201 | -} 202 | 203 | postulate 204 | -- Homogeneous equations between values are made useful just by asserting that 205 | -- predicates respect them. We recover the Leibniz property. 206 | RespTU : (X : TU) (P : ⟦ X ⟧TU → TU) 207 | (x x' : ⟦ X ⟧TU) → ⟦ Eq X x X x' ⟧TU → ⟦ P x ↔ P x' ⟧TU 208 | 209 | substTU : (X : TU) (P : ⟦ X ⟧TU → TU) 210 | (x x' : ⟦ X ⟧TU) → ⟦ Eq X x X x' ⟧TU → ⟦ P x ⟧TU → ⟦ P x' ⟧TU 211 | substTU X P x x' q = coe (P x) (P x') (RespTU X P x x' q) 212 | 213 | -- We can express the observation that all of our proofs belong to lazy types by 214 | -- splitting our universe into two Sorts, corresponding to sets and 215 | -- propositions, embedding the latter explicitly into the former with a new 216 | -- set-former, Prf'. 217 | 218 | data Sort : Set where set prop : Sort 219 | 220 | IsSet : Sort → Set 221 | IsSet set = One 222 | IsSet prop = Zero 223 | 224 | mutual 225 | data PU (u : Sort) : Set where 226 | Zero' One' : PU u 227 | Two' : {_ : IsSet u} → PU u 228 | Σ' : (S : PU u) (T : ⟦ S ⟧PU → PU u) → PU u 229 | Π' : (S : PU set) (T : ⟦ S ⟧PU → PU u) → PU u 230 | Tree' : {_ : IsSet u} 231 | (I : PU set) 232 | (F : ⟦ I ⟧PU → Σ (PU set) λ S → 233 | ⟦ S ⟧PU → Σ (PU set) λ P → 234 | ⟦ P ⟧PU → ⟦ I ⟧PU) 235 | (i : ⟦ I ⟧PU) → PU u 236 | Prf' : {_ : IsSet u} → PU prop → PU u 237 | 238 | ⟦_⟧PU : ∀ {u} → PU u → Set 239 | ⟦ Zero' ⟧PU = Zero 240 | ⟦ One' ⟧PU = One 241 | ⟦ Two' ⟧PU = Two 242 | ⟦ Σ' S T ⟧PU = Σ ⟦ S ⟧PU λ s → ⟦ T s ⟧PU 243 | ⟦ Π' S T ⟧PU = (s : ⟦ S ⟧PU) → ⟦ T s ⟧PU 244 | ⟦ Tree' I F i ⟧PU = ITree 245 | ((λ i → ⟦ fst (F i) ⟧PU) 246 | ◃ (λ i s → ⟦ fst (snd (F i) s) ⟧PU) 247 | $ (λ i s p → snd (snd (F i) s) p) 248 | ) i 249 | ⟦ Prf' P ⟧PU = ⟦ P ⟧PU 250 | 251 | 252 | -- Exercise 6.3 (observational propositional equality): Reconstruct the 253 | -- definition of observational equality in this more refined setting. Take 254 | -- equality of propositions to be mutual implication and equality of proofs to 255 | -- be trivial: after all, equality for proofs of the atomic Zero' and One' 256 | -- propositions are trivial. 257 | 258 | _∧_ : PU prop → PU prop → PU prop 259 | P ∧ Q = Σ' P λ _ → Q 260 | 261 | _⇒_ : PU prop → PU prop → PU prop 262 | P ⇒ Q = Π' (Prf' P) λ _ → Q 263 | 264 | mutual 265 | PEQ : (X Y : PU set) → PU prop × (⟦ X ⟧PU → ⟦ Y ⟧PU → PU prop) 266 | _⇔_ : PU set → PU set → PU prop 267 | X ⇔ Y = fst (PEQ X Y) 268 | 269 | PEq : (X : PU set) (x : ⟦ X ⟧PU) → (Y : PU set) (y : ⟦ Y ⟧PU) → PU prop 270 | PEq X x Y y = snd (PEQ X Y) x y 271 | 272 | PEQ (Prf' P) (Prf' Q) = ((P ⇒ Q) ∧ (Q ⇒ P)) , λ _ _ → One' 273 | -- more code goes here 274 | PEQ _ _ = Zero' , λ _ _ → One' 275 | -------------------------------------------------------------------------------- /Meta/Basics.agda: -------------------------------------------------------------------------------- 1 | module Meta.Basics where 2 | 3 | open import Agda.Primitive using (Level) 4 | 5 | id : ∀ {k}{X : Set k} → X → X 6 | id x = x 7 | 8 | const : ∀ {a b} {A : Set a} {B : Set b} → A → B → A 9 | const x _ = x 10 | 11 | _∘_ : ∀ {i j k} 12 | {A : Set i}{B : A → Set j}{C : (a : A) → B a → Set k} → 13 | (f : {a : A}(b : B a) → C a b) → 14 | (g : (a : A) → B a) → 15 | (a : A) → C a (g a) 16 | f ∘ g = λ a → f (g a) 17 | 18 | flip : ∀ {a b c} {A : Set a} {B : Set b} {C : A → B → Set c} → 19 | ((x : A) (y : B) → C x y) → ((y : B) (x : A) → C x y) 20 | flip f = λ y x → f x y 21 | 22 | infixr 9 _∘_ 23 | 24 | _$′_ : ∀ {a b} {A : Set a} {B : A → Set b} → (∀ x → B x) → ∀ x → B x 25 | f $′ x = f x 26 | 27 | infixr -20 _$′_ 28 | 29 | _∋_ : ∀ {a} (A : Set a) → A → A 30 | A ∋ x = x 31 | 32 | infixl 0 _∋_ 33 | 34 | data ℕ : Set where 35 | zero : ℕ 36 | suc : ℕ → ℕ 37 | 38 | {-# BUILTIN NATURAL ℕ #-} 39 | 40 | _+_ : ℕ → ℕ → ℕ 41 | zero + y = y 42 | suc x + y = suc (x + y) 43 | infixr 6 _+_ 44 | 45 | _*_ : ℕ → ℕ → ℕ 46 | zero * y = zero 47 | suc x * y = y + x * y 48 | infixr 7 _*_ 49 | 50 | _^_ : ℕ → ℕ → ℕ 51 | x ^ zero = 1 52 | x ^ suc n = (x ^ n) * x 53 | 54 | record Σ {l : Level}(S : Set l)(T : S → Set l) : Set l where 55 | constructor _,_ 56 | field 57 | fst : S 58 | snd : T fst 59 | open Σ public 60 | 61 | _×_ : {l : Level} → Set l → Set l → Set l 62 | _×_ S T = Σ S λ _ → T 63 | infixr 4 _,_ 64 | 65 | -- The induction principle for Σ types. 66 | vv_ : ∀ {k l}{S : Set k}{T : S → Set k}{P : Σ S T → Set l} → 67 | ((s : S)(t : T s) → P (s , t)) → -- Given a witness to the construction of P 68 | (p : Σ S T) → P p -- And a Σ, you get a p. 69 | (vv p) (s , t) = p s t 70 | infixr 1 vv_ 71 | 72 | record One {l : Level} : Set l where 73 | constructor <> 74 | open One public 75 | 76 | data _≃_ {l}{X : Set l}(x : X) : X → Set l where 77 | refl : x ≃ x 78 | infix 1 _≃_ 79 | {-# BUILTIN EQUALITY _≃_ #-} 80 | 81 | subst : ∀ {k l}{X : Set k}{s t : X} → s ≃ t → (P : X → Set l) → P s → P t 82 | subst refl P p = p 83 | 84 | _≐_ : ∀ {l}{S : Set l}{T : S → Set l} → (f g : (x : S) → T x) → Set l 85 | f ≐ g = ∀ x → f x ≃ g x 86 | infix 1 _≐_ 87 | 88 | begin_ : ∀ {l}{A : Set l}{x y : A} → x ≃ y → x ≃ y 89 | begin_ p = p 90 | 91 | _=[_⟩_ : ∀ {l}{X : Set l}(x : X){y z} → x ≃ y → y ≃ z → x ≃ z 92 | _ =[ refl ⟩ q = q 93 | 94 | _⟨_]=_ : ∀ {l}{X : Set l}(x : X){y z} → y ≃ x → y ≃ z → x ≃ z 95 | _ ⟨ refl ]= q = q 96 | 97 | _∎ : ∀ {l}{X : Set l}(x : X) → x ≃ x 98 | x ∎ = refl 99 | 100 | infix 3 _∎ 101 | infixr 2 _=[_⟩_ _⟨_]=_ 102 | infix 1 begin_ 103 | 104 | cong : ∀ {k l}{X : Set k}{Y : Set l}(f : X → Y){x y} → x ≃ y → f x ≃ f y 105 | cong f refl = refl 106 | 107 | sym : {X : Set} {s t : X} → s ≃ t → t ≃ s 108 | sym refl = refl 109 | 110 | trans : {X : Set} {r s t : X} → r ≃ s → s ≃ t → r ≃ t 111 | trans refl refl = refl 112 | 113 | data Two : Set where 114 | tt ff : Two 115 | 116 | _⟨?⟩_ : ∀ {l}{P : Two → Set l} → P tt → P ff → (b : Two) → P b 117 | (t ⟨?⟩ f) tt = t 118 | (t ⟨?⟩ f) ff = f 119 | 120 | _⊎_ : Set → Set → Set 121 | S ⊎ T = Σ Two (S ⟨?⟩ T) 122 | 123 | data Zero : Set where 124 | 125 | -- ex falso. 126 | magic : ∀ {l}{A : Set l} → Zero → A 127 | magic () 128 | 129 | Dec : Set → Set 130 | Dec X = X ⊎ (X → Zero) 131 | 132 | data List (X : Set) : Set where 133 | ⟨⟩ : List X 134 | _,_ : X → List X → List X 135 | 136 | {-# BUILTIN LIST List #-} 137 | {-# BUILTIN NIL ⟨⟩ #-} 138 | {-# BUILTIN CONS _,_ #-} 139 | 140 | length : {X : Set} → List X → ℕ 141 | length ⟨⟩ = zero 142 | length (x , xs) = suc (length xs) 143 | 144 | zip : {S T : Set} → List S → List T → List (S × T) 145 | zip ⟨⟩ ⟨⟩ = ⟨⟩ 146 | zip (x , xs) (y , ys) = (x , y) , zip xs ys 147 | zip _ _ = ⟨⟩ 148 | 149 | 150 | data _⁻¹_ { S T : Set }(f : S → T) : T → Set where 151 | from : (s : S) → f ⁻¹ f s 152 | 153 | 154 | +-assoc : ∀ a b c → a + (b + c) ≃ (a + b) + c -- hint: use cong 155 | +-assoc zero b c = refl 156 | +-assoc (suc a) b c = cong suc (+-assoc a b c) 157 | {- 158 | +-assoc zero b c = refl 159 | +-assoc (suc a) b c = cong suc (+-assoc a b c) 160 | -} 161 | -------------------------------------------------------------------------------- /Meta/Control/Applicative.agda: -------------------------------------------------------------------------------- 1 | module Meta.Control.Applicative where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | 6 | record Applicative (F : Set → Set) : Set₁ where 7 | infixl 4 _⍟_ 8 | field 9 | pure : ∀ { X } → X → F X 10 | _⍟_ : ∀ { S T } → F (S → T) → F S → F T 11 | 12 | applicativeEndoFunctor : EndoFunctor F 13 | applicativeEndoFunctor = record { map = _⍟_ ∘ pure } 14 | open Applicative {{...}} public 15 | 16 | record ApplicativeOKP F {{AF : Applicative F}} : Set₁ where 17 | field 18 | lawId : ∀ {X} (x : F X) → (pure {{AF}} id ⍟ x) ≃ x 19 | lawCo : ∀ {R S T} (f : F (S → T))(g : F (R → S))(r : F R) → (pure {{AF}} (λ f g → f ∘ g) ⍟ f ⍟ g ⍟ r) ≃ (f ⍟ (g ⍟ r)) 20 | lawHom : ∀ {S T} (f : S → T)(s : S) → (pure {{AF}} f ⍟ pure s) ≃ (pure (f s)) 21 | lawCom : ∀ {S T} (f : F (S → T))(s : S) → (f ⍟ pure s) ≃ (pure {{AF}} (λ f → f s) ⍟ f) 22 | 23 | applicativeEndoFunctorOKP : EndoFunctorOKP F {{applicativeEndoFunctor}} 24 | applicativeEndoFunctorOKP = record 25 | { endoFunctorId = lawId 26 | ; endoFunctorCo = λ f g r → 27 | begin 28 | pure {{AF}} f ⍟ (pure {{AF}} g ⍟ r) 29 | ⟨ lawCo (pure f) (pure g) r ]= 30 | pure {{AF}} (λ f g → f ∘ g) ⍟ pure f ⍟ pure g ⍟ r 31 | =[ cong (λ x → x ⍟ pure g ⍟ r) (lawHom (λ f g → f ∘ g) f) ⟩ 32 | pure {{AF}} (_∘_ f) ⍟ pure g ⍟ r 33 | =[ cong (λ x → x ⍟ r) (lawHom (_∘_ f) g) ⟩ 34 | pure {{AF}} (f ∘ g) ⍟ r 35 | ∎ 36 | } 37 | open ApplicativeOKP {{...}} public 38 | 39 | instance 40 | applicativeFun : ∀ { S } → Applicative λ X → S → X 41 | applicativeFun = record 42 | { pure = λ x _ → x 43 | ; _⍟_ = λ s f k → s k (f k) 44 | } 45 | 46 | applicativeId : Applicative id 47 | applicativeId = record 48 | { pure = id 49 | ; _⍟_ = id 50 | } 51 | 52 | applicativeComp : ∀ { F G } → Applicative F → Applicative G → Applicative (F ∘ G) 53 | applicativeComp {F} {G} f g = record 54 | { pure = λ z → Applicative.pure f (Applicative.pure g z) 55 | ; _⍟_ = λ k x → (f Applicative.⍟ (f Applicative.⍟ Applicative.pure f (Applicative._⍟_ g)) k) x 56 | } 57 | 58 | data Product (F G : Set → Set) (X : Set) : Set where 59 | prod : F X → G X → Product F G X 60 | 61 | applicativeProd : ∀ { F G } → Applicative F → Applicative G → Applicative (Product F G) 62 | applicativeProd { F } { G } aF aG = record 63 | { pure = λ z → prod (Applicative.pure aF z) (Applicative.pure aG z) 64 | ; _⍟_ = appProd 65 | } 66 | where 67 | appProd : ∀ { S T } → Product F G (S → T) → Product F G S → Product F G T 68 | appProd (prod f g) (prod x y) = prod ((aF Applicative.⍟ f) x) ((aG Applicative.⍟ g) y) 69 | 70 | _-:>_ : ∀ (F G : Set → Set) → Set₁ 71 | F -:> G = ∀ {X} → F X → G X 72 | 73 | record AppHom {F}{{AF : Applicative F}}{G}{{AG : Applicative G}} (k : F -:> G) : Set₁ where 74 | field 75 | respPure : ∀ {X} (x : X) → k (pure x) ≃ pure x 76 | resp-⍟ : ∀ {S T}(f : F (S → T))(s : F S) → k (f ⍟ s) ≃ k f ⍟ k s 77 | -------------------------------------------------------------------------------- /Meta/Control/Monad.agda: -------------------------------------------------------------------------------- 1 | module Meta.Control.Monad where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | open import Meta.Control.Applicative 6 | 7 | record Monad (F : Set → Set) : Set₁ where 8 | field 9 | return : ∀ { X } → X → F X 10 | _>>=_ : ∀ { S T } → F S → (S → F T) → F T 11 | monadApplicative : Applicative F 12 | monadApplicative = record 13 | { pure = return 14 | ; _⍟_ = λ ff fs → ff >>= λ f → fs >>= λ s → return (f s) 15 | } 16 | monadEndoFunctor : EndoFunctor F 17 | monadEndoFunctor = Applicative.applicativeEndoFunctor monadApplicative 18 | 19 | μᴹ : ∀ {X} → F (F X) → F X 20 | μᴹ x = x >>= id 21 | open Monad {{...}} public 22 | 23 | {- 24 | monadVec : { n : ℕ } → Monad λ X → Vec X n 25 | monadVec = record 26 | { return = vec 27 | ; _>>=_ = {! !} 28 | } 29 | -} 30 | -------------------------------------------------------------------------------- /Meta/Control/Traversable.agda: -------------------------------------------------------------------------------- 1 | module Meta.Control.Traversable where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | open import Meta.Control.Applicative 6 | open import Meta.Data.Monoid 7 | open import Meta.Data.Vector 8 | 9 | record Traversable (F : Set → Set) : Set₁ where 10 | field 11 | traverse : ∀ { G S T }{{ AG : Applicative G }} → (S → G T) → F S → G (F T) 12 | 13 | endofunctor : EndoFunctor F 14 | endofunctor = MkEndo traverse 15 | open Traversable {{...}} public 16 | 17 | crush : ∀ { F X Y }{{ TF : Traversable F }}{{ M : Monoid Y }} → (X → Y) → F X → Y 18 | crush {{ M = M }} = traverse { T = One }{{ AG = Meta.Data.Monoid.monoidApplicative {{ M }} }} where 19 | open Monoid M 20 | 21 | instance 22 | traversableId : Traversable id 23 | traversableId = record { traverse = id } 24 | 25 | traversableVec : ∀ {n} → Traversable (flip Vec n) 26 | traversableVec = record { traverse = vtr } 27 | where 28 | vtr : ∀ {A : Set}{n B F}⦃ _ : Applicative F ⦄ → (A → F B) → Vec A n → F (Vec B n) 29 | vtr ⦃ aF ⦄ f ⟨⟩ = pure ⦃ aF ⦄ ⟨⟩ 30 | vtr ⦃ aF ⦄ f (a , as) = pure ⦃ aF ⦄ _,_ ⍟ f a ⍟ vtr ⦃ aF ⦄ f as 31 | 32 | traversableComp : ∀ { F G } → Traversable F → Traversable G → Traversable (F ∘ G) 33 | traversableComp {F} {G} f g = record { traverse = traverse {{f}} ∘ traverse {{g}} } 34 | 35 | {- 36 | transpose : ∀ {m n}{X : Set} → Vec (Vec X n) m → Vec (Vec X m) n 37 | transpose = Traversable.traverse id 38 | -} 39 | 40 | {- 41 | record TraversableOKP F {{TF : Traversable F}} : Set₁ where 42 | field 43 | lawPId : ∀ {X} (xs : F X) → traverse {{TF}} id xs ≃ xs 44 | lawPCo : 45 | ∀ {G} {{AG : Applicative G}}{H}{{AH : Applicative H}} 46 | {R S T} (g : S → G T)(h : R → H S)(rs : F R) 47 | → let EH = EndoFunctor H ∋ applicativeEndoFunctor 48 | in map {H} (traverse {{TF}} g) (traverse {{TF}} h rs) 49 | ≃ 50 | traverse {{TF}} {{applicativeComp AH AG}} (map {H} g ∘ h) rs 51 | 52 | lawPHom : ∀ {G}{{AG : Applicative G}}{H}{{AH : Applicative H}} 53 | (h : G -:> H){S T}(g : S → G T) → AppHom h → 54 | (ss : F S) → 55 | traverse (h ∘ g) ss ≃ h (traverse g ss) 56 | open TraversableOKP {{...}} public 57 | -} 58 | -------------------------------------------------------------------------------- /Meta/Data/EndoFunctor.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.EndoFunctor where 2 | 3 | open import Meta.Basics 4 | 5 | record EndoFunctor (F : Set → Set) : Set₁ where 6 | constructor MkEndo 7 | field 8 | map : ∀ { S T } → (S → T) → F S → F T 9 | open EndoFunctor {{...}} public 10 | 11 | {- 12 | record EndoFunctorOK F {{FF : EndoFunctor F}} : Set₁ where 13 | field 14 | endoFunctorId : ∀ {X} → map {{FF}}{X} id ≃ id 15 | endoFunctorCo : ∀ {R S T} → (f : S → T)(g : R → S) → map {{FF}} f ∘ map g ≃ map (f ∘ g) 16 | open EndoFunctorOK {{...}} public 17 | -} 18 | 19 | record EndoFunctorOKP F {{FF : EndoFunctor F}} : Set₁ where 20 | field 21 | endoFunctorId : ∀ {X} → map {{FF}}{X} id ≐ id 22 | endoFunctorCo : ∀ {R S T} → (f : S → T)(g : R → S) → map {{FF}} f ∘ map g ≐ map (f ∘ g) 23 | open EndoFunctorOKP {{...}} public 24 | -------------------------------------------------------------------------------- /Meta/Data/Fin.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Fin where 2 | 3 | open import Meta.Basics 4 | 5 | data Fin : ℕ → Set where 6 | zero : {n : ℕ} → Fin (suc n) 7 | suc : {n : ℕ} → Fin n → Fin (suc n) 8 | -------------------------------------------------------------------------------- /Meta/Data/Functor/Container.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Functor.Container where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | 6 | record Con : Set₁ where 7 | constructor _◃_ 8 | field 9 | Sh : Set 10 | Po : Sh → Set 11 | ⟦_⟧◃ : Set → Set 12 | ⟦_⟧◃ X = Σ Sh λ s → Po s → X 13 | open Con public 14 | infixr 1 _◃_ 15 | 16 | K◃ : Set → Con 17 | K◃ A = A ◃ λ _ → Zero 18 | 19 | I◃ : Con 20 | I◃ = One ◃ λ _ → One 21 | 22 | _+◃_ : Con → Con → Con 23 | (S ◃ P) +◃ (S₁ ◃ P₁) = (S ⊎ S₁) ◃ vv P ⟨?⟩ P₁ 24 | 25 | _×◃_ : Con → Con → Con 26 | (S ◃ P) ×◃ (S₁ ◃ P₁) = (S × S₁) ◃ vv λ s s₁ → P s ⊎ P₁ s₁ 27 | 28 | Σ◃ : (A : Set)(C : A → Con) → Con 29 | Σ◃ A C = (Σ A λ a → Sh (C a)) ◃ vv λ a s → Po (C a) s 30 | 31 | Π◃ : (A : Set)(C : A → Con) → Con 32 | Π◃ A C = ((a : A) → Sh (C a)) ◃ (λ f → Σ A λ a → Po (C a) (f a)) 33 | 34 | _∘◃_ : Con → Con → Con 35 | (S ◃ P) ∘◃ C = Σ◃ S λ s → Π◃ (P s) λ p → C 36 | 37 | conEndoFunctor : {C : Con} → EndoFunctor ⟦ C ⟧◃ 38 | conEndoFunctor {S ◃ P} = record { map = λ f x → (fst x) , (λ z → f (snd x z)) } 39 | 40 | conEndoFunctorOKP : {C : Con} → EndoFunctorOKP ⟦ C ⟧◃ {{conEndoFunctor}} 41 | conEndoFunctorOKP {S ◃ P} = record 42 | { endoFunctorId = λ x → refl 43 | ; endoFunctorCo = λ f g x → refl 44 | } 45 | 46 | _⟶◃_ : Con → Con → Set 47 | (S ◃ P) ⟶◃ (S₁ ◃ P₁) = Σ (S → S₁) λ f → (s : S) → P₁ (f s) → P s 48 | 49 | _/◃_ : ∀ {C C₁} → C ⟶◃ C₁ → ∀ {X} → ⟦ C ⟧◃ X → ⟦ C₁ ⟧◃ X 50 | (to , fro) /◃ (s , k) = to s , k ∘ fro s 51 | 52 | morph◃ : ∀ {C C₁} → (∀ {X} → ⟦ C ⟧◃ X → ⟦ C₁ ⟧◃ X) → C ⟶◃ C₁ 53 | morph◃ f = (λ x → fst (f (x , id))) , (λ s → snd (f (s , id))) 54 | 55 | id⟶◃ : ∀ {C} → C ⟶◃ C 56 | id⟶◃ = id , (λ _ x → x) 57 | 58 | _∘⟶◃_ : ∀ {C D E} → (D ⟶◃ E) → (C ⟶◃ D) → (C ⟶◃ E) 59 | (toe , froe) ∘⟶◃ (tod , frod) = toe ∘ tod , (λ s z → frod s (froe (tod s) z)) 60 | -------------------------------------------------------------------------------- /Meta/Data/Functor/Container/Indexed.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Functor.Container.Indexed where 2 | 3 | open import Agda.Primitive 4 | open import Meta.Basics 5 | open import Meta.Language.LambdaCalculus 6 | 7 | record _▷_ (I J : Set) : Set₁ where 8 | constructor _◃_$_ 9 | field 10 | Shlx : J → Set 11 | Polx : (j : J) → Shlx j → Set 12 | rilx : (j : J) (s : Shlx j) (p : Polx j s) → I 13 | ⟦_⟧ᵢ : (I → Set) → (J → Set) 14 | ⟦_⟧ᵢ X j = Σ (Shlx j) λ s → (p : Polx j s) → X (rilx j s p) 15 | open _▷_ public 16 | 17 | _-:>_ : ∀ {k l}{I : Set k} → (I → Set l) → (I → Set l) → Set (l ⊔ k) 18 | X -:> Y = ∀ i → X i → Y i 19 | 20 | ixMap : ∀ {I J}{C : I ▷ J}{X Y} → (X -:> Y) → ⟦ C ⟧ᵢ X -:> ⟦ C ⟧ᵢ Y 21 | ixMap f j (s , k) = s , λ p → f _ (k p) 22 | -------------------------------------------------------------------------------- /Meta/Data/Functor/Normal.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Functor.Normal where 2 | 3 | open import Meta.Basics 4 | 5 | open import Meta.Data.Vector 6 | open import Meta.Data.Monoid 7 | open import Meta.Data.Fin 8 | 9 | open import Meta.Control.Applicative 10 | open import Meta.Control.Traversable 11 | 12 | record Normal : Set₁ where 13 | constructor _/_ 14 | field 15 | Shape : Set 16 | size : Shape → ℕ 17 | ⟦_⟧ℕ : Set → Set 18 | ⟦_⟧ℕ X = Σ Shape λ s → Vec X (size s) 19 | open Normal public 20 | infixr 0 _/_ 21 | 22 | VecN : ℕ → Normal 23 | VecN n = One / const n 24 | 25 | ListN : Normal 26 | ListN = ℕ / id 27 | 28 | Kℕ : Set → Normal 29 | Kℕ A = A / const 0 30 | 31 | IKℕ : Normal 32 | IKℕ = VecN 1 33 | 34 | _+N_ : Normal → Normal → Normal 35 | (ShF / szF) +N (ShG / szG) = (ShF ⊎ ShG) / vv szF ⟨?⟩ szG 36 | 37 | _×N_ : Normal → Normal → Normal 38 | (ShF / szF) ×N (ShG / szG) = (ShF × ShG) / vv λ f g → szF f + szG g 39 | 40 | nInj : ∀ { X } (F G : Normal) → ⟦ F ⟧ℕ X ⊎ ⟦ G ⟧ℕ X → ⟦ F +N G ⟧ℕ X 41 | nInj F G (tt , (ShF , xs)) = (tt , ShF) , xs 42 | nInj F G (ff , (ShG , xs)) = (ff , ShG) , xs 43 | 44 | data _^-1_ { S T : Set }(f : S → T) : T → Set where 45 | from : (s : S) → f ^-1 f s 46 | 47 | nCase : ∀ { X } F G (s :  ⟦ F +N G ⟧ℕ X) → nInj F G ^-1 s 48 | nCase F G ((tt , ShF) , xs) = from (tt , (ShF , xs)) 49 | nCase F G ((ff , ShG) , xs) = from (ff , (ShG , xs)) 50 | 51 | nOut : ∀ { X }(F G : Normal) → ⟦ F +N G ⟧ℕ X → ⟦ F ⟧ℕ X ⊎ ⟦ G ⟧ℕ X 52 | nOut F G xs' with nCase F G xs' 53 | nOut F G . (nInj F G xs) | from xs = xs 54 | 55 | nPair : ∀ { X }(F G : Normal) → ⟦ F ⟧ℕ X × ⟦ G ⟧ℕ X → ⟦ F ×N G ⟧ℕ X 56 | nPair F G ((ShFx , fs) , (ShGx , gs)) = (ShFx , ShGx) , (fs ++ gs) 57 | 58 | concatSurjectivity : ∀ {m n : ℕ} {X} → (x : Vec X (m + n)) → (vv λ (u : Vec X m) (v : Vec X n) → u ++ v) ⁻¹ x 59 | concatSurjectivity {zero} v = from (? , v) 60 | concatSurjectivity {suc m} (x , v) with concatSurjectivity {m} v 61 | concatSurjectivity {suc m} (x , .(u ++ w)) | from (u , w) = from ((x , u) , w) 62 | 63 | 64 | unAppend : ∀ m n {X} → (xsys : Vec X (m + n)) → (vv _++_ {m}{n}) ^-1 xsys 65 | unAppend zero n ys = from (⟨⟩ , ys) 66 | unAppend (suc m) n (x , xsys) with unAppend m n xsys 67 | unAppend (suc m) n (x , .(xs ++ ys)) | from (xs , ys) = from ((x , xs) , ys) 68 | 69 | listNMonoid : {X : Set} → Monoid (⟦ ListN ⟧ℕ X) 70 | listNMonoid = record 71 | { ε = (zero , ⟨⟩) 72 | ; _•_ = _++N_ 73 | } where 74 | _++N_ : {X : Set} → ⟦ ListN ⟧ℕ X → ⟦ ListN ⟧ℕ X → ⟦ ListN ⟧ℕ X 75 | (n , xs) ++N (m , ys) = (n + m) , (xs ++ ys) 76 | 77 | {- 78 | listNMonoidOK : {X : Set} → MonoidOK (⟦ ListN ⟧ℕ X) {{listNMonoid}} 79 | listNMonoidOK {X} = record 80 | { absorbL = λ _ → refl 81 | ; absorbR = vv ++<> 82 | ; assoc = λ _ _ _ → refl 83 | } where 84 | ++<> : {M : Monoid (⟦ ListN ⟧ℕ X)} → (n : ℕ) (xs : Vec X n) → (n , xs) • (zero , ⟨⟩) ≃ (n , xs) 85 | ++<> zero ⟨⟩ = refl 86 | ++<> (suc n) (x , xs) = cong (vv λ m ys → suc m , x , ys) (++<> n xs) 87 | -} 88 | 89 | normalTraversable : (F : Normal) → Traversable ⟦ F ⟧ℕ 90 | normalTraversable F = record 91 | { traverse = λ {{aG}} f → vv λ s xs → pure {{aG}} (_,_ s) ⍟ traverse {{traversableVec}} f xs } 92 | 93 | _◦N_ : Normal → Normal → Normal 94 | F ◦N (ShG / szG) = ⟦ F ⟧ℕ ShG / crush {{normalTraversable F}}{{sumMonoid}} szG 95 | 96 | sizeT : ∀ {F}{{TF : Traversable F}}{X} → F X → ℕ 97 | sizeT {F}{{TF}} = crush {{TF}}{{sumMonoid}} (λ _ → 1) 98 | 99 | normalT : ∀ F {{TF : Traversable F}} → Normal 100 | normalT F = F One / sizeT 101 | 102 | shapeT : ∀ {F}{{TF : Traversable F}}{X} → F X → F One 103 | shapeT = {!!} 104 | -- shapeT {F}{{TF}}{{AG}} xs = traverse {{TF}}{F}{{AG}} ? 105 | 106 | one : ∀ {X} → X → ⟦ ListN ⟧ℕ X 107 | one x = 1 , (x , ⟨⟩) 108 | 109 | contentsT : ∀ {F}{{TF : Traversable F}}{X} → F X → ⟦ ListN ⟧ℕ X 110 | contentsT {F} {{TF}} {X} = crush {{TF}}{{listNMonoid}} one 111 | 112 | _⟶N_ : Normal → Normal → Set 113 | F ⟶N G = (s : Shape F) → ⟦ G ⟧ℕ (Fin (size F s)) 114 | 115 | nMorph : ∀ {F G} → F ⟶N G → ∀ {X} → ⟦ F ⟧ℕ X → ⟦ G ⟧ℕ X 116 | nMorph f (s , xs) with f s 117 | nMorph f (x , xs) | (s' , is) = s' , vmap (proj xs) is 118 | 119 | morphN : ∀ {F G} → (∀ {X} → ⟦ F ⟧ℕ X → ⟦ G ⟧ℕ X) → F ⟶N G 120 | morphN f s = f (s , tabulate id) 121 | 122 | toNormal : ∀ {F}{{TF : Traversable F}}{X} → F X → ⟦ normalT F ⟧ℕ X 123 | toNormal fx = {!!} 124 | 125 | _⊗_ : Normal → Normal → Normal 126 | (ShF / szF) ⊗ (ShG / szG) = (ShF × ShG) / λ s → szF (fst s) * szG (snd s) 127 | 128 | fromMatrix : {m n : ℕ} {X : Set} → Vec (Vec X n) m → Vec X (m * n) 129 | fromMatrix ⟨⟩ = ⟨⟩ 130 | fromMatrix (v , vs) = {!!} 131 | 132 | unfromMatrix : (m n : ℕ) {X : Set} → (elts : Vec X (n * m)) → fromMatrix {n}{m}^-1 elts 133 | unfromMatrix _ _ _ = {!!} 134 | {- 135 | unfromMatrix m zero ⟨⟩ = from ⟨⟩ 136 | unfromMatrix m (suc n) elts with unAppend m (m * n) elts 137 | unfromMatrix m (suc n) .(col ++ rest) | from (col , rest) with unfromMatrix m n rest 138 | unfromMatrix m (suc n) .(col ++ fromMatrix cols) | from (col , .(fromMatrix cols)) | from cols = from (col , cols) 139 | -} 140 | 141 | toMatrix : {m n : ℕ} {X : Set} → Vec X (n * m) → Vec (Vec X m) n 142 | toMatrix {m} {n} elts with unfromMatrix m n elts 143 | toMatrix .(fromMatrix m) | from m = m 144 | 145 | swap : (F G : Normal) → (F ⊗ G) ⟶N (G ⊗ F) 146 | swap F G (ShF , ShG) = ((ShG , ShF) , fromMatrix {size G ShG} {size F ShF} (transpose (toMatrix (tabulate id)))) 147 | 148 | drop : (F G : Normal) → (F ⊗ G) ⟶N (F ◦N G) 149 | drop F G (ShF , ShG) = (ShF , (vec ShG)) , {!!} 150 | where 151 | k : {n m : ℕ} → n ≃ m → Vec (Fin m) n 152 | k {n} {.n} refl = tabulate id 153 | 154 | fstHom : ∀ {X} → MonoidHom {⟦ ListN ⟧ℕ X}{{listNMonoid}}{ℕ}{{sumMonoid}} fst 155 | fstHom = record 156 | { respε = refl 157 | ; resp• = λ _ _ → refl 158 | } 159 | 160 | Batch : Set → Set → Set 161 | Batch X Y = Σ ℕ λ n → Vec X n → Y 162 | 163 | ABatch : {X : Set} → Applicative (Batch X) 164 | ABatch {X} = 165 | record { 166 | pure = λ y → 0 , (λ x → y); 167 | _⍟_ = app 168 | } where 169 | app : {S T : Set} → Batch X (S → T) → Batch X S → Batch X T 170 | app (n , vs) (m , us) = (n + m) , (λ ws → splitAndApply ws) where 171 | splitAndApply : (ws : Vec _ (n + m)) → _ 172 | splitAndApply ws with concatSurjectivity {n} ws 173 | splitAndApply .(vi ++ ui) | from (vi , ui) = (vs vi) (us ui) 174 | 175 | unpack : ∀ {X} → Vec X 1 → X 176 | unpack (x , _) = x 177 | 178 | batcher : forall {F} {{TF : Traversable F}} → F One → forall {X} → Batch X (F X) 179 | batcher {F} {{TF}} sF {X} = traverse {F} {{TF}} {Batch X} {{ABatch}} (λ _ → 1 , unpack) sF 180 | 181 | open import Level 182 | 183 | coherence : ∀ {F} {{TF : Traversable F}} {X} → TraversableOKP F 184 | → (sF : F One) → 185 | fst (batcher {F} {{TF}} sF {X}) ≃ 186 | traverse {F} {{TF}} {λ _ → ℕ} {One} {One} {{monoidApplicative}} (λ _ → 1) sF 187 | coherence {F} {{TF}} {X} tokF u = 188 | fst (traverse TF (λ _ → 1 , unpack) u) 189 | ⟨ TraversableOKP.lawPHom {F} {TF} tokF {Batch X} {{ABatch}} {λ _ → ℕ} {{monoidApplicative}} fst (λ _ → 1 , unpack) 190 | (record { respPure = λ {X₁} x → refl 191 | ; resp-⍟ = λ f s → refl 192 | }) u ]= 193 | (traverse TF {λ _ → ℕ} {One {Level.zero}} {X} 194 | {{record { pure = λ {_} _ → 0; _<*>_ = λ {_} {_} → _+_ }}} 195 | (λ a → 1) u) 196 | =[ TraversableOKP.lawPCo {F} {TF} tokF 197 | {G = id} {{AG = applicativeId}} 198 | {H = λ _ → ℕ} {{AH = monoidApplicative}} 199 | (λ _ → <>) (λ _ → 1) u ⟩ 200 | (traverse TF {λ _ → ℕ} {One {Level.zero}} {One} 201 | {{record { pure = λ {_} _ → 0; _<*>_ = λ {_} {_} → _+_ }}} 202 | (λ a → 1) u) 203 | ∎ 204 | 205 | fromNormal : ∀ {F}{{TF : Traversable F}} → TraversableOKP F → 206 | ∀ {X} → ⟦ normalT F ⟧ℕ X → F X 207 | fromNormal {F} {{TF}} tokf {X} (sF , cF) with (coherence {F} {{TF}} {X} tokf sF) 208 | fromNormal {F} {{TF}} tokf {X} (sF , cF) | q with batcher {F} {{TF}} sF {X} 209 | fromNormal {F} {{TF}} tokf {X} (sF , cF) | q | n , csF = csF (subst (sym q) (λ u → Vec X u) cF) 210 | 211 | data Tree (N : Normal) : Set where 212 | ⟨_⟩ : ⟦ N ⟧ℕ (Tree N) → Tree N 213 | 214 | NatT : Normal 215 | NatT = Two / 0 ⟨?⟩ 1 216 | 217 | zeroT : Tree NatT 218 | zeroT = ⟨ tt , ⟨⟩ ⟩ 219 | 220 | sucT : Tree NatT → Tree NatT 221 | sucT n = ⟨ (ff , n) , ⟨⟩ ⟩ 222 | 223 | NatInd : ∀ {l} (P : Tree NatT → Set l) → 224 | P zeroT → 225 | ((n : Tree NatT) → P n → P (sucT n)) → 226 | (n : Tree NatT) → P n 227 | NatInd P z s ⟨ tt , ⟨⟩ ⟩ = z 228 | NatInd P z s ⟨ (ff , n) , ⟨⟩ ⟩ = s n (NatInd P z s n) 229 | 230 | All : ∀ {l X}(P : X → Set l) {n} → Vec X n → Set l 231 | All P ⟨⟩ = One 232 | All P (x , xs) = P x × All P xs 233 | 234 | induction : ∀ (N : Normal) {l} (P : Tree N → Set l) → 235 | ((s : Shape N)(ts : Vec (Tree N) (size N s)) → All P ts → P  ⟨ s , ts ⟩) → 236 | (t : Tree N) → P t 237 | induction N P p ⟨ s , ts ⟩ = p s ts (hyps ts) where 238 | hyps : ∀ {n} (ts : Vec (Tree N) n) → All P ts 239 | hyps ⟨⟩ = <> 240 | hyps (t , ts) = induction N P p t , hyps ts 241 | 242 | eq? : (N : Normal)(sheq? : (s s` : Shape N) → Dec (s ≃ s`)) → 243 | (t t` : Tree N) → Dec (t ≃ t`) 244 | eq? N sheq? t t' = tt , {!!} 245 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/Apocrypha/Roman.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.Apocrypha.Roman where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Functor.Container.Indexed 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container 7 | open import Meta.Data.Inductive.W 8 | 9 | record Roman (I J : Set) : Set₁ where 10 | constructor SPqr 11 | field 12 | S : Set 13 | P : S → Set 14 | q : S → J 15 | r : (s : S) → P s → I 16 | Plain : Con 17 | Plain = S ◃ P 18 | ⟦_⟧R : (I → Set) → (J → Set) 19 | ⟦_⟧R X j = Σ (Σ S λ s → q s ≃ j) (vv λ s _ → (p : P s) → X (r s p)) 20 | 21 | Plain = Roman.Plain 22 | ⟦_⟧R = Roman.⟦_⟧R 23 | 24 | 25 | FromRoman : ∀ {I J} → Roman I J → I ▷ J 26 | FromRoman (SPqr S P q r) = (λ j → Σ S λ s → q s ≃ j) 27 | ◃ (λ j → P ∘ fst) 28 | $ (λ f → r ∘ fst) 29 | 30 | onTheNose : ∀ {I J} (C : Roman I J) → ⟦ C ⟧R ≃ ⟦ FromRoman C ⟧ᵢ 31 | onTheNose C = refl 32 | 33 | ToRoman : ∀ {I J} → I ▷ J → Roman I J 34 | ToRoman {I}{J} (Shlx ◃ Polx $ rilx) = SPqr (Σ J Shlx) (vv Polx) fst (vv rilx) 35 | 36 | toRoman : ∀ {I J} (C : I ▷ J) → 37 | ∀ {X j} → ⟦ C ⟧ᵢ X j → ⟦ ToRoman C ⟧R X j 38 | toRoman C (fst , snd) = ((_ , fst) , refl) , snd 39 | 40 | fromRoman : ∀ {I J} (C : I ▷ J) → 41 | ∀ {X j} → ⟦ ToRoman C ⟧R X j → ⟦ C ⟧ᵢ X j 42 | fromRoman C (((j , snd) , refl) , snd₂) = snd , snd₂ 43 | 44 | toAndFromRoman : ∀ {I J} (C : I ▷ J){X j} 45 | → (∀ xs → toRoman C {X}{j} (fromRoman C {X}{j} xs) ≃ xs) 46 | × (∀ xs → fromRoman C {X}{j} (toRoman C {X}{j} xs) ≃ xs) 47 | toAndFromRoman C = (λ { (((_ , _) , refl) , _) → refl }) , (λ xs → refl) 48 | 49 | data RomanData {I} (C : Roman I I) : I → Set where 50 | _,_ : (s : Roman.S C) → ((p : Roman.P C s) → RomanData C (Roman.r C s p)) → RomanData C (Roman.q C s) 51 | 52 | ideology : ∀ {I} (C : Roman I I) → I → W (Plain C) → W (Plain C ×◃ K◃ I) 53 | ideology C i ⟨ s , k ⟩ = ⟨ ((s , i) , (λ { (tt , p) → ideology C (Roman.r C s p) (k p) ; (ff , ()) })) ⟩ 54 | 55 | phenomenology : ∀ {I} (C : Roman I I) → W (Plain C) → W (Plain C ×◃ K◃ I) 56 | phenomenology C ⟨ s , k ⟩ = ⟨ ((s , Roman.q C s) , (λ { (tt , snd) → phenomenology C (k snd) ; (ff , ()) })) ⟩ 57 | 58 | RomanW : ∀ {I} → Roman I I → I → Set 59 | RomanW C i = Σ (W (Plain C)) λ t → phenomenology C t ≃ ideology C i t 60 | 61 | fromRomanW : ∀ {I} (C : Roman I I) {i} → RomanW C i → RomanData C i 62 | fromRomanW C (t , good) = {!!} 63 | 64 | postulate 65 | extensionality : ∀ {S : Set}{T : S → Set} (f g : (s : S) → T s) → ((s : S) → f s ≃ g s) → f ≃ g 66 | 67 | toRomanW : ∀ {I} (C : Roman I I) {i} → RomanData C i → RomanW C i 68 | toRomanW C t = {!!} 69 | 70 | data _** {I : Set} (R : I × I → Set) : I × I → Set where 71 | ⟨⟩ : {i : I} → (R **) (i , i) 72 | _,_ : {i j k : I} → R (i , j) → (R **) (j , k) → (R **) (i , k) 73 | infix 1 _** 74 | 75 | NAT : Set 76 | NAT = (Loop **) _ where 77 | Loop : One × One → Set 78 | Loop _ = One 79 | 80 | one : ∀ {I}{R : I × I → Set} → R -:> (R **) 81 | one (fst , snd) x = x , ⟨⟩ 82 | 83 | join** : ∀ {I}{R : I × I → Set} → ((R **) **) -:> (R **) 84 | join** (_ , _) ⟨⟩ = ⟨⟩ 85 | join** (i , snd) (x , xs) = {!!} 86 | 87 | Pow : Set₁ → Set₁ 88 | Pow X = X → Set 89 | 90 | Fam : Set₁ → Set₁ 91 | Fam X = Σ Set λ I → I → X 92 | 93 | p2f : (Pow ∘ ⇑) -:> (Fam ∘ ⇑) 94 | p2f X P = (Σ X (P ∘ ↑)) , (↑ ∘ fst) 95 | 96 | f2p : (Fam ∘ ⇑) -:> (Pow ∘ ⇑) 97 | f2p X (I , f) (↑ x) = Σ I λ i → ↓ (f i) ≃ x 98 | 99 | _$p_ : ∀ {I J} → (J → I) → Pow I → Pow J 100 | f $p P = P ∘ f 101 | 102 | _$F_ : ∀ {I J} → (I → J) → Fam I → Fam J 103 | f $F (F , i) = F , (f ∘ i) 104 | 105 | ROMAN : Set → Set → Set1 106 | ROMAN I J = Σ (Fam (⇑ J)) λ { (S , q) → S → Fam (⇑ I) } 107 | 108 | HANCOCK : Set → Set → Set1 109 | HANCOCK I J = Σ (Pow (⇑ J)) λ S → Σ J (S ∘ ↑) → Fam (⇑ I) 110 | 111 | NOTTINGHAM : Set → Set → Set₁ 112 | NOTTINGHAM I J = Σ (Pow (⇑ J)) λ S → Σ J (S ∘ ↑) → Pow (⇑ I) 113 | 114 | HANCOCK' : Set → Set → Set₁ 115 | HANCOCK' I J = J → Fam (Fam (⇑ I)) 116 | 117 | NOTTINGHAM' : Set → Set → Set₁ 118 | NOTTINGHAM' I J = J → Fam (Pow (⇑ I)) 119 | 120 | 121 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/ITree.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.ITree where 2 | 3 | open import Agda.Primitive 4 | 5 | open import Meta.Basics 6 | open import Meta.Data.Functor.Container.Indexed 7 | open import Meta.Language.LambdaCalculus 8 | 9 | data ITree {J : Set} (C : J ▷ J)(j : J) : Set where 10 | ⟨_⟩ : ⟦ C ⟧ᵢ (ITree C) j → ITree C j 11 | 12 | NatC : One ▷ One 13 | NatC = (λ _ → Two) ◃ (λ _ → Zero ⟨?⟩ One) $ _ 14 | 15 | zeroC : ITree NatC <> 16 | zeroC = ⟨ tt , magic ⟩ 17 | 18 | sucC : ITree NatC <> → ITree NatC <> 19 | sucC n = ⟨ ff , (λ _ → n) ⟩ 20 | 21 | VecC : Set → ℕ ▷ ℕ 22 | VecC X = VS ◃ VP $ Vr where 23 | VS : ℕ → Set 24 | VS zero = One 25 | VS (suc n) = X 26 | 27 | VP : (n : ℕ) → VS n → Set 28 | VP zero _ = Zero 29 | VP (suc n) _ = One 30 | 31 | Vr : (n : ℕ)(s : VS n)(p : VP n s) → ℕ 32 | Vr zero <> () 33 | Vr (suc n) x <> = n 34 | 35 | vnil' : ∀ {X} → ITree (VecC X) 0 36 | vnil' = ⟨ <> , (λ ()) ⟩ 37 | 38 | vcons' : ∀ {X n} → X → ITree (VecC X) n → ITree (VecC X) (suc n) 39 | vcons' x xs = ⟨ (x , (λ _ → xs)) ⟩ 40 | 41 | IsArr : ⋆ → Set 42 | IsArr ι = Zero 43 | IsArr (_ ▹ _) = One 44 | 45 | STLC : (Cx ⋆ × ⋆) ▷ (Cx ⋆ × ⋆) 46 | STLC = (vv λ Γ T → (T ∈ Γ) ⊎ (⋆ ⊎ IsArr T)) ◃ (λ { (Γ , τ) (tt , _) → Zero 47 | ; (Γ , τ) (ff , (tt , l)) → Two 48 | ; (Γ , τ) (ff , (ff , l)) → One 49 | }) $ 50 | (λ { (Γ , τ) (tt , _) () 51 | ; (Γ , τ) (ff , (tt , σ)) tt → Γ , (σ ▹ τ) 52 | ; (Γ , τ) (ff , (tt , σ)) ff → Γ , σ 53 | ; (Γ , ι) (ff , (ff , ())) _ 54 | ; (Γ , (σ ▹ τ)) (ff , (ff , p)) _ → (Γ , σ) , τ 55 | }) 56 | Idlx : ∀ {I} → I ▷ I 57 | Idlx = (λ _ → One) ◃ (λ j _ → One) $ (λ j _ _ → j) 58 | 59 | Colx : ∀ {I J K} → J ▷ K → I ▷ J → I ▷ K 60 | Colx (S ◃ P $ r) (S₁ ◃ P₁ $ r₁) = (λ x → Σ (S x) λ s → (p : P x s) → S₁ (r x s p)) 61 | ◃ (λ x → vv λ s f → Σ (P x s) λ p → P₁ (r x s p) (f p)) 62 | $ (λ { j (s , f) (p , p₁) → r₁ (r j s p) (f p) p₁ }) 63 | 64 | data Desc {l} (I : Set l) : Set (lsuc l) where 65 | var : I → Desc I 66 | σ π : (A : Set l) (D : A → Desc I) → Desc I 67 | _×D_ : Desc I → Desc I → Desc I 68 | κ : Set l → Desc I 69 | infixr 4 _×D_ 70 | 71 | ⟦_⟧D : ∀ {l}{I : Set l} → Desc I → (I → Set l) → Set l 72 | ⟦ (var i) ⟧D X = X i 73 | ⟦ (σ A D) ⟧D X = Σ A λ a → ⟦ D a ⟧D X 74 | ⟦ (π A D) ⟧D X = (a : A) → ⟦ D a ⟧D X 75 | ⟦ (D ×D E) ⟧D X = ⟦ D ⟧D X × ⟦ E ⟧D X 76 | ⟦ (κ A) ⟧D X = A 77 | 78 | ixConDesc : ∀ {I J} → I ▷ J → J → Desc I 79 | ixConDesc (S ◃ P $ r) j = σ (S j) λ s → π (P j s) λ p → var (r j s p) 80 | 81 | DSh : {I : Set} → Desc I → Set 82 | DSh (var x) = One 83 | DSh (σ A D) = Σ A λ a → DSh (D a) 84 | DSh (π A D) = (a : A) → DSh (D a) 85 | DSh (D ×D D₁) = DSh D × DSh D₁ 86 | DSh (κ A) = A 87 | 88 | DPo : ∀ {I} (D : Desc I) → DSh D → Set 89 | DPo (var x) x₁ = One 90 | DPo (σ A D) (x , y) = DPo (D x) y 91 | DPo (π A D) f = Σ A λ a → DPo (D a) (f a) 92 | DPo (D ×D D₁) (x , y) = DPo D x ⊎ DPo D₁ y 93 | DPo (κ A) s = Zero 94 | 95 | Dri : ∀ {I}(D : Desc I)(s : DSh D) → DPo D s → I 96 | Dri (var x) s p = x 97 | Dri (σ A D) (x , y) p = Dri (D x) y p 98 | Dri (π A D) f (x , y) = Dri (D x) (f x) y 99 | Dri (D ×D D₁) (x , y) (tt , p) = Dri D x p 100 | Dri (D ×D D₁) (x , y) (ff , p) = Dri D₁ y p 101 | Dri (κ x) s () 102 | 103 | descIxCon : ∀ {I J} → (J → Desc I) → I ▷ J 104 | descIxCon F = (DSh ∘ F) ◃ (DPo ∘ F) $ (Dri ∘ F) 105 | 106 | {- 107 | vecD : Set → ℕ → Desc ℕ 108 | vecD X n = 109 | σ Two ( κ (n ≃ 0) 110 | ⟨?⟩ σ ℕ λ k → κ X ×D var k ×D κ (n ≃ ℕ.suc k) 111 | ) 112 | -} 113 | 114 | vecD : Set → ℕ → Desc ℕ 115 | vecD X zero = κ One 116 | vecD X (suc n) = κ X ×D var n 117 | 118 | data Data {l}{J : Set l} (F : J → Desc J)(j : J) : Set l where 119 | ⟨_⟩ : ⟦ F j ⟧D (Data F) → Data F j 120 | 121 | vnil : ∀ {X} → Data (vecD X) 0 122 | vnil = ⟨ <> ⟩ 123 | 124 | vcons : ∀ {X n} → X → Data (vecD X) n → Data (vecD X) (suc n) 125 | vcons x xs = ⟨ x , xs ⟩ 126 | 127 | record ⇑ {l}(X : Set l) : Set (lsuc l) where 128 | constructor ↑ 129 | field 130 | ↓ : X 131 | open ⇑ public 132 | 133 | {- 134 | data Desc⋆ {l} : Set l where 135 | varD σD πD ×D κD : Desc⋆ 136 | 137 | DescD : ∀ {l}(I : Set l) → One {lsuc l} → Desc (One {lsuc l}) 138 | DescD {l} I _ = Σ Desc⋆ (λ 139 | { varD → ? -- κD (⇑ I) 140 | ; σD → Σ (Set l) (λ A → π (⇑ A) λ _ → var <>) 141 | ; πD → Σ (Set l) (λ A → π (⇑ A) λ _ → var <>) 142 | ; ×D → ? 143 | ; κD → κD (Set l) 144 | }) 145 | -} 146 | 147 | Everywhere : ∀ {I J} (C : I ▷ J)(X : I → Set) → Σ I X ▷ Σ J (⟦ C ⟧ᵢ X) 148 | Everywhere (S ◃ P $ r) X 149 | = (λ _ → One) 150 | ◃ (λ {(j , (s , k)) _ → P j s }) 151 | $ (λ { (j , (s , k)) _ p → r j s p , k p }) 152 | 153 | allTrivial : ∀ {I J} (C : I ▷ J)(X : I → Set) jc → ⟦ Everywhere C X ⟧ᵢ (λ _ → One) jc 154 | allTrivial C X _ = <> , λ _ → <> 155 | 156 | Somewhere : ∀ {I J}(C : I ▷ J)(X : I → Set) → Σ I X ▷ Σ J (⟦ C ⟧ᵢ X) 157 | Somewhere (S ◃ P $ r) X 158 | = (λ { (j , (s , k)) → P j s}) 159 | ◃ (λ { (j , (s , k)) _ → One}) 160 | $ (λ { (j , (s , k)) p _ → r j s p , k p }) 161 | 162 | noMagic : ∀ {I J}(C : I ▷ J)(X : I → Set) jc → ⟦ Somewhere C X ⟧ᵢ (λ _ → Zero) jc → Zero 163 | noMagic C X _ (p , m) = m <> 164 | 165 | lookup : ∀ {I J}(C : I ▷ J)(X : I → Set) jc {Q R} → ⟦ Everywhere C X ⟧ᵢ Q jc → ⟦ Somewhere C X ⟧ᵢ R jc → ⟦ Somewhere C X ⟧ᵢ (λ ix → Q ix × R ix) jc 166 | lookup C X jc (_ , q) (p , r) = p , (λ _ → (q p) , (r <>)) 167 | 168 | treeInd : ∀ {I}(C : I ▷ I)(P : Σ I (ITree C) → Set) → 169 | (⟦ Everywhere C (ITree C) ⟧ᵢ P -:> 170 | (vv λ i ts → P (i , ⟨ ts ⟩))) → 171 | (i : I)(t : ITree C i) → P (i , t) 172 | treeInd C P m i ⟨ s , k ⟩ = m (i , (s , k)) (<> , λ p → treeInd C P m _ (k p)) 173 | 174 | treeFold : ∀ {I}(C : I ▷ I)(P : I → Set) → 175 | (⟦ C ⟧ᵢ P -:> P) → 176 | (ITree C -:> P) 177 | treeFold C P m i ⟨ s , k ⟩ = m i (s , λ p → treeFold C P m _ (k p)) 178 | 179 | Children : ∀ {I} (C : I ▷ I) → Σ I (ITree C) ▷ Σ I (ITree C) 180 | Children C = Colx (descIxCon (var ∘ buds)) (Everywhere C (ITree C)) 181 | where 182 | buds : ∀ {I}{C : I ▷ I} → Σ I (ITree C) → Σ I λ i → ⟦ C ⟧ᵢ (ITree C) i 183 | buds (i , ⟨ xs ⟩) = i , xs 184 | 185 | treeFoldInd : ∀ {I} (C : I ▷ I) P → 186 | (⟦ Children C ⟧ᵢ P -:> P) → 187 | ∀ it → P it 188 | treeFoldInd C P m (i , t) = treeFold (Children C) P m (i , t) (children C i t) 189 | where 190 | children : ∀ {I}(C : I ▷ I) i t → ITree (Children C) (i , t) 191 | children C i ⟨ s , k ⟩ = ⟨ _ , (vv (λ _ p → children C _ (k p))) ⟩ 192 | 193 | EverywhereD SomewhereD : {I : Set}(D : Desc I)(X : I → Set) → ⟦ D ⟧D X → Desc (Σ I X) 194 | EverywhereD (var x) X xs = var (x , xs) 195 | EverywhereD (σ A D) X (x , xs) = EverywhereD (D x) X xs 196 | EverywhereD (π A D) X f = π A λ a → EverywhereD (D a) X (f a) 197 | EverywhereD (D ×D D₁) X (x , xs) = EverywhereD D X x ×D EverywhereD D₁ X xs 198 | EverywhereD (κ x) X xs = κ One 199 | SomewhereD (var x) X xs = var (x , xs) 200 | SomewhereD (σ A D) X (x , xs) = SomewhereD (D x) X xs 201 | SomewhereD (π A D) X f = σ A λ a → SomewhereD (D a) X (f a) 202 | SomewhereD (D ×D D₁) X (x , xs) = σ Two (SomewhereD D X x ⟨?⟩ SomewhereD D₁ X xs) 203 | SomewhereD (κ x) X xs = κ Zero 204 | 205 | dataInd : ∀ {I : Set} (F : I → Desc I)(P : Σ I (Data F) → Set) → 206 | ((i : I)(ds : ⟦ F i ⟧D (Data F)) → 207 | ⟦ EverywhereD (F i) (Data F) ds ⟧D P → P (i , ⟨ ds ⟩)) → 208 | ∀ i d → P (i , d) 209 | dataInd F P m i ⟨ ds ⟩ = m i ds (lem (F i) ds) where 210 | lem : (D : Desc _)(ds : ⟦ D ⟧D (Data F)) → ⟦ EverywhereD D (Data F) ds ⟧D P 211 | lem (var x) d = dataInd F P m x d 212 | lem (σ A D) (l , r) = lem (D l) r 213 | lem (π A D) f a = lem (D a) (f a) 214 | lem (D ×D E) (l , r) = lem D l , lem E r 215 | lem (κ x) y = <> 216 | 217 | vecNodeIx : (One ⊎ ℕ) ▷ ℕ 218 | vecNodeIx = descIxCon {J = ℕ} λ 219 | { zero → κ One 220 | ; (suc n) → var (tt , <>) ×D var (ff , n) 221 | } 222 | 223 | μlx : ∀ {I J} → (I ⊎ J) ▷ J → I ▷ J 224 | μlx {I}{J} F = (ITree F₁ ∘ _,_ ff) ◃ (P₁ ∘ _,_ ff) $ (r₁ ∘ _,_ ff) where 225 | F₁ : (I ⊎ J) ▷ (I ⊎ J) 226 | F₁ = (vv (λ i → One) ⟨?⟩ Shlx F) 227 | ◃ (vv (λ _ _ → Zero) ⟨?⟩ Polx F) 228 | $ (vv (λ t s ()) ⟨?⟩ rilx F) 229 | P₁ : (x : I ⊎ J) → ITree F₁ x → Set 230 | P₁ (tt , i) _ = One 231 | P₁ (ff , j) ⟨ s , k ⟩ = Σ (Polx F j s) λ p → P₁ (rilx F j s p) (k p) 232 | 233 | r₁ : (x : I ⊎ J) (t : ITree F₁ x) → P₁ x t → I 234 | r₁ (tt , i) _ _ = i 235 | r₁ (ff , j) ⟨ s , k ⟩ (p , ps) = r₁ _ (k p) ps 236 | 237 | vecIx : One ▷ ℕ 238 | vecIx = μlx vecNodeIx 239 | 240 | Vec'' : Set → ℕ → Set 241 | Vec'' X = ⟦ vecIx ⟧ᵢ (λ _ → X) 242 | 243 | vnil'' : ∀ {X} → Vec'' X 0 244 | vnil'' = ⟨ (<> , λ ()) ⟩ , (vv λ ()) 245 | 246 | vcons'' : ∀ {X n} → X → Vec'' X n → Vec'' X (suc n) 247 | vcons'' x (s , k) 248 | = ⟨ _ , (λ { (tt , _) → ⟨ (_ , λ ()) ⟩ 249 | ; (ff , _) → s 250 | }) ⟩ 251 | , (λ { ((tt , _) , _) → x 252 | ; ((ff , _) , p) → k p 253 | }) 254 | 255 | data Descμ (I : Set) : Set₁ where 256 | var : I → Descμ I 257 | σ π : (A : Set)(D : A → Descμ I) → Descμ I 258 | _×D_ : Descμ I → Descμ I → Descμ I 259 | κ : Set → Descμ I 260 | μ : (J : Set) → (J → Descμ (I ⊎ J)) → J → Descμ I 261 | 262 | {- 263 | mutual 264 | ⟦_⟧Dμ : ∀ {I} → Descμ I → (I → Set) → Set 265 | ⟦ x ⟧Dμ X = ? 266 | 267 | data Dataμ {I J} (F : J → Descμ (I ⊎ J))(X : I → Set) (j : J) : Set where 268 | ⟨_⟩ : ⟦ F j ⟧Dμ (vv X ⟨?⟩ Dataμ F X) → Dataμ F X j 269 | -} 270 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/InductionRecursion.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.InductionRecursion where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Fin 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container.Indexed 7 | 8 | sum prod : (n : ℕ) → (Fin n → ℕ) → ℕ 9 | sum zero _ = 0 10 | sum (suc n) f = f zero + sum n (f ∘ suc) 11 | 12 | prod zero _ = 1 13 | prod (suc n) f = f zero * sum n (f ∘ suc) 14 | 15 | {- 16 | data FTy : Set where 17 | fin : ℕ → FTy 18 | σ π : (S : FTy) (T : Fin {!!} → FTy) → FTy 19 | -} 20 | 21 | {- 22 | mutual 23 | data FTy : Set where 24 | fin : ℕ → FTy 25 | σ π : (S : FTy) (T : Fin (# S) → FTy) → FTy 26 | 27 | # : FTy → ℕ 28 | # (fin x) = x 29 | # (σ S T) = sum (# S) λ s → # (T s) 30 | # (π S T) = prod (# S) λ s → # (T s) 31 | -} 32 | 33 | fog : ∀ {n} → Fin n → ℕ 34 | fog zero = zero 35 | fog (suc i) = suc (fog i) 36 | 37 | mutual 38 | data FTy : Set where 39 | fin : ℕ → FTy 40 | σ π : (S : FTy) (T : FEI S → FTy) → FTy 41 | 42 | FEI : FTy → Set 43 | FEI (fin x) = Fin x 44 | FEI (σ S T) = Σ (FEI S) λ s → FEI (T s) 45 | FEI (π S T) = (s : FEI S) → FEI (T s) 46 | 47 | module FRESHLIST (X : Set) (Xeq? : (x x₁ : X) → Dec (x ≃ x₁)) where 48 | mutual 49 | data FreshList : Set where 50 | [] : FreshList 51 | _,_ : (x : X)(xs : FreshList) {ok : Fresh x xs} → FreshList 52 | 53 | Fresh : X → FreshList → Set 54 | Fresh x [] = One 55 | Fresh x (x₁ , xs) with Xeq? x x₁ 56 | Fresh x (x₁ , xs) | tt , _ = Zero 57 | Fresh x (x₁ , xs) | ff , _ = Fresh x xs 58 | 59 | data RecR : Set₁ where 60 | ⟨⟩ : RecR 61 | _,_ : (A : Set) (R : A → RecR) → RecR 62 | 63 | ⟦_⟧RR : RecR → Set 64 | ⟦ ⟨⟩ ⟧RR = One 65 | ⟦ A , R ⟧RR = Σ A λ a → ⟦ R a ⟧RR 66 | 67 | sizeRR : (R : RecR) → ⟦ R ⟧RR → ℕ 68 | sizeRR ⟨⟩ r = 0 69 | sizeRR (A , R) (fst , snd) = suc (sizeRR (R fst) snd) 70 | 71 | TyRR : (R : RecR)(r : ⟦ R ⟧RR) → Fin (sizeRR R r) → Set 72 | TyRR ⟨⟩ _ () 73 | TyRR (A , R) (fst , snd) zero = A 74 | TyRR (A , R) (fst , snd) (suc i) = TyRR (R fst) snd i 75 | 76 | vaRR : (R : RecR)(r : ⟦ R ⟧RR)(i : Fin (sizeRR R r)) → TyRR R r i 77 | vaRR ⟨⟩ <> () 78 | vaRR (A , R) (fst , snd) zero = fst 79 | vaRR (A , R) (fst , snd) (suc i) = vaRR (R fst) snd i 80 | 81 | mutual 82 | data RecL : Set₁ where 83 | ε : RecL 84 | _,_ : (R : RecL)(A : ⟦ R ⟧RL → Set) → RecL 85 | 86 | ⟦_⟧RL : RecL → Set 87 | ⟦ ε ⟧RL = One 88 | ⟦ R , A ⟧RL = Σ ⟦ R ⟧RL A 89 | 90 | sizeRL : RecL → ℕ 91 | sizeRL ε = 0 92 | sizeRL (R , A) = suc (sizeRL R) 93 | 94 | TyRL : (R : RecL) → Fin (sizeRL R) → ⟦ R ⟧RL → Set 95 | TyRL ε () <> 96 | TyRL (R , A) zero (fst , snd) = A fst 97 | TyRL (R , A) (suc i) (fst , snd) = TyRL R i fst 98 | 99 | vaRL : (R : RecL)(i : Fin (sizeRL R))(r : ⟦ R ⟧RL) → TyRL R i r 100 | vaRL ε () <> 101 | vaRL (R , A) zero (fst , snd) = snd 102 | vaRL (R , A) (suc i) (fst , snd) = vaRL R i fst 103 | 104 | TruncRL : (R : RecL) → Fin (sizeRL R) → RecL 105 | TruncRL ε () 106 | TruncRL (R , A) zero = R 107 | TruncRL (R , A) (suc i) = TruncRL R i 108 | 109 | truncRL : (R : RecL)(i : Fin (sizeRL R)) → ⟦ R ⟧RL → ⟦ TruncRL R i ⟧RL 110 | truncRL ε () <> 111 | truncRL (R , A) zero (fst , snd) = fst 112 | truncRL (R , A) (suc i) (fst , snd) = truncRL R i fst 113 | 114 | data Manifest {A : Set} : A → Set where 115 | ⟨_⟩ : (a : A) → Manifest a 116 | 117 | mutual 118 | data RecM : ℕ → Set₁ where 119 | ε : RecM 0 120 | _,_ : {n : ℕ} (R : RecM n)(A : ⟦ R ⟧RM → Set) → RecM (suc n) 121 | _,_∋_ : {n : ℕ} (R : RecM n) (A : ⟦ R ⟧RM → Set)(a : (r : ⟦ R ⟧RM) → A r) → RecM (suc n) 122 | 123 | ⟦_⟧RM : {n : ℕ} → RecM n → Set 124 | ⟦ ε ⟧RM = One 125 | ⟦ R , A ⟧RM = Σ ⟦ R ⟧RM A 126 | ⟦ R , A ∋ a ⟧RM = Σ ⟦ R ⟧RM (Manifest ∘ a) 127 | 128 | TyRM : {n : ℕ} (R : RecM n) → Fin n → ⟦ R ⟧RM → Set 129 | TyRM ε () <> 130 | TyRM (R , A) zero (fst , snd) = A fst 131 | TyRM (R , A) (suc i) (fst , snd) = TyRM R i fst 132 | TyRM (R , A ∋ a) zero (fst , snd) = A fst 133 | TyRM (R , A ∋ a) (suc i) (fst , snd) = TyRM R i fst 134 | 135 | vaRM : {n : ℕ}(R : RecM n)(i : Fin n)(r : ⟦ R ⟧RM) → TyRM R i r 136 | vaRM ε () r 137 | vaRM (R , A) zero (fst , snd) = snd 138 | vaRM (R , A) (suc i) (fst , snd) = vaRM R i fst 139 | vaRM (R , A ∋ a) zero (fst , snd) = a fst 140 | vaRM (R , A ∋ a) (suc i) (fst , snd) = vaRM R i fst 141 | 142 | {- 143 | mutual 144 | data REx : {n m : ℕ} → RecM n → RecM m → Set₁ where 145 | ε : REx ε ε 146 | 147 | rfog : ∀ {n m}{R : RecM n}{R₁ : RecM m} (X : REx R R₁) → ⟦ R₁ ⟧RM → ⟦ R ⟧RM 148 | rfog ε <> = <> 149 | -} 150 | 151 | mutual 152 | data TU : Set where 153 | Zero' One' Two' : TU 154 | Σ' Π' : (S : TU)(R : ⟦ S ⟧TU → TU) → TU 155 | Tree' : (I : TU)(F : ⟦ I ⟧TU → Σ TU λ S → 156 | ⟦ S ⟧TU → Σ TU λ P → 157 | ⟦ P ⟧TU → ⟦ I ⟧TU) 158 | (i : ⟦ I ⟧TU) → TU 159 | ⟦_⟧TU : TU → Set 160 | ⟦ Zero' ⟧TU = Zero 161 | ⟦ One' ⟧TU = One 162 | ⟦ Two' ⟧TU = Two 163 | ⟦ Σ' S R ⟧TU = Σ ⟦ S ⟧TU λ s → ⟦ R s ⟧TU 164 | ⟦ Π' S R ⟧TU = (s : ⟦ S ⟧TU) → ⟦ R s ⟧TU 165 | ⟦ Tree' I F i ⟧TU = ITree 166 | ( (λ i → ⟦ fst (F i) ⟧TU) 167 | ◃ (λ i s → ⟦ fst (snd (F i) s) ⟧TU) 168 | $ (λ i s p → snd (snd (F i) s) p) 169 | ) i 170 | 171 | Pow : Set₁ → Set₁ 172 | Pow X = X → Set 173 | 174 | Fam : Set₁ → Set₁ 175 | Fam X = Σ Set λ I → I → X 176 | 177 | mutual 178 | data NU (X : Fam Set) : Set where 179 | U′ : NU X 180 | El′ : fst X → NU X 181 | Nat′ : NU X 182 | Π′ : (S : NU X) (T : ⟦ S ⟧NU → NU X) → NU X 183 | 184 | ⟦_⟧NU : ∀ {X} → NU X → Set 185 | ⟦_⟧NU { U , El } U′ = U 186 | ⟦_⟧NU { U , El } (El′ T) = El T 187 | ⟦ Nat′ ⟧NU = ℕ 188 | ⟦ Π′ S T ⟧NU = (s : ⟦ S ⟧NU) → ⟦ T s ⟧NU 189 | 190 | EMPTY : Fam Set 191 | EMPTY = Zero , λ () 192 | 193 | LEVEL : ℕ → Fam Set 194 | LEVEL zero = (NU EMPTY) , ⟦_⟧NU 195 | LEVEL (suc n) = (NU (LEVEL n)) , ⟦_⟧NU 196 | 197 | mutual 198 | data HU {n} (U : Fin n → Set) : Set where 199 | U′ : Fin n → HU U 200 | Nat′ : HU U 201 | Π′ : (S : HU U) (T : ⟦ S ⟧HU → HU U) → HU U 202 | ⟦_⟧HU : ∀ {n} {U : Fin n → Set} → HU U → Set 203 | ⟦_⟧HU {U = U} (U′ i) = U i 204 | ⟦ Nat′ ⟧HU = ℕ 205 | ⟦ Π′ S T ⟧HU = (s : ⟦ S ⟧HU) → ⟦ T s ⟧HU 206 | 207 | HPREDS : (n : ℕ) → Fin n → Set 208 | HPREDS zero () 209 | HPREDS (suc n) zero = HU (HPREDS n) 210 | HPREDS (suc n) (suc x) = HPREDS n x 211 | 212 | HSET : ℕ → Set 213 | HSET n = HU (HPREDS n) 214 | 215 | {- 216 | mutual 217 | data VV : Set where 218 | V′ : VV 219 | Π′ : (S : VV) (T : ⟦ S ⟧VV → VV) → VV 220 | ⟦_⟧VV : VV → Set 221 | ⟦ V′ ⟧VV = VV 222 | ⟦ Π′ S T ⟧VV = (s : ⟦ S ⟧VV) → ⟦ T s ⟧VV 223 | -} 224 | 225 | data DS (I J : Set₁) : Set₁ where 226 | ι : J → DS I J 227 | σ : (S : Set) (T : S → DS I J) → DS I J 228 | δ : (H : Set) (T : (H → I) → DS I J) → DS I J 229 | 230 | ⟦_⟧DS : ∀ {I J} → DS I J → Fam I → Fam J 231 | ⟦_⟧DS (ι x) x₁ 232 | = One 233 | , (λ {<> → x}) 234 | ⟦_⟧DS (σ S T) x 235 | = (Σ S λ s → fst (⟦ T s ⟧DS x)) 236 | , (λ { (s , t) → snd (⟦ T s ⟧DS x) t }) 237 | ⟦_⟧DS (δ H T) (X , xi) 238 | = (Σ (H → X) λ hx → fst (⟦ T (xi ∘ hx) ⟧DS (X , xi))) 239 | , (λ { (hx , t) → snd (⟦ T (xi ∘ hx) ⟧DS (X , xi)) t }) 240 | 241 | idDS : { I : Set₁ } → DS I I 242 | idDS = δ One λ f → ι (f <>) 243 | 244 | {- 245 | mutual 246 | data DataDS {I} (D : DS I I) : Set where 247 | ⟨_⟩ : fst (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) → DataDS D 248 | ⟦_⟧ds : {I : Set₁}{D : DS I I} → DataDS D → I 249 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = snd (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) ds 250 | -} 251 | 252 | mutual 253 | data DataDS {I} (D : DS I I) : Set where 254 | ⟨_⟩ : NoDS D D → DataDS D 255 | 256 | ⟦_⟧ds : {I : Set₁} {D : DS I I} → DataDS D → I 257 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = DeDS D D ds 258 | 259 | NoDS : {I : Set₁} (D D′ : DS I I) → Set 260 | NoDS D (ι i) = One 261 | NoDS D (σ S T) = Σ S λ s → NoDS D (T s) 262 | NoDS D (δ H T) = Σ (H → DataDS D) λ hd → NoDS D (T (λ h → ⟦ hd h ⟧ds)) 263 | 264 | DeDS : {I : Set₁} (D D′ : DS I I) → NoDS D D′ → I 265 | DeDS D (ι x) <> = x 266 | DeDS D (σ S T) (s , t) = DeDS D (T s) t 267 | DeDS D (δ H T) (hd , t) = DeDS D (T (λ h → ⟦ hd h ⟧ds)) t 268 | 269 | bindDS : ∀ {I J K} → DS I J → (J → DS I K) → DS I K 270 | bindDS (ι x) U = U x 271 | bindDS (σ S T) U = σ S (λ s → bindDS (T s) U) 272 | bindDS (δ H T) U = δ H (λ f → bindDS (T f) U) 273 | 274 | pairDS : ∀ {I J K} (T : DS I J) (U : J → DS I K) {X : Fam I} → 275 | (t : fst (⟦ T ⟧DS X)) (u : fst (⟦ U (snd (⟦ T ⟧DS X) t) ⟧DS X)) 276 | → fst (⟦ bindDS T U ⟧DS X) 277 | pairDS (ι x) U <> u = u 278 | pairDS (σ S T) U (s , t) u = s , pairDS (T s) U t u 279 | pairDS (δ H T) U {_ , d} (f , t) u = f , (pairDS (T (d ∘ f)) U t u) 280 | 281 | {- 282 | coDS : ∀ {I J K} → DS J K → DS I J → DS I K 283 | coDS E D = {!!} 284 | -} 285 | 286 | mutual 287 | data Irish (I : Set₁) : Set₁ where 288 | ι : Irish I 289 | κ : Set → Irish I 290 | π : (S : Set) (T : S → Irish I) → Irish I 291 | σ : (S : Irish I) (T : Info S → Irish I) → Irish I 292 | 293 | Info : ∀ {I} → Irish I → Set₁ 294 | Info {I} ι = I 295 | Info (κ A) = ⇑ A 296 | Info (π S T) = (s : S) → Info (T s) 297 | Info (σ S T) = Σ (Info S) λ s → Info (T s) 298 | 299 | ΠF : (S : Set) {J : S → Set₁} (T : (s : S) → Fam (J s)) → Fam ((s : S) → J s) 300 | ΠF S T = ((s : S) → fst (T s)) , (λ f s → snd (T s) (f s)) 301 | 302 | ΣF : {I : Set₁} (S : Fam I) {J : I → Set₁} (T : (i : I) → Fam (J i)) → Fam (Σ I J) 303 | ΣF S T = (Σ (fst S) (fst ∘ (T ∘ snd S))) , (λ { (s , t) → snd S s , snd (T (snd S s)) t }) 304 | 305 | Node : ∀ {I} (T : Irish I) → Fam I → Fam (Info T) 306 | Node ι X = X 307 | Node (κ A) X = A , ↑ 308 | Node (π S T) X = ΠF S λ s → Node (T s) X 309 | Node (σ S T) X = ΣF (Node S X) λ iS → Node (T iS) X 310 | 311 | IF : Set₁ → Set₁ → Set₁ 312 | IF I J = Σ (Irish I) λ T → Info T → J 313 | 314 | _$F_ : ∀ {I J} → (I → J) → Fam I → Fam J 315 | f $F (F , i) = F , (f ∘ i) 316 | 317 | ⟦_⟧IF : ∀ {I J} → IF I J → Fam I → Fam J 318 | ⟦ (T , d) ⟧IF X = d $F Node T X 319 | 320 | {- 321 | mutual 322 | data DataIF {I} (F : IF I I) : Set where 323 | ⟨_⟩ : fst (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) → DataIF F 324 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 325 | ⟦_⟧if {F = F} ⟨ ds ⟩ = snd (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) ds 326 | -} 327 | 328 | mutual 329 | data DataIF {I} (F : IF I I) : Set where 330 | ⟨_⟩ : NoIF F (fst F) → DataIF F 331 | 332 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 333 | ⟦_⟧if {F = F} ⟨ rs ⟩ = snd F (DelF F (fst F) rs) 334 | 335 | NoIF : ∀ {I} (F : IF I I) (T : Irish I) → Set 336 | NoIF F ι = DataIF F 337 | NoIF F (κ A) = A 338 | NoIF F (π S T) = (s : S) → NoIF F (T s) 339 | NoIF F (σ S T) = Σ (NoIF F S) λ s → NoIF F (T (DelF F S s)) 340 | 341 | DelF : ∀ {I} (F : IF I I) (T : Irish I) → NoIF F T → Info T 342 | DelF F ι r = ⟦ r ⟧if 343 | DelF F (κ A) r = ↑ r 344 | DelF F (π S T) f = λ s → DelF F (T s) (f s) 345 | DelF F (σ S T) (s , t) = let s′ = DelF F S s in s′ , DelF F (T s′) t 346 | 347 | DSIF : ∀ {I J} → DS I J → IF I J 348 | DSIF (ι x) = κ One , (λ _ → x) 349 | DSIF (σ S T) = (σ (κ S) (λ s → fst (DSIF (T (↓ s))))) , (λ { (s , t) → snd (DSIF (T (↓ s))) t }) 350 | DSIF (δ H T) = (σ (π H λ _ → ι) (λ f → fst (DSIF (T f)))) , (λ { (f , t) → snd (DSIF (T f)) t }) 351 | 352 | idIF : ∀ {I} → IF I I 353 | idIF = ι , id 354 | 355 | subIF : ∀ {I J} (T : Irish J) (F : IF I J) → IF I (Info T) 356 | subIF ι F = F 357 | subIF (κ A) F = κ A , (λ z → ↑ (↓ z)) 358 | subIF (π S T) F = (π S λ s → fst (subIF (T s) F)) , (λ f s → snd (subIF (T s) F) (f s)) 359 | subIF (σ S T) F with subIF S F 360 | ... | (SF , f) = (σ SF λ sf → fst (subIF (T (f sf)) F)) , (λ { (sf , tf) → f sf , snd (subIF (T (f sf)) F) tf }) 361 | 362 | colF : ∀ {I J K} → IF J K → IF I J → IF I K 363 | colF (S , x) F with subIF S F 364 | ... | TF , f = TF , (x ∘ f) 365 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/Jacobian.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.Jacobian where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Functor.Container.Indexed 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container 7 | open import Meta.Data.Inductive.W 8 | 9 | Jac : ∀ {I J} → I ▷ J → I ▷ (J × I) 10 | Jac (S ◃ P $ r) 11 | = (λ { (j , i) → Σ (S j) λ s → r j s ⁻¹ i }) 12 | ◃ (λ { (j , .(r j s p)) (s , from p) → P j s ─ p }) 13 | $ (λ { (j , .(r j s p)) (s , from p) (p' , _) → r j s p' }) 14 | 15 | plugIx : ∀ {I J} (C : I ▷ J) → ((j : J)(s : Shlx C j)(p p₁ : Polx C j s) → Dec (p ≃ p₁)) → 16 | ∀ {i j X} → ⟦ Jac C ⟧ᵢ X (j , i) → X i → ⟦ C ⟧ᵢ X j 17 | plugIx C eq? {X = X} ((s , from p) , k) x = {!!} 18 | 19 | Zipper : ∀ {I} → I ▷ I → (I × I) ▷ (I × I) 20 | Zipper {I} C = (λ { (fst , snd) → (fst ≃ snd) ⊎ Σ I λ ip → ⟦ Jac C ⟧ᵢ (ITree C) (ip , snd) }) 21 | ◃ (λ { _ (tt , _) → Zero ; _ (ff , _) → One }) 22 | $ (λ { _ (tt , snd) () ; (fst , snd) (ff , (fst₁ , _)) _ → fst , fst₁ }) 23 | 24 | zipOut : ∀ {I} (C : I ▷ I) {ir ih} → 25 | ((i : I)(s : Shlx C i)(p p₁ : Polx C i s) → Dec (p ≃ p₁)) → 26 | ITree (Zipper C) (ir , ih) → ITree C ih → ITree C ir 27 | zipOut C eq? ⟨ (tt , refl) , _ ⟩ t = t 28 | zipOut C eq? ⟨ (ff , (_ , c)) , cz ⟩ t = zipOut C eq? (cz <>) ⟨ plugIx C eq? c t ⟩ 29 | 30 | ▽ : {I : Set} → Desc I → I → Desc I 31 | ▽ (var x) h = κ (x ≃ h) 32 | ▽ (σ A D) h = σ A λ a → ▽ (D a) h 33 | ▽ (π A D) h = σ A λ a → ▽ (D a) h ×D π (A ─ a) λ { (a₁ , _) → D a₁ } 34 | ▽ (D ×D E) h = σ Two ((▽ D h ×D E) ⟨?⟩ (D ×D ▽ E h)) 35 | ▽ (κ x) h = κ Zero 36 | 37 | 38 | 39 | 40 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/Recursion.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.Recursion where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Fin 5 | open import Meta.Data.Inductive.ITree 6 | open import Meta.Data.Functor.Container.Indexed 7 | 8 | -- Finite sums are finite, so we can compute their sizes. 9 | sum prod : (n : ℕ) → (Fin n → ℕ) → ℕ 10 | sum zero _ = 0 11 | sum (suc n) f = f zero + sum n (f ∘ suc) 12 | 13 | -- Finite products are finite, so we can compute their sizes. 14 | prod zero _ = 1 15 | prod (suc n) f = f zero * sum n (f ∘ suc) 16 | 17 | {- 18 | data FTy : Set where 19 | fin : ℕ → FTy 20 | σ π : (S : FTy) (T : Fin {!!} → FTy) → FTy 21 | -} 22 | 23 | {- 24 | mutual 25 | data FTy : Set where 26 | fin : ℕ → FTy 27 | σ π : (S : FTy) (T : Fin (# S) → FTy) → FTy 28 | 29 | # : FTy → ℕ 30 | # (fin x) = x 31 | # (σ S T) = sum (# S) λ s → # (T s) 32 | # (π S T) = prod (# S) λ s → # (T s) 33 | 34 | -- A forgetful map from (Fin n) back to ℕ to illustrate induction-recursion. 35 | fog : ∀ {n} → Fin n → ℕ 36 | fog zero = zero 37 | fog (suc i) = suc (fog i) 38 | -} 39 | 40 | -- "Where an inductive definition tells us how to perform construction of data 41 | -- incrementally, induction-recursion tells us how to perform 42 | -- construction-with-interpretation incrementally." 43 | mutual 44 | -- Thus, Construction 45 | data FTy : Set where 46 | fin : ℕ → FTy 47 | σ π : (S : FTy) (T : FEI S → FTy) → FTy 48 | 49 | -- And interpretation. 50 | FEI : FTy → Set 51 | FEI (fin x) = Fin x 52 | FEI (σ S T) = Σ (FEI S) λ s → FEI (T s) 53 | FEI (π S T) = (s : FEI S) → FEI (T s) 54 | 55 | -- The above definition is an important step up in power. Rather than before 56 | -- where `FTy ∘ #` yielded an unstructured interpretation of a family of sets 57 | -- indexed by ℕ, this time `FTy ∘ FEI` yields a subset Types that have names in 58 | -- FTy while still being small enough to be a set. 59 | -- 60 | -- "IR stands for Incredible Ray that shrinks large sets to small encodings of 61 | -- subsets of them. 62 | 63 | -- Exercise 5.1 By means of a suitable choice of recursive interpretation, fill 64 | -- the ? with a condition which ensures that FreshLists have distinct elements. 65 | -- Try to make sure that, for any concrete FreshList, ok can be inferred 66 | -- trivially. 67 | 68 | module FRESHLIST (X : Set) (Xeq? : (x x₁ : X) → Dec (x ≃ x₁)) where 69 | mutual 70 | data FreshList : Set where 71 | [] : FreshList 72 | _,_ : (x : X)(xs : FreshList) {ok : Fresh x xs} → FreshList 73 | 74 | -- The distinctness (freshness) guarantee says that. 75 | Fresh : X → FreshList → Set 76 | Fresh x [] = One -- There's nothing to do with an x and the empty list. 77 | Fresh x (x₁ , xs) with Xeq? x x₁ -- Otherwise destruct a decidable eq on head. 78 | Fresh x (x₁ , xs) | tt , _ = Zero -- We have a match! It isn't fresh. 79 | Fresh x (x₁ , xs) | ff , _ = Fresh x xs -- Otherwise try again. 80 | 81 | -- "Randy Pollack identified the task of modelling record types as a key early 82 | -- use of induction-recursion , motivated to organise libraries for mathematical 83 | -- structure." 84 | -- In this case, we don't need the full power of the IR Ray because you can 85 | -- get there with Desc and right-nested Σ's. Essentially, the Σ's give us a way 86 | -- of creating a structure where later fields can depend on earlier fields. 87 | data RecR : Set₁ where 88 | ⟨⟩ : RecR 89 | _,_ : (A : Set) (R : A → RecR) → RecR 90 | 91 | -- Lower everything back down to Agda's level. 92 | ⟦_⟧RR : RecR → Set 93 | ⟦ ⟨⟩ ⟧RR = One 94 | ⟦ A , R ⟧RR = Σ A λ a → ⟦ R a ⟧RR 95 | 96 | -- Exercise 5.2: Show how to compute the size of a record, the define the 97 | -- projections, first of types, then of values. 98 | 99 | -- For the sizes of records: 100 | sizeRR : (R : RecR) → ⟦ R ⟧RR → ℕ 101 | sizeRR ⟨⟩ r = 0 -- If we have nothing we're done. 102 | sizeRR (A , R) (fld , rest) = suc (sizeRR (R fld) rest) -- Otherwise we add 1 field and try again. 103 | 104 | -- For types: 105 | TyRR : (R : RecR)(r : ⟦ R ⟧RR) → Fin (sizeRR R r) → Set 106 | TyRR ⟨⟩ _ () -- It makes no sense to project the type out of nothing. 107 | TyRR (A , R) (_ , _) zero = A -- If the structure has one field we just project the type of that thing. 108 | TyRR (A , R) (fld , rest) (suc i) = TyRR (R fld) rest i -- Else try again. 109 | 110 | -- For values: 111 | vaRR : (R : RecR)(r : ⟦ R ⟧RR)(i : Fin (sizeRR R r)) → TyRR R r i 112 | vaRR ⟨⟩ <> () -- It makes no sense to project a value out of nothing. 113 | vaRR (_ , _) (fld , rest) zero = fld -- If the structure has one field we just project the value out of that. 114 | vaRR (A , R) (fld , rest) (suc i) = vaRR (R fld) rest i -- Else try again. 115 | 116 | -- But if you want left-nesting of record types (i.e. Dependent Contexts), you 117 | -- do need the IR Ray. 118 | mutual 119 | data RecL : Set₁ where 120 | ε : RecL 121 | _,_ : (R : RecL)(A : ⟦ R ⟧RL → Set) → RecL 122 | 123 | ⟦_⟧RL : RecL → Set 124 | ⟦ ε ⟧RL = One 125 | ⟦ R , A ⟧RL = Σ ⟦ R ⟧RL A 126 | 127 | -- Exercise 5.3: Show how to compute the size of a RecL without knowing the 128 | -- individual record. Show how to interpret a projection as a function from a 129 | -- record, first for types, then values. 130 | sizeRL : RecL → ℕ 131 | sizeRL ε = 0 -- The size of the empty context is zero. 132 | sizeRL (R , A) = suc (sizeRL R) -- Else add one and try again. 133 | 134 | -- For Types: 135 | TyRL : (R : RecL) → Fin (sizeRL R) → ⟦ R ⟧RL → Set 136 | TyRL ε () <> -- It makes no sense to project a type out of the empty structure. 137 | TyRL (R , A) zero (idx , _) = A idx -- If the structure has one field we project it out by indexing with the field. 138 | TyRL (R , A) (suc i) (idx , _) = TyRL R i idx -- Else try again. 139 | 140 | -- For values: 141 | vaRL : (R : RecL)(i : Fin (sizeRL R))(r : ⟦ R ⟧RL) → TyRL R i r 142 | vaRL ε () <> -- It makes no sense to project a value out of the empty structure. 143 | vaRL (R , A) zero (_ , fld) = fld -- If the structure has one field just project the value out. 144 | vaRL (R , A) (suc i) (idx , snd) = vaRL R i idx 145 | 146 | -- Exercise 5.4: Show how to truncate a record signature from a given field and 147 | -- compute the corresponding projection on structures. 148 | 149 | -- To truncate a context: 150 | TruncRL : (R : RecL) → Fin (sizeRL R) → RecL 151 | TruncRL ε () -- It makes no sense to truncate the empty context. 152 | TruncRL (R , A) zero = R -- For a context with a single part, yield it. 153 | TruncRL (R , A) (suc i) = TruncRL R i -- Else try again. 154 | 155 | -- To project out of a truncation: 156 | truncRL : (R : RecL)(i : Fin (sizeRL R)) → ⟦ R ⟧RL → ⟦ TruncRL R i ⟧RL 157 | truncRL ε () <> -- How did we get here? 158 | truncRL (R , A) zero (fst , snd) = fst -- For a singleton context, give up the subpart. 159 | truncRL (R , A) (suc i) (fst , snd) = truncRL R i fst -- Else try again. 160 | 161 | -- A Manifest is a record whose values are computed from earlier fields. "It is 162 | -- rather like allowing definitions in contexts." 163 | data Manifest {A : Set} : A → Set where 164 | ⟨_⟩ : (a : A) → Manifest a 165 | 166 | mutual 167 | -- "I index by size, to save on measuring." 168 | data RecM : ℕ → Set₁ where 169 | ε : RecM 0 -- The empty manifest. 170 | _,_ : {n : ℕ} (R : RecM n)(A : ⟦ R ⟧RM → Set) → RecM (suc n) -- Old-school fields 171 | _,_∋_ : {n : ℕ} (R : RecM n) (A : ⟦ R ⟧RM → Set)(a : (r : ⟦ R ⟧RM) → A r) → RecM (suc n) -- A definition. 172 | 173 | ⟦_⟧RM : {n : ℕ} → RecM n → Set 174 | ⟦ ε ⟧RM = One 175 | ⟦ R , A ⟧RM = Σ ⟦ R ⟧RM A 176 | ⟦ R , A ∋ a ⟧RM = Σ ⟦ R ⟧RM (Manifest ∘ a) 177 | 178 | -- Exercise 5.5: Implement projection for RecM. 179 | 180 | -- For types: 181 | TyRM : {n : ℕ} (R : RecM n) → Fin n → ⟦ R ⟧RM → Set 182 | -- Makes no sense to project out of nothing. 183 | TyRM ε () <> 184 | -- If the structure has one field we project it out by indexing with the field 185 | TyRM (R , A) zero (idx , _) = A idx 186 | -- Else try again. 187 | TyRM (R , A) (suc i) (idx , _) = TyRM R i idx 188 | -- If the structure has one definition we project it out by indexing with the field. 189 | TyRM (R , A ∋ a) zero (idx , _) = A idx 190 | -- Else try again. 191 | TyRM (R , A ∋ a) (suc i) (idx , _) = TyRM R i idx 192 | 193 | -- For values: 194 | vaRM : {n : ℕ}(R : RecM n)(i : Fin n)(r : ⟦ R ⟧RM) → TyRM R i r 195 | -- Makes no sense to project out of nothing. 196 | vaRM ε () r 197 | -- If the structure has one field, project it out. 198 | vaRM (R , A) zero (_ , fld) = fld 199 | -- Else try again. 200 | vaRM (R , A) (suc i) (idx , _) = vaRM R i idx 201 | -- If the structure has one manifest field, project it out by feeding it what it depends on. 202 | vaRM (R , A ∋ a) zero (idx , _) = a idx 203 | -- Else try again. 204 | vaRM (R , A ∋ a) (suc i) (idx , _) = vaRM R i idx 205 | 206 | {- 207 | mutual 208 | data REx : {n m : ℕ} → RecM n → RecM m → Set₁ where 209 | ε : REx ε ε 210 | 211 | rfog : ∀ {n m}{R : RecM n}{R₁ : RecM m} (X : REx R R₁) → ⟦ R₁ ⟧RM → ⟦ R ⟧RM 212 | rfog ε <> = <> 213 | -} 214 | 215 | mutual 216 | -- Conor's Favorite Universe Featuring: 217 | data TU : Set where 218 | -- A scattering of base types. 219 | Zero' One' Two' : TU 220 | -- Dependent pairs (and functions). 221 | Σ' Π' : (S : TU)(R : ⟦ S ⟧TU → TU) → TU 222 | -- And Petersson-Synek Trees. 223 | Tree' : (I : TU)(F : ⟦ I ⟧TU → Σ TU λ S → 224 | ⟦ S ⟧TU → Σ TU λ P → 225 | ⟦ P ⟧TU → ⟦ I ⟧TU) 226 | (i : ⟦ I ⟧TU) → TU 227 | ⟦_⟧TU : TU → Set 228 | ⟦ Zero' ⟧TU = Zero 229 | ⟦ One' ⟧TU = One 230 | ⟦ Two' ⟧TU = Two 231 | ⟦ Σ' S R ⟧TU = Σ ⟦ S ⟧TU λ s → ⟦ R s ⟧TU 232 | ⟦ Π' S R ⟧TU = (s : ⟦ S ⟧TU) → ⟦ R s ⟧TU 233 | ⟦ Tree' I F i ⟧TU = ITree 234 | ( (λ i → ⟦ fst (F i) ⟧TU) 235 | ◃ (λ i s → ⟦ fst (snd (F i) s) ⟧TU) 236 | $ (λ i s p → snd (snd (F i) s) p) 237 | ) i 238 | 239 | -- Unfortunately, because this is an inductive-recursive structure in and of 240 | -- itself: you can't use the IR Ray on an IR Ray. 241 | 242 | Pow : Set₁ → Set₁ 243 | Pow X = X → Set 244 | 245 | Fam : Set₁ → Set₁ 246 | Fam X = Σ Set λ I → I → X 247 | 248 | mutual 249 | -- A predicative hierarchy of small universes built using induction-recursion. 250 | data NU (X : Fam Set) : Set where 251 | U′ : NU X 252 | El′ : fst X → NU X 253 | Nat′ : NU X 254 | Π′ : (S : NU X) (T : ⟦ S ⟧NU → NU X) → NU X 255 | 256 | ⟦_⟧NU : ∀ {X} → NU X → Set 257 | ⟦_⟧NU { U , El } U′ = U 258 | ⟦_⟧NU { U , El } (El′ T) = El T 259 | ⟦ Nat′ ⟧NU = ℕ 260 | ⟦ Π′ S T ⟧NU = (s : ⟦ S ⟧NU) → ⟦ T s ⟧NU 261 | 262 | -- Universes can be "jacked up" as far as we like. 263 | 264 | EMPTY : Fam Set 265 | EMPTY = Zero , λ () 266 | 267 | LEVEL : ℕ → Fam Set 268 | LEVEL zero = (NU EMPTY) , ⟦_⟧NU 269 | LEVEL (suc n) = (NU (LEVEL n)) , ⟦_⟧NU 270 | 271 | mutual 272 | -- To eliminate redudancy of lower universes, we parametrize the universe by a 273 | -- de Bruijin indexed collection of the previous universes. 274 | data HU {n} (U : Fin n → Set) : Set where 275 | U′ : Fin n → HU U 276 | Nat′ : HU U 277 | Π′ : (S : HU U) (T : ⟦ S ⟧HU → HU U) → HU U 278 | ⟦_⟧HU : ∀ {n} {U : Fin n → Set} → HU U → Set 279 | ⟦_⟧HU {U = U} (U′ i) = U i 280 | ⟦ Nat′ ⟧HU = ℕ 281 | ⟦ Π′ S T ⟧HU = (s : ⟦ S ⟧HU) → ⟦ T s ⟧HU 282 | 283 | -- "To finish the job, we must build the collections of levels to hand to HU. At 284 | -- each step, level zero is the new top level, built with a fresh appeal to HU, 285 | -- but lower levels can be projected from the previous collection." 286 | 287 | HPREDS : (n : ℕ) → Fin n → Set 288 | HPREDS zero () 289 | HPREDS (suc n) zero = HU (HPREDS n) 290 | HPREDS (suc n) (suc x) = HPREDS n x 291 | 292 | HSET : ℕ → Set 293 | HSET n = HU (HPREDS n) 294 | 295 | {- rejected. 296 | mutual 297 | data VV : Set where 298 | V′ : VV 299 | Π′ : (S : VV) (T : ⟦ S ⟧VV → VV) → VV 300 | ⟦_⟧VV : VV → Set 301 | ⟦ V′ ⟧VV = VV 302 | ⟦ Π′ S T ⟧VV = (s : ⟦ S ⟧VV) → ⟦ T s ⟧VV 303 | -} 304 | 305 | -- An encoding of allowed inductive-recursive things [Dybjer and Setzer 1999]. 306 | -- The encoding is as follows: 307 | -- ∘ Describe one node of inductive recursive types in the manner of RecR but 308 | -- where each J-value read gives us a way of reading off I values. 309 | data DS (I J : Set₁) : Set₁ where 310 | ι : J → DS I J -- no more fields 311 | σ : (S : Set) (T : S → DS I J) → DS I J -- ordinary field 312 | δ : (H : Set) (T : (H → I) → DS I J) → DS I J -- child field. 313 | 314 | -- A DS is a functor. 315 | -- 316 | -- "In each case, we must say which set is being encoded and how to read off a J 317 | -- from a value in that set. 318 | ⟦_⟧DS : ∀ {I J} → DS I J → Fam I → Fam J 319 | -- The ι constructor carries exactly the j required. The other two specify a 320 | -- field in the node structure, from which the computation of the J gains some 321 | -- information. 322 | ⟦_⟧DS (ι x) x₁ 323 | = One 324 | , (λ {<> → x}) 325 | -- The σ specifies a field of type S, and the rest of the structure may depend 326 | -- on a value of type S. 327 | ⟦_⟧DS (σ S T) x 328 | = (Σ S λ s → fst (⟦ T s ⟧DS x)) 329 | , (λ { (s , t) → snd (⟦ T s ⟧DS x) t }) 330 | -- The δ case is the clever bit. It specifies a place for an H-indexed bunch of 331 | -- children, and even through we do not fix what set represents the children, we 332 | -- do know that they allow us to read off an I. Correspondingly, the rest of 333 | -- the structure can at least depend on knowing a function in H → I which gives 334 | -- access to the interpretation of the children. Once we plugin a specific 335 | -- (X, xi) : Fam I, we represent the field by the small function space 336 | -- hx : H → X, then the composition xi ∘ hx tells us how to compute the large 337 | -- meaning of each child." 338 | ⟦_⟧DS (δ H T) (X , xi) 339 | = (Σ (H → X) λ hx → fst (⟦ T (xi ∘ hx) ⟧DS (X , xi))) 340 | , (λ { (hx , t) → snd (⟦ T (xi ∘ hx) ⟧DS (X , xi)) t }) 341 | 342 | -- Exercise 5.10: A morphism form (X, xi) to (Y, yi) in Fam I is a function 343 | -- f : X → Y such that xi = yi ∘ f. Construct a code for the identity functor 344 | -- on Fam I such that ⟦ idDS ⟧DS ≅ id. 345 | idDS : { I : Set₁ } → DS I I 346 | idDS = δ One λ f → ι (f <>) 347 | 348 | {- -- fails positivity check and termination check. 349 | mutual 350 | data DataDS {I} (D : DS I I) : Set where 351 | ⟨_⟩ : fst (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) → DataDS D 352 | ⟦_⟧ds : {I : Set₁}{D : DS I I} → DataDS D → I 353 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = snd (⟦ D ⟧DS (DataDS D , ⟦_⟧ds)) ds 354 | -} 355 | 356 | mutual 357 | -- We get out of this jam by inlining the interpretation. 358 | data DataDS {I} (D : DS I I) : Set where 359 | ⟨_⟩ : NoDS D D → DataDS D 360 | 361 | ⟦_⟧ds : {I : Set₁} {D : DS I I} → DataDS D → I 362 | ⟦_⟧ds {D = D} ⟨ ds ⟩ = DeDS D D ds 363 | 364 | NoDS : {I : Set₁} (D D′ : DS I I) → Set 365 | NoDS D (ι i) = One 366 | NoDS D (σ S T) = Σ S λ s → NoDS D (T s) 367 | NoDS D (δ H T) = Σ (H → DataDS D) λ hd → NoDS D (T (λ h → ⟦ hd h ⟧ds)) 368 | 369 | DeDS : {I : Set₁} (D D′ : DS I I) → NoDS D D′ → I 370 | DeDS D (ι x) <> = x 371 | DeDS D (σ S T) (s , t) = DeDS D (T s) t 372 | DeDS D (δ H T) (hd , t) = DeDS D (T (λ h → ⟦ hd h ⟧ds)) t 373 | 374 | bindDS : ∀ {I J K} → DS I J → (J → DS I K) → DS I K 375 | bindDS (ι x) U = U x 376 | bindDS (σ S T) U = σ S (λ s → bindDS (T s) U) 377 | bindDS (δ H T) U = δ H (λ f → bindDS (T f) U) 378 | 379 | pairDS : ∀ {I J K} (T : DS I J) (U : J → DS I K) {X : Fam I} → 380 | (t : fst (⟦ T ⟧DS X)) (u : fst (⟦ U (snd (⟦ T ⟧DS X) t) ⟧DS X)) 381 | → fst (⟦ bindDS T U ⟧DS X) 382 | pairDS (ι x) U <> u = u 383 | pairDS (σ S T) U (s , t) u = s , pairDS (T s) U t u 384 | pairDS (δ H T) U {_ , d} (f , t) u = f , (pairDS (T (d ∘ f)) U t u) 385 | 386 | {- 387 | coDS : ∀ {I J K} → DS J K → DS I J → DS I K 388 | coDS E D = {!!} 389 | -} 390 | 391 | mutual 392 | data Irish (I : Set₁) : Set₁ where 393 | ι : Irish I 394 | κ : Set → Irish I 395 | π : (S : Set) (T : S → Irish I) → Irish I 396 | σ : (S : Irish I) (T : Info S → Irish I) → Irish I 397 | 398 | Info : ∀ {I} → Irish I → Set₁ 399 | Info {I} ι = I 400 | Info (κ A) = ⇑ A 401 | Info (π S T) = (s : S) → Info (T s) 402 | Info (σ S T) = Σ (Info S) λ s → Info (T s) 403 | 404 | ΠF : (S : Set) {J : S → Set₁} (T : (s : S) → Fam (J s)) → Fam ((s : S) → J s) 405 | ΠF S T = ((s : S) → fst (T s)) , (λ f s → snd (T s) (f s)) 406 | 407 | ΣF : {I : Set₁} (S : Fam I) {J : I → Set₁} (T : (i : I) → Fam (J i)) → Fam (Σ I J) 408 | ΣF S T = (Σ (fst S) (fst ∘ (T ∘ snd S))) , (λ { (s , t) → snd S s , snd (T (snd S s)) t }) 409 | 410 | Node : ∀ {I} (T : Irish I) → Fam I → Fam (Info T) 411 | Node ι X = X 412 | Node (κ A) X = A , ↑ 413 | Node (π S T) X = ΠF S λ s → Node (T s) X 414 | Node (σ S T) X = ΣF (Node S X) λ iS → Node (T iS) X 415 | 416 | IF : Set₁ → Set₁ → Set₁ 417 | IF I J = Σ (Irish I) λ T → Info T → J 418 | 419 | _$F_ : ∀ {I J} → (I → J) → Fam I → Fam J 420 | f $F (F , i) = F , (f ∘ i) 421 | 422 | ⟦_⟧IF : ∀ {I J} → IF I J → Fam I → Fam J 423 | ⟦ (T , d) ⟧IF X = d $F Node T X 424 | 425 | {- 426 | mutual 427 | data DataIF {I} (F : IF I I) : Set where 428 | ⟨_⟩ : fst (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) → DataIF F 429 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 430 | ⟦_⟧if {F = F} ⟨ ds ⟩ = snd (⟦ F ⟧IF (DataIF F , ⟦_⟧if)) ds 431 | -} 432 | 433 | mutual 434 | data DataIF {I} (F : IF I I) : Set where 435 | ⟨_⟩ : NoIF F (fst F) → DataIF F 436 | 437 | ⟦_⟧if : ∀ {I} {F : IF I I} → DataIF F → I 438 | ⟦_⟧if {F = F} ⟨ rs ⟩ = snd F (DelF F (fst F) rs) 439 | 440 | NoIF : ∀ {I} (F : IF I I) (T : Irish I) → Set 441 | NoIF F ι = DataIF F 442 | NoIF F (κ A) = A 443 | NoIF F (π S T) = (s : S) → NoIF F (T s) 444 | NoIF F (σ S T) = Σ (NoIF F S) λ s → NoIF F (T (DelF F S s)) 445 | 446 | DelF : ∀ {I} (F : IF I I) (T : Irish I) → NoIF F T → Info T 447 | DelF F ι r = ⟦ r ⟧if 448 | DelF F (κ A) r = ↑ r 449 | DelF F (π S T) f = λ s → DelF F (T s) (f s) 450 | DelF F (σ S T) (s , t) = let s′ = DelF F S s in s′ , DelF F (T s′) t 451 | 452 | DSIF : ∀ {I J} → DS I J → IF I J 453 | DSIF (ι x) = κ One , (λ _ → x) 454 | DSIF (σ S T) = (σ (κ S) (λ s → fst (DSIF (T (↓ s))))) , (λ { (s , t) → snd (DSIF (T (↓ s))) t }) 455 | DSIF (δ H T) = (σ (π H λ _ → ι) (λ f → fst (DSIF (T f)))) , (λ { (f , t) → snd (DSIF (T f)) t }) 456 | 457 | idIF : ∀ {I} → IF I I 458 | idIF = ι , id 459 | 460 | subIF : ∀ {I J} (T : Irish J) (F : IF I J) → IF I (Info T) 461 | subIF ι F = F 462 | subIF (κ A) F = κ A , (λ z → ↑ (↓ z)) 463 | subIF (π S T) F = (π S λ s → fst (subIF (T s) F)) , (λ f s → snd (subIF (T s) F) (f s)) 464 | subIF (σ S T) F with subIF S F 465 | ... | (SF , f) = (σ SF λ sf → fst (subIF (T (f sf)) F)) , (λ { (sf , tf) → f sf , snd (subIF (T (f sf)) F) tf }) 466 | 467 | colF : ∀ {I J K} → IF J K → IF I J → IF I K 468 | colF (S , x) F with subIF S F 469 | ... | TF , f = TF , (x ∘ f) 470 | -------------------------------------------------------------------------------- /Meta/Data/Inductive/W.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Inductive.W where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Functor.Container 5 | open import Meta.Control.Monad 6 | 7 | data W (C : Con) : Set where 8 | ⟨_⟩ : ⟦ C ⟧◃ (W C) → W C 9 | 10 | NatW : Set 11 | NatW = W (Two ◃ Zero ⟨?⟩ One) 12 | 13 | zeroW : NatW 14 | zeroW = ⟨ tt , magic ⟩ 15 | 16 | sucW : NatW → NatW 17 | sucW n = ⟨ tt , (λ _ → n) ⟩ 18 | 19 | precW : ∀ {l}{T : Set l} → T → (NatW → T → T) → NatW → T 20 | precW z _ ⟨ tt , _ ⟩ = z 21 | precW z s ⟨ ff , k ⟩ = s (k <>) (precW z s (k <>)) 22 | 23 | addW : NatW → NatW → NatW 24 | addW x y = precW y (λ _ → sucW) x 25 | 26 | _*◃_ : Con → Set → Set 27 | C *◃ X = W (K◃ X +◃ C) 28 | 29 | freeMonad : (C : Con) → Monad (_*◃_ C) 30 | freeMonad C = record 31 | { return = λ x → ⟨ (tt , x) , magic ⟩ 32 | ; _>>=_ = bind 33 | } where 34 | bind : ∀ {S T : Set} → C *◃ S → (S → C *◃ T) → C *◃ T 35 | bind ⟨ (tt , a) , _ ⟩ f = f a 36 | bind ⟨ (ff , s) , k ⟩ f = ⟨ (ff , s) , (λ x → bind (k x) f) ⟩ 37 | 38 | _*◃ : Con → Con 39 | _*◃ C = C *◃ One ◃ Path where 40 | Path : C *◃ One → Set 41 | Path ⟨ (tt , _) , _ ⟩ = One 42 | Path ⟨ (ff , s) , k ⟩ = Σ (Po C s) λ p → Path (k p) 43 | 44 | call : ∀ {C} → (s : Sh C) → C *◃ Po C s 45 | call s = ⟨ (ff , s) , (λ x → ⟨ (tt , x) , magic ⟩) ⟩ 46 | 47 | Π⊥ : (S : Set) (T : S → Set) → Set 48 | Π⊥ S T = (s : S) → (S ◃ T) *◃ (T s) 49 | 50 | gas : ∀ {S T} → ℕ → Π⊥ S T → (s : S) → T s ⊎ One 51 | gas zero f s = ff , <> 52 | gas {S}{T} (suc n) f s = combust (f s) where 53 | combust : ∀ {X} → (S ◃ T) *◃ X → X ⊎ One 54 | combust ⟨ (tt , s) , l ⟩ = tt , s 55 | combust ⟨ (ff , s) , l ⟩ with gas n f s 56 | combust ⟨ (ff , _) , l ⟩ | tt , t = combust (l t) 57 | combust ⟨ (ff , _) , l ⟩ | ff , _ = ff , <> 58 | 59 | _─_ : (X : Set) (x : X) → Set 60 | X ─ x = Σ X λ x₁ → x₁ ≃ x → Zero 61 | 62 | δ : Con → Con 63 | δ (S ◃ P) = Σ S P ◃ vv λ s p → P s ─ p 64 | 65 | plug : ∀ {C} → ((s : Sh C)(p p₁ : Po C s) → Dec (p ≃ p₁)) → (δ C ×◃ I◃) ⟶◃ C 66 | plug {C} poeq? = (fst ∘ fst) , (λ { ((s , x) , _) x₁ → nplay s x x₁ }) where 67 | nplay : (s : Sh C)(x x₁ : Po C s) → (Po C s ─ x) ⊎ One 68 | nplay s x x₁ with poeq? s x₁ x 69 | nplay s x .x | tt , refl = ff , <> 70 | nplay s x x₁ | ff , pp = tt , x₁ , pp 71 | 72 | -------------------------------------------------------------------------------- /Meta/Data/Monoid.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Monoid where 2 | 3 | open import Meta.Basics 4 | open import Meta.Control.Applicative 5 | 6 | record Monoid (X : Set) : Set where 7 | infixr 4 _•_ 8 | field 9 | ε : X 10 | _•_ : X → X → X 11 | monoidApplicative : Applicative λ _ → X 12 | monoidApplicative = record 13 | { pure = λ _ → ε 14 | ; _⍟_ = _•_ 15 | } 16 | open Monoid {{...}} public 17 | 18 | instance 19 | sumMonoid : Monoid ℕ 20 | sumMonoid = record { ε = zero; _•_ = _+_ } 21 | 22 | record MonoidOK X {{M : Monoid X}} : Set where 23 | field 24 | absorbL : (x : X) → ε • x ≃ x 25 | absorbR : (x : X) → x • ε ≃ x 26 | assoc : (x y z : X) → (x • y) • z ≃ x • (y • z) 27 | open MonoidOK {{...}} public 28 | 29 | natMonoidOK : MonoidOK ℕ 30 | natMonoidOK = record 31 | { absorbL = λ _ → refl 32 | ; absorbR = _+zero 33 | ; assoc = assoc+ 34 | } where 35 | _+zero : ∀ x → x + zero ≃ x 36 | zero +zero = refl 37 | suc n +zero rewrite n +zero = refl 38 | 39 | assoc+ : ∀ x y z → (x + y) + z ≃ x + (y + z) 40 | assoc+ zero y z = refl 41 | assoc+ (suc x) y z rewrite assoc+ x y z = refl 42 | 43 | record MonoidHom {X}{{MX : Monoid X}}{Y}{{MY : Monoid Y}} (f : X → Y) : Set where 44 | field 45 | respε : f ε ≃ ε 46 | resp• : ∀ x x' → f (x • x') ≃ f x • f x' 47 | open MonoidHom {{...}} public 48 | 49 | monoidApplicativeHom : ∀ {X}{{MX : Monoid X}}{Y}{{MY : Monoid Y}} (f : X → Y) {{hf : MonoidHom f}} → 50 | AppHom {{monoidApplicative {{MX}}}} {{monoidApplicative {{MY}}}} f 51 | monoidApplicativeHom f {{hf}} = record 52 | { respPure = λ x → MonoidHom.respε hf 53 | ; resp-⍟ = MonoidHom.resp• hf 54 | } 55 | 56 | homSum : ∀ {F G}{{AF : Applicative F}}{{AG : Applicative G}} → (f : F -:> G) → Applicative λ X → F X ⊎ G X 57 | homSum {F}{G}{{AF}}{{AG}} f = record { 58 | pure = λ x → tt , (pure {{AF}} x) 59 | ; _⍟_ = app 60 | } where 61 | app : {S T : Set} → Σ Two (F (S → T) ⟨?⟩ G (S → T)) → Σ Two (F S ⟨?⟩ G S) → Σ Two (F T ⟨?⟩ G T) 62 | app (tt , k) (tt , g) = tt , (k ⍟ g) 63 | app (tt , k) (ff , g) = ff , (f k ⍟ g) 64 | app (ff , k) (tt , g) = ff , (k ⍟ (f g)) 65 | app (ff , k) (ff , g) = ff , (k ⍟ g) 66 | 67 | {- 68 | homSumOKP : ∀ {F G}{{AF : Applicative F}}{{AG : Applicative G}} → ApplicativeOKP F → ApplicativeOKP G → (f : F -:> G) → AppHom f → ApplicativeOKP _ {{homSum f}} 69 | homSumOKP {F}{G}{{AF}}{{AG}} FOK GOK f homf = record 70 | { lawId = lawIdLemma 71 | ; lawCo = lawCoLemma 72 | ; lawHom = {!!} 73 | ; lawCom = {!!} 74 | } where 75 | lawIdLemma : {X : Set} → (x : F X ⊎ G X) → 76 | _⍟_ {{homSum f}} (pure {{homSum f}} id) x ≃ x 77 | lawIdLemma (tt , u) = tt , (AF Applicative.Applicative.⍟ Applicative.Applicative.pure AF id) u 78 | =[ cong (_,_ tt) (ApplicativeOKP.lawId FOK u) ⟩ 79 | (tt , u) 80 | ∎ 81 | lawIdLemma (ff , u) = ff , (AG Applicative.Applicative.⍟ f (Applicative.Applicative.pure AF id)) u 82 | =[ cong (λ q → ff , (q ⍟ u)) (AppHom.respPure homf id) ⟩ 83 | ff , (AG Applicative.Applicative.⍟ Applicative.Applicative.pure AG id) u 84 | =[ cong (_,_ ff) (ApplicativeOKP.lawId GOK u) ⟩ 85 | (ff , u) 86 | ∎ 87 | comp : {A B C : Set} → (B → C) → (A → B) → A → C 88 | comp f g x = f (g x) 89 | 90 | lawCoLemma : {R S T : Set} (g : F (S → T) ⊎ G (S → T)) (h : F (R → S) ⊎ G (R → S)) (r : F R ⊎ G R) → 91 | _⍟_ {{homSum f}} 92 | (_⍟_ {{homSum f}} 93 | (_⍟_ {{homSum f}} (pure {{homSum f}} (λ g h → g ∘ h)) g) h) r ≃ _⍟_ {{homSum f}} g (_⍟_ {{homSum f}} h r) 94 | lawCoLemma (tt , g) (tt , h) (tt , r) = 95 | tt , _⍟_ {{AF}} ( _⍟_ {{AF}} (_⍟_ {{AF}} (pure {{AF}} (λ g h → g ∘ h)) g) h) r 96 | =[ cong (λ q → tt , q) (ApplicativeOKP.lawCo FOK g h r) ⟩ 97 | tt , (_⍟_ {{AF}} g) (_⍟_ {{AF}} h r) 98 | ∎ 99 | lawCoLemma (tt , g) (tt , h) (ff , r) = 100 | ff , (AG Applicative.Applicative.⍟ f ((AF Applicative.Applicative.⍟ (AF Applicative.Applicative.⍟ Applicative.Applicative.pure AF comp) g) h)) r 101 | =[ cong (λ q → ff , (_⍟_ {{AG}} q r)) (AppHom.resp-⍟ homf ((_⍟_ {{AF}} (pure {{AF}} (λ j k a → j (k a))) g)) h) ⟩ 102 | ff , (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ f ((AF Applicative.Applicative.⍟ Applicative.Applicative.pure AF comp) g)) (f h)) r 103 | =[ cong (λ q → ff , (AG Applicative.Applicative.⍟ (_⍟_ {{AG}} q) (f h)) r) (AppHom.resp-⍟ homf (Applicative.Applicative.pure AF comp) g) ⟩ 104 | ff , (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ (f ((Applicative.Applicative.pure AF comp)))) (f g)) (f h)) r 105 | =[ cong (λ q → ff , (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ q) (f g)) (f h)) r) (AppHom.respPure homf (λ g h → g ∘ h) ) ⟩ 106 | ff , (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ (AG Applicative.Applicative.⍟ Applicative.Applicative.pure AG comp) (f g)) (f h)) r 107 | =[ cong (λ q → ff , q) ((ApplicativeOKP.lawCo GOK (f g) (f h) r)) ⟩ 108 | ff , (AG Applicative.Applicative.⍟ f g) ((AG Applicative.Applicative.⍟ f h) r) 109 | ∎ 110 | lawCoLemma (tt , g) (ff , h) (tt , r) = {!!} 111 | lawCoLemma (tt , g) (ff , h) (ff , r) = {!!} 112 | lawCoLemma (ff , g) (tt , h) (tt , r) = {!!} 113 | lawCoLemma (ff , g) (tt , h) (ff , r) = {!!} 114 | lawCoLemma (ff , g) (ff , h) (tt , r) = {!!} 115 | lawCoLemma (ff , g) (ff , h) (ff , r) = {!!} 116 | -} 117 | -------------------------------------------------------------------------------- /Meta/Data/Nat.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Nat where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Monoid 5 | 6 | +-suc : ∀ x y → x + suc y ≃ suc (x + y) 7 | +-suc zero y = refl 8 | +-suc (suc x) y rewrite +-suc x y = refl 9 | 10 | +-commute : ∀ x y → x + y ≃ y + x 11 | +-commute zero y rewrite (MonoidOK.absorbR natMonoidOK y) = refl 12 | +-commute (suc x) y rewrite +-suc y x | +-commute x y = refl 13 | 14 | _*zero : ∀ x → x * zero ≃ zero 15 | zero *zero = refl 16 | suc n *zero rewrite n *zero = refl 17 | 18 | *-suc : ∀ x y → x * suc y ≃ x + x * y 19 | *-suc zero n = refl 20 | *-suc (suc x) y = begin 21 | suc x * suc y 22 | =[ refl ⟩ 23 | suc (y + x * suc y) 24 | =[ cong suc (lemma x y) ⟩ 25 | suc x + suc x * y 26 | ∎ where 27 | lemma : ∀ x y → y + x * suc y ≃ x + y + x * y 28 | lemma x y = begin 29 | y + x * suc y 30 | =[ cong (_+_ y) (*-suc x y) ⟩ 31 | y + (x + x * y) 32 | =[ +-assoc y x (x * y) ⟩ 33 | (y + x) + x * y 34 | =[ cong (flip _+_ (x * y)) (+-commute y x) ⟩ 35 | (x + y) + x * y 36 | =[ sym (+-assoc x y (x * y)) ⟩ 37 | x + y + x * y 38 | ∎ 39 | 40 | *-comm : ∀ m n → m * n ≃ n * m 41 | *-comm x zero = x *zero 42 | *-comm x (suc y) rewrite trans (*-suc x y) (cong (_+_ x) (*-comm x y)) = refl 43 | -------------------------------------------------------------------------------- /Meta/Data/Vector.agda: -------------------------------------------------------------------------------- 1 | module Meta.Data.Vector where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.EndoFunctor 5 | open import Meta.Control.Applicative 6 | open import Meta.Control.Monad 7 | open import Meta.Data.Fin 8 | 9 | data Vec (X : Set) : ℕ → Set where 10 | ⟨⟩ : Vec X zero 11 | _,_ : {n : ℕ} → X → Vec X n → Vec X (suc n) 12 | 13 | head : ∀ {n X} → Vec X (suc n) → X 14 | head (x , xs) = x 15 | 16 | tail : ∀ {n X} → Vec X (suc n) → Vec X n 17 | tail (x , xs) = xs 18 | 19 | zipv : ∀ {n S T} → Vec S n → Vec T n → Vec (S × T) n 20 | zipv ⟨⟩ ⟨⟩ = ⟨⟩ 21 | zipv (x , xs) (y , ys) = (x , y) , zipv xs ys 22 | 23 | vec : ∀ {n X} → X → Vec X n 24 | vec {zero} x = ⟨⟩ 25 | vec {suc n} x = x , (vec x) 26 | 27 | vapp : ∀ {n S T} → Vec (S → T) n → Vec S n → Vec T n 28 | vapp ⟨⟩ ⟨⟩ = ⟨⟩ 29 | vapp (f , fs) (x , xs) = (f x) , vapp fs xs 30 | 31 | proj : ∀ {n X} → Vec X n → Fin n → X 32 | proj ⟨⟩ () 33 | proj (x , xs) zero = x 34 | proj (x , xs) (suc i) = proj xs i 35 | 36 | vmap : ∀ {n S T} → (S → T) → Vec S n → Vec T n 37 | vmap f ⟨⟩ = ⟨⟩ 38 | vmap f (x , xs) = (f x) , vmap f xs 39 | 40 | concat : ∀ {n X} → Vec (Vec X n) n → Vec X n 41 | concat ⟨⟩ = ⟨⟩ 42 | concat ((x , xx) , xxs) = x , concat (vmap tail xxs) 43 | 44 | zipV : ∀ {n S T} → Vec S n → Vec T n → Vec (S × T) n 45 | zipV ss ts = vapp (vapp (vec _,_) ss) ts 46 | 47 | zipWithV : ∀ {n S T U} → (S → T → U) → Vec S n → Vec T n → Vec U n 48 | zipWithV f ⟨⟩ ⟨⟩ = ⟨⟩ 49 | zipWithV f (a , as) (b , bs) = f a b , zipWithV f as bs 50 | 51 | _++_ : ∀ {m n X} → Vec X m → Vec X n → Vec X (m + n) 52 | ⟨⟩ ++ ys = ys 53 | (x , xs) ++ ys = x , (xs ++ ys) 54 | 55 | vfoldl : ∀ {n S}{T : Set} → (T → S → T) → T → Vec S n → T 56 | vfoldl f i ⟨⟩ = i 57 | vfoldl f i (x , xs) = vfoldl f (f i x) xs 58 | 59 | chop : forall {X} m {n} → Vec X (m + n) → (Vec X m) × (Vec X n) 60 | chop zero xs = ⟨⟩ , xs 61 | chop (suc m) (x , xs) with chop m xs 62 | ... | ys , zs = (x , ys) , zs 63 | 64 | vreplicate : ∀ {X} → (n : ℕ) → (x : X) → Vec X n 65 | vreplicate zero x = ⟨⟩ 66 | vreplicate (suc k) x = x , vreplicate k x 67 | 68 | tabulate : ∀ {n X} → (Fin n → X) → Vec X n 69 | tabulate {zero} f = ⟨⟩ 70 | tabulate {suc n} f = f zero , tabulate (λ _ → f zero) 71 | 72 | instance 73 | monadVec : ∀ {n} → Monad λ X → Vec X n 74 | monadVec = record { return = vec; _>>=_ = λ m f → concat (vmap f m) } 75 | 76 | applicativeVec : ∀ {n} → Applicative λ X → Vec X n 77 | applicativeVec = record { pure = vec; _⍟_ = vapp } 78 | 79 | endoFunctorVec : ∀ {n} → EndoFunctor λ X → Vec X n 80 | endoFunctorVec = record { map = vmap } 81 | 82 | vecEndoFunctorOKP : ∀ {n} → EndoFunctorOKP (λ X → Vec X n) 83 | vecEndoFunctorOKP = record { endoFunctorId = vecId; endoFunctorCo = λ f g → vecComp f g } 84 | where 85 | vecId : ∀ {n X} → (x : Vec X n) → vmap id x ≃ x 86 | vecId ⟨⟩ = refl 87 | vecId (x , xs) rewrite vecId xs = refl 88 | 89 | vecComp : ∀ {n R S T} → (f : S → T)(g : R → S) → (x : Vec R n) → (map f ∘ map g) x ≃ map (f ∘ g) x 90 | vecComp f g ⟨⟩ = refl 91 | vecComp f g (x , xs) rewrite vecComp f g xs = refl 92 | 93 | allFin : ∀ n → Vec (Fin n) n 94 | allFin zero = ⟨⟩ 95 | allFin (suc n) = zero , map suc (allFin n) 96 | -------------------------------------------------------------------------------- /Meta/Language/LambdaCalculus.agda: -------------------------------------------------------------------------------- 1 | module Meta.Language.LambdaCalculus where 2 | 3 | open import Meta.Basics 4 | 5 | data ⋆ : Set where 6 | ι : ⋆ 7 | _▹_ : ⋆ → ⋆ → ⋆ 8 | infixl 5 _▹_ 9 | 10 | data Cx (X : Set) : Set where 11 | ε : Cx X 12 | _,_ : Cx X → X → Cx X 13 | infixl 4 _,_ 14 | 15 | data _∈_ (τ : ⋆) : Cx ⋆ → Set where 16 | zero : ∀ {Γ} → τ ∈ Γ , τ 17 | suc : ∀ {Γ σ} → τ ∈ Γ → τ ∈ Γ , σ 18 | infix 3 _∈_ 19 | 20 | data _⊢_ (Γ : Cx ⋆) : ⋆ → Set where 21 | var : ∀ {τ} 22 | → τ ∈ Γ 23 | ------- 24 | → Γ ⊢ τ 25 | lam : ∀ {σ τ} 26 | → Γ , σ ⊢ τ 27 | ------------ 28 | → Γ ⊢ σ ▹ τ 29 | app : ∀ {σ τ} 30 | → Γ ⊢ σ ▹ τ → Γ ⊢ σ 31 | ------------------- 32 | → Γ ⊢ τ 33 | infix 2 _⊢_ 34 | 35 | ⟦_⟧⋆ : ⋆ → Set 36 | ⟦ ι ⟧⋆ = ℕ 37 | ⟦ σ ▹ τ ⟧⋆ = ⟦ σ ⟧⋆ → ⟦ τ ⟧⋆ 38 | 39 | ⟦_⟧Cx : Cx ⋆ → (⋆ → Set) → Set 40 | ⟦ ε ⟧Cx V = One 41 | ⟦ (Γ , σ) ⟧Cx V = ⟦ Γ ⟧Cx V × V σ 42 | 43 | ⟦_⟧∈ : ∀ {Γ τ V} → τ ∈ Γ → ⟦ Γ ⟧Cx V → V τ 44 | ⟦ zero ⟧∈ (γ , t) = t 45 | ⟦ suc i ⟧∈ (γ , s) = ⟦ i ⟧∈ γ 46 | 47 | ⟦_⟧⊢ : ∀ {Γ τ} → Γ ⊢ τ → ⟦ Γ ⟧Cx ⟦_⟧⋆ → ⟦ τ ⟧⋆ 48 | ⟦ var i ⟧⊢ γ = ⟦ i ⟧∈ γ 49 | ⟦ lam t ⟧⊢ γ = λ s → ⟦ t ⟧⊢ (γ , s) 50 | ⟦ app f s ⟧⊢ γ = ⟦ f ⟧⊢ γ (⟦ s ⟧⊢ γ) 51 | 52 | Ren Sub : Cx ⋆ → Cx ⋆ → Set 53 | Ren Γ Δ = ∀ {τ} → τ ∈ Γ → τ ∈ Δ 54 | Sub Γ Δ = ∀ {τ} → τ ∈ Γ → Δ ⊢ τ 55 | 56 | _<><_ : ∀ { X } → Cx X → List X → Cx X 57 | xz <>< ⟨⟩ = xz 58 | xz <>< (x , xs) = (xz , x) <>< xs 59 | infixl 3 _<><_ 60 | 61 | Shub : Cx ⋆ → Cx ⋆ → Set 62 | Shub Γ Δ = ∀ Ξ → Sub (Γ <>< Ξ) (Δ <>< Ξ) 63 | 64 | _//_ : ∀ {Γ Δ}(θ : Shub Γ Δ) {τ} → Γ ⊢ τ → Δ ⊢ τ 65 | θ // var i = θ ⟨⟩ i 66 | θ // lam t = lam ((θ ∘ _,_ _) // t) 67 | θ // app f s = app (θ // f) (θ // s) 68 | 69 | wkr : ∀ {Γ Δ σ} → Ren Γ Δ → Ren (Γ , σ) (Δ , σ) 70 | wkr r zero = zero 71 | wkr r (suc n) = suc (r n) 72 | 73 | ren : ∀ {Γ Δ} → Ren Γ Δ → Shub Γ Δ 74 | ren r ⟨⟩ = var ∘ r 75 | ren r (_ , Ξ) = ren (wkr r) Ξ 76 | 77 | wks : ∀ {Γ Δ σ} → Sub Γ Δ → Sub (Γ , σ) (Δ , σ) 78 | wks s zero = var zero 79 | wks s (suc n) = ren suc // s n 80 | 81 | sub : ∀ {Γ Δ} → Sub Γ Δ → Shub Γ Δ 82 | sub s ⟨⟩ = s 83 | sub s (_ , Ξ) = sub (wks s) Ξ 84 | 85 | weak : ∀ {Γ} Ξ → Ren Γ (Γ <>< Ξ) 86 | weak ⟨⟩ i = i 87 | weak (_ , Ξ) i = weak Ξ (suc i) 88 | 89 | lambda' : ∀ {Γ σ τ} → ((∀ {Ξ} → Γ , σ <>< Ξ ⊢ σ) → Γ , σ ⊢ τ) → Γ ⊢ σ ▹ τ 90 | lambda' f = lam (f λ {Ξ} → var (weak Ξ zero)) 91 | 92 | _<>>_ : ∀ {X} → Cx X → List X → List X 93 | ε <>> ys = ys 94 | (xz , x) <>> ys = xz <>> (x , ys) 95 | 96 | open import Agda.Primitive 97 | 98 | lambda : ∀ {Γ σ τ} → ((∀ {Δ Ξ} {{_ : Δ <>> ⟨⟩ ≃ Γ <>> (σ , Ξ)}} → Δ ⊢ σ) → Γ , σ ⊢ τ) → Γ ⊢ σ ▹ τ 99 | lambda {Γ} f = lam ((f λ {Δ Ξ}{{q}} → subst (lem Δ Γ (_ , Ξ) q) (λ Γ → Γ ⊢ _) (var (weak Ξ zero)))) 100 | where 101 | {- This is Conor's. He wasn't kidding about the ugly. -} 102 | sucI : (a b : ℕ) → (_≃_ {lzero}{ℕ} (suc a) (suc b)) → a ≃ b 103 | sucI .b b refl = refl 104 | 105 | grr : (x y : ℕ) → suc x + y ≃ x + suc y 106 | grr zero y = refl 107 | grr (suc x) y rewrite grr x y = refl 108 | 109 | _+a_ : ℕ → ℕ → ℕ 110 | zero +a y = y 111 | suc x +a y = x +a suc y 112 | 113 | noc' : (x y : ℕ) → suc (x + y) ≃ y → {A : Set} → A 114 | noc' x zero () 115 | noc' x (suc y) q = noc' x y $′ 116 | begin 117 | suc (x + y) 118 | =[ grr x y ⟩ 119 | x + suc y 120 | =[ sucI _ _ q ⟩ 121 | y 122 | ∎ 123 | 124 | noc : (x k y : ℕ) → x +a (suc k + y) ≃ y → {A : Set} → A 125 | noc zero k y q = noc' k y q 126 | noc (suc x) k y q = noc x (suc k) y q 127 | 128 | len : ∀ {X} → Cx X → ℕ 129 | len ε = zero 130 | len (xz , x) = suc (len xz) 131 | 132 | lenlem : ∀ {X}(xz : Cx X)(xs : List X) → length (xz <>> xs) ≃ len xz +a length xs 133 | lenlem ε xs = refl 134 | lenlem (xz , x) xs = lenlem xz (x , xs) 135 | 136 | lem0 : ∀ {X}(xz yz : Cx X)(xs ys : List X) → length xs ≃ length ys → xz <>> xs ≃ yz <>> ys → (xz ≃ yz) × (xs ≃ ys) 137 | lem0 ε ε xs ys q q' = refl , q' 138 | lem0 ε (yz , x) .(yz <>> (x , ys)) ys q refl rewrite lenlem yz (x , ys) = noc (len yz) zero (length ys) q 139 | lem0 (xz , x) ε xs .(xz <>> (x , xs)) q refl rewrite lenlem xz (x , xs) = noc (len xz) zero (length xs) (_ ⟨ q ]= _ ∎) 140 | lem0 (xz , x) (yz , y) xs ys q q' with lem0 xz yz (x , xs) (y , ys) (cong suc q) q' 141 | lem0 (.yz , .y) (yz , y) .ys ys q q' | refl , refl = refl , refl 142 | 143 | lem : ∀ {X} (Δ Γ : Cx X) Ξ → Δ <>> ⟨⟩ ≃ Γ <>> Ξ → Γ <>< Ξ ≃ Δ 144 | lem Δ Γ ⟨⟩ q = 145 | begin 146 | Γ 147 | ⟨ fst (lem0 Δ Γ ⟨⟩ ⟨⟩ refl q) ]= 148 | Δ 149 | ∎ 150 | lem Δ Γ (x , Ξ) q = lem Δ (Γ , x) Ξ q 151 | 152 | {- 153 | myTest : ε ⊢ ι ▹ ι 154 | myTest = lambda λ x → x 155 | 156 | myTest2 : ε ⊢ (ι ▹ ι) ▹ (ι ▹ ι) 157 | myTest2 = lambda λ f → lambda λ x → app f (app f x) 158 | -} 159 | 160 | mutual 161 | data _⊨_ (Γ : Cx ⋆) : ⋆ → Set where 162 | lam : ∀ {σ τ} → Γ , σ ⊨ τ → Γ ⊨ σ ▹ τ 163 | _$_ : ∀ {τ} → τ ∈ Γ → Γ ⊨* τ → Γ ⊨ ι 164 | 165 | data _⊨*_ (Γ : Cx ⋆) : ⋆ → Set where 166 | ⟨⟩ : Γ ⊨* ι 167 | _,_ : ∀ {σ τ} → Γ ⊨ σ → Γ ⊨* τ → Γ ⊨* σ ▹ τ 168 | infix 3 _⊨_ _⊨*_ 169 | infix 3 _$_ 170 | 171 | _-_ : ∀ (Γ : Cx ⋆) {τ} (x : τ ∈ Γ) → Cx ⋆ 172 | ε - () 173 | (Γ , _) - zero = Γ 174 | (Γ , sg) - suc x = (Γ - x) , sg 175 | infixl 4 _-_ 176 | 177 | _≠_ : ∀ {Γ σ} (x : σ ∈ Γ) → Ren (Γ - x) Γ 178 | zero ≠ y = suc y 179 | suc x ≠ zero = zero 180 | suc x ≠ suc y = suc (x ≠ y) 181 | 182 | {- 183 | ⟨_↦_⟩_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ -x x ⊨ σ → Γ ⊨ τ → Γ -x x ⊨ τ 184 | ⟨ x ↦ s ⟩ lam t = lam (⟨ suc x ↦ ? ⟩ t) 185 | ⟨ x ↦ s ⟩ (x₁ $ x₂) = ? 186 | -} 187 | 188 | data Veq? {Γ σ}(x : σ ∈ Γ) : ∀ {τ} → τ ∈ Γ → Set where 189 | same : Veq? x x 190 | diff : ∀ {τ}(y : τ ∈ Γ - x) → Veq? x (x ≠ y) 191 | 192 | --Show that every |y| is discriminable with respect to a given |x|. 193 | 194 | veq? : ∀ {Γ σ τ}(x : σ ∈ Γ)(y : τ ∈ Γ) → Veq? x y 195 | veq? zero zero = same 196 | veq? zero (suc y) = diff y 197 | veq? (suc x) zero = diff zero 198 | veq? (suc x) (suc y) with veq? x y 199 | veq? (suc x) (suc .x) | same = same 200 | veq? (suc x) (suc .(x ≠ y)) | diff y = diff (suc y) 201 | 202 | --Show how to propagate a renaming through a normal form. 203 | mutual 204 | renNm : ∀ {Γ Δ τ} → Ren Γ Δ → Γ ⊨ τ → Δ ⊨ τ 205 | renNm ρ (lam n) = lam (renNm (wkr ρ) n) 206 | renNm ρ (f $ x) = ρ f $ (renSp ρ x) 207 | 208 | renSp : ∀ {Γ Δ τ} → Ren Γ Δ → Γ ⊨* τ → Δ ⊨* τ 209 | renSp ρ ⟨⟩ = ⟨⟩ 210 | renSp ρ (x , ss) = renNm ρ x , renSp ρ ss 211 | 212 | mutual 213 | ⟨_↦_⟩_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ - x ⊨ σ → Γ ⊨ τ → Γ - x ⊨ τ 214 | ⟨ x ↦ s ⟩ lam t = lam (⟨ suc x ↦ renNm (_≠_ zero) s ⟩ t) 215 | ⟨ x ↦ s ⟩ x₁ $ x₂ with veq? x x₁ 216 | ⟨ x ↦ s ⟩ .x $ xs | same = s $$ (⟨ x ↦ s ⟩* xs) 217 | ⟨ x ↦ s ⟩ .(x ≠ y) $ xs | diff y = y $ (⟨ x ↦ s ⟩* xs) 218 | 219 | ⟨_↦_⟩*_ : ∀ {Γ σ τ} → (x : σ ∈ Γ) → Γ - x ⊨ σ → Γ ⊨* τ → Γ - x ⊨* τ 220 | ⟨ x ↦ s ⟩* ⟨⟩ = ⟨⟩ 221 | ⟨ x ↦ s ⟩* (t , ts) = (⟨ x ↦ s ⟩ t) , (⟨ x ↦ s ⟩* ts) 222 | 223 | _$$_ : ∀ {Γ τ} → Γ ⊨ τ → Γ ⊨* τ → Γ ⊨ ι 224 | f $$ ⟨⟩ = f 225 | lam f $$ (s , ss) = (⟨ zero ↦ s ⟩ f) $$ ss 226 | infix 3 _$$_ 227 | infix 2 ⟨_↦_⟩_ 228 | 229 | η : ∀ {Γ σ}(x : σ ∈ Γ) τ → (∀ {Δ} → Ren Γ Δ → Δ ⊨* τ → Δ ⊨* σ) → Γ ⊨ τ 230 | η x ι f = x $ f id ⟨⟩ 231 | η x (σ ▹ τ) f = lam (η (suc x) τ λ ρ ss → f (ρ ∘ suc) ((η (ρ zero) σ (λ _ → id)) , ss)) 232 | 233 | normalize : ∀ {Γ τ} → Γ ⊢ τ → Γ ⊨ τ 234 | normalize (var x) = η x _ λ _ → id 235 | normalize (lam t) = lam (normalize t) 236 | normalize (app f s) with normalize f | normalize s 237 | normalize (app f s) | lam t | s2 = ⟨ zero ↦ s2 ⟩ t 238 | 239 | {- 240 | try₁ : ε ⊨ ((ι ▹ ι) ▹ (ι ▹ ι)) ▹ (ι ▹ ι) ▹ (ι ▹ ι) 241 | try₁ = normalize (lambda id) 242 | 243 | church₂ : ∀ {τ} → ε ⊢ (τ ▹ τ) ▹ τ ▹ τ 244 | church₂ = lambda λ f → lambda λ x → app f (app f x) 245 | 246 | try₂ : ε ⊨ (ι ▹ ι) ▹ (ι ▹ ι) 247 | try₂ = normalize (app (app church₂ church₂) church₂) 248 | -} 249 | 250 | data Stop (Γ : Cx ⋆) (τ : ⋆) : Set where 251 | var : τ ∈ Γ → Stop Γ τ 252 | _$_ : ∀ {σ} → Stop Γ (σ ▹ τ) → Γ ⊨ σ → Stop Γ τ 253 | 254 | renSt : ∀ {Γ Δ τ} → Ren Γ Δ → Stop Γ τ → Stop Δ τ 255 | renSt r (var x) = var (r x) 256 | renSt r (u $ s) = renSt r u $ renNm r s 257 | 258 | stopSp : ∀ {Γ τ} → Stop Γ τ → Γ ⊨* τ → Γ ⊨ ι 259 | stopSp (var x) ss = x $ ss 260 | stopSp (u $ x) ss = stopSp u (x , ss) 261 | 262 | mutual 263 | Val : Cx ⋆ → ⋆ → Set 264 | Val Γ τ = Go Γ τ ⊎ Stop Γ τ 265 | 266 | Go : Cx ⋆ → ⋆ → Set 267 | Go Γ ι = Zero 268 | Go Γ (σ ▹ τ) = ∀ {Δ} → Ren Γ Δ → Val Δ σ → Val Δ τ 269 | 270 | renVal : ∀ {Γ Δ} τ → Ren Γ Δ → Val Γ τ → Val Δ τ 271 | renVal τ r (ff , u) = ff , renSt r u 272 | renVal ι r (tt , ()) 273 | renVal (σ ▹ τ) r (tt , f) = tt , (λ r' s → f (r' ∘ r) s) 274 | 275 | renVals : ∀ Θ {Γ Δ} → Ren Γ Δ → ⟦ Θ ⟧Cx (Val Γ) → ⟦ Θ ⟧Cx (Val Δ) 276 | renVals ε r _ = <> 277 | renVals (Θ , σ) r (θ , τ) = (renVals Θ r θ) , renVal σ r τ 278 | 279 | idEnv : ∀ Γ → ⟦ Γ ⟧Cx (Val Γ) 280 | idEnv ε = <> 281 | idEnv (Γ , σ) = renVals Γ suc (idEnv Γ) , (ff , var zero) 282 | 283 | mutual 284 | apply : ∀ {Γ σ τ} → Val Γ (σ ▹ τ) → Val Γ σ → Val Γ τ 285 | apply (tt , f) s = f id s 286 | apply (ff , u) s = ff , (u $ quo _ s) 287 | 288 | quo : ∀ {Γ} τ → Val Γ τ → Γ ⊨ τ 289 | quo ι (tt , ()) 290 | quo ι (ff , u) = stopSp u ⟨⟩ 291 | quo (σ ▹ τ) v = lam (quo τ (apply (renVal _ suc v) (ff , var zero))) 292 | 293 | eval : ∀ {Γ Δ τ} → Γ ⊢ τ → ⟦ Γ ⟧Cx (Val Δ) → Val Δ τ 294 | eval (var x) γ = ⟦ x ⟧∈ γ 295 | eval {Γ}{_}{_} (lam t) γ = tt , λ r s → eval t (renVals Γ r γ , s) 296 | eval (app f s) γ = apply (eval f γ) (eval s γ) 297 | 298 | normByEval : ∀ {Γ τ} → Γ ⊢ τ → Γ ⊨ τ 299 | normByEval {Γ}{τ} t = quo τ (eval t (idEnv Γ)) 300 | 301 | {- 302 | zero : Γ ⊢ ι 303 | suc : Γ ⊢ ι → Γ ⊢ ι 304 | rec : ∀ {τ} → Γ ⊢ τ → Γ ⊢ (ι ▹ τ ▹ τ) → Γ ⊢ ι → Γ ⊢ τ 305 | -} 306 | -------------------------------------------------------------------------------- /Meta/Language/OTT.agda: -------------------------------------------------------------------------------- 1 | module Meta.Language.OTT where 2 | 3 | open import Meta.Basics 4 | open import Meta.Data.Inductive.InductionRecursion 5 | open import Meta.Data.Functor.Container.Indexed 6 | open import Meta.Data.Inductive.ITree 7 | 8 | {- 9 | data Favourite : (ℕ → ℕ) → Set where 10 | favourite : Favourite (λ x → zero + x) 11 | 12 | data Favourite (f : ℕ → ℕ) : Set where 13 | favourite : (λ x → zero + x) ≃ f → Favourite f 14 | -} 15 | 16 | mutual 17 | EQ : (X Y : TU) → TU × (⟦ X ⟧TU → ⟦ Y ⟧TU → TU) 18 | EQ Zero' Zero' = One' , λ _ _ → One' 19 | EQ One' One' = One' , λ _ _ → One' 20 | EQ Two' Two' = One' , λ 21 | { tt tt → One' 22 | ; ff ff → One' 23 | ; _ _ → Zero' 24 | } 25 | EQ (Σ' S T) (Σ' S' T') = (Σ' (S ↔ S') λ _ → 26 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 27 | T s ↔ T' s') 28 | , (λ { (s , t) (s' , t') → 29 | Σ' (Eq S s S' s') λ _ → Eq (T s) t (T' s') t' }) 30 | EQ (Π' S T) (Π' S' T') = (Σ' (S' ↔ S) λ _ → 31 | Π' S' λ s' → Π' S λ s → Π' (Eq S' s' S s) λ _ → 32 | T s ↔ T' s') 33 | , (λ { f f' → 34 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 35 | Eq (T s) (f s) (T' s') (f' s') }) 36 | {- EQ (Tree' I F i) (Tree' I' F' i') = (Σ' (I ↔ I') λ _ → Σ' (Eq I i I' i') λ _ → 37 | Π' I λ i → Π' I' λ i' → Π' (Eq I i I' i') λ _ → 38 | let (S , K) = F i ; S' , K' = F' i' 39 | in Σ' (S ↔ S') λ _ → 40 | Π' S λ s → Π' S' λ s' → Π' (Eq S s S' s') λ _ → 41 | let (P , r) = K s ; (P' , r') = K' s' 42 | in Σ' (P' ↔ P) λ _ → 43 | Π' P' λ p' → Π' P λ p → Π' (Eq P' p' P p) λ _ → 44 | Eq I (r p) I' (r' p')) 45 | , {!teq i i'!} where 46 | teq : (i : ⟦ I ⟧TU) → (i' : ⟦ I' ⟧TU) → ⟦ Tree' I F i ⟧TU → ⟦ Tree' I' F' i' ⟧TU → TU 47 | teq i i' ⟨ s , k ⟩ ⟨ s' , k' ⟩ = ? 48 | = let (S , K) = F i ; (S' , K') = F' i' 49 | (P , r) = K s ; (P' , r') = K' s' 50 | in Σ' (Eq S s S' s') λ _ → 51 | Π' P λ p → Π' P' λ p' → Π' (Eq P p P' p') λ _ → 52 | teq (r p) (r' p') (k p) (k' p') -} 53 | EQ _ _ = Zero' , λ _ _ → One' 54 | 55 | _↔_ : TU → TU → TU 56 | X ↔ Y = fst (EQ X Y) 57 | 58 | Eq : (X : TU) (x : ⟦ X ⟧TU) → (Y : TU) (y : ⟦ Y ⟧TU) → TU 59 | Eq X x Y y = snd (EQ X Y) x y 60 | 61 | coe : (X Y : TU) → ⟦ X ↔ Y ⟧TU → ⟦ X ⟧TU → ⟦ Y ⟧TU 62 | postulate 63 | coh : (X Y : TU) (Q : ⟦ X ↔ Y ⟧TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x Y (coe X Y Q x) ⟧TU 64 | coe X Y Q x = {!!} 65 | 66 | postulate 67 | reflTU : (X : TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x X x ⟧TU 68 | 69 | {- 70 | reflTU : (X : TU) (x : ⟦ X ⟧TU) → ⟦ Eq X x X x ⟧TU 71 | reflTU X x = ? 72 | -} 73 | 74 | postulate 75 | RespTU : (X : TU) (P : ⟦ X ⟧TU → TU) 76 | (x x' : ⟦ X ⟧TU) → ⟦ Eq X x X x' ⟧TU → ⟦ P x ↔ P x' ⟧TU 77 | 78 | substTU : (X : TU) (P : ⟦ X ⟧TU → TU) 79 | (x x' : ⟦ X ⟧TU) → ⟦ Eq X x X x' ⟧TU → ⟦ P x ⟧TU → ⟦ P x' ⟧TU 80 | substTU X P x x' q = coe (P x) (P x') (RespTU X P x x' q) 81 | 82 | data Sort : Set where set prop : Sort 83 | 84 | IsSet : Sort → Set 85 | IsSet set = One 86 | IsSet prop = Zero 87 | 88 | 89 | mutual 90 | data PU (u : Sort) : Set where 91 | Zero' One' : PU u 92 | Two' : {_ : IsSet u} → PU u 93 | Σ' : (S : PU u) (T : ⟦ S ⟧PU → PU u) → PU u 94 | Π' : (S : PU set) (T : ⟦ S ⟧PU → PU u) → PU u 95 | Tree' : {_ : IsSet u} 96 | (I : PU set) 97 | (F : ⟦ I ⟧PU → Σ (PU set) λ S → 98 | ⟦ S ⟧PU → Σ (PU set) λ P → 99 | ⟦ P ⟧PU → ⟦ I ⟧PU) 100 | (i : ⟦ I ⟧PU) → PU u 101 | Prf' : {_ : IsSet u} → PU prop → PU u 102 | 103 | ⟦_⟧PU : ∀ {u} → PU u → Set 104 | ⟦ Zero' ⟧PU = Zero 105 | ⟦ One' ⟧PU = One 106 | ⟦ Two' ⟧PU = Two 107 | ⟦ Σ' S T ⟧PU = Σ ⟦ S ⟧PU λ s → ⟦ T s ⟧PU 108 | ⟦ Π' S T ⟧PU = (s : ⟦ S ⟧PU) → ⟦ T s ⟧PU 109 | ⟦ Tree' I F i ⟧PU = ITree 110 | ((λ i → ⟦ fst (F i) ⟧PU) 111 | ◃ (λ i s → ⟦ fst (snd (F i) s) ⟧PU) 112 | $ (λ i s p → snd (snd (F i) s) p) 113 | ) i 114 | ⟦ Prf' P ⟧PU = ⟦ P ⟧PU 115 | 116 | _∧_ : PU prop → PU prop → PU prop 117 | P ∧ Q = Σ' P λ _ → Q 118 | 119 | _⇒_ : PU prop → PU prop → PU prop 120 | P ⇒ Q = Π' (Prf' P) λ _ → Q 121 | 122 | mutual 123 | PEQ : (X Y : PU set) → PU prop × (⟦ X ⟧PU → ⟦ Y ⟧PU → PU prop) 124 | _⇔_ : PU set → PU set → PU prop 125 | X ⇔ Y = fst (PEQ X Y) 126 | 127 | PEq : (X : PU set) (x : ⟦ X ⟧PU) → (Y : PU set) (y : ⟦ Y ⟧PU) → PU prop 128 | PEq X x Y y = snd (PEQ X Y) x y 129 | 130 | PEQ (Prf' P) (Prf' Q) = ((P ⇒ Q) ∧ (Q ⇒ P)) , λ _ _ → One' 131 | PEQ _ _ = Zero' , λ _ _ → One' 132 | 133 | -------------------------------------------------------------------------------- /meta.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/CodaFi/Agda-Metaprogramming/b3ad866f2a1c8a7879900dd296106996b58bdc95/meta.pdf --------------------------------------------------------------------------------