Monad.agda 6.7 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227
  1. ------------------------------------------------------------------------
  2. -- Parser monad
  3. ------------------------------------------------------------------------
  4. open import Relation.Binary
  5. open import Relation.Binary.OrderMorphism
  6. open import Relation.Binary.PropositionalEquality
  7. import Relation.Binary.Properties.StrictTotalOrder as STOProps
  8. open import Data.Product
  9. open import Level
  10. module Monad
  11. -- Input string positions.
  12. {Position : Set} {_<P_ : Rel Position zero}
  13. (posOrdered : IsStrictTotalOrder _≡_ _<P_)
  14. -- Input strings.
  15. (Input : Position -> Set)
  16. -- In order to be able to store results in a memo table (and avoid
  17. -- having to lift the table code to Set1) the result types have to
  18. -- come from the following universe:
  19. {Result : Set} (⟦_⟧ : Result -> Set)
  20. -- Memoisation keys. These keys must uniquely identify the
  21. -- computation that they are associated with, when paired up with
  22. -- the current input string position.
  23. {Key : let PosPoset = STOProps.poset
  24. (record { Carrier = _ ; _≈_ = _; _<_ = _
  25. ; isStrictTotalOrder = posOrdered })
  26. MonoFun = PosPoset ⇒-Poset PosPoset in
  27. MonoFun -> Result -> Set}
  28. {_≋_ _<_ : Rel (∃₂ Key) zero}
  29. (keyOrdered : IsStrictTotalOrder _≋_ _<_)
  30. -- Furthermore the underlying equality needs to be strong enough.
  31. (funsEqual : _≋_ =[ proj₁ ]⇒ _≡_)
  32. (resultsEqual : _≋_ =[ (\rfk -> proj₁ (proj₂ rfk)) ]⇒ _≡_)
  33. where
  34. open _⇒-Poset_
  35. open STOProps (record { Carrier = _ ; _≈_ = _; _<_ = _
  36. ; isStrictTotalOrder = posOrdered })
  37. import IndexedMap as Map -- renaming (Map to MemoTable)
  38. open import Category.Monad
  39. open import Category.Monad.State
  40. import Data.List as List; open List using (List)
  41. open import Data.Unit hiding (poset; _≤_)
  42. open import Function
  43. open import Data.Maybe hiding (Eq)
  44. open import Data.Product.Relation.Lex.Strict
  45. open import Data.Product.Relation.Pointwise.NonDependent
  46. import Relation.Binary.Construct.On as On
  47. ------------------------------------------------------------------------
  48. -- Monotone functions
  49. MonoFun : Set
  50. MonoFun = poset ⇒-Poset poset
  51. ------------------------------------------------------------------------
  52. -- Memo tables
  53. -- Indices and keys used by the memo table.
  54. Index : Set
  55. Index = Position × MonoFun × Result
  56. data MemoTableKey : Index -> Set where
  57. key : forall {f r} (key : Key f r) pos -> MemoTableKey (pos , f , r)
  58. -- Input strings of a certain maximum length.
  59. Input≤ : Position -> Set
  60. Input≤ pos = ∃ \pos′ -> pos′ ≤ pos × Input pos′
  61. -- Memo table values.
  62. Value : Index -> Set
  63. Value (pos , f , r) = List (⟦ r ⟧ × Input≤ (fun f pos))
  64. -- Shuffles the elements to simplify defining equality and order
  65. -- relations for the keys.
  66. shuffle : ∃ MemoTableKey -> Position × ∃₂ Key
  67. shuffle ((pos , f , r) , key k .pos) = (pos , f , r , k)
  68. -- Equality and order.
  69. Eq : Rel (∃ MemoTableKey) _
  70. Eq = Pointwise _≡_ _≋_ on shuffle
  71. Lt : Rel (∃ MemoTableKey) _
  72. Lt = ×-Lex _≡_ _<P_ _<_ on shuffle
  73. isOrdered : IsStrictTotalOrder Eq Lt
  74. isOrdered = On.isStrictTotalOrder shuffle
  75. (×-isStrictTotalOrder posOrdered keyOrdered)
  76. indicesEqual′ : Eq =[ proj₁ ]⇒ _≡_
  77. indicesEqual′ {((_ , _ , _) , key _ ._)}
  78. {((_ , _ , _) , key _ ._)} (eq₁ , eq₂) =
  79. cong₂ _,_ eq₁ (cong₂ _,_ (funsEqual eq₂) (resultsEqual eq₂))
  80. open Map isOrdered (\{k₁} {k₂} -> indicesEqual′ {k₁} {k₂}) Value
  81. {-
  82. ------------------------------------------------------------------------
  83. -- Parser monad
  84. -- The parser monad is built upon a list monad, for backtracking, and
  85. -- two state monads. One of the state monads stores a memo table, and
  86. -- is unaffected by backtracking. The other state monad, which /is/
  87. -- affected by backtracking, stores the remaining input string.
  88. -- The memo table state monad.
  89. module MemoState = RawMonadState (StateMonadState MemoTable)
  90. -- The list monad.
  91. module List = RawMonadPlus List.ListMonadPlus
  92. -- The inner monad (memo table plus list).
  93. module IM where
  94. Inner : Set -> Set
  95. Inner R = State MemoTable (List R)
  96. InnerMonadPlus : RawMonadPlus Inner
  97. InnerMonadPlus = record
  98. { monadZero = record
  99. { monad = record
  100. { return = \x -> return (List.return x)
  101. ; _>>=_ = \m f -> List.concat <$> (List.mapM monad f =<< m)
  102. }
  103. ; ∅ = return List.∅
  104. }
  105. ; _∣_ = \m₁ m₂ -> List._∣_ <$> m₁ ⊛ m₂
  106. }
  107. where
  108. open MemoState
  109. InnerMonadState : RawMonadState MemoTable Inner
  110. InnerMonadState = record
  111. { monad = RawMonadPlus.monad InnerMonadPlus
  112. ; get = List.return <$> get
  113. ; put = \s -> List.return <$> put s
  114. }
  115. where open MemoState
  116. open RawMonadPlus InnerMonadPlus public
  117. open RawMonadState InnerMonadState public
  118. using (get; put; modify)
  119. -- The complete parser monad.
  120. module PM where
  121. P : MonoFun -> Set -> Set
  122. P f A = forall {n} -> Input n -> IM.Inner (A × Input≤ (fun f n))
  123. -- Memoises the computation, assuming that the key is sufficiently
  124. -- unique.
  125. memoise : forall {f r} -> Key f r -> P f ⟦ r ⟧ -> P f ⟦ r ⟧
  126. memoise k p {pos} xs =
  127. let open IM in helper =<< lookup k′ <$> get
  128. where
  129. i = (pos , _)
  130. k′ : MemoTableKey i
  131. k′ = key k pos
  132. helper : Maybe (Value i) -> State MemoTable (Value i)
  133. helper (just ris) = return ris where open MemoState
  134. helper nothing = p xs >>= \ris ->
  135. modify (insert k′ ris) >>
  136. return ris
  137. where open MemoState
  138. -- Other monadic operations.
  139. return : forall {A} -> A -> P idM A
  140. return a = \xs -> IM.return (a , _ , refl , xs)
  141. _>>=_ : forall {A B f g} -> P f A -> (A -> P g B) -> P (g ∘M f) B
  142. _>>=_ {g = g} m₁ m₂ xs =
  143. m₁ xs ⟨ IM._>>=_ ⟩ \ays ->
  144. let a = proj₁ ays
  145. le = proj₁ $ proj₂ $ proj₂ ays
  146. ys = proj₂ $ proj₂ $ proj₂ ays in
  147. fix le ⟨ IM._<$>_ ⟩ m₂ a ys
  148. where
  149. lemma : forall {i j k} -> j ≤ k -> i ≤ fun g j -> i ≤ fun g k
  150. lemma j≤k i≤gj = trans i≤gj (monotone g j≤k)
  151. fix : forall {A i j} -> i ≤ j ->
  152. A × Input≤ (fun g i) ->
  153. A × Input≤ (fun g j)
  154. fix le = map-× id (map-Σ id (map-× (lemma le) id))
  155. ∅ : forall {A} -> P idM A
  156. ∅ = const IM.∅
  157. _∣_ : forall {A f} -> P f A -> P f A -> P f A
  158. m₁ ∣ m₂ = \xs -> IM._∣_ (m₁ xs) (m₂ xs)
  159. put : forall {n} -> Input n -> P (constM n) ⊤
  160. put xs = \_ -> IM.return (_ , _ , refl , xs)
  161. modify : forall {A f} ->
  162. (forall {n} -> Input n -> A × Input (fun f n)) ->
  163. P f A
  164. modify g xs = IM.return (proj₁ gxs , _ , refl , proj₂ gxs)
  165. where gxs = g xs
  166. -}