Talk.agda 6.3 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186
  1. {-
  2. Data, Deliberately
  3. Conor McBride
  4. Workshop in Dependently Typed Programming
  5. Nottingham, 2008
  6. (Agda rendering by Ulf Norell)
  7. -}
  8. module Talk where
  9. open import SomeBasicStuff
  10. -- Codes for (first order) inductive families.
  11. data Code (I : Set) : Set1 where
  12. arg : (A : Set) -> (A -> Code I) -> Code I
  13. rec : I -> Code I -> Code I
  14. out : I -> Code I
  15. -- The semantics of a code is a functor.
  16. ⟦_⟧ : {I : Set} -> Code I -> (I -> Set) -> I -> Set
  17. ⟦ out i ⟧ X j = i == j
  18. ⟦ arg A B ⟧ X j = Σ A \a -> ⟦ B a ⟧ X j
  19. ⟦ rec i C ⟧ X j = X i × ⟦ C ⟧ X j
  20. map : {I : Set}{X Y : I -> Set}(C : Code I) ->
  21. ({i : I} -> X i -> Y i) ->
  22. {i : I} -> ⟦ C ⟧ X i -> ⟦ C ⟧ Y i
  23. map (out i) f x = x
  24. map (arg A B) f (a , b) = a , map (B a) f b
  25. map (rec i C) f (a , b) = f a , map C f b
  26. -- Tying the recursive knot. The positivity checker won't spot that
  27. -- ⟦ C ⟧ X i is strictly positive i X, so we've switched it off.
  28. -- ASR (08 August 2014): It's not necessary disable the strictly
  29. -- positive checher.
  30. data μ {I : Set}(C : Code I) : I -> Set where
  31. <_> : forall {i} -> ⟦ C ⟧ (μ C) i -> μ C i
  32. -- Who needs a primitive case-construct anyway?
  33. infix 0 case_of_
  34. case_of_ : {A : Set1}{ts : Enumeration} -> Enum ts -> Table A ts -> A
  35. case t of tbl = lookup tbl t
  36. -- The code for lists
  37. `List` : Set -> Code True
  38. `List` X = arg _ \t ->
  39. case t of
  40. "nil" ↦ out _
  41. ∣ "cons" ↦ arg X (\_ -> rec _ (out _))
  42. ∣ []
  43. -- The actual list type
  44. List : Set -> Set
  45. List A = μ (`List` A) _
  46. -- We can define the code for vectors directly. However, the point is
  47. -- that we won't have to.
  48. `VecD` : Set -> Code Nat
  49. `VecD` X = arg _ \t ->
  50. case t of
  51. "nil" ↦ out zero
  52. ∣ "cons" ↦ ( arg Nat \n ->
  53. arg X \_ ->
  54. rec n (out (suc n))
  55. )
  56. ∣ []
  57. -- An ornamentation of a datatype adds some new indexing.
  58. data Orn {I : Set}(J : I -> Set) : Code I -> Set1 where
  59. arg : forall {A B} -> ((a : A) -> Orn J (B a)) -> Orn J (arg A B)
  60. rec : forall {i C} -> J i -> Orn J C -> Orn J (rec i C)
  61. out : forall {i} -> J i -> Orn J (out i)
  62. new : forall {C} -> (A : Set) -> (A -> Orn J C) -> Orn J C
  63. -- An ornamented datatype is indexed by pairs of the old and the new index.
  64. orn : {I : Set}{J : I -> Set}{C : Code I} -> Orn J C -> Code (Σ I J)
  65. orn (out j) = out (_ , j)
  66. orn (arg B) = arg _ \a -> orn (B a)
  67. orn (new A B) = arg A \a -> orn (B a)
  68. orn (rec j C) = rec (_ , j) (orn C)
  69. -- We can forget the ornamentation and recover an element of the original type.
  70. forget' : {I : Set}{J : I -> Set}{C : Code I}{i : I}{j : J i}
  71. {X : Σ I J -> Set}{Y : I -> Set} ->
  72. ({i : I}{j : J i} -> X (i , j) -> Y i) ->
  73. (ΔC : Orn J C) -> ⟦ orn ΔC ⟧ X (i , j) -> ⟦ C ⟧ Y i
  74. forget' φ (out j) refl = refl
  75. forget' φ (arg B) (a , b) = a , forget' φ (B a) b
  76. forget' φ (new A B) (a , b) = forget' φ (B a) b
  77. forget' φ (rec j C) (a , b) = φ a , forget' φ C b
  78. -- The termination checker runs into the same problem as the positivity
  79. -- checker--it can't tell that forget' φ C x is only applying φ to
  80. -- things smaller than x.
  81. -- ASR (08 August 2014): Added the {-# TERMINATING #-} pragma.
  82. {-# TERMINATING #-}
  83. forget : {I : Set}{J : I -> Set}{C : Code I}{i : I}{j : J i}
  84. (ΔC : Orn J C) -> μ (orn ΔC) (i , j) -> μ C i
  85. forget ΔC < x > = < forget' (forget ΔC) ΔC x >
  86. -- A C-algebra over X takes us from ⟦ C ⟧ X i to X i.
  87. Alg : {I : Set} -> Code I -> (I -> Set) -> Set
  88. Alg C X = forall i -> ⟦ C ⟧ X i -> X i
  89. -- We can fold by an algebra.
  90. -- ASR (08 August 2014): Added the {-# TERMINATING #-} pragma.
  91. {-# TERMINATING #-}
  92. fold : {I : Set}{X : I -> Set}{C : Code I} ->
  93. Alg C X -> {i : I} -> μ C i -> X i
  94. fold {C = C} φ < x > = φ _ (map C (fold φ) x)
  95. -- A type can be ornamented an arbitrary algebra over its functor.
  96. decorate : {I : Set}{X : I -> Set}(C : Code I)
  97. (φ : Alg C X) -> Orn X C
  98. decorate (out i) φ = out (φ i refl)
  99. decorate (arg A B) φ = arg \a -> decorate (B a) (\i b -> φ i (a , b))
  100. decorate {X = X} (rec i C) φ = new (X i) \x -> rec x (decorate C \i b -> φ i (x , b))
  101. -- Main theorem: If you have an element in a type decorated by φ, you
  102. -- can recover forgotten decorations by folding with φ. Specialised to
  103. -- lists and vectors we get
  104. -- ∀ xs : Vec A n. length (forget xs) == n.
  105. -- Two-level definition as usual.
  106. thm' : {I : Set}{X J : I -> Set}{Y : Σ I J -> Set}
  107. (C : Code I){i : I}{j : J i}(φ : Alg C J)
  108. (F : {i : I} -> X i -> J i)
  109. (ψ : {i : I}{j : J i} -> Y (i , j) -> X i) ->
  110. ({i : I}{j : J i}(z : Y (i , j)) -> F (ψ z) == j) ->
  111. let ΔC = decorate C φ in
  112. (x : ⟦ orn ΔC ⟧ Y (i , j)) ->
  113. φ i (map C F (forget' ψ ΔC x)) == j
  114. thm' (out i) φ F ψ ih refl = refl
  115. thm' (arg A B) φ F ψ ih (a , b) = thm' (B a) (\i b -> φ i (a , b)) F ψ ih b
  116. thm' (rec i C) {i = i0}{j = j0} φ F ψ ih (j , x , c)
  117. with F (ψ x) | ih x | thm' C (\i b -> φ i (j , b)) F ψ ih c
  118. ... | .j | refl | rest = rest
  119. -- ASR (08 August 2014): Added the {-# TERMINATING #-}
  120. -- pragma.
  121. {-# TERMINATING #-}
  122. thm : {I : Set}{J : I -> Set}(C : Code I){i : I}{j : J i}(φ : Alg C J) ->
  123. (x : μ (orn (decorate C φ)) (i , j)) ->
  124. fold φ (forget (decorate C φ) x) == j
  125. thm C φ < x > = thm' C φ (fold φ) (forget (decorate C φ)) (thm C φ) x
  126. -- Vectors as decorated lists.
  127. lengthAlg : {A : Set} -> Alg (`List` A) (\_ -> Nat)
  128. lengthAlg _ (enum ."nil" zero , _) = zero
  129. lengthAlg _ (enum ."cons" (suc zero) , x , n , _) = suc n
  130. lengthAlg _ (enum _ (suc (suc ())) , _)
  131. length : {A : Set} -> List A -> Nat
  132. length = fold lengthAlg
  133. -- Now vectors are really lists decorated by their length.
  134. `Vec` : (A : Set) -> Orn (\_ -> Nat) (`List` A)
  135. `Vec` A = decorate (`List` A) lengthAlg
  136. Vec : Set -> Nat -> Set
  137. Vec A n = μ (orn (`Vec` A)) (_ , n)
  138. nil : {A : Set} -> Vec A zero
  139. nil = < enum "nil" zero , refl >
  140. cons : {A : Set}{n : Nat} -> A -> Vec A n -> Vec A (suc n)
  141. cons {n = n} x xs = < enum "cons" (suc zero) , x , n , xs , refl >
  142. -- The proof that the index of the vector is really the length follows directly
  143. -- from our main theorem.
  144. corollary : {A : Set}{n : Nat}(xs : Vec A n) ->
  145. length (forget (`Vec` _) xs) == n
  146. corollary = thm (`List` _) lengthAlg