Vec.agda 4.2 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168
  1. module Vec where
  2. {- Computed datatypes -}
  3. data One : Set where
  4. unit : One
  5. data Nat : Set where
  6. zero : Nat
  7. suc : Nat -> Nat
  8. data _*_ (A B : Set) : Set where
  9. pair : A -> B -> A * B
  10. infixr 20 _=>_
  11. data _=>_ (A B : Set) : Set where
  12. lam : (A -> B) -> A => B
  13. lam2 : {A B C : Set} -> (A -> B -> C) -> (A => B => C)
  14. lam2 f = lam (\x -> lam (f x))
  15. app : {A B : Set} -> (A => B) -> A -> B
  16. app (lam f) x = f x
  17. Vec : Nat -> Set -> Set
  18. Vec zero X = One
  19. Vec (suc n) X = X * Vec n X
  20. {- ... construct the vectors of a given length -}
  21. vHead : {X : Set} -> (n : Nat)-> Vec (suc n) X -> X
  22. vHead n (pair a b) = a
  23. vTail : {X : Set} -> (n : Nat)-> Vec (suc n) X -> Vec n X
  24. vTail n (pair a b) = b
  25. {- safe destructors for nonempty vectors -}
  26. {- useful vector programming operators -}
  27. vec : {n : Nat}{X : Set} -> X -> Vec n X
  28. vec {zero } x = unit
  29. vec {suc n} x = pair x (vec x)
  30. vapp : {n : Nat}{S T : Set} -> Vec n (S => T) -> Vec n S -> Vec n T
  31. vapp {zero } unit unit = unit
  32. vapp {suc n} (pair f fs) (pair s ss) = pair (app f s) (vapp fs ss)
  33. {- mapping and zipping come from these -}
  34. vMap : {n : Nat}{S T : Set} -> (S -> T) -> Vec n S -> Vec n T
  35. vMap f ss = vapp (vec (lam f)) ss
  36. {- transposition gets the type it deserves -}
  37. transpose : {m n : Nat}{X : Set} -> Vec m (Vec n X) -> Vec n (Vec m X)
  38. transpose {zero } xss = vec unit
  39. transpose {suc m} (pair xs xss) =
  40. vapp (vapp (vec (lam2 pair)) xs)
  41. (transpose xss)
  42. {- Sets of a given finite size may be computed as follows... -}
  43. {- Resist the temptation to mention idioms. -}
  44. data Zero : Set where
  45. data _+_ (A B : Set) : Set where
  46. inl : A -> A + B
  47. inr : B -> A + B
  48. Fin : Nat -> Set
  49. Fin zero = Zero
  50. Fin (suc n) = One + Fin n
  51. {- We can use these sets to index vectors safely. -}
  52. vProj : {n : Nat}{X : Set} -> Vec n X -> Fin n -> X
  53. vProj {zero } _ ()
  54. vProj {suc n} (pair x xs) (inl unit) = x
  55. vProj {suc n} (pair x xs) (inr i) = vProj xs i
  56. {- We can also tabulate a function as a vector. Resist
  57. the temptation to mention logarithms. -}
  58. vTab : {n : Nat}{X : Set} -> (Fin n -> X) -> Vec n X
  59. vTab {zero } _ = unit
  60. vTab {suc n} f = pair (f (inl unit)) (vTab (\x -> f (inr x)))
  61. {- Question to ponder in your own time:
  62. if we use functional vectors what are vec and vapp -}
  63. {- Answer: K and S -}
  64. {- Inductive datatypes of the unfocused variety -}
  65. {- Every constructor must target the whole family rather
  66. than focusing on specific indices. -}
  67. data Tm (n : Nat) : Set where
  68. evar : Fin n -> Tm n
  69. eapp : Tm n -> Tm n -> Tm n
  70. elam : Tm (suc n) -> Tm n
  71. {- Renamings -}
  72. Ren : Nat -> Nat -> Set
  73. Ren m n = Vec m (Fin n)
  74. _`Ren`_ = Ren
  75. {- identity and composition -}
  76. idR : {n : Nat} -> n `Ren` n
  77. idR = vTab (\i -> i)
  78. coR : {l m n : Nat} -> m `Ren` n -> l `Ren` m -> l `Ren` n
  79. coR m2n l2m = vMap (vProj m2n) l2m
  80. {- what theorems should we prove -}
  81. {- the lifting functor for Ren -}
  82. liftR : {m n : Nat} -> m `Ren` n -> suc m `Ren` suc n
  83. liftR m2n = pair (inl unit) (vMap inr m2n)
  84. {- what theorems should we prove -}
  85. {- the functor from Ren to Tm-arrows -}
  86. rename : {m n : Nat} -> (m `Ren` n) -> Tm m -> Tm n
  87. rename m2n (evar i) = evar (vProj m2n i)
  88. rename m2n (eapp f s) = eapp (rename m2n f) (rename m2n s)
  89. rename m2n (elam t) = elam (rename (liftR m2n) t)
  90. {- Substitutions -}
  91. Sub : Nat -> Nat -> Set
  92. Sub m n = Vec m (Tm n)
  93. _`Sub`_ = Sub
  94. {- identity; composition must wait; why -}
  95. idS : {n : Nat} -> n `Sub` n
  96. idS = vTab evar
  97. {- functor from renamings to substitutions -}
  98. Ren2Sub : {m n : Nat} -> m `Ren` n -> m `Sub` n
  99. Ren2Sub m2n = vMap evar m2n
  100. {- lifting functor for substitution -}
  101. liftS : {m n : Nat} -> m `Sub` n -> suc m `Sub` suc n
  102. liftS m2n = pair (evar (inl unit))
  103. (vMap (rename (vMap inr idR)) m2n)
  104. {- functor from Sub to Tm-arrows -}
  105. subst : {m n : Nat} -> m `Sub` n -> Tm m -> Tm n
  106. subst m2n (evar i) = vProj m2n i
  107. subst m2n (eapp f s) = eapp (subst m2n f) (subst m2n s)
  108. subst m2n (elam t) = elam (subst (liftS m2n) t)
  109. {- and now we can define composition -}
  110. coS : {l m n : Nat} -> m `Sub` n -> l `Sub` m -> l `Sub` n
  111. coS m2n l2m = vMap (subst m2n) l2m