MonadPostulates.agda 6.7 KB

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