src/Tools/Haskell/Haskell.thy
changeset 69278 30f6e8d2cd96
parent 69248 9f21381600e3
child 69280 e1d01b351724
equal deleted inserted replaced
69277:258bef08b31e 69278:30f6e8d2cd96
   512 
   512 
   513 type Output = (String, String)
   513 type Output = (String, String)
   514 
   514 
   515 no_output :: Output
   515 no_output :: Output
   516 no_output = ("", "")
   516 no_output = ("", "")
       
   517 \<close>
       
   518 
       
   519 generate_haskell_file "File.hs" = \<open>
       
   520 {-  Title:      Tools/Haskell/File.hs
       
   521     Author:     Makarius
       
   522     LICENSE:    BSD 3-clause (Isabelle)
       
   523 
       
   524 File-system operations
       
   525 -}
       
   526 
       
   527 module Isabelle.File (setup, read, write, append) where
       
   528 
       
   529 import Prelude hiding (read)
       
   530 import System.IO (IO)
       
   531 import qualified System.IO as IO
       
   532 
       
   533 setup :: IO.Handle -> IO ()
       
   534 setup h = do
       
   535   IO.hSetEncoding h IO.utf8
       
   536   IO.hSetNewlineMode h IO.noNewlineTranslation
       
   537 
       
   538 read :: IO.FilePath -> IO String
       
   539 read path =
       
   540   IO.withFile path IO.ReadMode (\h ->
       
   541     do setup h; IO.hGetContents h >>= \s -> length s `seq` return s)
       
   542 
       
   543 write :: IO.FilePath -> String -> IO ()
       
   544 write path s =
       
   545   IO.withFile path IO.WriteMode (\h -> do setup h; IO.hPutStr h s)
       
   546 
       
   547 append :: IO.FilePath -> String -> IO ()
       
   548 append path s =
       
   549   IO.withFile path IO.AppendMode (\h -> do setup h; IO.hPutStr h s)
   517 \<close>
   550 \<close>
   518 
   551 
   519 generate_haskell_file "XML.hs" = \<open>
   552 generate_haskell_file "XML.hs" = \<open>
   520 {-  Title:      Tools/Haskell/XML.hs
   553 {-  Title:      Tools/Haskell/XML.hs
   521     Author:     Makarius
   554     Author:     Makarius