| author | paulson <lp15@cam.ac.uk> | 
| Wed, 27 Oct 2021 11:47:42 +0100 | |
| changeset 74598 | 5d91897a8e54 | 
| parent 74327 | 9dca3df78b6a | 
| child 74871 | 0597884e6e91 | 
| permissions | -rw-r--r-- | 
| 69222 | 1 | (* Title: Tools/Haskell/Haskell.thy | 
| 2 | Author: Makarius | |
| 69225 | 3 | |
| 4 | Support for Isabelle tools in Haskell. | |
| 69222 | 5 | *) | 
| 6 | ||
| 7 | theory Haskell | |
| 74105 | 8 | imports Main | 
| 69222 | 9 | begin | 
| 10 | ||
| 74084 | 11 | generate_file "Isabelle/Bytes.hs" = \<open> | 
| 12 | {-  Title:      Isabelle/Bytes.hs
 | |
| 13 | Author: Makarius | |
| 14 | LICENSE: BSD 3-clause (Isabelle) | |
| 15 | ||
| 16 | Compact byte strings. | |
| 17 | ||
| 18 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/bytes.ML\<close> | |
| 19 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/General/bytes.scala\<close>. | |
| 20 | -} | |
| 21 | ||
| 74132 | 22 | {-# LANGUAGE TypeApplications #-}
 | 
| 23 | ||
| 74084 | 24 | module Isabelle.Bytes ( | 
| 25 | Bytes, | |
| 26 | make, unmake, pack, unpack, | |
| 27 | empty, null, length, index, all, any, | |
| 74216 | 28 | head, last, take, drop, isPrefixOf, isSuffixOf, try_unprefix, try_unsuffix, | 
| 74137 | 29 | concat, space, spaces, char, all_char, any_char, byte, singleton | 
| 74084 | 30 | ) | 
| 31 | where | |
| 32 | ||
| 33 | import Prelude hiding (null, length, all, any, head, last, take, drop, concat) | |
| 34 | ||
| 35 | import qualified Data.ByteString.Short as ShortByteString | |
| 36 | import Data.ByteString.Short (ShortByteString) | |
| 74090 | 37 | import qualified Data.ByteString as ByteString | 
| 74084 | 38 | import Data.ByteString (ByteString) | 
| 39 | import qualified Data.List as List | |
| 40 | import Data.Word (Word8) | |
| 74090 | 41 | import Data.Array (Array, array, (!)) | 
| 74084 | 42 | |
| 43 | ||
| 44 | type Bytes = ShortByteString | |
| 45 | ||
| 46 | make :: ByteString -> Bytes | |
| 47 | make = ShortByteString.toShort | |
| 48 | ||
| 49 | unmake :: Bytes -> ByteString | |
| 50 | unmake = ShortByteString.fromShort | |
| 51 | ||
| 52 | pack :: [Word8] -> Bytes | |
| 53 | pack = ShortByteString.pack | |
| 54 | ||
| 55 | unpack :: Bytes -> [Word8] | |
| 56 | unpack = ShortByteString.unpack | |
| 57 | ||
| 58 | empty :: Bytes | |
| 59 | empty = ShortByteString.empty | |
| 60 | ||
| 61 | null :: Bytes -> Bool | |
| 62 | null = ShortByteString.null | |
| 63 | ||
| 64 | length :: Bytes -> Int | |
| 65 | length = ShortByteString.length | |
| 66 | ||
| 67 | index :: Bytes -> Int -> Word8 | |
| 68 | index = ShortByteString.index | |
| 69 | ||
| 70 | all :: (Word8 -> Bool) -> Bytes -> Bool | |
| 71 | all p = List.all p . unpack | |
| 72 | ||
| 73 | any :: (Word8 -> Bool) -> Bytes -> Bool | |
| 74 | any p = List.any p . unpack | |
| 75 | ||
| 76 | head :: Bytes -> Word8 | |
| 77 | head bytes = index bytes 0 | |
| 78 | ||
| 79 | last :: Bytes -> Word8 | |
| 80 | last bytes = index bytes (length bytes - 1) | |
| 81 | ||
| 82 | take :: Int -> Bytes -> Bytes | |
| 74216 | 83 | take n bs | 
| 84 | | n == 0 = empty | |
| 85 | | n >= length bs = bs | |
| 86 | | otherwise = pack (List.take n (unpack bs)) | |
| 74084 | 87 | |
| 88 | drop :: Int -> Bytes -> Bytes | |
| 74216 | 89 | drop n bs | 
| 90 | | n == 0 = bs | |
| 91 | | n >= length bs = empty | |
| 92 | | otherwise = pack (List.drop n (unpack bs)) | |
| 74084 | 93 | |
| 74099 | 94 | isPrefixOf :: Bytes -> Bytes -> Bool | 
| 95 | isPrefixOf bs1 bs2 = | |
| 96 | n1 <= n2 && List.all (\i -> index bs1 i == index bs2 i) [0 .. n1 - 1] | |
| 97 | where n1 = length bs1; n2 = length bs2 | |
| 98 | ||
| 99 | isSuffixOf :: Bytes -> Bytes -> Bool | |
| 100 | isSuffixOf bs1 bs2 = | |
| 101 | n1 <= n2 && List.all (\i -> index bs1 i == index bs2 (i + k)) [0 .. n1 - 1] | |
| 102 | where n1 = length bs1; n2 = length bs2; k = n2 - n1 | |
| 103 | ||
| 74216 | 104 | try_unprefix :: Bytes -> Bytes -> Maybe Bytes | 
| 105 | try_unprefix bs1 bs2 = | |
| 106 | if isPrefixOf bs1 bs2 then Just (drop (length bs1) bs2) | |
| 107 | else Nothing | |
| 108 | ||
| 109 | try_unsuffix :: Bytes -> Bytes -> Maybe Bytes | |
| 110 | try_unsuffix bs1 bs2 = | |
| 111 | if isSuffixOf bs1 bs2 then Just (take (length bs2 - length bs1) bs2) | |
| 112 | else Nothing | |
| 113 | ||
| 74084 | 114 | concat :: [Bytes] -> Bytes | 
| 115 | concat = mconcat | |
| 74088 | 116 | |
| 74095 | 117 | space :: Word8 | 
| 118 | space = 32 | |
| 119 | ||
| 120 | small_spaces :: Array Int Bytes | |
| 121 | small_spaces = array (0, 64) [(i, pack (replicate i space)) | i <- [0 .. 64]] | |
| 122 | ||
| 123 | spaces :: Int -> Bytes | |
| 124 | spaces n = | |
| 125 | if n < 64 then small_spaces ! n | |
| 126 | else concat ((small_spaces ! (n `mod` 64)) : replicate (n `div` 64) (small_spaces ! 64)) | |
| 127 | ||
| 128 | char :: Word8 -> Char | |
| 129 | char = toEnum . fromEnum | |
| 130 | ||
| 74135 | 131 | all_char :: (Char -> Bool) -> Bytes -> Bool | 
| 132 | all_char pred = all (pred . char) | |
| 133 | ||
| 134 | any_char :: (Char -> Bool) -> Bytes -> Bool | |
| 135 | any_char pred = any (pred . char) | |
| 136 | ||
| 74095 | 137 | byte :: Char -> Word8 | 
| 138 | byte = toEnum . fromEnum | |
| 139 | ||
| 74132 | 140 | singletons :: Array Word8 Bytes | 
| 141 | singletons = | |
| 74137 | 142 | array (minBound, maxBound) | 
| 143 | [(i, make (ByteString.singleton i)) | i <- [minBound .. maxBound]] | |
| 74132 | 144 | |
| 145 | singleton :: Word8 -> Bytes | |
| 146 | singleton b = singletons ! b | |
| 74084 | 147 | \<close> | 
| 148 | ||
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 149 | generate_file "Isabelle/UTF8.hs" = \<open> | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 150 | {-  Title:      Isabelle/UTF8.hs
 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 151 | Author: Makarius | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 152 | LICENSE: BSD 3-clause (Isabelle) | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 153 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 154 | Variations on UTF-8. | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 155 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 156 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/utf8.ML\<close> | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 157 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/General/utf8.scala\<close>. | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 158 | -} | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 159 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 160 | {-# LANGUAGE MultiParamTypeClasses #-}
 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 161 | {-# LANGUAGE TypeSynonymInstances #-}
 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 162 | {-# LANGUAGE FlexibleInstances #-}
 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 163 | {-# LANGUAGE InstanceSigs #-}
 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 164 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 165 | module Isabelle.UTF8 ( | 
| 74098 | 166 | setup, setup3, | 
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 167 | Recode (..) | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 168 | ) | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 169 | where | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 170 | |
| 74098 | 171 | import qualified System.IO as IO | 
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 172 | import Data.Text (Text) | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 173 | import qualified Data.Text as Text | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 174 | import qualified Data.Text.Encoding as Encoding | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 175 | import qualified Data.Text.Encoding.Error as Error | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 176 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 177 | import Data.ByteString (ByteString) | 
| 74084 | 178 | |
| 179 | import qualified Isabelle.Bytes as Bytes | |
| 180 | import Isabelle.Bytes (Bytes) | |
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 181 | |
| 74098 | 182 | setup :: IO.Handle -> IO () | 
| 183 | setup h = do | |
| 184 | IO.hSetEncoding h IO.utf8 | |
| 185 | IO.hSetNewlineMode h IO.noNewlineTranslation | |
| 186 | ||
| 187 | setup3 :: IO.Handle -> IO.Handle -> IO.Handle -> IO () | |
| 188 | setup3 h1 h2 h3 = do | |
| 189 | setup h1 | |
| 190 | setup h2 | |
| 191 | setup h3 | |
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 192 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 193 | class Recode a b where | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 194 | encode :: a -> b | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 195 | decode :: b -> a | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 196 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 197 | instance Recode Text ByteString where | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 198 | encode :: Text -> ByteString | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 199 | encode = Encoding.encodeUtf8 | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 200 | decode :: ByteString -> Text | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 201 | decode = Encoding.decodeUtf8With Error.lenientDecode | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 202 | |
| 74084 | 203 | instance Recode Text Bytes where | 
| 204 | encode :: Text -> Bytes | |
| 205 | encode = Bytes.make . encode | |
| 206 | decode :: Bytes -> Text | |
| 207 | decode = decode . Bytes.unmake | |
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 208 | |
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 209 | instance Recode String ByteString where | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 210 | encode :: String -> ByteString | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 211 | encode = encode . Text.pack | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 212 | decode :: ByteString -> String | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 213 | decode = Text.unpack . decode | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 214 | |
| 74084 | 215 | instance Recode String Bytes where | 
| 216 | encode :: String -> Bytes | |
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 217 | encode = encode . Text.pack | 
| 74084 | 218 | decode :: Bytes -> String | 
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 219 | decode = Text.unpack . decode | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 220 | \<close> | 
| 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 221 | |
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 222 | generate_file "Isabelle/Library.hs" = \<open> | 
| 69445 | 223 | {-  Title:      Isabelle/Library.hs
 | 
| 69225 | 224 | Author: Makarius | 
| 225 | LICENSE: BSD 3-clause (Isabelle) | |
| 226 | ||
| 227 | Basic library of Isabelle idioms. | |
| 69280 | 228 | |
| 74178 | 229 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/basics.ML\<close> | 
| 230 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/library.ML\<close>. | |
| 69225 | 231 | -} | 
| 232 | ||
| 74086 | 233 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 74093 | 234 | {-# LANGUAGE TypeSynonymInstances #-}
 | 
| 235 | {-# LANGUAGE FlexibleInstances #-}
 | |
| 74134 | 236 | {-# LANGUAGE InstanceSigs #-}
 | 
| 74086 | 237 | |
| 69240 | 238 | module Isabelle.Library ( | 
| 239 | (|>), (|->), (#>), (#->), | |
| 240 | ||
| 74327 | 241 | fold, fold_rev, fold_map, single, the_single, singletonM, | 
| 242 | map_index, get_index, separate, | |
| 69240 | 243 | |
| 74093 | 244 | StringLike, STRING (..), TEXT (..), BYTES (..), | 
| 74128 | 245 | show_bytes, show_text, | 
| 74093 | 246 | |
| 74185 | 247 | proper_string, enclose, quote, space_implode, commas, commas_quote, cat_lines, | 
| 74213 | 248 | space_explode, split_lines, trim_line, trim_split_lines, | 
| 74211 | 249 | |
| 250 | getenv, getenv_strict) | |
| 69225 | 251 | where | 
| 252 | ||
| 74211 | 253 | import System.Environment (lookupEnv) | 
| 254 | import Data.Maybe (fromMaybe) | |
| 74093 | 255 | import qualified Data.Text as Text | 
| 256 | import Data.Text (Text) | |
| 74096 | 257 | import qualified Data.Text.Lazy as Lazy | 
| 74086 | 258 | import Data.String (IsString) | 
| 69453 | 259 | import qualified Data.List.Split as Split | 
| 69491 | 260 | import qualified Isabelle.Symbol as Symbol | 
| 74130 | 261 | import qualified Isabelle.Bytes as Bytes | 
| 74093 | 262 | import Isabelle.Bytes (Bytes) | 
| 263 | import qualified Isabelle.UTF8 as UTF8 | |
| 69234 | 264 | |
| 265 | ||
| 69225 | 266 | {- functions -}
 | 
| 267 | ||
| 268 | (|>) :: a -> (a -> b) -> b | |
| 269 | x |> f = f x | |
| 270 | ||
| 271 | (|->) :: (a, b) -> (a -> b -> c) -> c | |
| 272 | (x, y) |-> f = f x y | |
| 273 | ||
| 274 | (#>) :: (a -> b) -> (b -> c) -> a -> c | |
| 275 | (f #> g) x = x |> f |> g | |
| 276 | ||
| 277 | (#->) :: (a -> (c, b)) -> (c -> b -> d) -> a -> d | |
| 278 | (f #-> g) x = x |> f |-> g | |
| 279 | ||
| 280 | ||
| 281 | {- lists -}
 | |
| 282 | ||
| 283 | fold :: (a -> b -> b) -> [a] -> b -> b | |
| 284 | fold _ [] y = y | |
| 285 | fold f (x : xs) y = fold f xs (f x y) | |
| 286 | ||
| 287 | fold_rev :: (a -> b -> b) -> [a] -> b -> b | |
| 288 | fold_rev _ [] y = y | |
| 289 | fold_rev f (x : xs) y = f x (fold_rev f xs y) | |
| 290 | ||
| 74327 | 291 | fold_map :: (a -> b -> (c, b)) -> [a] -> b -> ([c], b) | 
| 292 | fold_map _ [] y = ([], y) | |
| 293 | fold_map f (x : xs) y = | |
| 294 | let | |
| 295 | (x', y') = f x y | |
| 296 | (xs', y'') = fold_map f xs y' | |
| 297 | in (x' : xs', y'') | |
| 298 | ||
| 69225 | 299 | single :: a -> [a] | 
| 300 | single x = [x] | |
| 301 | ||
| 74203 | 302 | the_single :: [a] -> a | 
| 303 | the_single [x] = x | |
| 304 | the_single _ = undefined | |
| 305 | ||
| 74204 | 306 | singletonM :: Monad m => ([a] -> m [b]) -> a -> m b | 
| 307 | singletonM f x = the_single <$> f [x] | |
| 308 | ||
| 69240 | 309 | map_index :: ((Int, a) -> b) -> [a] -> [b] | 
| 310 | map_index f = map_aux 0 | |
| 311 | where | |
| 312 | map_aux _ [] = [] | |
| 313 | map_aux i (x : xs) = f (i, x) : map_aux (i + 1) xs | |
| 314 | ||
| 315 | get_index :: (a -> Maybe b) -> [a] -> Maybe (Int, b) | |
| 316 | get_index f = get_aux 0 | |
| 317 | where | |
| 318 | get_aux _ [] = Nothing | |
| 319 | get_aux i (x : xs) = | |
| 320 | case f x of | |
| 321 | Nothing -> get_aux (i + 1) xs | |
| 322 | Just y -> Just (i, y) | |
| 323 | ||
| 74086 | 324 | separate :: a -> [a] -> [a] | 
| 74093 | 325 | separate s (x : (xs @ (_ : _))) = x : s : separate s xs | 
| 74086 | 326 | separate _ xs = xs; | 
| 327 | ||
| 69225 | 328 | |
| 74093 | 329 | {- string-like interfaces -}
 | 
| 330 | ||
| 74132 | 331 | class (IsString a, Monoid a, Eq a, Ord a) => StringLike a where | 
| 332 | space_explode :: Char -> a -> [a] | |
| 74134 | 333 | trim_line :: a -> a | 
| 74130 | 334 | |
| 74136 | 335 | gen_trim_line :: Int -> (Int -> Char) -> (Int -> a -> a) -> a -> a | 
| 336 | gen_trim_line n at trim s = | |
| 337 | if n >= 2 && at (n - 2) == '\r' && at (n - 1) == '\n' then trim (n - 2) s | |
| 338 | else if n >= 1 && Symbol.is_ascii_line_terminator (at (n - 1)) then trim (n - 1) s | |
| 339 | else s | |
| 340 | ||
| 74130 | 341 | instance StringLike String where | 
| 74134 | 342 | space_explode :: Char -> String -> [String] | 
| 74132 | 343 | space_explode c = Split.split (Split.dropDelims (Split.whenElt (== c))) | 
| 74134 | 344 | trim_line :: String -> String | 
| 74136 | 345 | trim_line s = gen_trim_line (length s) ((!!) s) take s | 
| 74130 | 346 | |
| 347 | instance StringLike Text where | |
| 74134 | 348 | space_explode :: Char -> Text -> [Text] | 
| 74132 | 349 | space_explode c str = | 
| 74130 | 350 | if Text.null str then [] | 
| 74132 | 351 | else if Text.all (/= c) str then [str] | 
| 352 | else map Text.pack $ space_explode c $ Text.unpack str | |
| 74134 | 353 | trim_line :: Text -> Text | 
| 74136 | 354 | trim_line s = gen_trim_line (Text.length s) (Text.index s) Text.take s | 
| 74130 | 355 | |
| 356 | instance StringLike Lazy.Text where | |
| 74134 | 357 | space_explode :: Char -> Lazy.Text -> [Lazy.Text] | 
| 74132 | 358 | space_explode c str = | 
| 74130 | 359 | if Lazy.null str then [] | 
| 74132 | 360 | else if Lazy.all (/= c) str then [str] | 
| 361 | else map Lazy.pack $ space_explode c $ Lazy.unpack str | |
| 74134 | 362 | trim_line :: Lazy.Text -> Lazy.Text | 
| 363 | trim_line = Lazy.fromStrict . trim_line . Lazy.toStrict | |
| 74130 | 364 | |
| 365 | instance StringLike Bytes where | |
| 74134 | 366 | space_explode :: Char -> Bytes -> [Bytes] | 
| 74132 | 367 | space_explode c str = | 
| 74130 | 368 | if Bytes.null str then [] | 
| 74135 | 369 | else if Bytes.all_char (/= c) str then [str] | 
| 74132 | 370 | else | 
| 371 | explode (Bytes.unpack str) | |
| 372 | where | |
| 373 | explode rest = | |
| 374 | case span (/= (Bytes.byte c)) rest of | |
| 375 | (_, []) -> [Bytes.pack rest] | |
| 376 | (prfx, _ : rest') -> Bytes.pack prfx : explode rest' | |
| 74134 | 377 | trim_line :: Bytes -> Bytes | 
| 74136 | 378 | trim_line s = gen_trim_line (Bytes.length s) (Bytes.char . Bytes.index s) Bytes.take s | 
| 74093 | 379 | |
| 380 | class StringLike a => STRING a where make_string :: a -> String | |
| 381 | instance STRING String where make_string = id | |
| 382 | instance STRING Text where make_string = Text.unpack | |
| 74096 | 383 | instance STRING Lazy.Text where make_string = Lazy.unpack | 
| 74093 | 384 | instance STRING Bytes where make_string = UTF8.decode | 
| 385 | ||
| 386 | class StringLike a => TEXT a where make_text :: a -> Text | |
| 387 | instance TEXT String where make_text = Text.pack | |
| 388 | instance TEXT Text where make_text = id | |
| 74096 | 389 | instance TEXT Lazy.Text where make_text = Lazy.toStrict | 
| 74093 | 390 | instance TEXT Bytes where make_text = UTF8.decode | 
| 391 | ||
| 392 | class StringLike a => BYTES a where make_bytes :: a -> Bytes | |
| 393 | instance BYTES String where make_bytes = UTF8.encode | |
| 394 | instance BYTES Text where make_bytes = UTF8.encode | |
| 74096 | 395 | instance BYTES Lazy.Text where make_bytes = UTF8.encode . Lazy.toStrict | 
| 74093 | 396 | instance BYTES Bytes where make_bytes = id | 
| 397 | ||
| 74128 | 398 | show_bytes :: Show a => a -> Bytes | 
| 399 | show_bytes = make_bytes . show | |
| 400 | ||
| 401 | show_text :: Show a => a -> Text | |
| 402 | show_text = make_text . show | |
| 403 | ||
| 74093 | 404 | |
| 69225 | 405 | {- strings -}
 | 
| 406 | ||
| 74093 | 407 | proper_string :: StringLike a => a -> Maybe a | 
| 74086 | 408 | proper_string s = if s == "" then Nothing else Just s | 
| 409 | ||
| 74185 | 410 | enclose :: StringLike a => a -> a -> a -> a | 
| 411 | enclose lpar rpar str = lpar <> str <> rpar | |
| 412 | ||
| 74093 | 413 | quote :: StringLike a => a -> a | 
| 74185 | 414 | quote = enclose "\"" "\"" | 
| 69225 | 415 | |
| 74093 | 416 | space_implode :: StringLike a => a -> [a] -> a | 
| 74086 | 417 | space_implode s = mconcat . separate s | 
| 418 | ||
| 74093 | 419 | commas, commas_quote :: StringLike a => [a] -> a | 
| 69453 | 420 | commas = space_implode ", " | 
| 421 | commas_quote = commas . map quote | |
| 422 | ||
| 74132 | 423 | split_lines :: StringLike a => a -> [a] | 
| 424 | split_lines = space_explode '\n' | |
| 425 | ||
| 74093 | 426 | cat_lines :: StringLike a => [a] -> a | 
| 69453 | 427 | cat_lines = space_implode "\n" | 
| 74211 | 428 | |
| 74213 | 429 | trim_split_lines :: StringLike a => a -> [a] | 
| 430 | trim_split_lines = trim_line #> split_lines #> map trim_line | |
| 431 | ||
| 74211 | 432 | |
| 433 | {- getenv -}
 | |
| 434 | ||
| 435 | getenv :: Bytes -> IO Bytes | |
| 436 | getenv x = do | |
| 437 | y <- lookupEnv (make_string x) | |
| 438 | return $ make_bytes $ fromMaybe "" y | |
| 439 | ||
| 440 | getenv_strict :: Bytes -> IO Bytes | |
| 441 | getenv_strict x = do | |
| 442 | y <- getenv x | |
| 443 | if Bytes.null y then | |
| 444 |     errorWithoutStackTrace $ make_string ("Undefined Isabelle environment variable: " <> quote x)
 | |
| 445 | else return y | |
| 69225 | 446 | \<close> | 
| 447 | ||
| 74091 | 448 | |
| 449 | generate_file "Isabelle/Symbol.hs" = \<open> | |
| 450 | {-  Title:      Isabelle/Symbols.hs
 | |
| 451 | Author: Makarius | |
| 452 | LICENSE: BSD 3-clause (Isabelle) | |
| 453 | ||
| 454 | Isabelle text symbols. | |
| 74172 | 455 | |
| 456 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/symbol.ML\<close> | |
| 457 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/General/symbol_explode.ML\<close>. | |
| 74091 | 458 | -} | 
| 459 | ||
| 74172 | 460 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 461 | ||
| 462 | module Isabelle.Symbol ( | |
| 74177 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 463 | Symbol, eof, is_eof, not_eof, | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 464 | |
| 74172 | 465 | is_ascii_letter, is_ascii_digit, is_ascii_hex, is_ascii_quasi, | 
| 466 | is_ascii_blank, is_ascii_line_terminator, is_ascii_letdig, | |
| 467 | is_ascii_identifier, | |
| 468 | ||
| 74177 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 469 | explode | 
| 74172 | 470 | ) | 
| 471 | where | |
| 472 | ||
| 473 | import Data.Word (Word8) | |
| 474 | import qualified Isabelle.Bytes as Bytes | |
| 475 | import Isabelle.Bytes (Bytes) | |
| 476 | ||
| 74091 | 477 | |
| 74177 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 478 | {- type -}
 | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 479 | |
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 480 | type Symbol = Bytes | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 481 | |
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 482 | eof :: Symbol | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 483 | eof = "" | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 484 | |
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 485 | is_eof, not_eof :: Symbol -> Bool | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 486 | is_eof = Bytes.null | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 487 | not_eof = not . is_eof | 
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 488 | |
| 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 489 | |
| 74091 | 490 | {- ASCII characters -}
 | 
| 491 | ||
| 492 | is_ascii_letter :: Char -> Bool | |
| 493 | is_ascii_letter c = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' | |
| 494 | ||
| 495 | is_ascii_digit :: Char -> Bool | |
| 496 | is_ascii_digit c = '0' <= c && c <= '9' | |
| 497 | ||
| 498 | is_ascii_hex :: Char -> Bool | |
| 499 | is_ascii_hex c = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' | |
| 500 | ||
| 501 | is_ascii_quasi :: Char -> Bool | |
| 502 | is_ascii_quasi c = c == '_' || c == '\'' | |
| 503 | ||
| 504 | is_ascii_blank :: Char -> Bool | |
| 74172 | 505 | is_ascii_blank c = c `elem` (" \t\n\11\f\r" :: String)
 | 
| 74091 | 506 | |
| 507 | is_ascii_line_terminator :: Char -> Bool | |
| 508 | is_ascii_line_terminator c = c == '\r' || c == '\n' | |
| 509 | ||
| 510 | is_ascii_letdig :: Char -> Bool | |
| 511 | is_ascii_letdig c = is_ascii_letter c || is_ascii_digit c || is_ascii_quasi c | |
| 512 | ||
| 513 | is_ascii_identifier :: String -> Bool | |
| 514 | is_ascii_identifier s = | |
| 515 | not (null s) && is_ascii_letter (head s) && all is_ascii_letdig s | |
| 74172 | 516 | |
| 517 | ||
| 518 | {- explode symbols: ASCII, UTF8, named -}
 | |
| 519 | ||
| 520 | is_utf8 :: Word8 -> Bool | |
| 521 | is_utf8 b = b >= 128 | |
| 522 | ||
| 523 | is_utf8_trailer :: Word8 -> Bool | |
| 524 | is_utf8_trailer b = 128 <= b && b < 192 | |
| 525 | ||
| 526 | is_utf8_control :: Word8 -> Bool | |
| 527 | is_utf8_control b = 128 <= b && b < 160 | |
| 528 | ||
| 529 | (|>) :: a -> (a -> b) -> b | |
| 530 | x |> f = f x | |
| 531 | ||
| 532 | explode :: Bytes -> [Symbol] | |
| 533 | explode string = scan 0 | |
| 534 | where | |
| 535 | byte = Bytes.index string | |
| 536 | substring i j = | |
| 537 | if i == j - 1 then Bytes.singleton (byte i) | |
| 538 | else Bytes.pack (map byte [i .. j - 1]) | |
| 539 | ||
| 540 | n = Bytes.length string | |
| 541 | test pred i = i < n && pred (byte i) | |
| 542 | test_char pred i = i < n && pred (Bytes.char (byte i)) | |
| 543 | many pred i = if test pred i then many pred (i + 1) else i | |
| 544 | maybe_char c i = if test_char (== c) i then i + 1 else i | |
| 545 | maybe_ascii_id i = | |
| 546 | if test_char is_ascii_letter i | |
| 547 | then many (is_ascii_letdig . Bytes.char) (i + 1) | |
| 548 | else i | |
| 549 | ||
| 550 | scan i = | |
| 551 | if i < n then | |
| 552 | let | |
| 553 | b = byte i | |
| 554 | c = Bytes.char b | |
| 555 | in | |
| 556 |           {-encoded newline-}
 | |
| 557 | if c == '\r' then "\n" : scan (maybe_char '\n' (i + 1)) | |
| 558 |           {-pseudo utf8: encoded ascii control-}
 | |
| 559 | else if b == 192 && test is_utf8_control (i + 1) && not (test is_utf8 (i + 2)) | |
| 560 | then Bytes.singleton (byte (i + 1) - 128) : scan (i + 2) | |
| 561 |           {-utf8-}
 | |
| 562 | else if is_utf8 b then | |
| 563 | let j = many is_utf8_trailer (i + 1) | |
| 564 | in substring i j : scan j | |
| 565 |           {-named symbol-}
 | |
| 566 | else if c == '\\' && test_char (== '<') (i + 1) then | |
| 567 | let j = (i + 2) |> maybe_char '^' |> maybe_ascii_id |> maybe_char '>' | |
| 568 | in substring i j : scan j | |
| 569 |           {-single character-}
 | |
| 570 | else Bytes.singleton b : scan (i + 1) | |
| 571 | else [] | |
| 74091 | 572 | \<close> | 
| 573 | ||
| 74095 | 574 | generate_file "Isabelle/Buffer.hs" = \<open> | 
| 575 | {-  Title:      Isabelle/Buffer.hs
 | |
| 576 | Author: Makarius | |
| 577 | LICENSE: BSD 3-clause (Isabelle) | |
| 578 | ||
| 579 | Efficient buffer of byte strings. | |
| 580 | ||
| 74178 | 581 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/buffer.ML\<close>. | 
| 74095 | 582 | -} | 
| 583 | ||
| 74231 | 584 | module Isabelle.Buffer (T, empty, add, content, build, build_content) | 
| 74095 | 585 | where | 
| 586 | ||
| 587 | import qualified Isabelle.Bytes as Bytes | |
| 588 | import Isabelle.Bytes (Bytes) | |
| 74231 | 589 | import Isabelle.Library | 
| 74095 | 590 | |
| 591 | ||
| 592 | newtype T = Buffer [Bytes] | |
| 593 | ||
| 594 | empty :: T | |
| 595 | empty = Buffer [] | |
| 596 | ||
| 597 | add :: Bytes -> T -> T | |
| 598 | add b (Buffer bs) = Buffer (if Bytes.null b then bs else b : bs) | |
| 599 | ||
| 600 | content :: T -> Bytes | |
| 601 | content (Buffer bs) = Bytes.concat (reverse bs) | |
| 74231 | 602 | |
| 603 | build :: (T -> T) -> T | |
| 604 | build f = f empty | |
| 605 | ||
| 606 | build_content :: (T -> T) -> Bytes | |
| 607 | build_content f = build f |> content | |
| 74095 | 608 | \<close> | 
| 609 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 610 | generate_file "Isabelle/Value.hs" = \<open> | 
| 73246 | 611 | {-  Title:      Isabelle/Value.hs
 | 
| 69233 | 612 | Author: Makarius | 
| 613 | LICENSE: BSD 3-clause (Isabelle) | |
| 614 | ||
| 615 | Plain values, represented as string. | |
| 69280 | 616 | |
| 74178 | 617 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/value.ML\<close>. | 
| 69233 | 618 | -} | 
| 619 | ||
| 74095 | 620 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 621 | ||
| 69233 | 622 | module Isabelle.Value | 
| 69452 | 623 | (print_bool, parse_bool, parse_nat, print_int, parse_int, print_real, parse_real) | 
| 69233 | 624 | where | 
| 625 | ||
| 626 | import qualified Data.List as List | |
| 627 | import qualified Text.Read as Read | |
| 74095 | 628 | import Isabelle.Bytes (Bytes) | 
| 629 | import Isabelle.Library | |
| 69233 | 630 | |
| 631 | ||
| 632 | {- bool -}
 | |
| 633 | ||
| 74095 | 634 | print_bool :: Bool -> Bytes | 
| 69233 | 635 | print_bool True = "true" | 
| 636 | print_bool False = "false" | |
| 637 | ||
| 74095 | 638 | parse_bool :: Bytes -> Maybe Bool | 
| 69233 | 639 | parse_bool "true" = Just True | 
| 640 | parse_bool "false" = Just False | |
| 641 | parse_bool _ = Nothing | |
| 642 | ||
| 643 | ||
| 69452 | 644 | {- nat -}
 | 
| 645 | ||
| 74095 | 646 | parse_nat :: Bytes -> Maybe Int | 
| 69452 | 647 | parse_nat s = | 
| 74095 | 648 | case Read.readMaybe (make_string s) of | 
| 69452 | 649 | Just n | n >= 0 -> Just n | 
| 650 | _ -> Nothing | |
| 651 | ||
| 652 | ||
| 69233 | 653 | {- int -}
 | 
| 654 | ||
| 74095 | 655 | print_int :: Int -> Bytes | 
| 74128 | 656 | print_int = show_bytes | 
| 74095 | 657 | |
| 658 | parse_int :: Bytes -> Maybe Int | |
| 659 | parse_int = Read.readMaybe . make_string | |
| 69233 | 660 | |
| 661 | ||
| 662 | {- real -}
 | |
| 663 | ||
| 74095 | 664 | print_real :: Double -> Bytes | 
| 69233 | 665 | print_real x = | 
| 666 | let s = show x in | |
| 667 | case span (/= '.') s of | |
| 74095 | 668 | (a, '.' : b) | List.all (== '0') b -> make_bytes a | 
| 669 | _ -> make_bytes s | |
| 670 | ||
| 671 | parse_real :: Bytes -> Maybe Double | |
| 672 | parse_real = Read.readMaybe . make_string | |
| 69225 | 673 | \<close> | 
| 674 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 675 | generate_file "Isabelle/Properties.hs" = \<open> | 
| 69445 | 676 | {-  Title:      Isabelle/Properties.hs
 | 
| 69225 | 677 | Author: Makarius | 
| 678 | LICENSE: BSD 3-clause (Isabelle) | |
| 679 | ||
| 680 | Property lists. | |
| 69280 | 681 | |
| 74178 | 682 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/properties.ML\<close>. | 
| 69225 | 683 | -} | 
| 684 | ||
| 69477 | 685 | module Isabelle.Properties (Entry, T, defined, get, get_value, put, remove) | 
| 69225 | 686 | where | 
| 687 | ||
| 688 | import qualified Data.List as List | |
| 74095 | 689 | import Isabelle.Bytes (Bytes) | 
| 690 | ||
| 691 | ||
| 692 | type Entry = (Bytes, Bytes) | |
| 69225 | 693 | type T = [Entry] | 
| 694 | ||
| 74095 | 695 | defined :: T -> Bytes -> Bool | 
| 69225 | 696 | defined props name = any (\(a, _) -> a == name) props | 
| 697 | ||
| 74095 | 698 | get :: T -> Bytes -> Maybe Bytes | 
| 69225 | 699 | get props name = List.lookup name props | 
| 700 | ||
| 74095 | 701 | get_value :: (Bytes -> Maybe a) -> T -> Bytes -> Maybe a | 
| 69477 | 702 | get_value parse props name = | 
| 703 | case get props name of | |
| 704 | Nothing -> Nothing | |
| 705 | Just s -> parse s | |
| 706 | ||
| 69225 | 707 | put :: Entry -> T -> T | 
| 708 | put entry props = entry : remove (fst entry) props | |
| 709 | ||
| 74095 | 710 | remove :: Bytes -> T -> T | 
| 69225 | 711 | remove name props = | 
| 712 | if defined props name then filter (\(a, _) -> a /= name) props | |
| 713 | else props | |
| 714 | \<close> | |
| 715 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 716 | generate_file "Isabelle/Markup.hs" = \<open> | 
| 73246 | 717 | {-  Title:      Isabelle/Markup.hs
 | 
| 69225 | 718 | Author: Makarius | 
| 719 | LICENSE: BSD 3-clause (Isabelle) | |
| 720 | ||
| 721 | Quasi-abstract markup elements. | |
| 69280 | 722 | |
| 74178 | 723 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/markup.ML\<close>. | 
| 69225 | 724 | -} | 
| 725 | ||
| 74095 | 726 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 727 | {-# OPTIONS_GHC -fno-warn-missing-signatures #-}
 | 
| 728 | ||
| 69234 | 729 | module Isabelle.Markup ( | 
| 730 | T, empty, is_empty, properties, | |
| 731 | ||
| 732 | nameN, name, xnameN, xname, kindN, | |
| 733 | ||
| 69315 | 734 | bindingN, binding, entityN, entity, defN, refN, | 
| 735 | ||
| 69288 | 736 | completionN, completion, no_completionN, no_completion, | 
| 737 | ||
| 69234 | 738 | lineN, end_lineN, offsetN, end_offsetN, fileN, idN, positionN, position, | 
| 74182 | 739 | position_properties, def_name, | 
| 69234 | 740 | |
| 69291 | 741 | expressionN, expression, | 
| 742 | ||
| 743 | citationN, citation, | |
| 744 | ||
| 71489 | 745 | pathN, path, urlN, url, docN, doc, | 
| 69291 | 746 | |
| 69248 | 747 | markupN, consistentN, unbreakableN, indentN, widthN, | 
| 748 | blockN, block, breakN, break, fbreakN, fbreak, itemN, item, | |
| 749 | ||
| 69968 
1a400b14fd3a
clarified spell-checking (see also 30233285270a);
 wenzelm parents: 
69794diff
changeset | 750 | wordsN, words, | 
| 69234 | 751 | |
| 752 | tfreeN, tfree, tvarN, tvar, freeN, free, skolemN, skolem, boundN, bound, varN, var, | |
| 753 | numeralN, numeral, literalN, literal, delimiterN, delimiter, inner_stringN, inner_string, | |
| 69320 | 754 | inner_cartoucheN, inner_cartouche, | 
| 69234 | 755 | token_rangeN, token_range, | 
| 756 | sortingN, sorting, typingN, typing, class_parameterN, class_parameter, | |
| 757 | ||
| 758 | antiquotedN, antiquoted, antiquoteN, antiquote, | |
| 759 | ||
| 760 | paragraphN, paragraph, text_foldN, text_fold, | |
| 761 | ||
| 762 | keyword1N, keyword1, keyword2N, keyword2, keyword3N, keyword3, quasi_keywordN, quasi_keyword, | |
| 763 | improperN, improper, operatorN, operator, stringN, string, alt_stringN, alt_string, | |
| 69320 | 764 | verbatimN, verbatim, cartoucheN, cartouche, commentN, comment, comment1N, comment1, | 
| 765 | comment2N, comment2, comment3N, comment3, | |
| 69234 | 766 | |
| 70667 | 767 | forkedN, forked, joinedN, joined, runningN, running, finishedN, finished, | 
| 69794 | 768 | failedN, failed, canceledN, canceled, initializedN, initialized, finalizedN, finalized, | 
| 769 | consolidatedN, consolidated, | |
| 770 | ||
| 69234 | 771 | writelnN, writeln, stateN, state, informationN, information, tracingN, tracing, | 
| 772 | warningN, warning, legacyN, legacy, errorN, error, reportN, report, no_reportN, no_report, | |
| 773 | ||
| 774 | intensifyN, intensify, | |
| 775 | Output, no_output) | |
| 69225 | 776 | where | 
| 777 | ||
| 69248 | 778 | import Prelude hiding (words, error, break) | 
| 74182 | 779 | import Data.Map.Strict (Map) | 
| 780 | import qualified Data.Map.Strict as Map | |
| 69234 | 781 | |
| 782 | import Isabelle.Library | |
| 69225 | 783 | import qualified Isabelle.Properties as Properties | 
| 69248 | 784 | import qualified Isabelle.Value as Value | 
| 74095 | 785 | import qualified Isabelle.Bytes as Bytes | 
| 786 | import Isabelle.Bytes (Bytes) | |
| 69225 | 787 | |
| 788 | ||
| 69234 | 789 | {- basic markup -}
 | 
| 790 | ||
| 74095 | 791 | type T = (Bytes, Properties.T) | 
| 69225 | 792 | |
| 793 | empty :: T | |
| 794 | empty = ("", [])
 | |
| 795 | ||
| 796 | is_empty :: T -> Bool | |
| 797 | is_empty ("", _) = True
 | |
| 798 | is_empty _ = False | |
| 799 | ||
| 69234 | 800 | properties :: Properties.T -> T -> T | 
| 801 | properties more_props (elem, props) = | |
| 802 | (elem, fold_rev Properties.put more_props props) | |
| 803 | ||
| 74095 | 804 | markup_elem :: Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 805 | markup_elem name = (name, []) | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 806 | |
| 74095 | 807 | markup_string :: Bytes -> Bytes -> Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 808 | markup_string name prop = \s -> (name, [(prop, s)]) | 
| 69234 | 809 | |
| 810 | ||
| 811 | {- misc properties -}
 | |
| 812 | ||
| 74095 | 813 | nameN :: Bytes | 
| 69234 | 814 | nameN = \<open>Markup.nameN\<close> | 
| 815 | ||
| 74095 | 816 | name :: Bytes -> T -> T | 
| 69234 | 817 | name a = properties [(nameN, a)] | 
| 818 | ||
| 74095 | 819 | xnameN :: Bytes | 
| 69234 | 820 | xnameN = \<open>Markup.xnameN\<close> | 
| 821 | ||
| 74095 | 822 | xname :: Bytes -> T -> T | 
| 69234 | 823 | xname a = properties [(xnameN, a)] | 
| 824 | ||
| 74095 | 825 | kindN :: Bytes | 
| 69234 | 826 | kindN = \<open>Markup.kindN\<close> | 
| 827 | ||
| 828 | ||
| 69315 | 829 | {- formal entities -}
 | 
| 830 | ||
| 74095 | 831 | bindingN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 832 | bindingN = \<open>Markup.bindingN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 833 | binding :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 834 | binding = markup_elem bindingN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 835 | |
| 74095 | 836 | entityN :: Bytes | 
| 69315 | 837 | entityN = \<open>Markup.entityN\<close> | 
| 74095 | 838 | entity :: Bytes -> Bytes -> T | 
| 69315 | 839 | entity kind name = | 
| 840 | (entityN, | |
| 74095 | 841 | (if Bytes.null name then [] else [(nameN, name)]) <> | 
| 842 | (if Bytes.null kind then [] else [(kindN, kind)])) | |
| 843 | ||
| 844 | defN :: Bytes | |
| 69315 | 845 | defN = \<open>Markup.defN\<close> | 
| 846 | ||
| 74095 | 847 | refN :: Bytes | 
| 69315 | 848 | refN = \<open>Markup.refN\<close> | 
| 849 | ||
| 850 | ||
| 69288 | 851 | {- completion -}
 | 
| 852 | ||
| 74095 | 853 | completionN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 854 | completionN = \<open>Markup.completionN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 855 | completion :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 856 | completion = markup_elem completionN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 857 | |
| 74095 | 858 | no_completionN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 859 | no_completionN = \<open>Markup.no_completionN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 860 | no_completion :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 861 | no_completion = markup_elem no_completionN | 
| 69288 | 862 | |
| 863 | ||
| 69234 | 864 | {- position -}
 | 
| 865 | ||
| 74095 | 866 | lineN, end_lineN :: Bytes | 
| 69234 | 867 | lineN = \<open>Markup.lineN\<close> | 
| 868 | end_lineN = \<open>Markup.end_lineN\<close> | |
| 869 | ||
| 74095 | 870 | offsetN, end_offsetN :: Bytes | 
| 69234 | 871 | offsetN = \<open>Markup.offsetN\<close> | 
| 872 | end_offsetN = \<open>Markup.end_offsetN\<close> | |
| 873 | ||
| 74095 | 874 | fileN, idN :: Bytes | 
| 69234 | 875 | fileN = \<open>Markup.fileN\<close> | 
| 876 | idN = \<open>Markup.idN\<close> | |
| 877 | ||
| 74095 | 878 | positionN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 879 | positionN = \<open>Markup.positionN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 880 | position :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 881 | position = markup_elem positionN | 
| 69234 | 882 | |
| 74182 | 883 | position_properties :: [Bytes] | 
| 884 | position_properties = [lineN, offsetN, end_offsetN, fileN, idN] | |
| 885 | ||
| 886 | ||
| 887 | {- position "def" names -}
 | |
| 888 | ||
| 889 | make_def :: Bytes -> Bytes | |
| 890 | make_def a = "def_" <> a | |
| 891 | ||
| 892 | def_names :: Map Bytes Bytes | |
| 893 | def_names = Map.fromList $ map (\a -> (a, make_def a)) position_properties | |
| 894 | ||
| 895 | def_name :: Bytes -> Bytes | |
| 896 | def_name a = | |
| 897 | case Map.lookup a def_names of | |
| 898 | Just b -> b | |
| 899 | Nothing -> make_def a | |
| 900 | ||
| 69234 | 901 | |
| 69291 | 902 | {- expression -}
 | 
| 903 | ||
| 74095 | 904 | expressionN :: Bytes | 
| 69291 | 905 | expressionN = \<open>Markup.expressionN\<close> | 
| 906 | ||
| 74095 | 907 | expression :: Bytes -> T | 
| 69291 | 908 | expression kind = (expressionN, if kind == "" then [] else [(kindN, kind)]) | 
| 909 | ||
| 910 | ||
| 911 | {- citation -}
 | |
| 912 | ||
| 74095 | 913 | citationN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 914 | citationN = \<open>Markup.citationN\<close> | 
| 74095 | 915 | citation :: Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 916 | citation = markup_string nameN citationN | 
| 69291 | 917 | |
| 918 | ||
| 919 | {- external resources -}
 | |
| 920 | ||
| 74095 | 921 | pathN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 922 | pathN = \<open>Markup.pathN\<close> | 
| 74095 | 923 | path :: Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 924 | path = markup_string pathN nameN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 925 | |
| 74095 | 926 | urlN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 927 | urlN = \<open>Markup.urlN\<close> | 
| 74095 | 928 | url :: Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 929 | url = markup_string urlN nameN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 930 | |
| 74095 | 931 | docN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 932 | docN = \<open>Markup.docN\<close> | 
| 74095 | 933 | doc :: Bytes -> T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 934 | doc = markup_string docN nameN | 
| 69291 | 935 | |
| 936 | ||
| 69248 | 937 | {- pretty printing -}
 | 
| 938 | ||
| 74095 | 939 | markupN, consistentN, unbreakableN, indentN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 940 | markupN = \<open>Markup.markupN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 941 | consistentN = \<open>Markup.consistentN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 942 | unbreakableN = \<open>Markup.unbreakableN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 943 | indentN = \<open>Markup.indentN\<close> | 
| 69248 | 944 | |
| 74095 | 945 | widthN :: Bytes | 
| 69248 | 946 | widthN = \<open>Markup.widthN\<close> | 
| 947 | ||
| 74095 | 948 | blockN :: Bytes | 
| 69248 | 949 | blockN = \<open>Markup.blockN\<close> | 
| 950 | block :: Bool -> Int -> T | |
| 951 | block c i = | |
| 952 | (blockN, | |
| 74081 | 953 | (if c then [(consistentN, Value.print_bool c)] else []) <> | 
| 69248 | 954 | (if i /= 0 then [(indentN, Value.print_int i)] else [])) | 
| 955 | ||
| 74095 | 956 | breakN :: Bytes | 
| 69248 | 957 | breakN = \<open>Markup.breakN\<close> | 
| 958 | break :: Int -> Int -> T | |
| 959 | break w i = | |
| 960 | (breakN, | |
| 74081 | 961 | (if w /= 0 then [(widthN, Value.print_int w)] else []) <> | 
| 69248 | 962 | (if i /= 0 then [(indentN, Value.print_int i)] else [])) | 
| 963 | ||
| 74095 | 964 | fbreakN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 965 | fbreakN = \<open>Markup.fbreakN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 966 | fbreak :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 967 | fbreak = markup_elem fbreakN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 968 | |
| 74095 | 969 | itemN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 970 | itemN = \<open>Markup.itemN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 971 | item :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 972 | item = markup_elem itemN | 
| 69248 | 973 | |
| 974 | ||
| 69234 | 975 | {- text properties -}
 | 
| 976 | ||
| 74095 | 977 | wordsN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 978 | wordsN = \<open>Markup.wordsN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 979 | words :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 980 | words = markup_elem wordsN | 
| 69234 | 981 | |
| 982 | ||
| 983 | {- inner syntax -}
 | |
| 984 | ||
| 74095 | 985 | tfreeN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 986 | tfreeN = \<open>Markup.tfreeN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 987 | tfree :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 988 | tfree = markup_elem tfreeN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 989 | |
| 74095 | 990 | tvarN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 991 | tvarN = \<open>Markup.tvarN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 992 | tvar :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 993 | tvar = markup_elem tvarN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 994 | |
| 74095 | 995 | freeN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 996 | freeN = \<open>Markup.freeN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 997 | free :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 998 | free = markup_elem freeN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 999 | |
| 74095 | 1000 | skolemN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1001 | skolemN = \<open>Markup.skolemN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1002 | skolem :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1003 | skolem = markup_elem skolemN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1004 | |
| 74095 | 1005 | boundN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1006 | boundN = \<open>Markup.boundN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1007 | bound :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1008 | bound = markup_elem boundN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1009 | |
| 74095 | 1010 | varN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1011 | varN = \<open>Markup.varN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1012 | var :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1013 | var = markup_elem varN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1014 | |
| 74095 | 1015 | numeralN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1016 | numeralN = \<open>Markup.numeralN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1017 | numeral :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1018 | numeral = markup_elem numeralN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1019 | |
| 74095 | 1020 | literalN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1021 | literalN = \<open>Markup.literalN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1022 | literal :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1023 | literal = markup_elem literalN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1024 | |
| 74095 | 1025 | delimiterN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1026 | delimiterN = \<open>Markup.delimiterN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1027 | delimiter :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1028 | delimiter = markup_elem delimiterN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1029 | |
| 74095 | 1030 | inner_stringN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1031 | inner_stringN = \<open>Markup.inner_stringN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1032 | inner_string :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1033 | inner_string = markup_elem inner_stringN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1034 | |
| 74095 | 1035 | inner_cartoucheN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1036 | inner_cartoucheN = \<open>Markup.inner_cartoucheN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1037 | inner_cartouche :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1038 | inner_cartouche = markup_elem inner_cartoucheN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1039 | |
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1040 | |
| 74095 | 1041 | token_rangeN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1042 | token_rangeN = \<open>Markup.token_rangeN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1043 | token_range :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1044 | token_range = markup_elem token_rangeN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1045 | |
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1046 | |
| 74095 | 1047 | sortingN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1048 | sortingN = \<open>Markup.sortingN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1049 | sorting :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1050 | sorting = markup_elem sortingN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1051 | |
| 74095 | 1052 | typingN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1053 | typingN = \<open>Markup.typingN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1054 | typing :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1055 | typing = markup_elem typingN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1056 | |
| 74095 | 1057 | class_parameterN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1058 | class_parameterN = \<open>Markup.class_parameterN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1059 | class_parameter :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1060 | class_parameter = markup_elem class_parameterN | 
| 69234 | 1061 | |
| 1062 | ||
| 1063 | {- antiquotations -}
 | |
| 1064 | ||
| 74095 | 1065 | antiquotedN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1066 | antiquotedN = \<open>Markup.antiquotedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1067 | antiquoted :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1068 | antiquoted = markup_elem antiquotedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1069 | |
| 74095 | 1070 | antiquoteN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1071 | antiquoteN = \<open>Markup.antiquoteN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1072 | antiquote :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1073 | antiquote = markup_elem antiquoteN | 
| 69234 | 1074 | |
| 1075 | ||
| 1076 | {- text structure -}
 | |
| 1077 | ||
| 74095 | 1078 | paragraphN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1079 | paragraphN = \<open>Markup.paragraphN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1080 | paragraph :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1081 | paragraph = markup_elem paragraphN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1082 | |
| 74095 | 1083 | text_foldN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1084 | text_foldN = \<open>Markup.text_foldN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1085 | text_fold :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1086 | text_fold = markup_elem text_foldN | 
| 69234 | 1087 | |
| 1088 | ||
| 1089 | {- outer syntax -}
 | |
| 1090 | ||
| 74095 | 1091 | keyword1N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1092 | keyword1N = \<open>Markup.keyword1N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1093 | keyword1 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1094 | keyword1 = markup_elem keyword1N | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1095 | |
| 74095 | 1096 | keyword2N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1097 | keyword2N = \<open>Markup.keyword2N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1098 | keyword2 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1099 | keyword2 = markup_elem keyword2N | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1100 | |
| 74095 | 1101 | keyword3N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1102 | keyword3N = \<open>Markup.keyword3N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1103 | keyword3 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1104 | keyword3 = markup_elem keyword3N | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1105 | |
| 74095 | 1106 | quasi_keywordN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1107 | quasi_keywordN = \<open>Markup.quasi_keywordN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1108 | quasi_keyword :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1109 | quasi_keyword = markup_elem quasi_keywordN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1110 | |
| 74095 | 1111 | improperN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1112 | improperN = \<open>Markup.improperN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1113 | improper :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1114 | improper = markup_elem improperN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1115 | |
| 74095 | 1116 | operatorN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1117 | operatorN = \<open>Markup.operatorN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1118 | operator :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1119 | operator = markup_elem operatorN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1120 | |
| 74095 | 1121 | stringN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1122 | stringN = \<open>Markup.stringN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1123 | string :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1124 | string = markup_elem stringN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1125 | |
| 74095 | 1126 | alt_stringN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1127 | alt_stringN = \<open>Markup.alt_stringN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1128 | alt_string :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1129 | alt_string = markup_elem alt_stringN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1130 | |
| 74095 | 1131 | verbatimN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1132 | verbatimN = \<open>Markup.verbatimN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1133 | verbatim :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1134 | verbatim = markup_elem verbatimN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1135 | |
| 74095 | 1136 | cartoucheN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1137 | cartoucheN = \<open>Markup.cartoucheN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1138 | cartouche :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1139 | cartouche = markup_elem cartoucheN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1140 | |
| 74095 | 1141 | commentN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1142 | commentN = \<open>Markup.commentN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1143 | comment :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1144 | comment = markup_elem commentN | 
| 69234 | 1145 | |
| 1146 | ||
| 69320 | 1147 | {- comments -}
 | 
| 1148 | ||
| 74095 | 1149 | comment1N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1150 | comment1N = \<open>Markup.comment1N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1151 | comment1 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1152 | comment1 = markup_elem comment1N | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1153 | |
| 74095 | 1154 | comment2N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1155 | comment2N = \<open>Markup.comment2N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1156 | comment2 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1157 | comment2 = markup_elem comment2N | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1158 | |
| 74095 | 1159 | comment3N :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1160 | comment3N = \<open>Markup.comment3N\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1161 | comment3 :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1162 | comment3 = markup_elem comment3N | 
| 69320 | 1163 | |
| 1164 | ||
| 69793 | 1165 | {- command status -}
 | 
| 1166 | ||
| 70667 | 1167 | forkedN, joinedN, runningN, finishedN, failedN, canceledN, | 
| 74095 | 1168 | initializedN, finalizedN, consolidatedN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1169 | forkedN = \<open>Markup.forkedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1170 | joinedN = \<open>Markup.joinedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1171 | runningN = \<open>Markup.runningN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1172 | finishedN = \<open>Markup.finishedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1173 | failedN = \<open>Markup.failedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1174 | canceledN = \<open>Markup.canceledN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1175 | initializedN = \<open>Markup.initializedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1176 | finalizedN = \<open>Markup.finalizedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1177 | consolidatedN = \<open>Markup.consolidatedN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1178 | |
| 70667 | 1179 | forked, joined, running, finished, failed, canceled, | 
| 69793 | 1180 | initialized, finalized, consolidated :: T | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1181 | forked = markup_elem forkedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1182 | joined = markup_elem joinedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1183 | running = markup_elem runningN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1184 | finished = markup_elem finishedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1185 | failed = markup_elem failedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1186 | canceled = markup_elem canceledN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1187 | initialized = markup_elem initializedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1188 | finalized = markup_elem finalizedN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1189 | consolidated = markup_elem consolidatedN | 
| 69793 | 1190 | |
| 1191 | ||
| 69234 | 1192 | {- messages -}
 | 
| 1193 | ||
| 74095 | 1194 | writelnN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1195 | writelnN = \<open>Markup.writelnN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1196 | writeln :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1197 | writeln = markup_elem writelnN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1198 | |
| 74095 | 1199 | stateN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1200 | stateN = \<open>Markup.stateN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1201 | state :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1202 | state = markup_elem stateN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1203 | |
| 74095 | 1204 | informationN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1205 | informationN = \<open>Markup.informationN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1206 | information :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1207 | information = markup_elem informationN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1208 | |
| 74095 | 1209 | tracingN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1210 | tracingN = \<open>Markup.tracingN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1211 | tracing :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1212 | tracing = markup_elem tracingN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1213 | |
| 74095 | 1214 | warningN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1215 | warningN = \<open>Markup.warningN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1216 | warning :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1217 | warning = markup_elem warningN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1218 | |
| 74095 | 1219 | legacyN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1220 | legacyN = \<open>Markup.legacyN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1221 | legacy :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1222 | legacy = markup_elem legacyN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1223 | |
| 74095 | 1224 | errorN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1225 | errorN = \<open>Markup.errorN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1226 | error :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1227 | error = markup_elem errorN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1228 | |
| 74095 | 1229 | reportN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1230 | reportN = \<open>Markup.reportN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1231 | report :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1232 | report = markup_elem reportN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1233 | |
| 74095 | 1234 | no_reportN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1235 | no_reportN = \<open>Markup.no_reportN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1236 | no_report :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1237 | no_report = markup_elem no_reportN | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1238 | |
| 74095 | 1239 | intensifyN :: Bytes | 
| 73199 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1240 | intensifyN = \<open>Markup.intensifyN\<close> | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1241 | intensify :: T | 
| 
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
 wenzelm parents: 
73178diff
changeset | 1242 | intensify = markup_elem intensifyN | 
| 69234 | 1243 | |
| 1244 | ||
| 1245 | {- output -}
 | |
| 69225 | 1246 | |
| 74095 | 1247 | type Output = (Bytes, Bytes) | 
| 69225 | 1248 | |
| 1249 | no_output :: Output | |
| 1250 | no_output = ("", "")
 | |
| 69222 | 1251 | \<close> | 
| 1252 | ||
| 74167 | 1253 | generate_file "Isabelle/Position.hs" = \<open> | 
| 1254 | {-  Title:      Isabelle/Position.hs
 | |
| 1255 | Author: Makarius | |
| 1256 | LICENSE: BSD 3-clause (Isabelle) | |
| 1257 | ||
| 74173 | 1258 | Source positions starting from 1; values <= 0 mean "absent". Count Isabelle | 
| 1259 | symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a | |
| 1260 | right-open interval offset .. end_offset (exclusive). | |
| 74167 | 1261 | |
| 74178 | 1262 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/position.ML\<close>. | 
| 74167 | 1263 | -} | 
| 1264 | ||
| 1265 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 1266 | ||
| 1267 | module Isabelle.Position ( | |
| 1268 | T, line_of, column_of, offset_of, end_offset_of, file_of, id_of, | |
| 74185 | 1269 | start, none, put_file, file, file_only, put_id, id, id_only, | 
| 74174 | 1270 | symbol, symbol_explode, symbol_explode_string, shift_offsets, | 
| 74184 | 1271 | of_properties, properties_of, def_properties_of, entity_markup, make_entity_markup, | 
| 74187 | 1272 | Report, Report_Text, is_reported, is_reported_range, here, | 
| 74167 | 1273 | Range, no_range, no_range_position, range_position, range | 
| 74184 | 1274 | ) | 
| 1275 | where | |
| 74167 | 1276 | |
| 74185 | 1277 | import Prelude hiding (id) | 
| 74168 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1278 | import Data.Maybe (isJust, fromMaybe) | 
| 74182 | 1279 | import Data.Bifunctor (first) | 
| 74167 | 1280 | import qualified Isabelle.Properties as Properties | 
| 1281 | import qualified Isabelle.Bytes as Bytes | |
| 1282 | import qualified Isabelle.Value as Value | |
| 1283 | import Isabelle.Bytes (Bytes) | |
| 1284 | import qualified Isabelle.Markup as Markup | |
| 1285 | import qualified Isabelle.YXML as YXML | |
| 1286 | import Isabelle.Library | |
| 74173 | 1287 | import qualified Isabelle.Symbol as Symbol | 
| 1288 | import Isabelle.Symbol (Symbol) | |
| 74167 | 1289 | |
| 1290 | ||
| 1291 | {- position -}
 | |
| 1292 | ||
| 1293 | data T = | |
| 1294 |   Position {
 | |
| 1295 | _line :: Int, | |
| 1296 | _column :: Int, | |
| 1297 | _offset :: Int, | |
| 1298 | _end_offset :: Int, | |
| 1299 | _file :: Bytes, | |
| 1300 | _id :: Bytes } | |
| 1301 | deriving (Eq, Ord) | |
| 1302 | ||
| 74173 | 1303 | valid, invalid :: Int -> Bool | 
| 1304 | valid i = i > 0 | |
| 1305 | invalid = not . valid | |
| 74167 | 1306 | |
| 1307 | maybe_valid :: Int -> Maybe Int | |
| 74173 | 1308 | maybe_valid i = if valid i then Just i else Nothing | 
| 1309 | ||
| 1310 | if_valid :: Int -> Int -> Int | |
| 1311 | if_valid i i' = if valid i then i' else i | |
| 74167 | 1312 | |
| 1313 | ||
| 1314 | {- fields -}
 | |
| 1315 | ||
| 1316 | line_of, column_of, offset_of, end_offset_of :: T -> Maybe Int | |
| 1317 | line_of = maybe_valid . _line | |
| 1318 | column_of = maybe_valid . _column | |
| 1319 | offset_of = maybe_valid . _offset | |
| 1320 | end_offset_of = maybe_valid . _end_offset | |
| 1321 | ||
| 1322 | file_of :: T -> Maybe Bytes | |
| 1323 | file_of = proper_string . _file | |
| 1324 | ||
| 1325 | id_of :: T -> Maybe Bytes | |
| 1326 | id_of = proper_string . _id | |
| 1327 | ||
| 1328 | ||
| 1329 | {- make position -}
 | |
| 1330 | ||
| 1331 | start :: T | |
| 1332 | start = Position 1 1 1 0 Bytes.empty Bytes.empty | |
| 1333 | ||
| 1334 | none :: T | |
| 1335 | none = Position 0 0 0 0 Bytes.empty Bytes.empty | |
| 1336 | ||
| 1337 | put_file :: Bytes -> T -> T | |
| 1338 | put_file file pos = pos { _file = file }
 | |
| 1339 | ||
| 1340 | file :: Bytes -> T | |
| 1341 | file file = put_file file start | |
| 1342 | ||
| 1343 | file_only :: Bytes -> T | |
| 1344 | file_only file = put_file file none | |
| 1345 | ||
| 1346 | put_id :: Bytes -> T -> T | |
| 1347 | put_id id pos = pos { _id = id }
 | |
| 1348 | ||
| 74185 | 1349 | id :: Bytes -> T | 
| 1350 | id id = put_id id start | |
| 1351 | ||
| 1352 | id_only :: Bytes -> T | |
| 1353 | id_only id = put_id id none | |
| 1354 | ||
| 74167 | 1355 | |
| 74176 | 1356 | {- count position -}
 | 
| 1357 | ||
| 1358 | count_line :: Symbol -> Int -> Int | |
| 1359 | count_line "\n" line = if_valid line (line + 1) | |
| 1360 | count_line _ line = line | |
| 1361 | ||
| 1362 | count_column :: Symbol -> Int -> Int | |
| 1363 | count_column "\n" column = if_valid column 1 | |
| 74177 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 1364 | count_column s column = if Symbol.not_eof s then if_valid column (column + 1) else column | 
| 74176 | 1365 | |
| 1366 | count_offset :: Symbol -> Int -> Int | |
| 74177 
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
 wenzelm parents: 
74176diff
changeset | 1367 | count_offset s offset = if Symbol.not_eof s then if_valid offset (offset + 1) else offset | 
| 74173 | 1368 | |
| 74174 | 1369 | symbol :: Symbol -> T -> T | 
| 1370 | symbol s pos = | |
| 74167 | 1371 |   pos {
 | 
| 74176 | 1372 | _line = count_line s (_line pos), | 
| 1373 | _column = count_column s (_column pos), | |
| 1374 | _offset = count_offset s (_offset pos) } | |
| 74173 | 1375 | |
| 74174 | 1376 | symbol_explode :: BYTES a => a -> T -> T | 
| 1377 | symbol_explode = fold symbol . Symbol.explode . make_bytes | |
| 1378 | ||
| 1379 | symbol_explode_string :: String -> T -> T | |
| 1380 | symbol_explode_string = symbol_explode | |
| 74167 | 1381 | |
| 1382 | ||
| 1383 | {- shift offsets -}
 | |
| 1384 | ||
| 1385 | shift_offsets :: Int -> T -> T | |
| 1386 | shift_offsets shift pos = pos { _offset = offset', _end_offset = end_offset' }
 | |
| 1387 | where | |
| 1388 | offset = _offset pos | |
| 1389 | end_offset = _end_offset pos | |
| 1390 | offset' = if invalid offset || invalid shift then offset else offset + shift | |
| 1391 | end_offset' = if invalid end_offset || invalid shift then end_offset else end_offset + shift | |
| 1392 | ||
| 1393 | ||
| 1394 | {- markup properties -}
 | |
| 1395 | ||
| 1396 | get_string :: Properties.T -> Bytes -> Bytes | |
| 74168 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1397 | get_string props name = fromMaybe "" (Properties.get_value Just props name) | 
| 74167 | 1398 | |
| 1399 | get_int :: Properties.T -> Bytes -> Int | |
| 74168 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1400 | get_int props name = fromMaybe 0 (Properties.get_value Value.parse_int props name) | 
| 74167 | 1401 | |
| 1402 | of_properties :: Properties.T -> T | |
| 1403 | of_properties props = | |
| 1404 |   none {
 | |
| 1405 | _line = get_int props Markup.lineN, | |
| 1406 | _offset = get_int props Markup.offsetN, | |
| 1407 | _end_offset = get_int props Markup.end_offsetN, | |
| 1408 | _file = get_string props Markup.fileN, | |
| 1409 | _id = get_string props Markup.idN } | |
| 1410 | ||
| 1411 | string_entry :: Bytes -> Bytes -> Properties.T | |
| 1412 | string_entry k s = if Bytes.null s then [] else [(k, s)] | |
| 1413 | ||
| 1414 | int_entry :: Bytes -> Int -> Properties.T | |
| 1415 | int_entry k i = if invalid i then [] else [(k, Value.print_int i)] | |
| 1416 | ||
| 1417 | properties_of :: T -> Properties.T | |
| 1418 | properties_of pos = | |
| 1419 | int_entry Markup.lineN (_line pos) ++ | |
| 1420 | int_entry Markup.offsetN (_offset pos) ++ | |
| 1421 | int_entry Markup.end_offsetN (_end_offset pos) ++ | |
| 1422 | string_entry Markup.fileN (_file pos) ++ | |
| 1423 | string_entry Markup.idN (_id pos) | |
| 1424 | ||
| 1425 | def_properties_of :: T -> Properties.T | |
| 74182 | 1426 | def_properties_of = properties_of #> map (first Markup.def_name) | 
| 74167 | 1427 | |
| 1428 | entity_markup :: Bytes -> (Bytes, T) -> Markup.T | |
| 1429 | entity_markup kind (name, pos) = | |
| 1430 | Markup.entity kind name |> Markup.properties (def_properties_of pos) | |
| 1431 | ||
| 74184 | 1432 | make_entity_markup :: Bool -> Int -> Bytes -> (Bytes, T) -> Markup.T | 
| 1433 | make_entity_markup def serial kind (name, pos) = | |
| 1434 | let | |
| 1435 | props = | |
| 1436 | if def then (Markup.defN, Value.print_int serial) : properties_of pos | |
| 1437 | else (Markup.refN, Value.print_int serial) : def_properties_of pos | |
| 1438 | in Markup.entity kind name |> Markup.properties props | |
| 74167 | 1439 | |
| 1440 | ||
| 1441 | {- reports -}
 | |
| 1442 | ||
| 74187 | 1443 | type Report = (T, Markup.T) | 
| 1444 | type Report_Text = (Report, Bytes) | |
| 1445 | ||
| 74167 | 1446 | is_reported :: T -> Bool | 
| 1447 | is_reported pos = isJust (offset_of pos) && isJust (id_of pos) | |
| 1448 | ||
| 1449 | is_reported_range :: T -> Bool | |
| 1450 | is_reported_range pos = is_reported pos && isJust (end_offset_of pos) | |
| 1451 | ||
| 1452 | ||
| 1453 | {- here: user output -}
 | |
| 1454 | ||
| 1455 | here :: T -> Bytes | |
| 1456 | here pos = if Bytes.null s2 then "" else s1 <> m1 <> s2 <> m2 | |
| 1457 | where | |
| 1458 | props = properties_of pos | |
| 1459 | (m1, m2) = YXML.output_markup (Markup.properties props Markup.position) | |
| 1460 | (s1, s2) = | |
| 1461 | case (line_of pos, file_of pos) of | |
| 1462 |         (Just i, Nothing) -> (" ", "(line " <> Value.print_int i <> ")")
 | |
| 1463 |         (Just i, Just name) -> (" ", "(line " <> Value.print_int i <> " of " <> quote name <> ")")
 | |
| 1464 |         (Nothing, Just name) -> (" ", "(file " <> quote name <> ")")
 | |
| 1465 |         _ -> if is_reported pos then ("", "\092<^here>") else ("", "")
 | |
| 1466 | ||
| 1467 | ||
| 1468 | {- range -}
 | |
| 1469 | ||
| 1470 | type Range = (T, T) | |
| 1471 | ||
| 1472 | no_range :: Range | |
| 1473 | no_range = (none, none) | |
| 1474 | ||
| 1475 | no_range_position :: T -> T | |
| 1476 | no_range_position pos = pos { _end_offset = 0 }
 | |
| 1477 | ||
| 1478 | range_position :: Range -> T | |
| 1479 | range_position (pos, pos') = pos { _end_offset = _offset pos' }
 | |
| 1480 | ||
| 1481 | range :: Range -> Range | |
| 1482 | range (pos, pos') = (range_position (pos, pos'), no_range_position pos') | |
| 1483 | \<close> | |
| 1484 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 1485 | generate_file "Isabelle/XML.hs" = \<open> | 
| 69445 | 1486 | {-  Title:      Isabelle/XML.hs
 | 
| 69225 | 1487 | Author: Makarius | 
| 1488 | LICENSE: BSD 3-clause (Isabelle) | |
| 1489 | ||
| 1490 | Untyped XML trees and representation of ML values. | |
| 69280 | 1491 | |
| 74178 | 1492 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>. | 
| 69225 | 1493 | -} | 
| 1494 | ||
| 74095 | 1495 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 1496 | {-# OPTIONS_GHC -fno-warn-missing-signatures #-}
 | 
| 1497 | ||
| 69225 | 1498 | module Isabelle.XML (Attributes, Body, Tree(..), wrap_elem, unwrap_elem, content_of) | 
| 1499 | where | |
| 1500 | ||
| 1501 | import Isabelle.Library | |
| 1502 | import qualified Isabelle.Properties as Properties | |
| 1503 | import qualified Isabelle.Markup as Markup | |
| 1504 | import qualified Isabelle.Buffer as Buffer | |
| 74095 | 1505 | import qualified Isabelle.Bytes as Bytes | 
| 1506 | import Isabelle.Bytes (Bytes) | |
| 69225 | 1507 | |
| 1508 | ||
| 1509 | {- types -}
 | |
| 1510 | ||
| 1511 | type Attributes = Properties.T | |
| 1512 | type Body = [Tree] | |
| 74095 | 1513 | data Tree = Elem (Markup.T, Body) | Text Bytes | 
| 69225 | 1514 | |
| 1515 | ||
| 1516 | {- wrapped elements -}
 | |
| 1517 | ||
| 69482 | 1518 | wrap_elem :: ((Markup.T, Body), [Tree]) -> Tree | 
| 69236 | 1519 | wrap_elem (((a, atts), body1), body2) = | 
| 69290 | 1520 | Elem ((\<open>XML.xml_elemN\<close>, (\<open>XML.xml_nameN\<close>, a) : atts), Elem ((\<open>XML.xml_bodyN\<close>, []), body1) : body2) | 
| 69225 | 1521 | |
| 69482 | 1522 | unwrap_elem :: Tree -> Maybe ((Markup.T, Body), [Tree]) | 
| 69236 | 1523 | unwrap_elem | 
| 69290 | 1524 | (Elem ((\<open>XML.xml_elemN\<close>, (\<open>XML.xml_nameN\<close>, a) : atts), Elem ((\<open>XML.xml_bodyN\<close>, []), body1) : body2)) = | 
| 69236 | 1525 | Just (((a, atts), body1), body2) | 
| 69225 | 1526 | unwrap_elem _ = Nothing | 
| 1527 | ||
| 1528 | ||
| 1529 | {- text content -}
 | |
| 1530 | ||
| 69482 | 1531 | add_content :: Tree -> Buffer.T -> Buffer.T | 
| 69225 | 1532 | add_content tree = | 
| 1533 | case unwrap_elem tree of | |
| 1534 | Just (_, ts) -> fold add_content ts | |
| 1535 | Nothing -> | |
| 1536 | case tree of | |
| 69290 | 1537 | Elem (_, ts) -> fold add_content ts | 
| 69225 | 1538 | Text s -> Buffer.add s | 
| 1539 | ||
| 74095 | 1540 | content_of :: Body -> Bytes | 
| 74231 | 1541 | content_of = Buffer.build_content . fold add_content | 
| 69225 | 1542 | |
| 1543 | ||
| 1544 | {- string representation -}
 | |
| 1545 | ||
| 74095 | 1546 | encode_char :: Char -> String | 
| 1547 | encode_char '<' = "<" | |
| 1548 | encode_char '>' = ">" | |
| 1549 | encode_char '&' = "&" | |
| 1550 | encode_char '\'' = "'" | |
| 1551 | encode_char '\"' = """ | |
| 1552 | encode_char c = [c] | |
| 1553 | ||
| 1554 | encode_text :: Bytes -> Bytes | |
| 1555 | encode_text = make_bytes . concatMap (encode_char . Bytes.char) . Bytes.unpack | |
| 69225 | 1556 | |
| 1557 | instance Show Tree where | |
| 1558 | show tree = | |
| 74231 | 1559 | make_string $ Buffer.build_content (show_tree tree) | 
| 69225 | 1560 | where | 
| 69290 | 1561 | show_tree (Elem ((name, atts), [])) = | 
| 69225 | 1562 | Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add "/>" | 
| 69290 | 1563 | show_tree (Elem ((name, atts), ts)) = | 
| 69225 | 1564 | Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add ">" #> | 
| 1565 | fold show_tree ts #> | |
| 1566 | Buffer.add "</" #> Buffer.add name #> Buffer.add ">" | |
| 74095 | 1567 | show_tree (Text s) = Buffer.add (encode_text s) | 
| 69225 | 1568 | |
| 1569 | show_elem name atts = | |
| 74095 | 1570 | space_implode " " (name : map (\(a, x) -> a <> "=\"" <> encode_text x <> "\"") atts) | 
| 69225 | 1571 | \<close> | 
| 1572 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 1573 | generate_file "Isabelle/XML/Encode.hs" = \<open> | 
| 69445 | 1574 | {-  Title:      Isabelle/XML/Encode.hs
 | 
| 69240 | 1575 | Author: Makarius | 
| 1576 | LICENSE: BSD 3-clause (Isabelle) | |
| 1577 | ||
| 1578 | XML as data representation language. | |
| 69280 | 1579 | |
| 74178 | 1580 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>. | 
| 69240 | 1581 | -} | 
| 1582 | ||
| 74095 | 1583 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 1584 | {-# OPTIONS_GHC -fno-warn-missing-signatures #-}
 | 
| 1585 | ||
| 69240 | 1586 | module Isabelle.XML.Encode ( | 
| 70845 | 1587 | A, T, V, P, | 
| 69240 | 1588 | |
| 1589 | int_atom, bool_atom, unit_atom, | |
| 1590 | ||
| 71490 | 1591 | tree, properties, string, int, bool, unit, pair, triple, list, option, variant | 
| 69240 | 1592 | ) | 
| 1593 | where | |
| 1594 | ||
| 74168 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1595 | import Data.Maybe (fromJust) | 
| 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1596 | |
| 69240 | 1597 | import Isabelle.Library | 
| 74095 | 1598 | import Isabelle.Bytes (Bytes) | 
| 69240 | 1599 | import qualified Isabelle.Value as Value | 
| 1600 | import qualified Isabelle.Properties as Properties | |
| 1601 | import qualified Isabelle.XML as XML | |
| 1602 | ||
| 1603 | ||
| 74095 | 1604 | type A a = a -> Bytes | 
| 69240 | 1605 | type T a = a -> XML.Body | 
| 74095 | 1606 | type V a = a -> Maybe ([Bytes], XML.Body) | 
| 1607 | type P a = a -> [Bytes] | |
| 69240 | 1608 | |
| 1609 | ||
| 1610 | -- atomic values | |
| 1611 | ||
| 1612 | int_atom :: A Int | |
| 1613 | int_atom = Value.print_int | |
| 1614 | ||
| 1615 | bool_atom :: A Bool | |
| 1616 | bool_atom False = "0" | |
| 1617 | bool_atom True = "1" | |
| 1618 | ||
| 1619 | unit_atom :: A () | |
| 1620 | unit_atom () = "" | |
| 1621 | ||
| 1622 | ||
| 1623 | -- structural nodes | |
| 1624 | ||
| 69290 | 1625 | node ts = XML.Elem ((":", []), ts)
 | 
| 69240 | 1626 | |
| 1627 | vector = map_index (\(i, x) -> (int_atom i, x)) | |
| 1628 | ||
| 69290 | 1629 | tagged (tag, (xs, ts)) = XML.Elem ((int_atom tag, vector xs), ts) | 
| 69240 | 1630 | |
| 1631 | ||
| 1632 | -- representation of standard types | |
| 1633 | ||
| 1634 | tree :: T XML.Tree | |
| 1635 | tree t = [t] | |
| 1636 | ||
| 1637 | properties :: T Properties.T | |
| 69290 | 1638 | properties props = [XML.Elem ((":", props), [])]
 | 
| 69240 | 1639 | |
| 74095 | 1640 | string :: T Bytes | 
| 69240 | 1641 | string "" = [] | 
| 1642 | string s = [XML.Text s] | |
| 1643 | ||
| 1644 | int :: T Int | |
| 1645 | int = string . int_atom | |
| 1646 | ||
| 1647 | bool :: T Bool | |
| 1648 | bool = string . bool_atom | |
| 1649 | ||
| 1650 | unit :: T () | |
| 1651 | unit = string . unit_atom | |
| 1652 | ||
| 1653 | pair :: T a -> T b -> T (a, b) | |
| 1654 | pair f g (x, y) = [node (f x), node (g y)] | |
| 1655 | ||
| 1656 | triple :: T a -> T b -> T c -> T (a, b, c) | |
| 1657 | triple f g h (x, y, z) = [node (f x), node (g y), node (h z)] | |
| 1658 | ||
| 1659 | list :: T a -> T [a] | |
| 1660 | list f xs = map (node . f) xs | |
| 1661 | ||
| 71490 | 1662 | option :: T a -> T (Maybe a) | 
| 1663 | option _ Nothing = [] | |
| 1664 | option f (Just x) = [node (f x)] | |
| 1665 | ||
| 69240 | 1666 | variant :: [V a] -> T a | 
| 74168 
f0b2136e2204
tuned signature: prefer existing Haskell operations;
 wenzelm parents: 
74167diff
changeset | 1667 | variant fs x = [tagged (fromJust (get_index (\f -> f x) fs))] | 
| 69240 | 1668 | \<close> | 
| 1669 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 1670 | generate_file "Isabelle/XML/Decode.hs" = \<open> | 
| 69445 | 1671 | {-  Title:      Isabelle/XML/Decode.hs
 | 
| 69240 | 1672 | Author: Makarius | 
| 1673 | LICENSE: BSD 3-clause (Isabelle) | |
| 1674 | ||
| 1675 | XML as data representation language. | |
| 69280 | 1676 | |
| 74178 | 1677 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>. | 
| 69240 | 1678 | -} | 
| 1679 | ||
| 74095 | 1680 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 1681 | {-# OPTIONS_GHC -fno-warn-missing-signatures #-}
 | 
| 1682 | ||
| 69240 | 1683 | module Isabelle.XML.Decode ( | 
| 70845 | 1684 | A, T, V, P, | 
| 69240 | 1685 | |
| 1686 | int_atom, bool_atom, unit_atom, | |
| 1687 | ||
| 71490 | 1688 | tree, properties, string, int, bool, unit, pair, triple, list, option, variant | 
| 69240 | 1689 | ) | 
| 1690 | where | |
| 1691 | ||
| 1692 | import Isabelle.Library | |
| 74095 | 1693 | import Isabelle.Bytes (Bytes) | 
| 69240 | 1694 | import qualified Isabelle.Value as Value | 
| 1695 | import qualified Isabelle.Properties as Properties | |
| 1696 | import qualified Isabelle.XML as XML | |
| 1697 | ||
| 1698 | ||
| 74095 | 1699 | type A a = Bytes -> a | 
| 69240 | 1700 | type T a = XML.Body -> a | 
| 74095 | 1701 | type V a = ([Bytes], XML.Body) -> a | 
| 1702 | type P a = [Bytes] -> a | |
| 69240 | 1703 | |
| 1704 | err_atom = error "Malformed XML atom" | |
| 1705 | err_body = error "Malformed XML body" | |
| 1706 | ||
| 1707 | ||
| 1708 | {- atomic values -}
 | |
| 1709 | ||
| 1710 | int_atom :: A Int | |
| 1711 | int_atom s = | |
| 1712 | case Value.parse_int s of | |
| 1713 | Just i -> i | |
| 1714 | Nothing -> err_atom | |
| 1715 | ||
| 1716 | bool_atom :: A Bool | |
| 1717 | bool_atom "0" = False | |
| 1718 | bool_atom "1" = True | |
| 1719 | bool_atom _ = err_atom | |
| 1720 | ||
| 1721 | unit_atom :: A () | |
| 1722 | unit_atom "" = () | |
| 1723 | unit_atom _ = err_atom | |
| 1724 | ||
| 1725 | ||
| 1726 | {- structural nodes -}
 | |
| 1727 | ||
| 69290 | 1728 | node (XML.Elem ((":", []), ts)) = ts
 | 
| 69240 | 1729 | node _ = err_body | 
| 1730 | ||
| 1731 | vector atts = | |
| 1732 | map_index (\(i, (a, x)) -> if int_atom a == i then x else err_atom) atts | |
| 1733 | ||
| 69290 | 1734 | tagged (XML.Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts)) | 
| 69240 | 1735 | tagged _ = err_body | 
| 1736 | ||
| 1737 | ||
| 1738 | {- representation of standard types -}
 | |
| 1739 | ||
| 1740 | tree :: T XML.Tree | |
| 1741 | tree [t] = t | |
| 1742 | tree _ = err_body | |
| 1743 | ||
| 1744 | properties :: T Properties.T | |
| 69290 | 1745 | properties [XML.Elem ((":", props), [])] = props
 | 
| 69240 | 1746 | properties _ = err_body | 
| 1747 | ||
| 74095 | 1748 | string :: T Bytes | 
| 69240 | 1749 | string [] = "" | 
| 1750 | string [XML.Text s] = s | |
| 1751 | string _ = err_body | |
| 1752 | ||
| 1753 | int :: T Int | |
| 1754 | int = int_atom . string | |
| 1755 | ||
| 1756 | bool :: T Bool | |
| 1757 | bool = bool_atom . string | |
| 1758 | ||
| 1759 | unit :: T () | |
| 1760 | unit = unit_atom . string | |
| 1761 | ||
| 1762 | pair :: T a -> T b -> T (a, b) | |
| 1763 | pair f g [t1, t2] = (f (node t1), g (node t2)) | |
| 1764 | pair _ _ _ = err_body | |
| 1765 | ||
| 1766 | triple :: T a -> T b -> T c -> T (a, b, c) | |
| 1767 | triple f g h [t1, t2, t3] = (f (node t1), g (node t2), h (node t3)) | |
| 1768 | triple _ _ _ _ = err_body | |
| 1769 | ||
| 1770 | list :: T a -> T [a] | |
| 1771 | list f ts = map (f . node) ts | |
| 1772 | ||
| 1773 | option :: T a -> T (Maybe a) | |
| 1774 | option _ [] = Nothing | |
| 1775 | option f [t] = Just (f (node t)) | |
| 1776 | option _ _ = err_body | |
| 1777 | ||
| 1778 | variant :: [V a] -> T a | |
| 1779 | variant fs [t] = (fs !! tag) (xs, ts) | |
| 1780 | where (tag, (xs, ts)) = tagged t | |
| 1781 | variant _ _ = err_body | |
| 1782 | \<close> | |
| 1783 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 1784 | generate_file "Isabelle/YXML.hs" = \<open> | 
| 69445 | 1785 | {-  Title:      Isabelle/YXML.hs
 | 
| 69288 | 1786 | Author: Makarius | 
| 1787 | LICENSE: BSD 3-clause (Isabelle) | |
| 1788 | ||
| 1789 | Efficient text representation of XML trees. Suitable for direct | |
| 1790 | inlining into plain text. | |
| 1791 | ||
| 74178 | 1792 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/yxml.ML\<close>. | 
| 69288 | 1793 | -} | 
| 1794 | ||
| 74095 | 1795 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 1796 | {-# OPTIONS_GHC -fno-warn-missing-signatures -fno-warn-incomplete-patterns #-}
 | 
| 1797 | ||
| 69288 | 1798 | module Isabelle.YXML (charX, charY, strX, strY, detect, output_markup, | 
| 1799 | buffer_body, buffer, string_of_body, string_of, parse_body, parse) | |
| 1800 | where | |
| 1801 | ||
| 1802 | import qualified Data.List as List | |
| 74095 | 1803 | import Data.Word (Word8) | 
| 69288 | 1804 | |
| 1805 | import Isabelle.Library | |
| 74095 | 1806 | import qualified Isabelle.Bytes as Bytes | 
| 1807 | import Isabelle.Bytes (Bytes) | |
| 69288 | 1808 | import qualified Isabelle.Markup as Markup | 
| 1809 | import qualified Isabelle.XML as XML | |
| 1810 | import qualified Isabelle.Buffer as Buffer | |
| 1811 | ||
| 1812 | ||
| 1813 | {- markers -}
 | |
| 1814 | ||
| 74095 | 1815 | charX, charY :: Word8 | 
| 1816 | charX = 5 | |
| 1817 | charY = 6 | |
| 1818 | ||
| 1819 | strX, strY, strXY, strXYX :: Bytes | |
| 1820 | strX = Bytes.singleton charX | |
| 1821 | strY = Bytes.singleton charY | |
| 74081 | 1822 | strXY = strX <> strY | 
| 1823 | strXYX = strXY <> strX | |
| 69288 | 1824 | |
| 74095 | 1825 | detect :: Bytes -> Bool | 
| 1826 | detect = Bytes.any (\c -> c == charX || c == charY) | |
| 69288 | 1827 | |
| 1828 | ||
| 1829 | {- output -}
 | |
| 1830 | ||
| 1831 | output_markup :: Markup.T -> Markup.Output | |
| 1832 | output_markup markup@(name, atts) = | |
| 1833 | if Markup.is_empty markup then Markup.no_output | |
| 74095 | 1834 | else (strXY <> name <> Bytes.concat (map (\(a, x) -> strY <> a <> "=" <> x) atts) <> strX, strXYX) | 
| 69288 | 1835 | |
| 1836 | buffer_attrib (a, x) = | |
| 1837 | Buffer.add strY #> Buffer.add a #> Buffer.add "=" #> Buffer.add x | |
| 1838 | ||
| 1839 | buffer_body :: XML.Body -> Buffer.T -> Buffer.T | |
| 1840 | buffer_body = fold buffer | |
| 1841 | ||
| 1842 | buffer :: XML.Tree -> Buffer.T -> Buffer.T | |
| 69290 | 1843 | buffer (XML.Elem ((name, atts), ts)) = | 
| 69288 | 1844 | Buffer.add strXY #> Buffer.add name #> fold buffer_attrib atts #> Buffer.add strX #> | 
| 1845 | buffer_body ts #> | |
| 1846 | Buffer.add strXYX | |
| 1847 | buffer (XML.Text s) = Buffer.add s | |
| 1848 | ||
| 74095 | 1849 | string_of_body :: XML.Body -> Bytes | 
| 74231 | 1850 | string_of_body = Buffer.build_content . buffer_body | 
| 69288 | 1851 | |
| 74095 | 1852 | string_of :: XML.Tree -> Bytes | 
| 69288 | 1853 | string_of = string_of_body . single | 
| 1854 | ||
| 1855 | ||
| 1856 | {- parse -}
 | |
| 1857 | ||
| 1858 | -- split: fields or non-empty tokens | |
| 1859 | ||
| 74095 | 1860 | split :: Bool -> Word8 -> [Word8] -> [[Word8]] | 
| 69288 | 1861 | split _ _ [] = [] | 
| 1862 | split fields sep str = splitting str | |
| 1863 | where | |
| 1864 | splitting rest = | |
| 1865 | case span (/= sep) rest of | |
| 1866 | (_, []) -> cons rest [] | |
| 1867 | (prfx, _ : rest') -> cons prfx (splitting rest') | |
| 1868 | cons item = if fields || not (null item) then (:) item else id | |
| 1869 | ||
| 1870 | ||
| 1871 | -- structural errors | |
| 1872 | ||
| 74095 | 1873 | err :: Bytes -> a | 
| 1874 | err msg = error (make_string ("Malformed YXML: " <> msg))
 | |
| 1875 | ||
| 69288 | 1876 | err_attribute = err "bad attribute" | 
| 1877 | err_element = err "bad element" | |
| 74095 | 1878 | |
| 1879 | err_unbalanced :: Bytes -> a | |
| 1880 | err_unbalanced name = | |
| 1881 | if Bytes.null name then err "unbalanced element" | |
| 1882 |   else err ("unbalanced element " <> quote name)
 | |
| 69288 | 1883 | |
| 1884 | ||
| 1885 | -- stack operations | |
| 1886 | ||
| 1887 | add x ((elem, body) : pending) = (elem, x : body) : pending | |
| 1888 | ||
| 74095 | 1889 | push name atts pending = | 
| 1890 | if Bytes.null name then err_element | |
| 1891 | else ((name, atts), []) : pending | |
| 1892 | ||
| 1893 | pop (((name, atts), body) : pending) = | |
| 1894 | if Bytes.null name then err_unbalanced name | |
| 1895 | else add (XML.Elem ((name, atts), reverse body)) pending | |
| 69288 | 1896 | |
| 1897 | ||
| 1898 | -- parsing | |
| 1899 | ||
| 1900 | parse_attrib s = | |
| 74095 | 1901 | case List.elemIndex (Bytes.byte '=') s of | 
| 1902 | Just i | i > 0 -> (Bytes.pack $ take i s, Bytes.pack $ drop (i + 1) s) | |
| 69288 | 1903 | _ -> err_attribute | 
| 1904 | ||
| 74095 | 1905 | parse_chunk [[], []] = pop | 
| 1906 | parse_chunk ([] : name : atts) = push (Bytes.pack name) (map parse_attrib atts) | |
| 1907 | parse_chunk txts = fold (add . XML.Text . Bytes.pack) txts | |
| 1908 | ||
| 1909 | parse_body :: Bytes -> XML.Body | |
| 69288 | 1910 | parse_body source = | 
| 74095 | 1911 | case fold parse_chunk chunks [((Bytes.empty, []), [])] of | 
| 1912 | [((name, _), result)] | Bytes.null name -> reverse result | |
| 69288 | 1913 | ((name, _), _) : _ -> err_unbalanced name | 
| 74095 | 1914 | where chunks = source |> Bytes.unpack |> split False charX |> map (split True charY) | 
| 1915 | ||
| 1916 | parse :: Bytes -> XML.Tree | |
| 69288 | 1917 | parse source = | 
| 1918 | case parse_body source of | |
| 1919 | [result] -> result | |
| 1920 | [] -> XML.Text "" | |
| 1921 | _ -> err "multiple results" | |
| 1922 | \<close> | |
| 1923 | ||
| 74091 | 1924 | generate_file "Isabelle/Completion.hs" = \<open> | 
| 1925 | {-  Title:      Isabelle/Completion.hs
 | |
| 1926 | Author: Makarius | |
| 1927 | LICENSE: BSD 3-clause (Isabelle) | |
| 1928 | ||
| 1929 | Completion of names. | |
| 1930 | ||
| 74178 | 1931 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/completion.ML\<close>. | 
| 74091 | 1932 | -} | 
| 1933 | ||
| 74095 | 1934 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 1935 | ||
| 74091 | 1936 | module Isabelle.Completion ( | 
| 1937 | Name, T, names, none, make, markup_element, markup_report, make_report | |
| 1938 | ) where | |
| 1939 | ||
| 74095 | 1940 | import qualified Isabelle.Bytes as Bytes | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1941 | import qualified Isabelle.Name as Name | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1942 | import Isabelle.Name (Name) | 
| 74091 | 1943 | import qualified Isabelle.Properties as Properties | 
| 1944 | import qualified Isabelle.Markup as Markup | |
| 74129 | 1945 | import Isabelle.XML.Classes | 
| 74091 | 1946 | import qualified Isabelle.XML as XML | 
| 1947 | import qualified Isabelle.YXML as YXML | |
| 1948 | ||
| 1949 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1950 | type Names = [(Name, (Name, Name))] -- external name, kind, internal name | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1951 | data T = Completion Properties.T Int Names -- position, total length, names | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1952 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1953 | names :: Int -> Properties.T -> Names -> T | 
| 74091 | 1954 | names limit props names = Completion props (length names) (take limit names) | 
| 1955 | ||
| 1956 | none :: T | |
| 1957 | none = names 0 [] [] | |
| 1958 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1959 | make :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> T | 
| 74091 | 1960 | make limit (name, props) make_names = | 
| 74095 | 1961 | if name /= "" && name /= "_" then | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1962 | names limit props (make_names (Bytes.isPrefixOf (Name.clean name))) | 
| 74091 | 1963 | else none | 
| 1964 | ||
| 1965 | markup_element :: T -> (Markup.T, XML.Body) | |
| 1966 | markup_element (Completion props total names) = | |
| 1967 | if not (null names) then | |
| 74129 | 1968 | (Markup.properties props Markup.completion, encode (total, names)) | 
| 74091 | 1969 | else (Markup.empty, []) | 
| 1970 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1971 | markup_report :: [T] -> Name | 
| 74095 | 1972 | markup_report [] = Bytes.empty | 
| 74091 | 1973 | markup_report elems = | 
| 1974 | YXML.string_of $ XML.Elem (Markup.report, map (XML.Elem . markup_element) elems) | |
| 1975 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 1976 | make_report :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> Name | 
| 74091 | 1977 | make_report limit name_props make_names = | 
| 1978 | markup_report [make limit name_props make_names] | |
| 1979 | \<close> | |
| 1980 | ||
| 1981 | generate_file "Isabelle/File.hs" = \<open> | |
| 1982 | {-  Title:      Isabelle/File.hs
 | |
| 1983 | Author: Makarius | |
| 1984 | LICENSE: BSD 3-clause (Isabelle) | |
| 1985 | ||
| 1986 | File-system operations. | |
| 1987 | ||
| 74178 | 1988 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/file.ML\<close>. | 
| 74091 | 1989 | -} | 
| 1990 | ||
| 74098 | 1991 | module Isabelle.File (read, write, append) where | 
| 74091 | 1992 | |
| 1993 | import Prelude hiding (read) | |
| 1994 | import qualified System.IO as IO | |
| 74098 | 1995 | import qualified Data.ByteString as ByteString | 
| 1996 | import qualified Isabelle.Bytes as Bytes | |
| 1997 | import Isabelle.Bytes (Bytes) | |
| 1998 | ||
| 1999 | read :: IO.FilePath -> IO Bytes | |
| 2000 | read path = Bytes.make <$> IO.withFile path IO.ReadMode ByteString.hGetContents | |
| 2001 | ||
| 2002 | write :: IO.FilePath -> Bytes -> IO () | |
| 2003 | write path bs = IO.withFile path IO.WriteMode (\h -> ByteString.hPut h (Bytes.unmake bs)) | |
| 2004 | ||
| 2005 | append :: IO.FilePath -> Bytes -> IO () | |
| 2006 | append path bs = IO.withFile path IO.AppendMode (\h -> ByteString.hPut h (Bytes.unmake bs)) | |
| 74091 | 2007 | \<close> | 
| 2008 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 2009 | generate_file "Isabelle/Pretty.hs" = \<open> | 
| 69445 | 2010 | {-  Title:      Isabelle/Pretty.hs
 | 
| 69248 | 2011 | Author: Makarius | 
| 2012 | LICENSE: BSD 3-clause (Isabelle) | |
| 2013 | ||
| 2014 | Generic pretty printing module. | |
| 69280 | 2015 | |
| 74178 | 2016 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/pretty.ML\<close>. | 
| 69248 | 2017 | -} | 
| 2018 | ||
| 74095 | 2019 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 2020 | {-# OPTIONS_GHC -fno-warn-missing-signatures #-}
 | 
| 2021 | ||
| 69248 | 2022 | module Isabelle.Pretty ( | 
| 2023 | T, symbolic, formatted, unformatted, | |
| 2024 | ||
| 2025 | str, brk_indent, brk, fbrk, breaks, fbreaks, blk, block, strs, markup, mark, mark_str, marks_str, | |
| 2026 | item, text_fold, keyword1, keyword2, text, paragraph, para, quote, cartouche, separate, | |
| 2027 | commas, enclose, enum, list, str_list, big_list) | |
| 2028 | where | |
| 2029 | ||
| 74095 | 2030 | import qualified Data.List as List | 
| 2031 | ||
| 2032 | import qualified Isabelle.Bytes as Bytes | |
| 2033 | import Isabelle.Bytes (Bytes) | |
| 74185 | 2034 | import Isabelle.Library hiding (enclose, quote, separate, commas) | 
| 69248 | 2035 | import qualified Isabelle.Buffer as Buffer | 
| 2036 | import qualified Isabelle.Markup as Markup | |
| 2037 | import qualified Isabelle.XML as XML | |
| 2038 | import qualified Isabelle.YXML as YXML | |
| 2039 | ||
| 2040 | ||
| 2041 | data T = | |
| 2042 | Block Markup.T Bool Int [T] | |
| 2043 | | Break Int Int | |
| 74095 | 2044 | | Str Bytes | 
| 69248 | 2045 | |
| 2046 | ||
| 2047 | {- output -}
 | |
| 2048 | ||
| 74095 | 2049 | symbolic_text s = if Bytes.null s then [] else [XML.Text s] | 
| 69248 | 2050 | |
| 2051 | symbolic_markup markup body = | |
| 2052 | if Markup.is_empty markup then body | |
| 69290 | 2053 | else [XML.Elem (markup, body)] | 
| 69248 | 2054 | |
| 2055 | symbolic :: T -> XML.Body | |
| 2056 | symbolic (Block markup consistent indent prts) = | |
| 2057 | concatMap symbolic prts | |
| 2058 | |> symbolic_markup block_markup | |
| 2059 | |> symbolic_markup markup | |
| 2060 | where block_markup = if null prts then Markup.empty else Markup.block consistent indent | |
| 74095 | 2061 | symbolic (Break wd ind) = [XML.Elem (Markup.break wd ind, symbolic_text (Bytes.spaces wd))] | 
| 69248 | 2062 | symbolic (Str s) = symbolic_text s | 
| 2063 | ||
| 74095 | 2064 | formatted :: T -> Bytes | 
| 69248 | 2065 | formatted = YXML.string_of_body . symbolic | 
| 2066 | ||
| 74095 | 2067 | unformatted :: T -> Bytes | 
| 74231 | 2068 | unformatted = Buffer.build_content . out | 
| 69248 | 2069 | where | 
| 2070 | out (Block markup _ _ prts) = | |
| 2071 | let (bg, en) = YXML.output_markup markup | |
| 2072 | in Buffer.add bg #> fold out prts #> Buffer.add en | |
| 74095 | 2073 | out (Break _ wd) = Buffer.add (Bytes.spaces wd) | 
| 69248 | 2074 | out (Str s) = Buffer.add s | 
| 2075 | ||
| 2076 | ||
| 2077 | {- derived operations to create formatting expressions -}
 | |
| 2078 | ||
| 2079 | force_nat n | n < 0 = 0 | |
| 2080 | force_nat n = n | |
| 2081 | ||
| 74095 | 2082 | str :: BYTES a => a -> T | 
| 2083 | str = Str . make_bytes | |
| 69248 | 2084 | |
| 2085 | brk_indent :: Int -> Int -> T | |
| 2086 | brk_indent wd ind = Break (force_nat wd) ind | |
| 2087 | ||
| 2088 | brk :: Int -> T | |
| 2089 | brk wd = brk_indent wd 0 | |
| 2090 | ||
| 2091 | fbrk :: T | |
| 74095 | 2092 | fbrk = Str "\n" | 
| 69248 | 2093 | |
| 2094 | breaks, fbreaks :: [T] -> [T] | |
| 2095 | breaks = List.intersperse (brk 1) | |
| 2096 | fbreaks = List.intersperse fbrk | |
| 2097 | ||
| 2098 | blk :: (Int, [T]) -> T | |
| 2099 | blk (indent, es) = Block Markup.empty False (force_nat indent) es | |
| 2100 | ||
| 2101 | block :: [T] -> T | |
| 2102 | block prts = blk (2, prts) | |
| 2103 | ||
| 74095 | 2104 | strs :: BYTES a => [a] -> T | 
| 69248 | 2105 | strs = block . breaks . map str | 
| 2106 | ||
| 2107 | markup :: Markup.T -> [T] -> T | |
| 2108 | markup m = Block m False 0 | |
| 2109 | ||
| 2110 | mark :: Markup.T -> T -> T | |
| 2111 | mark m prt = if m == Markup.empty then prt else markup m [prt] | |
| 2112 | ||
| 74095 | 2113 | mark_str :: BYTES a => (Markup.T, a) -> T | 
| 69248 | 2114 | mark_str (m, s) = mark m (str s) | 
| 2115 | ||
| 74095 | 2116 | marks_str :: BYTES a => ([Markup.T], a) -> T | 
| 69248 | 2117 | marks_str (ms, s) = fold_rev mark ms (str s) | 
| 2118 | ||
| 2119 | item :: [T] -> T | |
| 2120 | item = markup Markup.item | |
| 2121 | ||
| 2122 | text_fold :: [T] -> T | |
| 2123 | text_fold = markup Markup.text_fold | |
| 2124 | ||
| 74095 | 2125 | keyword1, keyword2 :: BYTES a => a -> T | 
| 69248 | 2126 | keyword1 name = mark_str (Markup.keyword1, name) | 
| 2127 | keyword2 name = mark_str (Markup.keyword2, name) | |
| 2128 | ||
| 74095 | 2129 | text :: BYTES a => a -> [T] | 
| 74132 | 2130 | text = breaks . map str . filter (not . Bytes.null) . space_explode ' ' . make_bytes | 
| 69248 | 2131 | |
| 2132 | paragraph :: [T] -> T | |
| 2133 | paragraph = markup Markup.paragraph | |
| 2134 | ||
| 74095 | 2135 | para :: BYTES a => a -> T | 
| 69248 | 2136 | para = paragraph . text | 
| 2137 | ||
| 2138 | quote :: T -> T | |
| 74095 | 2139 | quote prt = blk (1, [Str "\"", prt, Str "\""]) | 
| 69248 | 2140 | |
| 2141 | cartouche :: T -> T | |
| 74095 | 2142 | cartouche prt = blk (1, [Str "\92<open>", prt, Str "\92<close>"]) | 
| 2143 | ||
| 2144 | separate :: BYTES a => a -> [T] -> [T] | |
| 69248 | 2145 | separate sep = List.intercalate [str sep, brk 1] . map single | 
| 2146 | ||
| 2147 | commas :: [T] -> [T] | |
| 74095 | 2148 | commas = separate ("," :: Bytes)
 | 
| 2149 | ||
| 2150 | enclose :: BYTES a => a -> a -> [T] -> T | |
| 74081 | 2151 | enclose lpar rpar prts = block (str lpar : prts <> [str rpar]) | 
| 69248 | 2152 | |
| 74095 | 2153 | enum :: BYTES a => a -> a -> a -> [T] -> T | 
| 69248 | 2154 | enum sep lpar rpar = enclose lpar rpar . separate sep | 
| 2155 | ||
| 74095 | 2156 | list :: BYTES a => a -> a -> [T] -> T | 
| 69248 | 2157 | list = enum "," | 
| 2158 | ||
| 74095 | 2159 | str_list :: BYTES a => a -> a -> [a] -> T | 
| 69248 | 2160 | str_list lpar rpar = list lpar rpar . map str | 
| 2161 | ||
| 74095 | 2162 | big_list :: BYTES a => a -> [T] -> T | 
| 69248 | 2163 | big_list name prts = block (fbreaks (str name : prts)) | 
| 2164 | \<close> | |
| 2165 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2166 | generate_file "Isabelle/Name.hs" = \<open> | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2167 | {-  Title:      Isabelle/Name.hs
 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2168 | Author: Makarius | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2169 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2170 | Names of basic logical entities (variables etc.). | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2171 | |
| 74178 | 2172 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/name.ML\<close>. | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2173 | -} | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2174 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2175 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2176 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2177 | module Isabelle.Name ( | 
| 74194 | 2178 | Name, | 
| 2179 | uu, uu_, aT, | |
| 74217 | 2180 | clean_index, clean, internal, skolem, is_internal, is_skolem, dest_internal, dest_skolem, | 
| 74327 | 2181 | Context, declare, declare_renaming, is_declared, declared, context, make_context, | 
| 2182 | variant, variant_list | |
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2183 | ) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2184 | where | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2185 | |
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2186 | import Data.Maybe (fromMaybe) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2187 | import Data.Map.Strict (Map) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2188 | import qualified Data.Map.Strict as Map | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2189 | import Data.Word (Word8) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2190 | import qualified Isabelle.Bytes as Bytes | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2191 | import Isabelle.Bytes (Bytes) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2192 | import qualified Isabelle.Symbol as Symbol | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2193 | import Isabelle.Library | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2194 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2195 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2196 | type Name = Bytes | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2197 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2198 | |
| 74194 | 2199 | {- common defaults -}
 | 
| 2200 | ||
| 2201 | uu, uu_, aT :: Name | |
| 2202 | uu = "uu" | |
| 2203 | uu_ = "uu_" | |
| 2204 | aT = "'a" | |
| 2205 | ||
| 2206 | ||
| 74196 | 2207 | {- internal names -- NB: internal subsumes skolem -}
 | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2208 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2209 | underscore :: Word8 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2210 | underscore = Bytes.byte '_' | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2211 | |
| 74196 | 2212 | internal, skolem :: Name -> Name | 
| 2213 | internal x = x <> "_" | |
| 2214 | skolem x = x <> "__" | |
| 2215 | ||
| 2216 | is_internal, is_skolem :: Name -> Bool | |
| 2217 | is_internal = Bytes.isSuffixOf "_" | |
| 2218 | is_skolem = Bytes.isSuffixOf "__" | |
| 2219 | ||
| 74217 | 2220 | dest_internal, dest_skolem :: Name -> Maybe Name | 
| 2221 | dest_internal = Bytes.try_unsuffix "_" | |
| 2222 | dest_skolem = Bytes.try_unsuffix "__" | |
| 2223 | ||
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2224 | clean_index :: (Name, Int) -> (Name, Int) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2225 | clean_index (x, i) = | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2226 | case dest_internal x of | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2227 | Nothing -> (x, i) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2228 | Just x' -> clean_index (x', i + 1) | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2229 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2230 | clean :: Name -> Name | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2231 | clean x = fst (clean_index (x, 0)) | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2232 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2233 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2234 | {- context for used names -}
 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2235 | |
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2236 | newtype Context = Context (Map Name (Maybe Name))  {-declared names with latest renaming-}
 | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2237 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2238 | declare :: Name -> Context -> Context | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2239 | declare x (Context names) = | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2240 | Context ( | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2241 | let a = clean x | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2242 | in if Map.member a names then names else Map.insert a Nothing names) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2243 | |
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2244 | declare_renaming :: (Name, Name) -> Context -> Context | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2245 | declare_renaming (x, x') (Context names) = | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2246 | Context (Map.insert (clean x) (Just (clean x')) names) | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2247 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2248 | is_declared :: Context -> Name -> Bool | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2249 | is_declared (Context names) x = Map.member x names | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2250 | |
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2251 | declared :: Context -> Name -> Maybe (Maybe Name) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2252 | declared (Context names) a = Map.lookup a names | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2253 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2254 | context :: Context | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2255 | context = Context Map.empty |> fold declare ["", "'"] | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2256 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2257 | make_context :: [Name] -> Context | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2258 | make_context used = fold declare used context | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2259 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2260 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2261 | {- generating fresh names -}
 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2262 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2263 | bump_init :: Name -> Name | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2264 | bump_init str = str <> "a" | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2265 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2266 | bump_string :: Name -> Name | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2267 | bump_string str = | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2268 | let | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2269 | a = Bytes.byte 'a' | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2270 | z = Bytes.byte 'z' | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2271 | bump (b : bs) | b == z = a : bump bs | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2272 | bump (b : bs) | a <= b && b < z = b + 1 : bs | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2273 | bump bs = a : bs | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2274 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2275 | rev = reverse (Bytes.unpack str) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2276 | part2 = reverse (takeWhile (Symbol.is_ascii_quasi . Bytes.char) rev) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2277 | part1 = reverse (bump (drop (length part2) rev)) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2278 | in Bytes.pack (part1 <> part2) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2279 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2280 | variant :: Name -> Context -> (Name, Context) | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2281 | variant name ctxt = | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2282 | let | 
| 74315 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2283 | vary x = | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2284 | case declared ctxt x of | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2285 | Nothing -> x | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2286 | Just x' -> vary (bump_string (fromMaybe x x')) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2287 | |
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2288 | (x, n) = clean_index (name, 0) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2289 | (x', ctxt') = | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2290 | if not (is_declared ctxt x) then (x, declare x ctxt) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2291 | else | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2292 | let | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2293 | x0 = bump_init x | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2294 | x' = vary x0 | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2295 | ctxt' = ctxt | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2296 | |> (if x0 /= x' then declare_renaming (x0, x') else id) | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2297 | |> declare x' | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2298 | in (x', ctxt') | 
| 
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
 wenzelm parents: 
74310diff
changeset | 2299 | in (x' <> Bytes.pack (replicate n underscore), ctxt') | 
| 74327 | 2300 | |
| 2301 | variant_list :: [Name] -> [Name] -> [Name] | |
| 2302 | variant_list used names = fst (make_context used |> fold_map variant names) | |
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2303 | \<close> | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2304 | |
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 2305 | generate_file "Isabelle/Term.hs" = \<open> | 
| 69445 | 2306 | {-  Title:      Isabelle/Term.hs
 | 
| 69240 | 2307 | Author: Makarius | 
| 2308 | LICENSE: BSD 3-clause (Isabelle) | |
| 2309 | ||
| 2310 | Lambda terms, types, sorts. | |
| 69280 | 2311 | |
| 74178 | 2312 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/term.scala\<close>. | 
| 69240 | 2313 | -} | 
| 2314 | ||
| 74095 | 2315 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 2316 | ||
| 69240 | 2317 | module Isabelle.Term ( | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2318 | Indexname, Sort, Typ(..), Term(..), | 
| 74214 | 2319 | Free, lambda, declare_frees, incr_boundvars, subst_bound, dest_lambda, strip_lambda, | 
| 74197 | 2320 | type_op0, type_op1, op0, op1, op2, typed_op0, typed_op1, typed_op2, binder, | 
| 74105 | 2321 | dummyS, dummyT, is_dummyT, propT, is_propT, (-->), dest_funT, (--->), | 
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2322 | aconv, list_comb, strip_comb, head_of | 
| 74105 | 2323 | ) | 
| 69240 | 2324 | where | 
| 2325 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2326 | import Isabelle.Library | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2327 | import qualified Isabelle.Name as Name | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2328 | import Isabelle.Name (Name) | 
| 74095 | 2329 | |
| 74105 | 2330 | infixr 5 --> | 
| 2331 | infixr ---> | |
| 2332 | ||
| 2333 | ||
| 2334 | {- types and terms -}
 | |
| 2335 | ||
| 2336 | type Indexname = (Name, Int) | |
| 2337 | ||
| 2338 | type Sort = [Name] | |
| 2339 | ||
| 2340 | data Typ = | |
| 2341 | Type (Name, [Typ]) | |
| 2342 | | TFree (Name, Sort) | |
| 2343 | | TVar (Indexname, Sort) | |
| 2344 | deriving (Show, Eq, Ord) | |
| 2345 | ||
| 2346 | data Term = | |
| 2347 | Const (Name, [Typ]) | |
| 2348 | | Free (Name, Typ) | |
| 2349 | | Var (Indexname, Typ) | |
| 2350 | | Bound Int | |
| 2351 | | Abs (Name, Typ, Term) | |
| 2352 | | App (Term, Term) | |
| 2353 | deriving (Show, Eq, Ord) | |
| 2354 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2355 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2356 | {- free and bound variables -}
 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2357 | |
| 74105 | 2358 | type Free = (Name, Typ) | 
| 2359 | ||
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2360 | lambda :: Free -> Term -> Term | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2361 | lambda (name, typ) body = Abs (name, typ, abstract 0 body) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2362 | where | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2363 | abstract lev (Free (x, ty)) | name == x && typ == ty = Bound lev | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2364 | abstract lev (Abs (a, ty, t)) = Abs (a, ty, abstract (lev + 1) t) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2365 | abstract lev (App (t, u)) = App (abstract lev t, abstract lev u) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2366 | abstract _ t = t | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2367 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2368 | declare_frees :: Term -> Name.Context -> Name.Context | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2369 | declare_frees (Free (x, _)) = Name.declare x | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2370 | declare_frees (Abs (_, _, b)) = declare_frees b | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2371 | declare_frees (App (t, u)) = declare_frees t #> declare_frees u | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2372 | declare_frees _ = id | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2373 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2374 | incr_boundvars :: Int -> Term -> Term | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2375 | incr_boundvars inc = if inc == 0 then id else incr 0 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2376 | where | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2377 | incr lev (Bound i) = if i >= lev then Bound (i + inc) else Bound i | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2378 | incr lev (Abs (a, ty, b)) = Abs (a, ty, incr (lev + 1) b) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2379 | incr lev (App (t, u)) = App (incr lev t, incr lev u) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2380 | incr _ t = t | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2381 | |
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2382 | subst_bound :: Term -> Term -> Term | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2383 | subst_bound arg = subst 0 | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2384 | where | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2385 | subst lev (Bound i) = | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2386 | if i < lev then Bound i | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2387 | else if i == lev then incr_boundvars lev arg | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2388 | else Bound (i - 1) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2389 | subst lev (Abs (a, ty, b)) = Abs (a, ty, subst (lev + 1) b) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2390 | subst lev (App (t, u)) = App (subst lev t, subst lev u) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2391 | subst _ t = t | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2392 | |
| 74214 | 2393 | dest_lambda :: Name.Context -> Term -> Maybe (Free, Term) | 
| 2394 | dest_lambda names (Abs (x, ty, b)) = | |
| 74106 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2395 | let | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2396 | (x', _) = Name.variant x (declare_frees b names) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2397 | v = (x', ty) | 
| 
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
 wenzelm parents: 
74105diff
changeset | 2398 | in Just (v, subst_bound (Free v) b) | 
| 74214 | 2399 | dest_lambda _ _ = Nothing | 
| 2400 | ||
| 2401 | strip_lambda :: Name.Context -> Term -> ([Free], Term) | |
| 2402 | strip_lambda names tm = | |
| 2403 | case dest_lambda names tm of | |
| 74124 | 2404 | Just (v, t) -> | 
| 74214 | 2405 | let (vs, t') = strip_lambda names t' | 
| 74124 | 2406 | in (v : vs, t') | 
| 2407 | Nothing -> ([], tm) | |
| 2408 | ||
| 74105 | 2409 | |
| 2410 | {- type and term operators -}
 | |
| 2411 | ||
| 2412 | type_op0 :: Name -> (Typ, Typ -> Bool) | |
| 2413 | type_op0 name = (mk, is) | |
| 2414 | where | |
| 2415 | mk = Type (name, []) | |
| 74205 | 2416 | is (Type (c, _)) = c == name | 
| 74105 | 2417 | is _ = False | 
| 2418 | ||
| 2419 | type_op1 :: Name -> (Typ -> Typ, Typ -> Maybe Typ) | |
| 2420 | type_op1 name = (mk, dest) | |
| 2421 | where | |
| 2422 | mk ty = Type (name, [ty]) | |
| 74205 | 2423 | dest (Type (c, [ty])) | c == name = Just ty | 
| 74105 | 2424 | dest _ = Nothing | 
| 2425 | ||
| 2426 | type_op2 :: Name -> (Typ -> Typ -> Typ, Typ -> Maybe (Typ, Typ)) | |
| 2427 | type_op2 name = (mk, dest) | |
| 2428 | where | |
| 2429 | mk ty1 ty2 = Type (name, [ty1, ty2]) | |
| 74205 | 2430 | dest (Type (c, [ty1, ty2])) | c == name = Just (ty1, ty2) | 
| 74105 | 2431 | dest _ = Nothing | 
| 2432 | ||
| 2433 | op0 :: Name -> (Term, Term -> Bool) | |
| 2434 | op0 name = (mk, is) | |
| 2435 | where | |
| 2436 | mk = Const (name, []) | |
| 2437 | is (Const (c, _)) = c == name | |
| 2438 | is _ = False | |
| 2439 | ||
| 2440 | op1 :: Name -> (Term -> Term, Term -> Maybe Term) | |
| 2441 | op1 name = (mk, dest) | |
| 2442 | where | |
| 2443 | mk t = App (Const (name, []), t) | |
| 2444 | dest (App (Const (c, _), t)) | c == name = Just t | |
| 2445 | dest _ = Nothing | |
| 2446 | ||
| 2447 | op2 :: Name -> (Term -> Term -> Term, Term -> Maybe (Term, Term)) | |
| 2448 | op2 name = (mk, dest) | |
| 2449 | where | |
| 2450 | mk t u = App (App (Const (name, []), t), u) | |
| 2451 | dest (App (App (Const (c, _), t), u)) | c == name = Just (t, u) | |
| 2452 | dest _ = Nothing | |
| 2453 | ||
| 74197 | 2454 | typed_op0 :: Name -> (Typ -> Term, Term -> Maybe Typ) | 
| 2455 | typed_op0 name = (mk, dest) | |
| 2456 | where | |
| 2457 | mk ty = Const (name, [ty]) | |
| 2458 | dest (Const (c, [ty])) | c == name = Just ty | |
| 2459 | dest _ = Nothing | |
| 2460 | ||
| 74188 | 2461 | typed_op1 :: Name -> (Typ -> Term -> Term, Term -> Maybe (Typ, Term)) | 
| 2462 | typed_op1 name = (mk, dest) | |
| 2463 | where | |
| 2464 | mk ty t = App (Const (name, [ty]), t) | |
| 2465 | dest (App (Const (c, [ty]), t)) | c == name = Just (ty, t) | |
| 2466 | dest _ = Nothing | |
| 2467 | ||
| 74105 | 2468 | typed_op2 :: Name -> (Typ -> Term -> Term -> Term, Term -> Maybe (Typ, Term, Term)) | 
| 2469 | typed_op2 name = (mk, dest) | |
| 2470 | where | |
| 2471 | mk ty t u = App (App (Const (name, [ty]), t), u) | |
| 2472 | dest (App (App (Const (c, [ty]), t), u)) | c == name = Just (ty, t, u) | |
| 2473 | dest _ = Nothing | |
| 2474 | ||
| 74124 | 2475 | binder :: Name -> (Free -> Term -> Term, Name.Context -> Term -> Maybe (Free, Term)) | 
| 2476 | binder name = (mk, dest) | |
| 2477 | where | |
| 2478 | mk (a, ty) b = App (Const (name, [ty]), lambda (a, ty) b) | |
| 74214 | 2479 | dest names (App (Const (c, _), t)) | c == name = dest_lambda names t | 
| 74124 | 2480 | dest _ _ = Nothing | 
| 74105 | 2481 | |
| 2482 | ||
| 2483 | {- type operations -}
 | |
| 69240 | 2484 | |
| 2485 | dummyS :: Sort | |
| 2486 | dummyS = [""] | |
| 2487 | ||
| 74105 | 2488 | dummyT :: Typ; is_dummyT :: Typ -> Bool | 
| 2489 | (dummyT, is_dummyT) = type_op0 \<open>\<^type_name>\<open>dummy\<close>\<close> | |
| 2490 | ||
| 2491 | propT :: Typ; is_propT :: Typ -> Bool | |
| 2492 | (propT, is_propT) = type_op0 \<open>\<^type_name>\<open>prop\<close>\<close> | |
| 2493 | ||
| 2494 | (-->) :: Typ -> Typ -> Typ; dest_funT :: Typ -> Maybe (Typ, Typ) | |
| 2495 | ((-->), dest_funT) = type_op2 \<open>\<^type_name>\<open>fun\<close>\<close> | |
| 2496 | ||
| 2497 | (--->) :: [Typ] -> Typ -> Typ | |
| 2498 | [] ---> b = b | |
| 2499 | (a : as) ---> b = a --> (as ---> b) | |
| 2500 | ||
| 2501 | ||
| 2502 | {- term operations -}
 | |
| 74100 | 2503 | |
| 2504 | aconv :: Term -> Term -> Bool | |
| 2505 | aconv (App (t1, u1)) (App (t2, u2)) = aconv t1 t2 && aconv u1 u2 | |
| 2506 | aconv (Abs (_, ty1, t1)) (Abs (_, ty2, t2)) = aconv t1 t2 && ty1 == ty2 | |
| 2507 | aconv a1 a2 = a1 == a2 | |
| 2508 | ||
| 2509 | list_comb :: Term -> [Term] -> Term | |
| 2510 | list_comb f [] = f | |
| 2511 | list_comb f (t : ts) = list_comb (App (f, t)) ts | |
| 2512 | ||
| 2513 | strip_comb :: Term -> (Term, [Term]) | |
| 2514 | strip_comb tm = strip (tm, []) | |
| 2515 | where | |
| 2516 | strip (App (f, t), ts) = strip (f, t : ts) | |
| 2517 | strip x = x | |
| 2518 | ||
| 2519 | head_of :: Term -> Term | |
| 2520 | head_of (App (f, _)) = head_of f | |
| 2521 | head_of u = u | |
| 69240 | 2522 | \<close> | 
| 2523 | ||
| 74105 | 2524 | generate_file "Isabelle/Pure.hs" = \<open> | 
| 2525 | {-  Title:      Isabelle/Term.hs
 | |
| 2526 | Author: Makarius | |
| 2527 | LICENSE: BSD 3-clause (Isabelle) | |
| 2528 | ||
| 2529 | Support for Isabelle/Pure logic. | |
| 2530 | ||
| 74178 | 2531 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/logic.ML\<close>. | 
| 74105 | 2532 | -} | 
| 2533 | ||
| 2534 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 2535 | ||
| 2536 | module Isabelle.Pure ( | |
| 74188 | 2537 | mk_forall_op, dest_forall_op, mk_forall, dest_forall, | 
| 2538 | mk_equals, dest_equals, mk_implies, dest_implies | |
| 74105 | 2539 | ) | 
| 2540 | where | |
| 2541 | ||
| 74124 | 2542 | import qualified Isabelle.Name as Name | 
| 74105 | 2543 | import Isabelle.Term | 
| 2544 | ||
| 74188 | 2545 | mk_forall_op :: Typ -> Term -> Term; dest_forall_op :: Term -> Maybe (Typ, Term) | 
| 2546 | (mk_forall_op, dest_forall_op) = typed_op1 \<open>\<^const_name>\<open>Pure.all\<close>\<close> | |
| 2547 | ||
| 74124 | 2548 | mk_forall :: Free -> Term -> Term; dest_forall :: Name.Context -> Term -> Maybe (Free, Term) | 
| 2549 | (mk_forall, dest_forall) = binder \<open>\<^const_name>\<open>Pure.all\<close>\<close> | |
| 74105 | 2550 | |
| 2551 | mk_equals :: Typ -> Term -> Term -> Term; dest_equals :: Term -> Maybe (Typ, Term, Term) | |
| 2552 | (mk_equals, dest_equals) = typed_op2 \<open>\<^const_name>\<open>Pure.eq\<close>\<close> | |
| 2553 | ||
| 2554 | mk_implies :: Term -> Term -> Term; dest_implies :: Term -> Maybe (Term, Term) | |
| 2555 | (mk_implies, dest_implies) = op2 \<open>\<^const_name>\<open>Pure.imp\<close>\<close> | |
| 2556 | \<close> | |
| 2557 | ||
| 2558 | generate_file "Isabelle/HOL.hs" = \<open> | |
| 2559 | {-  Title:      Isabelle/Term.hs
 | |
| 2560 | Author: Makarius | |
| 2561 | LICENSE: BSD 3-clause (Isabelle) | |
| 2562 | ||
| 2563 | Support for Isabelle/HOL logic. | |
| 2564 | ||
| 74178 | 2565 | See \<^file>\<open>$ISABELLE_HOME/src/HOL/Tools/hologic.ML\<close>. | 
| 74105 | 2566 | -} | 
| 2567 | ||
| 2568 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 2569 | ||
| 2570 | module Isabelle.HOL ( | |
| 2571 | boolT, is_boolT, mk_trueprop, dest_trueprop, | |
| 2572 | mk_setT, dest_setT, mk_mem, dest_mem, | |
| 2573 | mk_eq, dest_eq, true, is_true, false, is_false, | |
| 2574 | mk_not, dest_not, mk_conj, dest_conj, mk_disj, dest_disj, | |
| 2575 | mk_imp, dest_imp, mk_iff, dest_iff, | |
| 74188 | 2576 | mk_all_op, dest_all_op, mk_ex_op, dest_ex_op, | 
| 74197 | 2577 | mk_all, dest_all, mk_ex, dest_ex, | 
| 2578 | mk_undefined, dest_undefined | |
| 74105 | 2579 | ) | 
| 2580 | where | |
| 2581 | ||
| 74124 | 2582 | import qualified Isabelle.Name as Name | 
| 74105 | 2583 | import Isabelle.Term | 
| 2584 | ||
| 2585 | ||
| 2586 | boolT :: Typ; is_boolT :: Typ -> Bool | |
| 2587 | (boolT, is_boolT) = type_op0 \<open>\<^type_name>\<open>bool\<close>\<close> | |
| 2588 | ||
| 2589 | mk_trueprop :: Term -> Term; dest_trueprop :: Term -> Maybe Term | |
| 2590 | (mk_trueprop, dest_trueprop) = op1 \<open>\<^const_name>\<open>Trueprop\<close>\<close> | |
| 2591 | ||
| 2592 | mk_setT :: Typ -> Typ; dest_setT :: Typ -> Maybe Typ | |
| 2593 | (mk_setT, dest_setT) = type_op1 \<open>\<^type_name>\<open>set\<close>\<close> | |
| 2594 | ||
| 2595 | mk_mem :: Typ -> Term -> Term -> Term; dest_mem :: Term -> Maybe (Typ, Term, Term) | |
| 2596 | (mk_mem, dest_mem) = typed_op2 \<open>\<^const_name>\<open>Set.member\<close>\<close> | |
| 2597 | ||
| 2598 | mk_eq :: Typ -> Term -> Term -> Term; dest_eq :: Term -> Maybe (Typ, Term, Term) | |
| 2599 | (mk_eq, dest_eq) = typed_op2 \<open>\<^const_name>\<open>HOL.eq\<close>\<close> | |
| 2600 | ||
| 2601 | true :: Term; is_true :: Term -> Bool | |
| 2602 | (true, is_true) = op0 \<open>\<^const_name>\<open>True\<close>\<close> | |
| 2603 | ||
| 2604 | false :: Term; is_false :: Term -> Bool | |
| 2605 | (false, is_false) = op0 \<open>\<^const_name>\<open>False\<close>\<close> | |
| 2606 | ||
| 2607 | mk_not :: Term -> Term; dest_not :: Term -> Maybe Term | |
| 2608 | (mk_not, dest_not) = op1 \<open>\<^const_name>\<open>Not\<close>\<close> | |
| 2609 | ||
| 2610 | mk_conj :: Term -> Term -> Term; dest_conj :: Term -> Maybe (Term, Term) | |
| 2611 | (mk_conj, dest_conj) = op2 \<open>\<^const_name>\<open>conj\<close>\<close> | |
| 2612 | ||
| 2613 | mk_disj :: Term -> Term -> Term; dest_disj :: Term -> Maybe (Term, Term) | |
| 2614 | (mk_disj, dest_disj) = op2 \<open>\<^const_name>\<open>disj\<close>\<close> | |
| 2615 | ||
| 2616 | mk_imp :: Term -> Term -> Term; dest_imp :: Term -> Maybe (Term, Term) | |
| 2617 | (mk_imp, dest_imp) = op2 \<open>\<^const_name>\<open>implies\<close>\<close> | |
| 2618 | ||
| 2619 | mk_iff :: Term -> Term -> Term | |
| 2620 | mk_iff = mk_eq boolT | |
| 2621 | ||
| 2622 | dest_iff :: Term -> Maybe (Term, Term) | |
| 2623 | dest_iff tm = | |
| 2624 | case dest_eq tm of | |
| 2625 | Just (ty, t, u) | ty == boolT -> Just (t, u) | |
| 2626 | _ -> Nothing | |
| 2627 | ||
| 74188 | 2628 | mk_all_op :: Typ -> Term -> Term; dest_all_op :: Term -> Maybe (Typ, Term) | 
| 2629 | (mk_all_op, dest_all_op) = typed_op1 \<open>\<^const_name>\<open>All\<close>\<close> | |
| 2630 | ||
| 2631 | mk_ex_op :: Typ -> Term -> Term; dest_ex_op :: Term -> Maybe (Typ, Term) | |
| 2632 | (mk_ex_op, dest_ex_op) = typed_op1 \<open>\<^const_name>\<open>Ex\<close>\<close> | |
| 2633 | ||
| 74124 | 2634 | mk_all :: Free -> Term -> Term; dest_all :: Name.Context -> Term -> Maybe (Free, Term) | 
| 2635 | (mk_all, dest_all) = binder \<open>\<^const_name>\<open>All\<close>\<close> | |
| 2636 | ||
| 2637 | mk_ex :: Free -> Term -> Term; dest_ex :: Name.Context -> Term -> Maybe (Free, Term) | |
| 2638 | (mk_ex, dest_ex) = binder \<open>\<^const_name>\<open>Ex\<close>\<close> | |
| 74197 | 2639 | |
| 2640 | mk_undefined :: Typ -> Term; dest_undefined :: Term -> Maybe Typ | |
| 2641 | (mk_undefined, dest_undefined) = typed_op0 \<open>\<^const_name>\<open>undefined\<close>\<close> | |
| 74105 | 2642 | \<close> | 
| 2643 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 2644 | generate_file "Isabelle/Term_XML/Encode.hs" = \<open> | 
| 69445 | 2645 | {-  Title:      Isabelle/Term_XML/Encode.hs
 | 
| 69240 | 2646 | Author: Makarius | 
| 2647 | LICENSE: BSD 3-clause (Isabelle) | |
| 2648 | ||
| 2649 | XML data representation of lambda terms. | |
| 69280 | 2650 | |
| 74178 | 2651 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>. | 
| 69240 | 2652 | -} | 
| 2653 | ||
| 2654 | {-# LANGUAGE LambdaCase #-}
 | |
| 2655 | ||
| 70845 | 2656 | module Isabelle.Term_XML.Encode (indexname, sort, typ, typ_body, term) | 
| 69240 | 2657 | where | 
| 2658 | ||
| 2659 | import Isabelle.Library | |
| 2660 | import Isabelle.XML.Encode | |
| 2661 | import Isabelle.Term | |
| 2662 | ||
| 70845 | 2663 | indexname :: P Indexname | 
| 2664 | indexname (a, b) = if b == 0 then [a] else [a, int_atom b] | |
| 69240 | 2665 | |
| 2666 | sort :: T Sort | |
| 2667 | sort = list string | |
| 2668 | ||
| 2669 | typ :: T Typ | |
| 2670 | typ ty = | |
| 2671 | ty |> variant | |
| 2672 |    [\case { Type (a, b) -> Just ([a], list typ b); _ -> Nothing },
 | |
| 2673 |     \case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing },
 | |
| 70845 | 2674 |     \case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }]
 | 
| 2675 | ||
| 2676 | typ_body :: T Typ | |
| 2677 | typ_body ty = if is_dummyT ty then [] else typ ty | |
| 69240 | 2678 | |
| 2679 | term :: T Term | |
| 2680 | term t = | |
| 2681 | t |> variant | |
| 70784 
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
 wenzelm parents: 
70667diff
changeset | 2682 |    [\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing },
 | 
| 70845 | 2683 |     \case { Free (a, b) -> Just ([a], typ_body b); _ -> Nothing },
 | 
| 2684 |     \case { Var (a, b) -> Just (indexname a, typ_body b); _ -> Nothing },
 | |
| 2685 |     \case { Bound a -> Just ([], int a); _ -> Nothing },
 | |
| 69240 | 2686 |     \case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing },
 | 
| 2687 |     \case { App a -> Just ([], pair term term a); _ -> Nothing }]
 | |
| 2688 | \<close> | |
| 2689 | ||
| 69444 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 wenzelm parents: 
69381diff
changeset | 2690 | generate_file "Isabelle/Term_XML/Decode.hs" = \<open> | 
| 69445 | 2691 | {-  Title:      Isabelle/Term_XML/Decode.hs
 | 
| 69240 | 2692 | Author: Makarius | 
| 2693 | LICENSE: BSD 3-clause (Isabelle) | |
| 2694 | ||
| 2695 | XML data representation of lambda terms. | |
| 69280 | 2696 | |
| 74178 | 2697 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>. | 
| 69240 | 2698 | -} | 
| 2699 | ||
| 73177 | 2700 | {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-}
 | 
| 2701 | ||
| 70845 | 2702 | module Isabelle.Term_XML.Decode (indexname, sort, typ, typ_body, term) | 
| 69240 | 2703 | where | 
| 2704 | ||
| 2705 | import Isabelle.Library | |
| 2706 | import Isabelle.XML.Decode | |
| 2707 | import Isabelle.Term | |
| 2708 | ||
| 2709 | ||
| 70845 | 2710 | indexname :: P Indexname | 
| 2711 | indexname [a] = (a, 0) | |
| 2712 | indexname [a, b] = (a, int_atom b) | |
| 2713 | ||
| 69240 | 2714 | sort :: T Sort | 
| 2715 | sort = list string | |
| 2716 | ||
| 2717 | typ :: T Typ | |
| 2718 | typ ty = | |
| 2719 | ty |> variant | |
| 2720 | [\([a], b) -> Type (a, list typ b), | |
| 2721 | \([a], b) -> TFree (a, sort b), | |
| 70845 | 2722 | \(a, b) -> TVar (indexname a, sort b)] | 
| 2723 | ||
| 2724 | typ_body :: T Typ | |
| 2725 | typ_body [] = dummyT | |
| 2726 | typ_body body = typ body | |
| 69240 | 2727 | |
| 2728 | term :: T Term | |
| 2729 | term t = | |
| 2730 | t |> variant | |
| 70784 
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
 wenzelm parents: 
70667diff
changeset | 2731 | [\([a], b) -> Const (a, list typ b), | 
| 70845 | 2732 | \([a], b) -> Free (a, typ_body b), | 
| 2733 | \(a, b) -> Var (indexname a, typ_body b), | |
| 2734 | \([], a) -> Bound (int a), | |
| 69240 | 2735 | \([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d), | 
| 2736 | \([], a) -> App (pair term term a)] | |
| 2737 | \<close> | |
| 2738 | ||
| 74129 | 2739 | generate_file "Isabelle/XML/Classes.hs" = \<open> | 
| 2740 | {- generated by Isabelle -}
 | |
| 2741 | ||
| 2742 | {-  Title:      Isabelle/XML/Classes.hs
 | |
| 2743 | Author: Makarius | |
| 2744 | LICENSE: BSD 3-clause (Isabelle) | |
| 2745 | ||
| 2746 | Type classes for XML data representation. | |
| 2747 | -} | |
| 2748 | ||
| 2749 | {-# LANGUAGE TypeSynonymInstances #-}
 | |
| 2750 | {-# LANGUAGE FlexibleInstances #-}
 | |
| 2751 | ||
| 2752 | module Isabelle.XML.Classes | |
| 2753 | (Encode_Atom(..), Decode_Atom(..), Encode (..), Decode (..)) | |
| 2754 | where | |
| 2755 | ||
| 2756 | import qualified Isabelle.XML as XML | |
| 2757 | import qualified Isabelle.XML.Encode as Encode | |
| 2758 | import qualified Isabelle.XML.Decode as Decode | |
| 2759 | import qualified Isabelle.Term_XML.Encode as Encode | |
| 2760 | import qualified Isabelle.Term_XML.Decode as Decode | |
| 2761 | import qualified Isabelle.Properties as Properties | |
| 2762 | import Isabelle.Bytes (Bytes) | |
| 2763 | import Isabelle.Term (Typ, Term) | |
| 2764 | ||
| 2765 | ||
| 2766 | class Encode_Atom a where encode_atom :: Encode.A a | |
| 2767 | class Decode_Atom a where decode_atom :: Decode.A a | |
| 2768 | ||
| 2769 | instance Encode_Atom Int where encode_atom = Encode.int_atom | |
| 2770 | instance Decode_Atom Int where decode_atom = Decode.int_atom | |
| 2771 | ||
| 2772 | instance Encode_Atom Bool where encode_atom = Encode.bool_atom | |
| 2773 | instance Decode_Atom Bool where decode_atom = Decode.bool_atom | |
| 2774 | ||
| 2775 | instance Encode_Atom () where encode_atom = Encode.unit_atom | |
| 2776 | instance Decode_Atom () where decode_atom = Decode.unit_atom | |
| 2777 | ||
| 2778 | ||
| 2779 | class Encode a where encode :: Encode.T a | |
| 2780 | class Decode a where decode :: Decode.T a | |
| 2781 | ||
| 2782 | instance Encode Bytes where encode = Encode.string | |
| 2783 | instance Decode Bytes where decode = Decode.string | |
| 2784 | ||
| 2785 | instance Encode Int where encode = Encode.int | |
| 2786 | instance Decode Int where decode = Decode.int | |
| 2787 | ||
| 2788 | instance Encode Bool where encode = Encode.bool | |
| 2789 | instance Decode Bool where decode = Decode.bool | |
| 2790 | ||
| 2791 | instance Encode () where encode = Encode.unit | |
| 2792 | instance Decode () where decode = Decode.unit | |
| 2793 | ||
| 2794 | instance (Encode a, Encode b) => Encode (a, b) | |
| 2795 | where encode = Encode.pair encode encode | |
| 2796 | instance (Decode a, Decode b) => Decode (a, b) | |
| 2797 | where decode = Decode.pair decode decode | |
| 2798 | ||
| 2799 | instance (Encode a, Encode b, Encode c) => Encode (a, b, c) | |
| 2800 | where encode = Encode.triple encode encode encode | |
| 2801 | instance (Decode a, Decode b, Decode c) => Decode (a, b, c) | |
| 2802 | where decode = Decode.triple decode decode decode | |
| 2803 | ||
| 2804 | instance Encode a => Encode [a] where encode = Encode.list encode | |
| 2805 | instance Decode a => Decode [a] where decode = Decode.list decode | |
| 2806 | ||
| 2807 | instance Encode a => Encode (Maybe a) where encode = Encode.option encode | |
| 2808 | instance Decode a => Decode (Maybe a) where decode = Decode.option decode | |
| 2809 | ||
| 2810 | instance Encode XML.Tree where encode = Encode.tree | |
| 2811 | instance Decode XML.Tree where decode = Decode.tree | |
| 2812 | ||
| 2813 | instance Encode Properties.T where encode = Encode.properties | |
| 2814 | instance Decode Properties.T where decode = Decode.properties | |
| 2815 | ||
| 2816 | instance Encode Typ where encode = Encode.typ | |
| 2817 | instance Decode Typ where decode = Decode.typ | |
| 2818 | ||
| 2819 | instance Encode Term where encode = Encode.term | |
| 2820 | instance Decode Term where decode = Decode.term | |
| 2821 | \<close> | |
| 2822 | ||
| 69459 | 2823 | generate_file "Isabelle/UUID.hs" = \<open> | 
| 2824 | {-  Title:      Isabelle/UUID.hs
 | |
| 2825 | Author: Makarius | |
| 2826 | LICENSE: BSD 3-clause (Isabelle) | |
| 2827 | ||
| 2828 | Universally unique identifiers. | |
| 2829 | ||
| 2830 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/uuid.scala\<close>. | |
| 2831 | -} | |
| 2832 | ||
| 74095 | 2833 | module Isabelle.UUID (T, random, print, parse) | 
| 69459 | 2834 | where | 
| 2835 | ||
| 74095 | 2836 | import Prelude hiding (print) | 
| 2837 | ||
| 69459 | 2838 | import Data.UUID (UUID) | 
| 2839 | import qualified Data.UUID as UUID | |
| 2840 | import Data.UUID.V4 (nextRandom) | |
| 2841 | ||
| 74084 | 2842 | import qualified Isabelle.Bytes as Bytes | 
| 2843 | import Isabelle.Bytes (Bytes) | |
| 2844 | ||
| 69459 | 2845 | |
| 69462 | 2846 | type T = UUID | 
| 2847 | ||
| 2848 | random :: IO T | |
| 69459 | 2849 | random = nextRandom | 
| 2850 | ||
| 74095 | 2851 | print :: T -> Bytes | 
| 2852 | print = Bytes.make . UUID.toASCIIBytes | |
| 2853 | ||
| 2854 | parse :: Bytes -> Maybe T | |
| 2855 | parse = UUID.fromASCIIBytes . Bytes.unmake | |
| 69459 | 2856 | \<close> | 
| 2857 | ||
| 69448 | 2858 | generate_file "Isabelle/Byte_Message.hs" = \<open> | 
| 2859 | {-  Title:      Isabelle/Byte_Message.hs
 | |
| 69446 | 2860 | Author: Makarius | 
| 2861 | LICENSE: BSD 3-clause (Isabelle) | |
| 2862 | ||
| 69448 | 2863 | Byte-oriented messages. | 
| 2864 | ||
| 2865 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.ML\<close> | |
| 2866 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.scala\<close>. | |
| 69446 | 2867 | -} | 
| 2868 | ||
| 74083 | 2869 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 73177 | 2870 | {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-}
 | 
| 2871 | ||
| 69452 | 2872 | module Isabelle.Byte_Message ( | 
| 69467 | 2873 | write, write_line, | 
| 74088 | 2874 | read, read_block, read_line, | 
| 74187 | 2875 | make_message, write_message, read_message, | 
| 2876 | exchange_message, exchange_message0, | |
| 73178 | 2877 | make_line_message, write_line_message, read_line_message, | 
| 2878 | read_yxml, write_yxml | |
| 69452 | 2879 | ) | 
| 69446 | 2880 | where | 
| 2881 | ||
| 69449 | 2882 | import Prelude hiding (read) | 
| 69454 | 2883 | import Data.Maybe | 
| 69446 | 2884 | import qualified Data.ByteString as ByteString | 
| 74084 | 2885 | import qualified Isabelle.Bytes as Bytes | 
| 2886 | import Isabelle.Bytes (Bytes) | |
| 74135 | 2887 | import qualified Isabelle.Symbol as Symbol | 
| 74080 
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
 wenzelm parents: 
73246diff
changeset | 2888 | import qualified Isabelle.UTF8 as UTF8 | 
| 73178 | 2889 | import qualified Isabelle.XML as XML | 
| 2890 | import qualified Isabelle.YXML as YXML | |
| 73177 | 2891 | |
| 69446 | 2892 | import Network.Socket (Socket) | 
| 74082 | 2893 | import qualified Network.Socket.ByteString as Socket | 
| 69446 | 2894 | |
| 74088 | 2895 | import Isabelle.Library | 
| 69446 | 2896 | import qualified Isabelle.Value as Value | 
| 2897 | ||
| 2898 | ||
| 69452 | 2899 | {- output operations -}
 | 
| 2900 | ||
| 74084 | 2901 | write :: Socket -> [Bytes] -> IO () | 
| 2902 | write socket = Socket.sendMany socket . map Bytes.unmake | |
| 2903 | ||
| 2904 | write_line :: Socket -> Bytes -> IO () | |
| 74083 | 2905 | write_line socket s = write socket [s, "\n"] | 
| 69467 | 2906 | |
| 69452 | 2907 | |
| 2908 | {- input operations -}
 | |
| 2909 | ||
| 74084 | 2910 | read :: Socket -> Int -> IO Bytes | 
| 69452 | 2911 | read socket n = read_body 0 [] | 
| 69449 | 2912 | where | 
| 74084 | 2913 | result = Bytes.concat . reverse | 
| 69452 | 2914 | read_body len ss = | 
| 69449 | 2915 | if len >= n then return (result ss) | 
| 2916 | else | |
| 2917 | (do | |
| 74082 | 2918 | s <- Socket.recv socket (min (n - len) 8192) | 
| 69449 | 2919 | case ByteString.length s of | 
| 2920 | 0 -> return (result ss) | |
| 74084 | 2921 | m -> read_body (len + m) (Bytes.make s : ss)) | 
| 2922 | ||
| 2923 | read_block :: Socket -> Int -> IO (Maybe Bytes, Int) | |
| 69452 | 2924 | read_block socket n = do | 
| 2925 | msg <- read socket n | |
| 74084 | 2926 | let len = Bytes.length msg | 
| 69452 | 2927 | return (if len == n then Just msg else Nothing, len) | 
| 69449 | 2928 | |
| 74084 | 2929 | read_line :: Socket -> IO (Maybe Bytes) | 
| 69452 | 2930 | read_line socket = read_body [] | 
| 69446 | 2931 | where | 
| 74134 | 2932 | result = trim_line . Bytes.pack . reverse | 
| 69452 | 2933 | read_body bs = do | 
| 74082 | 2934 | s <- Socket.recv socket 1 | 
| 69446 | 2935 | case ByteString.length s of | 
| 2936 | 0 -> return (if null bs then Nothing else Just (result bs)) | |
| 2937 | 1 -> | |
| 2938 | case ByteString.head s of | |
| 2939 | 10 -> return (Just (result bs)) | |
| 69452 | 2940 | b -> read_body (b : bs) | 
| 69446 | 2941 | |
| 69448 | 2942 | |
| 69454 | 2943 | {- messages with multiple chunks (arbitrary content) -}
 | 
| 2944 | ||
| 74084 | 2945 | make_header :: [Int] -> [Bytes] | 
| 74095 | 2946 | make_header ns = [space_implode "," (map Value.print_int ns), "\n"] | 
| 69454 | 2947 | |
| 74084 | 2948 | make_message :: [Bytes] -> [Bytes] | 
| 2949 | make_message chunks = make_header (map Bytes.length chunks) <> chunks | |
| 2950 | ||
| 2951 | write_message :: Socket -> [Bytes] -> IO () | |
| 69476 | 2952 | write_message socket = write socket . make_message | 
| 69454 | 2953 | |
| 74084 | 2954 | parse_header :: Bytes -> [Int] | 
| 69454 | 2955 | parse_header line = | 
| 2956 | let | |
| 74132 | 2957 | res = map Value.parse_nat (space_explode ',' line) | 
| 69454 | 2958 | in | 
| 2959 | if all isJust res then map fromJust res | |
| 74081 | 2960 |     else error ("Malformed message header: " <> quote (UTF8.decode line))
 | 
| 69454 | 2961 | |
| 74084 | 2962 | read_chunk :: Socket -> Int -> IO Bytes | 
| 69454 | 2963 | read_chunk socket n = do | 
| 2964 | res <- read_block socket n | |
| 2965 | return $ | |
| 2966 | case res of | |
| 2967 | (Just chunk, _) -> chunk | |
| 2968 | (Nothing, len) -> | |
| 74081 | 2969 |         error ("Malformed message chunk: unexpected EOF after " <>
 | 
| 2970 | show len <> " of " <> show n <> " bytes") | |
| 69454 | 2971 | |
| 74084 | 2972 | read_message :: Socket -> IO (Maybe [Bytes]) | 
| 69454 | 2973 | read_message socket = do | 
| 2974 | res <- read_line socket | |
| 2975 | case res of | |
| 2976 | Just line -> Just <$> mapM (read_chunk socket) (parse_header line) | |
| 2977 | Nothing -> return Nothing | |
| 2978 | ||
| 74186 | 2979 | exchange_message :: Socket -> [Bytes] -> IO (Maybe [Bytes]) | 
| 2980 | exchange_message socket msg = do | |
| 2981 | write_message socket msg | |
| 2982 | read_message socket | |
| 2983 | ||
| 74187 | 2984 | exchange_message0 :: Socket -> [Bytes] -> IO () | 
| 2985 | exchange_message0 socket msg = do | |
| 2986 | _ <- exchange_message socket msg | |
| 2987 | return () | |
| 2988 | ||
| 69454 | 2989 | |
| 69448 | 2990 | -- hybrid messages: line or length+block (with content restriction) | 
| 2991 | ||
| 74084 | 2992 | is_length :: Bytes -> Bool | 
| 69452 | 2993 | is_length msg = | 
| 74135 | 2994 | not (Bytes.null msg) && Bytes.all_char (\c -> '0' <= c && c <= '9') msg | 
| 74084 | 2995 | |
| 2996 | is_terminated :: Bytes -> Bool | |
| 69452 | 2997 | is_terminated msg = | 
| 74135 | 2998 | not (Bytes.null msg) && Symbol.is_ascii_line_terminator (Bytes.char $ Bytes.last msg) | 
| 74084 | 2999 | |
| 3000 | make_line_message :: Bytes -> [Bytes] | |
| 69476 | 3001 | make_line_message msg = | 
| 74084 | 3002 | let n = Bytes.length msg in | 
| 69476 | 3003 | if is_length msg || is_terminated msg then | 
| 74081 | 3004 |       error ("Bad content for line message:\n" <> take 100 (UTF8.decode msg))
 | 
| 69476 | 3005 | else | 
| 74135 | 3006 | (if n > 100 || Bytes.any_char (== '\n') msg then make_header [n + 1] else []) <> [msg, "\n"] | 
| 74084 | 3007 | |
| 3008 | write_line_message :: Socket -> Bytes -> IO () | |
| 69476 | 3009 | write_line_message socket = write socket . make_line_message | 
| 69448 | 3010 | |
| 74084 | 3011 | read_line_message :: Socket -> IO (Maybe Bytes) | 
| 69448 | 3012 | read_line_message socket = do | 
| 69446 | 3013 | opt_line <- read_line socket | 
| 3014 | case opt_line of | |
| 3015 | Nothing -> return Nothing | |
| 3016 | Just line -> | |
| 74095 | 3017 | case Value.parse_nat line of | 
| 69446 | 3018 | Nothing -> return $ Just line | 
| 74134 | 3019 | Just n -> fmap trim_line . fst <$> read_block socket n | 
| 73178 | 3020 | |
| 3021 | ||
| 3022 | read_yxml :: Socket -> IO (Maybe XML.Body) | |
| 3023 | read_yxml socket = do | |
| 3024 | res <- read_line_message socket | |
| 74095 | 3025 | return (YXML.parse_body <$> res) | 
| 73178 | 3026 | |
| 3027 | write_yxml :: Socket -> XML.Body -> IO () | |
| 3028 | write_yxml socket body = | |
| 74095 | 3029 | write_line_message socket (YXML.string_of_body body) | 
| 69446 | 3030 | \<close> | 
| 3031 | ||
| 71692 | 3032 | generate_file "Isabelle/Isabelle_Thread.hs" = \<open> | 
| 3033 | {-  Title:      Isabelle/Isabelle_Thread.hs
 | |
| 69473 | 3034 | Author: Makarius | 
| 3035 | LICENSE: BSD 3-clause (Isabelle) | |
| 3036 | ||
| 71692 | 3037 | Isabelle-specific thread management. | 
| 3038 | ||
| 3039 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\<close> | |
| 3040 | and \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\<close>. | |
| 69473 | 3041 | -} | 
| 3042 | ||
| 3043 | {-# LANGUAGE NamedFieldPuns #-}
 | |
| 3044 | ||
| 71692 | 3045 | module Isabelle.Isabelle_Thread ( | 
| 69473 | 3046 | ThreadId, Result, | 
| 69494 | 3047 | find_id, | 
| 69473 | 3048 | properties, change_properties, | 
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3049 | add_resource, del_resource, bracket_resource, | 
| 69497 | 3050 | is_stopped, expose_stopped, stop, | 
| 69496 | 3051 | my_uuid, stop_uuid, | 
| 69494 | 3052 | Fork, fork_finally, fork) | 
| 69473 | 3053 | where | 
| 3054 | ||
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3055 | import Data.Unique | 
| 69473 | 3056 | import Data.IORef | 
| 3057 | import System.IO.Unsafe | |
| 3058 | ||
| 69494 | 3059 | import qualified Data.List as List | 
| 69497 | 3060 | import Control.Monad (when, forM_) | 
| 69473 | 3061 | import Data.Map.Strict (Map) | 
| 3062 | import qualified Data.Map.Strict as Map | |
| 3063 | import Control.Exception as Exception | |
| 3064 | import Control.Concurrent (ThreadId) | |
| 3065 | import qualified Control.Concurrent as Concurrent | |
| 3066 | import Control.Concurrent.Thread (Result) | |
| 3067 | import qualified Control.Concurrent.Thread as Thread | |
| 69494 | 3068 | import qualified Isabelle.UUID as UUID | 
| 69473 | 3069 | import qualified Isabelle.Properties as Properties | 
| 3070 | ||
| 3071 | ||
| 3072 | {- thread info -}
 | |
| 3073 | ||
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3074 | type Resources = Map Unique (IO ()) | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3075 | data Info = Info {uuid :: UUID.T, props :: Properties.T, stopped :: Bool, resources :: Resources}
 | 
| 69473 | 3076 | type Infos = Map ThreadId Info | 
| 3077 | ||
| 3078 | lookup_info :: Infos -> ThreadId -> Maybe Info | |
| 69494 | 3079 | lookup_info infos id = Map.lookup id infos | 
| 3080 | ||
| 3081 | init_info :: ThreadId -> UUID.T -> Infos -> (Infos, ()) | |
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3082 | init_info id uuid infos = (Map.insert id (Info uuid [] False Map.empty) infos, ()) | 
| 69473 | 3083 | |
| 3084 | ||
| 3085 | {- global state -}
 | |
| 3086 | ||
| 3087 | {-# NOINLINE global_state #-}
 | |
| 3088 | global_state :: IORef Infos | |
| 3089 | global_state = unsafePerformIO (newIORef Map.empty) | |
| 3090 | ||
| 69494 | 3091 | find_id :: UUID.T -> IO (Maybe ThreadId) | 
| 3092 | find_id uuid = do | |
| 3093 | state <- readIORef global_state | |
| 3094 |   return $ fst <$> List.find (\(_, Info{uuid = uuid'}) -> uuid == uuid') (Map.assocs state)
 | |
| 3095 | ||
| 69473 | 3096 | get_info :: ThreadId -> IO (Maybe Info) | 
| 3097 | get_info id = do | |
| 3098 | state <- readIORef global_state | |
| 3099 | return $ lookup_info state id | |
| 3100 | ||
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3101 | map_info :: ThreadId -> (Info -> Info) -> IO (Maybe Info) | 
| 69473 | 3102 | map_info id f = | 
| 3103 | atomicModifyIORef' global_state | |
| 3104 | (\infos -> | |
| 3105 | case lookup_info infos id of | |
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3106 | Nothing -> (infos, Nothing) | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3107 | Just info -> | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3108 | let info' = f info | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3109 | in (Map.insert id info' infos, Just info')) | 
| 69473 | 3110 | |
| 69495 | 3111 | delete_info :: ThreadId -> IO () | 
| 3112 | delete_info id = | |
| 3113 | atomicModifyIORef' global_state (\infos -> (Map.delete id infos, ())) | |
| 3114 | ||
| 69473 | 3115 | |
| 3116 | {- thread properties -}
 | |
| 3117 | ||
| 69498 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3118 | my_info :: IO (Maybe Info) | 
| 69473 | 3119 | my_info = do | 
| 3120 | id <- Concurrent.myThreadId | |
| 69498 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3121 | get_info id | 
| 69473 | 3122 | |
| 3123 | properties :: IO Properties.T | |
| 69498 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3124 | properties = maybe [] props <$> my_info | 
| 69473 | 3125 | |
| 3126 | change_properties :: (Properties.T -> Properties.T) -> IO () | |
| 3127 | change_properties f = do | |
| 3128 | id <- Concurrent.myThreadId | |
| 3129 |   map_info id (\info -> info {props = f (props info)})
 | |
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3130 | return () | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3131 | |
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3132 | |
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3133 | {- managed resources -}
 | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3134 | |
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3135 | add_resource :: IO () -> IO Unique | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3136 | add_resource resource = do | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3137 | id <- Concurrent.myThreadId | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3138 | u <- newUnique | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3139 |   map_info id (\info -> info {resources = Map.insert u resource (resources info)})
 | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3140 | return u | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3141 | |
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3142 | del_resource :: Unique -> IO () | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3143 | del_resource u = do | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3144 | id <- Concurrent.myThreadId | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3145 |   map_info id (\info -> info {resources = Map.delete u (resources info)})
 | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3146 | return () | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3147 | |
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3148 | bracket_resource :: IO () -> IO a -> IO a | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3149 | bracket_resource resource body = | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3150 | Exception.bracket (add_resource resource) del_resource (const body) | 
| 69473 | 3151 | |
| 3152 | ||
| 3153 | {- stop -}
 | |
| 3154 | ||
| 3155 | is_stopped :: IO Bool | |
| 69498 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3156 | is_stopped = maybe False stopped <$> my_info | 
| 69473 | 3157 | |
| 69497 | 3158 | expose_stopped :: IO () | 
| 3159 | expose_stopped = do | |
| 3160 | stopped <- is_stopped | |
| 3161 | when stopped $ throw ThreadKilled | |
| 3162 | ||
| 69473 | 3163 | stop :: ThreadId -> IO () | 
| 69499 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3164 | stop id = do | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3165 |   info <- map_info id (\info -> info {stopped = True})
 | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3166 |   let ops = case info of Nothing -> []; Just Info{resources} -> map snd (Map.toDescList resources)
 | 
| 
638fdbbc7d1f
more Haskell operations: managed resources for threads;
 wenzelm parents: 
69498diff
changeset | 3167 | sequence_ ops | 
| 69473 | 3168 | |
| 3169 | ||
| 69496 | 3170 | {- UUID -}
 | 
| 3171 | ||
| 69498 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3172 | my_uuid :: IO (Maybe UUID.T) | 
| 
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
 wenzelm parents: 
69497diff
changeset | 3173 | my_uuid = fmap uuid <$> my_info | 
| 69496 | 3174 | |
| 3175 | stop_uuid :: UUID.T -> IO () | |
| 3176 | stop_uuid uuid = do | |
| 3177 | id <- find_id uuid | |
| 3178 | forM_ id stop | |
| 3179 | ||
| 3180 | ||
| 69473 | 3181 | {- fork -}
 | 
| 3182 | ||
| 69494 | 3183 | type Fork a = (ThreadId, UUID.T, IO (Result a)) | 
| 3184 | ||
| 3185 | fork_finally :: IO a -> (Either SomeException a -> IO b) -> IO (Fork b) | |
| 69473 | 3186 | fork_finally body finally = do | 
| 69494 | 3187 | uuid <- UUID.random | 
| 69473 | 3188 | (id, result) <- | 
| 3189 | Exception.mask (\restore -> | |
| 3190 | Thread.forkIO | |
| 3191 | (Exception.try | |
| 3192 | (do | |
| 3193 | id <- Concurrent.myThreadId | |
| 69494 | 3194 | atomicModifyIORef' global_state (init_info id uuid) | 
| 69495 | 3195 | restore body) | 
| 3196 | >>= (\res -> do id <- Concurrent.myThreadId; delete_info id; finally res))) | |
| 69494 | 3197 | return (id, uuid, result) | 
| 3198 | ||
| 3199 | fork :: IO a -> IO (Fork a) | |
| 69473 | 3200 | fork body = fork_finally body Thread.result | 
| 3201 | \<close> | |
| 3202 | ||
| 69455 | 3203 | generate_file "Isabelle/Server.hs" = \<open> | 
| 3204 | {-  Title:      Isabelle/Server.hs
 | |
| 3205 | Author: Makarius | |
| 3206 | LICENSE: BSD 3-clause (Isabelle) | |
| 3207 | ||
| 3208 | TCP server on localhost. | |
| 3209 | -} | |
| 3210 | ||
| 74127 | 3211 | {-# LANGUAGE OverloadedStrings #-}
 | 
| 3212 | ||
| 69465 | 3213 | module Isabelle.Server ( | 
| 74216 | 3214 | localhost_name, localhost_prefix, localhost, publish_text, publish_stdout, | 
| 73178 | 3215 | server, connection | 
| 69465 | 3216 | ) | 
| 3217 | where | |
| 69455 | 3218 | |
| 69465 | 3219 | import Control.Monad (forever, when) | 
| 69455 | 3220 | import qualified Control.Exception as Exception | 
| 3221 | import Network.Socket (Socket) | |
| 3222 | import qualified Network.Socket as Socket | |
| 69472 | 3223 | import qualified System.IO as IO | 
| 74127 | 3224 | import qualified Data.ByteString.Char8 as Char8 | 
| 69455 | 3225 | |
| 69462 | 3226 | import Isabelle.Library | 
| 74127 | 3227 | import qualified Isabelle.Bytes as Bytes | 
| 74084 | 3228 | import Isabelle.Bytes (Bytes) | 
| 69462 | 3229 | import qualified Isabelle.UUID as UUID | 
| 69465 | 3230 | import qualified Isabelle.Byte_Message as Byte_Message | 
| 71692 | 3231 | import qualified Isabelle.Isabelle_Thread as Isabelle_Thread | 
| 69462 | 3232 | |
| 3233 | ||
| 3234 | {- server address -}
 | |
| 69455 | 3235 | |
| 74127 | 3236 | localhost_name :: Bytes | 
| 69463 | 3237 | localhost_name = "127.0.0.1" | 
| 3238 | ||
| 74211 | 3239 | localhost_prefix :: Bytes | 
| 3240 | localhost_prefix = localhost_name <> ":" | |
| 3241 | ||
| 69455 | 3242 | localhost :: Socket.HostAddress | 
| 3243 | localhost = Socket.tupleToHostAddress (127, 0, 0, 1) | |
| 3244 | ||
| 74127 | 3245 | publish_text :: Bytes -> Bytes -> UUID.T -> Bytes | 
| 69462 | 3246 | publish_text name address password = | 
| 74127 | 3247 | "server " <> quote name <> " = " <> address <> | 
| 74128 | 3248 | " (password " <> quote (show_bytes password) <> ")" | 
| 74127 | 3249 | |
| 3250 | publish_stdout :: Bytes -> Bytes -> UUID.T -> IO () | |
| 74095 | 3251 | publish_stdout name address password = | 
| 74127 | 3252 | Char8.putStrLn (Bytes.unmake $ publish_text name address password) | 
| 69462 | 3253 | |
| 3254 | ||
| 3255 | {- server -}
 | |
| 3256 | ||
| 74127 | 3257 | server :: (Bytes -> UUID.T -> IO ()) -> (Socket -> IO ()) -> IO () | 
| 69455 | 3258 | server publish handle = | 
| 69462 | 3259 | Socket.withSocketsDo $ Exception.bracket open (Socket.close . fst) (uncurry loop) | 
| 69455 | 3260 | where | 
| 74084 | 3261 | open :: IO (Socket, Bytes) | 
| 69455 | 3262 | open = do | 
| 69466 | 3263 | server_socket <- Socket.socket Socket.AF_INET Socket.Stream Socket.defaultProtocol | 
| 3264 | Socket.bind server_socket (Socket.SockAddrInet 0 localhost) | |
| 3265 | Socket.listen server_socket 50 | |
| 69462 | 3266 | |
| 69466 | 3267 | port <- Socket.socketPort server_socket | 
| 74128 | 3268 | let address = localhost_name <> ":" <> show_bytes port | 
| 69462 | 3269 | password <- UUID.random | 
| 3270 | publish address password | |
| 69455 | 3271 | |
| 74095 | 3272 | return (server_socket, UUID.print password) | 
| 69462 | 3273 | |
| 74084 | 3274 | loop :: Socket -> Bytes -> IO () | 
| 69466 | 3275 | loop server_socket password = forever $ do | 
| 69480 | 3276 | (connection, _) <- Socket.accept server_socket | 
| 71692 | 3277 | Isabelle_Thread.fork_finally | 
| 69465 | 3278 | (do | 
| 3279 | line <- Byte_Message.read_line connection | |
| 3280 | when (line == Just password) $ handle connection) | |
| 69473 | 3281 | (\finally -> do | 
| 69472 | 3282 | Socket.close connection | 
| 69473 | 3283 | case finally of | 
| 69472 | 3284 | Left exn -> IO.hPutStrLn IO.stderr $ Exception.displayException exn | 
| 3285 | Right () -> return ()) | |
| 69455 | 3286 | return () | 
| 73178 | 3287 | |
| 3288 | ||
| 3289 | {- client connection -}
 | |
| 3290 | ||
| 74095 | 3291 | connection :: String -> Bytes -> (Socket -> IO a) -> IO a | 
| 73178 | 3292 | connection port password client = | 
| 3293 | Socket.withSocketsDo $ do | |
| 3294 | addr <- resolve | |
| 3295 | Exception.bracket (open addr) Socket.close body | |
| 3296 | where | |
| 3297 | resolve = do | |
| 3298 | let hints = | |
| 3299 |             Socket.defaultHints {
 | |
| 3300 | Socket.addrFlags = [Socket.AI_NUMERICHOST, Socket.AI_NUMERICSERV], | |
| 3301 | Socket.addrSocketType = Socket.Stream } | |
| 74211 | 3302 | head <$> Socket.getAddrInfo (Just hints) (Just $ make_string localhost_name) (Just port) | 
| 73178 | 3303 | |
| 3304 | open addr = do | |
| 3305 | socket <- Socket.socket (Socket.addrFamily addr) (Socket.addrSocketType addr) | |
| 3306 | (Socket.addrProtocol addr) | |
| 3307 | Socket.connect socket $ Socket.addrAddress addr | |
| 3308 | return socket | |
| 3309 | ||
| 3310 | body socket = do | |
| 74095 | 3311 | Byte_Message.write_line socket password | 
| 73178 | 3312 | client socket | 
| 69455 | 3313 | \<close> | 
| 3314 | ||
| 74161 | 3315 | generate_file "Isabelle/Time.hs" = \<open> | 
| 3316 | {-  Title:      Isabelle/Time.hs
 | |
| 3317 | Author: Makarius | |
| 3318 | LICENSE: BSD 3-clause (Isabelle) | |
| 3319 | ||
| 3320 | Time based on milliseconds. | |
| 3321 | ||
| 3322 | See \<^file>\<open>~~/src/Pure/General/time.scala\<close> | |
| 3323 | -} | |
| 3324 | ||
| 3325 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 3326 | ||
| 3327 | module Isabelle.Time ( | |
| 3328 | Time, seconds, minutes, ms, zero, is_zero, is_relevant, | |
| 74209 | 3329 | get_seconds, get_minutes, get_ms, message, now | 
| 74161 | 3330 | ) | 
| 3331 | where | |
| 3332 | ||
| 3333 | import Text.Printf (printf) | |
| 74209 | 3334 | import Data.Time.Clock.POSIX (getPOSIXTime) | 
| 74161 | 3335 | import Isabelle.Bytes (Bytes) | 
| 3336 | import Isabelle.Library | |
| 3337 | ||
| 3338 | ||
| 74218 | 3339 | newtype Time = Time Int | 
| 74161 | 3340 | |
| 3341 | instance Eq Time where Time ms1 == Time ms2 = ms1 == ms2 | |
| 3342 | instance Ord Time where compare (Time ms1) (Time ms2) = compare ms1 ms2 | |
| 3343 | ||
| 3344 | seconds :: Double -> Time | |
| 3345 | seconds s = Time (round (s * 1000.0)) | |
| 3346 | ||
| 3347 | minutes :: Double -> Time | |
| 3348 | minutes m = Time (round (m * 60000.0)) | |
| 3349 | ||
| 3350 | ms :: Int -> Time | |
| 3351 | ms = Time | |
| 3352 | ||
| 3353 | zero :: Time | |
| 3354 | zero = ms 0 | |
| 3355 | ||
| 3356 | is_zero :: Time -> Bool | |
| 3357 | is_zero (Time ms) = ms == 0 | |
| 3358 | ||
| 3359 | is_relevant :: Time -> Bool | |
| 3360 | is_relevant (Time ms) = ms >= 1 | |
| 3361 | ||
| 3362 | get_seconds :: Time -> Double | |
| 3363 | get_seconds (Time ms) = fromIntegral ms / 1000.0 | |
| 3364 | ||
| 3365 | get_minutes :: Time -> Double | |
| 3366 | get_minutes (Time ms) = fromIntegral ms / 60000.0 | |
| 3367 | ||
| 3368 | get_ms :: Time -> Int | |
| 3369 | get_ms (Time ms) = ms | |
| 3370 | ||
| 3371 | instance Show Time where | |
| 3372 | show t = printf "%.3f" (get_seconds t) | |
| 3373 | ||
| 3374 | message :: Time -> Bytes | |
| 3375 | message t = make_bytes (show t) <> "s" | |
| 74209 | 3376 | |
| 3377 | now :: IO Time | |
| 3378 | now = do | |
| 3379 | t <- getPOSIXTime | |
| 3380 | return $ Time (round (realToFrac t * 1000.0 :: Double)) | |
| 74161 | 3381 | \<close> | 
| 3382 | ||
| 3383 | generate_file "Isabelle/Timing.hs" = \<open> | |
| 3384 | {-  Title:      Isabelle/Timing.hs
 | |
| 3385 | Author: Makarius | |
| 3386 | LICENSE: BSD 3-clause (Isabelle) | |
| 3387 | ||
| 3388 | Support for time measurement. | |
| 3389 | ||
| 3390 | See \<^file>\<open>~~/src/Pure/General/timing.ML\<close> | |
| 3391 | and \<^file>\<open>~~/src/Pure/General/timing.scala\<close> | |
| 3392 | -} | |
| 3393 | ||
| 3394 | module Isabelle.Timing ( | |
| 3395 | Timing (..), zero, is_zero, is_relevant | |
| 3396 | ) | |
| 3397 | where | |
| 3398 | ||
| 3399 | import qualified Isabelle.Time as Time | |
| 3400 | import Isabelle.Time (Time) | |
| 3401 | ||
| 3402 | data Timing = Timing {elapsed :: Time, cpu :: Time, gc :: Time}
 | |
| 3403 | deriving (Show, Eq) | |
| 3404 | ||
| 3405 | zero :: Timing | |
| 3406 | zero = Timing Time.zero Time.zero Time.zero | |
| 3407 | ||
| 3408 | is_zero :: Timing -> Bool | |
| 3409 | is_zero t = Time.is_zero (elapsed t) && Time.is_zero (cpu t) && Time.is_zero (gc t) | |
| 3410 | ||
| 3411 | is_relevant :: Timing -> Bool | |
| 3412 | is_relevant t = Time.is_relevant (elapsed t) || Time.is_relevant (cpu t) || Time.is_relevant (gc t) | |
| 3413 | \<close> | |
| 3414 | ||
| 3415 | generate_file "Isabelle/Bash.hs" = \<open> | |
| 3416 | {-  Title:      Isabelle/Bash.hs
 | |
| 3417 | Author: Makarius | |
| 3418 | LICENSE: BSD 3-clause (Isabelle) | |
| 3419 | ||
| 3420 | Support for GNU bash. | |
| 3421 | ||
| 3422 | See \<^file>\<open>$ISABELLE_HOME/src/Pure/System/bash.ML\<close> | |
| 3423 | -} | |
| 3424 | ||
| 3425 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 3426 | ||
| 3427 | module Isabelle.Bash ( | |
| 3428 | string, strings, | |
| 3429 | ||
| 3430 | Params, | |
| 3431 | get_script, get_input, get_cwd, get_putenv, get_redirect, | |
| 3432 | get_timeout, get_description, | |
| 3433 | script, input, cwd, putenv, redirect, timeout, description, | |
| 74211 | 3434 | server_run, server_kill, | 
| 3435 | server_uuid, server_interrupt, server_failure, server_result | |
| 74161 | 3436 | ) | 
| 3437 | where | |
| 3438 | ||
| 3439 | import Text.Printf (printf) | |
| 3440 | import qualified Isabelle.Symbol as Symbol | |
| 3441 | import qualified Isabelle.Bytes as Bytes | |
| 3442 | import Isabelle.Bytes (Bytes) | |
| 3443 | import qualified Isabelle.Time as Time | |
| 3444 | import Isabelle.Time (Time) | |
| 3445 | import Isabelle.Library | |
| 3446 | ||
| 3447 | ||
| 3448 | {- concrete syntax -}
 | |
| 3449 | ||
| 3450 | string :: Bytes -> Bytes | |
| 3451 | string str = | |
| 3452 | if Bytes.null str then "\"\"" | |
| 3453 | else str |> Bytes.unpack |> map trans |> Bytes.concat | |
| 3454 | where | |
| 3455 | trans b = | |
| 3456 | case Bytes.char b of | |
| 3457 | '\t' -> "$'\\t'" | |
| 3458 | '\n' -> "$'\\n'" | |
| 3459 | '\f' -> "$'\\f'" | |
| 3460 | '\r' -> "$'\\r'" | |
| 3461 | c -> | |
| 74169 | 3462 |             if Symbol.is_ascii_letter c || Symbol.is_ascii_digit c || c `elem` ("+,-./:_" :: String)
 | 
| 74161 | 3463 | then Bytes.singleton b | 
| 3464 | else if b < 32 || b >= 127 then make_bytes (printf "$'\\x%02x'" b :: String) | |
| 3465 | else "\\" <> Bytes.singleton b | |
| 3466 | ||
| 3467 | strings :: [Bytes] -> Bytes | |
| 3468 | strings = space_implode " " . map string | |
| 3469 | ||
| 3470 | ||
| 3471 | {- server parameters -}
 | |
| 3472 | ||
| 3473 | data Params = Params {
 | |
| 3474 | _script :: Bytes, | |
| 3475 | _input :: Bytes, | |
| 3476 | _cwd :: Maybe Bytes, | |
| 3477 | _putenv :: [(Bytes, Bytes)], | |
| 3478 | _redirect :: Bool, | |
| 3479 | _timeout :: Time, | |
| 3480 | _description :: Bytes} | |
| 3481 | deriving (Show, Eq) | |
| 3482 | ||
| 3483 | get_script :: Params -> Bytes | |
| 3484 | get_script = _script | |
| 3485 | ||
| 3486 | get_input :: Params -> Bytes | |
| 3487 | get_input = _input | |
| 3488 | ||
| 3489 | get_cwd :: Params -> Maybe Bytes | |
| 3490 | get_cwd = _cwd | |
| 3491 | ||
| 3492 | get_putenv :: Params -> [(Bytes, Bytes)] | |
| 3493 | get_putenv = _putenv | |
| 3494 | ||
| 3495 | get_redirect :: Params -> Bool | |
| 3496 | get_redirect = _redirect | |
| 3497 | ||
| 3498 | get_timeout :: Params -> Time | |
| 3499 | get_timeout = _timeout | |
| 3500 | ||
| 3501 | get_description :: Params -> Bytes | |
| 3502 | get_description = _description | |
| 3503 | ||
| 3504 | script :: Bytes -> Params | |
| 3505 | script script = Params script "" Nothing [] False Time.zero "" | |
| 3506 | ||
| 3507 | input :: Bytes -> Params -> Params | |
| 3508 | input input params = params { _input = input }
 | |
| 3509 | ||
| 3510 | cwd :: Bytes -> Params -> Params | |
| 3511 | cwd cwd params = params { _cwd = Just cwd }
 | |
| 3512 | ||
| 3513 | putenv :: [(Bytes, Bytes)] -> Params -> Params | |
| 3514 | putenv putenv params = params { _putenv = putenv }
 | |
| 3515 | ||
| 3516 | redirect :: Params -> Params | |
| 3517 | redirect params = params { _redirect = True }
 | |
| 3518 | ||
| 3519 | timeout :: Time -> Params -> Params | |
| 3520 | timeout timeout params = params { _timeout = timeout }
 | |
| 3521 | ||
| 3522 | description :: Bytes -> Params -> Params | |
| 3523 | description description params = params { _description = description }
 | |
| 74211 | 3524 | |
| 3525 | ||
| 3526 | {- server messages -}
 | |
| 3527 | ||
| 3528 | server_run, server_kill :: Bytes | |
| 3529 | server_run = \<open>Bash.server_run\<close>; | |
| 3530 | server_kill = \<open>Bash.server_kill\<close>; | |
| 3531 | ||
| 3532 | server_uuid, server_interrupt, server_failure, server_result :: Bytes | |
| 3533 | server_uuid = \<open>Bash.server_uuid\<close>; | |
| 3534 | server_interrupt = \<open>Bash.server_interrupt\<close>; | |
| 3535 | server_failure = \<open>Bash.server_failure\<close>; | |
| 3536 | server_result = \<open>Bash.server_result\<close>; | |
| 74161 | 3537 | \<close> | 
| 3538 | ||
| 3539 | generate_file "Isabelle/Process_Result.hs" = \<open> | |
| 3540 | {-  Title:      Isabelle/Process_Result.hs
 | |
| 3541 | Author: Makarius | |
| 3542 | LICENSE: BSD 3-clause (Isabelle) | |
| 3543 | ||
| 3544 | Result of system process. | |
| 3545 | ||
| 3546 | See \<^file>\<open>~~/src/Pure/System/process_result.ML\<close> | |
| 3547 | and \<^file>\<open>~~/src/Pure/System/process_result.scala\<close> | |
| 3548 | -} | |
| 3549 | ||
| 3550 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 3551 | ||
| 3552 | module Isabelle.Process_Result ( | |
| 74310 | 3553 | ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc, | 
| 74161 | 3554 | |
| 3555 | T, make, rc, out_lines, err_lines, timing, timing_elapsed, out, err, ok, check | |
| 3556 | ) | |
| 3557 | where | |
| 3558 | ||
| 3559 | import Isabelle.Time (Time) | |
| 3560 | import qualified Isabelle.Timing as Timing | |
| 3561 | import Isabelle.Timing (Timing) | |
| 3562 | import Isabelle.Bytes (Bytes) | |
| 3563 | import Isabelle.Library | |
| 3564 | ||
| 3565 | ||
| 74310 | 3566 | ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc :: Int | 
| 3567 | ok_rc = 0 | |
| 3568 | error_rc = 1 | |
| 3569 | failure_rc = 2 | |
| 74161 | 3570 | interrupt_rc = 130 | 
| 3571 | timeout_rc = 142 | |
| 3572 | ||
| 3573 | data T = | |
| 3574 |   Process_Result {
 | |
| 3575 | _rc :: Int, | |
| 3576 | _out_lines :: [Bytes], | |
| 3577 | _err_lines :: [Bytes], | |
| 3578 | _timing :: Timing} | |
| 3579 | deriving (Show, Eq) | |
| 3580 | ||
| 3581 | make :: Int -> [Bytes] -> [Bytes] -> Timing -> T | |
| 3582 | make = Process_Result | |
| 3583 | ||
| 3584 | rc :: T -> Int | |
| 3585 | rc = _rc | |
| 3586 | ||
| 3587 | out_lines :: T -> [Bytes] | |
| 3588 | out_lines = _out_lines | |
| 3589 | ||
| 3590 | err_lines :: T -> [Bytes] | |
| 3591 | err_lines = _err_lines | |
| 3592 | ||
| 3593 | timing :: T -> Timing | |
| 3594 | timing = _timing | |
| 3595 | ||
| 3596 | timing_elapsed :: T -> Time | |
| 3597 | timing_elapsed = Timing.elapsed . timing | |
| 3598 | ||
| 3599 | out :: T -> Bytes | |
| 3600 | out = trim_line . cat_lines . out_lines | |
| 3601 | ||
| 3602 | err :: T -> Bytes | |
| 3603 | err = trim_line . cat_lines . err_lines | |
| 3604 | ||
| 3605 | ok :: T -> Bool | |
| 74310 | 3606 | ok result = rc result == ok_rc | 
| 74161 | 3607 | |
| 3608 | check :: T -> T | |
| 3609 | check result = if ok result then result else error (make_string $ err result) | |
| 3610 | \<close> | |
| 3611 | ||
| 3612 | generate_file "Isabelle/Options.hs" = \<open> | |
| 3613 | {-  Title:      Isabelle/Options.hs
 | |
| 3614 | Author: Makarius | |
| 3615 | LICENSE: BSD 3-clause (Isabelle) | |
| 3616 | ||
| 3617 | System options with external string representation. | |
| 3618 | ||
| 3619 | See \<^file>\<open>~~/src/Pure/System/options.ML\<close> | |
| 3620 | and \<^file>\<open>~~/src/Pure/System/options.scala\<close> | |
| 3621 | -} | |
| 3622 | ||
| 3623 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 3624 | {-# LANGUAGE InstanceSigs #-}
 | |
| 3625 | ||
| 3626 | module Isabelle.Options ( | |
| 3627 | boolT, intT, realT, stringT, unknownT, | |
| 3628 | ||
| 3629 | T, typ, bool, int, real, seconds, string, | |
| 3630 | decode | |
| 3631 | ) | |
| 3632 | where | |
| 3633 | ||
| 3634 | import qualified Data.Map.Strict as Map | |
| 3635 | import Data.Map.Strict (Map) | |
| 3636 | import qualified Isabelle.Properties as Properties | |
| 3637 | import Isabelle.Bytes (Bytes) | |
| 3638 | import qualified Isabelle.Value as Value | |
| 3639 | import qualified Isabelle.Time as Time | |
| 3640 | import Isabelle.Time (Time) | |
| 3641 | import Isabelle.Library | |
| 3642 | import qualified Isabelle.XML.Decode as Decode | |
| 3643 | import Isabelle.XML.Classes (Decode (..)) | |
| 3644 | ||
| 3645 | ||
| 3646 | {- representation -}
 | |
| 3647 | ||
| 3648 | boolT :: Bytes | |
| 3649 | boolT = "bool" | |
| 3650 | ||
| 3651 | intT :: Bytes | |
| 3652 | intT = "int" | |
| 3653 | ||
| 3654 | realT :: Bytes | |
| 3655 | realT = "real" | |
| 3656 | ||
| 3657 | stringT :: Bytes | |
| 3658 | stringT = "string" | |
| 3659 | ||
| 3660 | unknownT :: Bytes | |
| 3661 | unknownT = "unknown" | |
| 3662 | ||
| 3663 | data Opt = Opt {
 | |
| 3664 | _pos :: Properties.T, | |
| 3665 | _name :: Bytes, | |
| 3666 | _typ :: Bytes, | |
| 3667 | _value :: Bytes } | |
| 3668 | ||
| 3669 | data T = Options (Map Bytes Opt) | |
| 3670 | ||
| 3671 | ||
| 3672 | {- check -}
 | |
| 3673 | ||
| 3674 | check_name :: T -> Bytes -> Opt | |
| 3675 | check_name (Options map) name = | |
| 3676 | case Map.lookup name map of | |
| 3677 | Just opt | _typ opt /= unknownT -> opt | |
| 3678 |     _ -> error (make_string ("Unknown system option " <> quote name))
 | |
| 3679 | ||
| 3680 | check_type :: T -> Bytes -> Bytes -> Opt | |
| 3681 | check_type options name typ = | |
| 3682 | let | |
| 3683 | opt = check_name options name | |
| 3684 | t = _typ opt | |
| 3685 | in | |
| 3686 | if t == typ then opt | |
| 3687 |     else error (make_string ("Ill-typed system option " <> quote name <> " : " <> t <> " vs. " <> typ))
 | |
| 3688 | ||
| 3689 | ||
| 3690 | {- get typ -}
 | |
| 3691 | ||
| 3692 | typ :: T -> Bytes -> Bytes | |
| 3693 | typ options name = _typ (check_name options name) | |
| 3694 | ||
| 3695 | ||
| 3696 | {- get value -}
 | |
| 3697 | ||
| 3698 | get :: Bytes -> (Bytes -> Maybe a) -> T -> Bytes -> a | |
| 3699 | get typ parse options name = | |
| 3700 | let opt = check_type options name typ in | |
| 3701 | case parse (_value opt) of | |
| 3702 | Just x -> x | |
| 3703 | Nothing -> | |
| 3704 |         error (make_string ("Malformed value for system option " <> quote name <>
 | |
| 3705 | " : " <> typ <> " =\n" <> quote (_value opt))) | |
| 3706 | ||
| 3707 | bool :: T -> Bytes -> Bool | |
| 3708 | bool = get boolT Value.parse_bool | |
| 3709 | ||
| 3710 | int :: T -> Bytes -> Int | |
| 3711 | int = get intT Value.parse_int | |
| 3712 | ||
| 3713 | real :: T -> Bytes -> Double | |
| 3714 | real = get realT Value.parse_real | |
| 3715 | ||
| 3716 | seconds :: T -> Bytes -> Time | |
| 3717 | seconds options = Time.seconds . real options | |
| 3718 | ||
| 3719 | string :: T -> Bytes -> Bytes | |
| 3720 | string = get stringT Just | |
| 3721 | ||
| 3722 | ||
| 3723 | {- decode -}
 | |
| 3724 | ||
| 3725 | instance Decode T where | |
| 3726 | decode :: Decode.T T | |
| 3727 | decode = | |
| 3728 | let | |
| 3729 | decode_entry :: Decode.T (Bytes, Opt) | |
| 3730 | decode_entry body = | |
| 3731 | let | |
| 3732 | (pos, (name, (typ, value))) = | |
| 3733 | Decode.pair Decode.properties (Decode.pair Decode.string (Decode.pair Decode.string Decode.string)) body | |
| 3734 |         in (name, Opt { _pos = pos, _name = name, _typ = typ, _value = value })
 | |
| 3735 | in Options . Map.fromList . Decode.list decode_entry | |
| 3736 | \<close> | |
| 3737 | ||
| 74211 | 3738 | generate_file "Isabelle/Isabelle_System.hs" = \<open> | 
| 3739 | {-  Title:      Isabelle/Isabelle_System.hs
 | |
| 3740 | Author: Makarius | |
| 3741 | LICENSE: BSD 3-clause (Isabelle) | |
| 3742 | ||
| 3743 | Isabelle system support. | |
| 3744 | ||
| 3745 | See \<^file>\<open>~~/src/Pure/System/isabelle_system.ML\<close> | |
| 3746 | and \<^file>\<open>~~/src/Pure/System/isabelle_system.scala\<close> | |
| 3747 | -} | |
| 3748 | ||
| 3749 | {-# LANGUAGE OverloadedStrings #-}
 | |
| 3750 | ||
| 3751 | module Isabelle.Isabelle_System ( | |
| 3752 | bash_process, bash_process0 | |
| 3753 | ) | |
| 3754 | where | |
| 3755 | ||
| 3756 | import Data.Maybe (fromMaybe) | |
| 3757 | import Control.Exception (throw, AsyncException (UserInterrupt)) | |
| 3758 | import Network.Socket (Socket) | |
| 74216 | 3759 | import qualified Isabelle.Bytes as Bytes | 
| 74211 | 3760 | import Isabelle.Bytes (Bytes) | 
| 3761 | import qualified Isabelle.Byte_Message as Byte_Message | |
| 3762 | import qualified Isabelle.Time as Time | |
| 3763 | import Isabelle.Timing (Timing (..)) | |
| 3764 | import qualified Isabelle.Options as Options | |
| 3765 | import qualified Isabelle.Bash as Bash | |
| 3766 | import qualified Isabelle.Process_Result as Process_Result | |
| 3767 | import qualified Isabelle.XML.Encode as Encode | |
| 3768 | import qualified Isabelle.YXML as YXML | |
| 3769 | import qualified Isabelle.Value as Value | |
| 3770 | import qualified Isabelle.Server as Server | |
| 3771 | import qualified Isabelle.Isabelle_Thread as Isabelle_Thread | |
| 3772 | import Isabelle.Library | |
| 3773 | ||
| 3774 | ||
| 3775 | {- bash_process -}
 | |
| 3776 | ||
| 3777 | absolute_path :: Bytes -> Bytes -- FIXME dummy | |
| 3778 | absolute_path = id | |
| 3779 | ||
| 3780 | bash_process :: Options.T -> Bash.Params -> IO Process_Result.T | |
| 3781 | bash_process options = bash_process0 address password | |
| 3782 | where | |
| 3783 | address = Options.string options "bash_process_address" | |
| 3784 | password = Options.string options "bash_process_password" | |
| 3785 | ||
| 3786 | bash_process0 :: Bytes -> Bytes -> Bash.Params -> IO Process_Result.T | |
| 3787 | bash_process0 address password params = do | |
| 3788 | Server.connection port password | |
| 3789 | (\socket -> do | |
| 3790 | isabelle_tmp <- getenv "ISABELLE_TMP" | |
| 3791 | Byte_Message.write_message socket (run isabelle_tmp) | |
| 3792 | loop Nothing socket) | |
| 3793 | where | |
| 3794 | port = | |
| 74216 | 3795 | case Bytes.try_unprefix Server.localhost_prefix address of | 
| 74211 | 3796 | Just port -> make_string port | 
| 3797 | Nothing -> errorWithoutStackTrace "Bad bash_process server address" | |
| 3798 | ||
| 3799 | script = Bash.get_script params | |
| 3800 | input = Bash.get_input params | |
| 3801 | cwd = Bash.get_cwd params | |
| 3802 | putenv = Bash.get_putenv params | |
| 3803 | redirect = Bash.get_redirect params | |
| 3804 | timeout = Bash.get_timeout params | |
| 3805 | description = Bash.get_description params | |
| 3806 | ||
| 3807 | run :: Bytes -> [Bytes] | |
| 3808 | run isabelle_tmp = | |
| 3809 | [Bash.server_run, script, input, | |
| 3810 | YXML.string_of_body (Encode.option (Encode.string . absolute_path) cwd), | |
| 3811 | YXML.string_of_body | |
| 3812 | (Encode.list (Encode.pair Encode.string Encode.string) | |
| 3813 |         (("ISABELLE_TMP", isabelle_tmp) : putenv)),
 | |
| 3814 | Value.print_bool redirect, | |
| 3815 | Value.print_real (Time.get_seconds timeout), | |
| 3816 | description] | |
| 3817 | ||
| 3818 | kill :: Maybe Bytes -> IO () | |
| 3819 | kill maybe_uuid = do | |
| 3820 | case maybe_uuid of | |
| 3821 | Just uuid -> | |
| 3822 | Server.connection port password (\socket -> | |
| 3823 | Byte_Message.write_message socket [Bash.server_kill, uuid]) | |
| 3824 | Nothing -> return () | |
| 3825 | ||
| 3826 | err = errorWithoutStackTrace "Malformed result from bash_process server" | |
| 3827 | the = fromMaybe err | |
| 3828 | ||
| 3829 | loop :: Maybe Bytes -> Socket -> IO Process_Result.T | |
| 3830 | loop maybe_uuid socket = do | |
| 3831 | result <- Isabelle_Thread.bracket_resource (kill maybe_uuid) (Byte_Message.read_message socket) | |
| 3832 | case result of | |
| 3833 | Just [head, uuid] | head == Bash.server_uuid -> loop (Just uuid) socket | |
| 3834 | Just [head] | head == Bash.server_interrupt -> throw UserInterrupt | |
| 3835 | Just [head, msg] | head == Bash.server_failure -> errorWithoutStackTrace $ make_string msg | |
| 3836 | Just (head : a : b : c : d : lines) | head == Bash.server_result -> | |
| 3837 | let | |
| 3838 | rc = the $ Value.parse_int a | |
| 3839 | elapsed = Time.ms $ the $ Value.parse_int b | |
| 3840 | cpu = Time.ms $ the $ Value.parse_int c | |
| 3841 | timing = Timing elapsed cpu Time.zero | |
| 3842 | n = the $ Value.parse_int d | |
| 3843 | out_lines = take n lines | |
| 3844 | err_lines = drop n lines | |
| 3845 | in return $ Process_Result.make rc out_lines err_lines timing | |
| 3846 | _ -> err | |
| 3847 | \<close> | |
| 3848 | ||
| 70047 
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
 wenzelm parents: 
69968diff
changeset | 3849 | export_generated_files _ | 
| 69628 | 3850 | |
| 69222 | 3851 | end |