|
| 1 | +module Text.Parser.Stateful |
| 2 | + |
| 3 | +import Control.Monad.State |
| 4 | +import Control.ST |
| 5 | +import Text.Parser |
| 6 | + |
| 7 | +import public Control.Delayed |
| 8 | + |
| 9 | +%access public export |
| 10 | +%default total |
| 11 | + |
| 12 | +get : StateT stateType (Grammar tok False) stateType |
| 13 | +get = ST (\x => pure (x, x)) |
| 14 | + |
| 15 | +put : stateType -> StateT stateType (Grammar tok False) () |
| 16 | +put x = ST (\y => pure ((), x)) |
| 17 | + |
| 18 | +modify : (stateType -> stateType) -> StateT stateType (Grammar tok False) () |
| 19 | +modify f = ST (\x => pure ((), f x)) |
| 20 | + |
| 21 | +lemma : (a, stateType) -> Inf (a -> StateT stateType (Grammar tok c2) b) -> Grammar tok c2 (b, stateType) |
| 22 | +lemma (v, st') k = let ST kv = k v in |
| 23 | + kv st' |
| 24 | + |
| 25 | +(>>=) : StateT stateType (Grammar tok c1) a -> |
| 26 | + inf c1 (a -> StateT stateType (Grammar tok c2) b) -> |
| 27 | + StateT stateType (Grammar tok (c1 || c2)) b |
| 28 | +(>>=) {c1 = False} (ST f) k = ST (\st => do (v, st') <- f st |
| 29 | + let ST kv = k v |
| 30 | + kv st') |
| 31 | +(>>=) {c1 = True} (ST f) k = ST (\st => do x <- f st |
| 32 | + lemma x k) |
| 33 | + |
| 34 | +pure : a -> StateT stateType (Grammar tok False) a |
| 35 | +pure x = ST (\st => pure (x, st)) |
| 36 | + |
| 37 | + |
| 38 | +map : (a -> b) -> StateT stateType (Grammar tok c) a -> |
| 39 | + StateT stateType (Grammar tok c) b |
| 40 | +map f (ST g) = ST (\st => map (mapFst f) (g st)) |
| 41 | + where mapFst : (a -> x) -> (a, s) -> (x, s) |
| 42 | + mapFst fn (a, b) = (fn a, b) |
| 43 | + |
| 44 | +-- (>>=) {c1 = True} {c2 = True} (ST f) k = ST (\st => (f st) >>= Delay (\x => ?whatNow3)) |
| 45 | + -- kv st') |
| 46 | + |
| 47 | + |
| 48 | +-- gets : (stateType -> a) -> StateT stateType (Grammar tok False) a |
| 49 | +-- gets f = do s <- get |
| 50 | +-- pure (f s) |
| 51 | + |
| 52 | + |
| 53 | +GrammarTrans : Type -> Type -> Bool -> Type -> Type |
| 54 | +GrammarTrans st tok consumes result = st -> Grammar tok consumes (result, st) |
| 55 | + |
| 56 | +identity : Grammar tok True a -> GrammarTrans st tok True a |
| 57 | +identity g state = map (\x => (x, state)) g |
| 58 | + |
| 59 | +comma : Grammar Char True () |
| 60 | +comma = terminal (\c => if c == ',' then Just () else Nothing) |
| 61 | + |
| 62 | +commaCounter : GrammarTrans Nat Char True () |
| 63 | +commaCounter n = do comma |
| 64 | + commaCounter (S n) <|> pure ((), S n) |
| 65 | + |
| 66 | +countCommas : Grammar Char True Nat |
| 67 | +countCommas = map (\(_, n) => n) (commaCounter Z) |
| 68 | + |
0 commit comments