{-# OPTIONS --guardedness #-}
module README.Foreign.Haskell where
open import Level using (Level)
open import Agda.Builtin.Int
open import Agda.Builtin.Nat
open import Data.Bool.Base using (Bool; if_then_else_)
open import Data.Char as Char
open import Data.List.Base as List using (List; _∷_; []; takeWhile; dropWhile)
open import Data.Maybe.Base using (Maybe; just; nothing)
open import Data.Product.Base using (_×_; _,_)
open import Function
open import Relation.Nullary.Decidable
import Foreign.Haskell as FFI
open import Foreign.Haskell.Coerce
private
variable
a : Level
A : Set a
postulate
primUncons : List A → Maybe (FFI.Pair A (List A))
primCatMaybes : List (Maybe A) → List A
primTestChar : Char → Bool
primIntEq : Int → Int → Bool
{-# COMPILE GHC primUncons = \ _ _ xs -> case xs of
{ [] -> Nothing
; (x : xs) -> Just (x, xs)
}
#-}
{-# FOREIGN GHC import Data.Maybe #-}
{-# COMPILE GHC primCatMaybes = \ _ _ -> catMaybes #-}
{-# COMPILE GHC primTestChar = ('-' /=) #-}
{-# COMPILE GHC primIntEq = (==) #-}
uncons : List A → Maybe (A × List A)
uncons = coerce primUncons
catMaybes : List (Maybe A) → List A
catMaybes = primCatMaybes
testChar : Char → Bool
testChar = coerce primTestChar
eqNat : Nat → Nat → Bool
eqNat = coerce primIntEq
open import IO
open import Codata.Musical.Notation
open import Data.String.Base using (toList; fromList; unlines; _++_)
open import Relation.Nullary.Negation
main = run $ do
content ← readFiniteFile "README/Foreign/Haskell.agda"
let chars = toList content
let cleanup = catMaybes ∘ List.map (λ c → if testChar c then just c else nothing)
let cleaned = dropWhile ('\n' ≟_) $ cleanup chars
case uncons cleaned of λ where
nothing → putStrLn "I cannot believe this file is filed with dashes only!"
(just (c , cs)) → putStrLn $ unlines
$ ("First (non dash) character: " ++ Char.show c)
∷ ("Rest (dash free) of the line: " ++ fromList (takeWhile (¬? ∘ ('\n' ≟_)) cs))
∷ []