author | wenzelm |
Wed, 22 Jan 2025 22:22:19 +0100 | |
changeset 81954 | 6f2bcdfa9a19 |
parent 80913 | 46f59511b7bb |
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:
73246
diff
changeset
|
149 |
generate_file "Isabelle/UTF8.hs" = \<open> |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
150 |
{- Title: Isabelle/UTF8.hs |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
151 |
Author: Makarius |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
152 |
LICENSE: BSD 3-clause (Isabelle) |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
153 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
154 |
Variations on UTF-8. |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
155 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
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:
73246
diff
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:
73246
diff
changeset
|
158 |
-} |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
159 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
160 |
{-# LANGUAGE MultiParamTypeClasses #-} |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
161 |
{-# LANGUAGE TypeSynonymInstances #-} |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
162 |
{-# LANGUAGE FlexibleInstances #-} |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
163 |
{-# LANGUAGE InstanceSigs #-} |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
164 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
165 |
module Isabelle.UTF8 ( |
74098 | 166 |
setup, setup3, |
74080
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
167 |
Recode (..) |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
168 |
) |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
169 |
where |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
170 |
|
74098 | 171 |
import qualified System.IO as IO |
74080
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
172 |
import Data.Text (Text) |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
173 |
import qualified Data.Text as Text |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
174 |
import qualified Data.Text.Encoding as Encoding |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
175 |
import qualified Data.Text.Encoding.Error as Error |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
176 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
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:
73246
diff
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:
73246
diff
changeset
|
192 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
193 |
class Recode a b where |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
194 |
encode :: a -> b |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
195 |
decode :: b -> a |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
196 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
197 |
instance Recode Text ByteString where |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
198 |
encode :: Text -> ByteString |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
199 |
encode = Encoding.encodeUtf8 |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
200 |
decode :: ByteString -> Text |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
201 |
decode = Encoding.decodeUtf8With Error.lenientDecode |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
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:
73246
diff
changeset
|
208 |
|
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
209 |
instance Recode String ByteString where |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
210 |
encode :: String -> ByteString |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
211 |
encode = encode . Text.pack |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
212 |
decode :: ByteString -> String |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
213 |
decode = Text.unpack . decode |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
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:
73246
diff
changeset
|
217 |
encode = encode . Text.pack |
74084 | 218 |
decode :: Bytes -> String |
74080
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
219 |
decode = Text.unpack . decode |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
220 |
\<close> |
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
221 |
|
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
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 |
|
80913 | 247 |
proper_string, enclose, quote, space_implode, implode_space, commas, commas_quote, |
248 |
cat_lines, 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] |
74871 | 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 |
74871 | 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 |
||
80913 | 419 |
implode_space :: StringLike a => [a] -> a |
420 |
implode_space = space_implode " " |
|
421 |
||
74093 | 422 |
commas, commas_quote :: StringLike a => [a] -> a |
69453 | 423 |
commas = space_implode ", " |
424 |
commas_quote = commas . map quote |
|
425 |
||
74132 | 426 |
split_lines :: StringLike a => a -> [a] |
427 |
split_lines = space_explode '\n' |
|
428 |
||
74093 | 429 |
cat_lines :: StringLike a => [a] -> a |
69453 | 430 |
cat_lines = space_implode "\n" |
74211 | 431 |
|
74213 | 432 |
trim_split_lines :: StringLike a => a -> [a] |
433 |
trim_split_lines = trim_line #> split_lines #> map trim_line |
|
434 |
||
74211 | 435 |
|
436 |
{- getenv -} |
|
437 |
||
438 |
getenv :: Bytes -> IO Bytes |
|
439 |
getenv x = do |
|
440 |
y <- lookupEnv (make_string x) |
|
441 |
return $ make_bytes $ fromMaybe "" y |
|
442 |
||
443 |
getenv_strict :: Bytes -> IO Bytes |
|
444 |
getenv_strict x = do |
|
445 |
y <- getenv x |
|
446 |
if Bytes.null y then |
|
447 |
errorWithoutStackTrace $ make_string ("Undefined Isabelle environment variable: " <> quote x) |
|
448 |
else return y |
|
69225 | 449 |
\<close> |
450 |
||
74091 | 451 |
|
452 |
generate_file "Isabelle/Symbol.hs" = \<open> |
|
453 |
{- Title: Isabelle/Symbols.hs |
|
454 |
Author: Makarius |
|
455 |
LICENSE: BSD 3-clause (Isabelle) |
|
456 |
||
457 |
Isabelle text symbols. |
|
74172 | 458 |
|
459 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/symbol.ML\<close> |
|
460 |
and \<^file>\<open>$ISABELLE_HOME/src/Pure/General/symbol_explode.ML\<close>. |
|
74091 | 461 |
-} |
462 |
||
74172 | 463 |
{-# LANGUAGE OverloadedStrings #-} |
464 |
||
465 |
module Isabelle.Symbol ( |
|
74177
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
466 |
Symbol, eof, is_eof, not_eof, |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
467 |
|
74172 | 468 |
is_ascii_letter, is_ascii_digit, is_ascii_hex, is_ascii_quasi, |
469 |
is_ascii_blank, is_ascii_line_terminator, is_ascii_letdig, |
|
470 |
is_ascii_identifier, |
|
471 |
||
74177
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
472 |
explode |
74172 | 473 |
) |
474 |
where |
|
475 |
||
476 |
import Data.Word (Word8) |
|
477 |
import qualified Isabelle.Bytes as Bytes |
|
478 |
import Isabelle.Bytes (Bytes) |
|
479 |
||
74091 | 480 |
|
74177
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
481 |
{- type -} |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
482 |
|
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
483 |
type Symbol = Bytes |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
484 |
|
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
485 |
eof :: Symbol |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
486 |
eof = "" |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
487 |
|
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
488 |
is_eof, not_eof :: Symbol -> Bool |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
489 |
is_eof = Bytes.null |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
490 |
not_eof = not . is_eof |
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
491 |
|
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
492 |
|
74091 | 493 |
{- ASCII characters -} |
494 |
||
495 |
is_ascii_letter :: Char -> Bool |
|
496 |
is_ascii_letter c = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' |
|
497 |
||
498 |
is_ascii_digit :: Char -> Bool |
|
499 |
is_ascii_digit c = '0' <= c && c <= '9' |
|
500 |
||
501 |
is_ascii_hex :: Char -> Bool |
|
502 |
is_ascii_hex c = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' |
|
503 |
||
504 |
is_ascii_quasi :: Char -> Bool |
|
505 |
is_ascii_quasi c = c == '_' || c == '\'' |
|
506 |
||
507 |
is_ascii_blank :: Char -> Bool |
|
74172 | 508 |
is_ascii_blank c = c `elem` (" \t\n\11\f\r" :: String) |
74091 | 509 |
|
510 |
is_ascii_line_terminator :: Char -> Bool |
|
511 |
is_ascii_line_terminator c = c == '\r' || c == '\n' |
|
512 |
||
513 |
is_ascii_letdig :: Char -> Bool |
|
514 |
is_ascii_letdig c = is_ascii_letter c || is_ascii_digit c || is_ascii_quasi c |
|
515 |
||
516 |
is_ascii_identifier :: String -> Bool |
|
517 |
is_ascii_identifier s = |
|
518 |
not (null s) && is_ascii_letter (head s) && all is_ascii_letdig s |
|
74172 | 519 |
|
520 |
||
521 |
{- explode symbols: ASCII, UTF8, named -} |
|
522 |
||
523 |
is_utf8 :: Word8 -> Bool |
|
524 |
is_utf8 b = b >= 128 |
|
525 |
||
526 |
is_utf8_trailer :: Word8 -> Bool |
|
527 |
is_utf8_trailer b = 128 <= b && b < 192 |
|
528 |
||
529 |
is_utf8_control :: Word8 -> Bool |
|
530 |
is_utf8_control b = 128 <= b && b < 160 |
|
531 |
||
532 |
(|>) :: a -> (a -> b) -> b |
|
533 |
x |> f = f x |
|
534 |
||
535 |
explode :: Bytes -> [Symbol] |
|
536 |
explode string = scan 0 |
|
537 |
where |
|
538 |
byte = Bytes.index string |
|
539 |
substring i j = |
|
540 |
if i == j - 1 then Bytes.singleton (byte i) |
|
541 |
else Bytes.pack (map byte [i .. j - 1]) |
|
542 |
||
543 |
n = Bytes.length string |
|
544 |
test pred i = i < n && pred (byte i) |
|
545 |
test_char pred i = i < n && pred (Bytes.char (byte i)) |
|
546 |
many pred i = if test pred i then many pred (i + 1) else i |
|
547 |
maybe_char c i = if test_char (== c) i then i + 1 else i |
|
548 |
maybe_ascii_id i = |
|
549 |
if test_char is_ascii_letter i |
|
550 |
then many (is_ascii_letdig . Bytes.char) (i + 1) |
|
551 |
else i |
|
552 |
||
553 |
scan i = |
|
554 |
if i < n then |
|
555 |
let |
|
556 |
b = byte i |
|
557 |
c = Bytes.char b |
|
558 |
in |
|
559 |
{-encoded newline-} |
|
560 |
if c == '\r' then "\n" : scan (maybe_char '\n' (i + 1)) |
|
561 |
{-pseudo utf8: encoded ascii control-} |
|
562 |
else if b == 192 && test is_utf8_control (i + 1) && not (test is_utf8 (i + 2)) |
|
563 |
then Bytes.singleton (byte (i + 1) - 128) : scan (i + 2) |
|
564 |
{-utf8-} |
|
565 |
else if is_utf8 b then |
|
566 |
let j = many is_utf8_trailer (i + 1) |
|
567 |
in substring i j : scan j |
|
568 |
{-named symbol-} |
|
569 |
else if c == '\\' && test_char (== '<') (i + 1) then |
|
570 |
let j = (i + 2) |> maybe_char '^' |> maybe_ascii_id |> maybe_char '>' |
|
571 |
in substring i j : scan j |
|
572 |
{-single character-} |
|
573 |
else Bytes.singleton b : scan (i + 1) |
|
574 |
else [] |
|
74091 | 575 |
\<close> |
576 |
||
74095 | 577 |
generate_file "Isabelle/Buffer.hs" = \<open> |
578 |
{- Title: Isabelle/Buffer.hs |
|
579 |
Author: Makarius |
|
580 |
LICENSE: BSD 3-clause (Isabelle) |
|
581 |
||
582 |
Efficient buffer of byte strings. |
|
583 |
||
74178 | 584 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/buffer.ML\<close>. |
74095 | 585 |
-} |
586 |
||
74231 | 587 |
module Isabelle.Buffer (T, empty, add, content, build, build_content) |
74095 | 588 |
where |
589 |
||
590 |
import qualified Isabelle.Bytes as Bytes |
|
591 |
import Isabelle.Bytes (Bytes) |
|
74231 | 592 |
import Isabelle.Library |
74095 | 593 |
|
594 |
||
595 |
newtype T = Buffer [Bytes] |
|
596 |
||
597 |
empty :: T |
|
598 |
empty = Buffer [] |
|
599 |
||
600 |
add :: Bytes -> T -> T |
|
601 |
add b (Buffer bs) = Buffer (if Bytes.null b then bs else b : bs) |
|
602 |
||
603 |
content :: T -> Bytes |
|
604 |
content (Buffer bs) = Bytes.concat (reverse bs) |
|
74231 | 605 |
|
606 |
build :: (T -> T) -> T |
|
607 |
build f = f empty |
|
608 |
||
609 |
build_content :: (T -> T) -> Bytes |
|
610 |
build_content f = build f |> content |
|
74095 | 611 |
\<close> |
612 |
||
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
changeset
|
613 |
generate_file "Isabelle/Value.hs" = \<open> |
73246 | 614 |
{- Title: Isabelle/Value.hs |
69233 | 615 |
Author: Makarius |
616 |
LICENSE: BSD 3-clause (Isabelle) |
|
617 |
||
618 |
Plain values, represented as string. |
|
69280 | 619 |
|
74178 | 620 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/value.ML\<close>. |
69233 | 621 |
-} |
622 |
||
74095 | 623 |
{-# LANGUAGE OverloadedStrings #-} |
624 |
||
69233 | 625 |
module Isabelle.Value |
69452 | 626 |
(print_bool, parse_bool, parse_nat, print_int, parse_int, print_real, parse_real) |
69233 | 627 |
where |
628 |
||
629 |
import qualified Data.List as List |
|
630 |
import qualified Text.Read as Read |
|
74095 | 631 |
import Isabelle.Bytes (Bytes) |
632 |
import Isabelle.Library |
|
69233 | 633 |
|
634 |
||
635 |
{- bool -} |
|
636 |
||
74095 | 637 |
print_bool :: Bool -> Bytes |
69233 | 638 |
print_bool True = "true" |
639 |
print_bool False = "false" |
|
640 |
||
74095 | 641 |
parse_bool :: Bytes -> Maybe Bool |
69233 | 642 |
parse_bool "true" = Just True |
643 |
parse_bool "false" = Just False |
|
644 |
parse_bool _ = Nothing |
|
645 |
||
646 |
||
69452 | 647 |
{- nat -} |
648 |
||
74095 | 649 |
parse_nat :: Bytes -> Maybe Int |
69452 | 650 |
parse_nat s = |
74095 | 651 |
case Read.readMaybe (make_string s) of |
69452 | 652 |
Just n | n >= 0 -> Just n |
653 |
_ -> Nothing |
|
654 |
||
655 |
||
69233 | 656 |
{- int -} |
657 |
||
74095 | 658 |
print_int :: Int -> Bytes |
74128 | 659 |
print_int = show_bytes |
74095 | 660 |
|
661 |
parse_int :: Bytes -> Maybe Int |
|
662 |
parse_int = Read.readMaybe . make_string |
|
69233 | 663 |
|
664 |
||
665 |
{- real -} |
|
666 |
||
74095 | 667 |
print_real :: Double -> Bytes |
69233 | 668 |
print_real x = |
669 |
let s = show x in |
|
670 |
case span (/= '.') s of |
|
74095 | 671 |
(a, '.' : b) | List.all (== '0') b -> make_bytes a |
672 |
_ -> make_bytes s |
|
673 |
||
674 |
parse_real :: Bytes -> Maybe Double |
|
675 |
parse_real = Read.readMaybe . make_string |
|
69225 | 676 |
\<close> |
677 |
||
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
changeset
|
678 |
generate_file "Isabelle/Properties.hs" = \<open> |
69445 | 679 |
{- Title: Isabelle/Properties.hs |
69225 | 680 |
Author: Makarius |
681 |
LICENSE: BSD 3-clause (Isabelle) |
|
682 |
||
683 |
Property lists. |
|
69280 | 684 |
|
74178 | 685 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/properties.ML\<close>. |
69225 | 686 |
-} |
687 |
||
69477 | 688 |
module Isabelle.Properties (Entry, T, defined, get, get_value, put, remove) |
69225 | 689 |
where |
690 |
||
691 |
import qualified Data.List as List |
|
74095 | 692 |
import Isabelle.Bytes (Bytes) |
693 |
||
694 |
||
695 |
type Entry = (Bytes, Bytes) |
|
69225 | 696 |
type T = [Entry] |
697 |
||
74095 | 698 |
defined :: T -> Bytes -> Bool |
69225 | 699 |
defined props name = any (\(a, _) -> a == name) props |
700 |
||
74095 | 701 |
get :: T -> Bytes -> Maybe Bytes |
69225 | 702 |
get props name = List.lookup name props |
703 |
||
74095 | 704 |
get_value :: (Bytes -> Maybe a) -> T -> Bytes -> Maybe a |
74871 | 705 |
get_value parse props name = maybe Nothing parse (get props name) |
69477 | 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:
69381
diff
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 |
||
78021 | 738 |
lineN, end_lineN, offsetN, end_offsetN, labelN, fileN, idN, positionN, position, |
74182 | 739 |
position_properties, def_name, |
69234 | 740 |
|
69291 | 741 |
expressionN, expression, |
742 |
||
71489 | 743 |
pathN, path, urlN, url, docN, doc, |
69291 | 744 |
|
69248 | 745 |
markupN, consistentN, unbreakableN, indentN, widthN, |
746 |
blockN, block, breakN, break, fbreakN, fbreak, itemN, item, |
|
747 |
||
69968
1a400b14fd3a
clarified spell-checking (see also 30233285270a);
wenzelm
parents:
69794
diff
changeset
|
748 |
wordsN, words, |
69234 | 749 |
|
750 |
tfreeN, tfree, tvarN, tvar, freeN, free, skolemN, skolem, boundN, bound, varN, var, |
|
751 |
numeralN, numeral, literalN, literal, delimiterN, delimiter, inner_stringN, inner_string, |
|
69320 | 752 |
inner_cartoucheN, inner_cartouche, |
69234 | 753 |
token_rangeN, token_range, |
754 |
sortingN, sorting, typingN, typing, class_parameterN, class_parameter, |
|
755 |
||
756 |
antiquotedN, antiquoted, antiquoteN, antiquote, |
|
757 |
||
758 |
paragraphN, paragraph, text_foldN, text_fold, |
|
759 |
||
760 |
keyword1N, keyword1, keyword2N, keyword2, keyword3N, keyword3, quasi_keywordN, quasi_keyword, |
|
761 |
improperN, improper, operatorN, operator, stringN, string, alt_stringN, alt_string, |
|
69320 | 762 |
verbatimN, verbatim, cartoucheN, cartouche, commentN, comment, comment1N, comment1, |
763 |
comment2N, comment2, comment3N, comment3, |
|
69234 | 764 |
|
70667 | 765 |
forkedN, forked, joinedN, joined, runningN, running, finishedN, finished, |
69794 | 766 |
failedN, failed, canceledN, canceled, initializedN, initialized, finalizedN, finalized, |
767 |
consolidatedN, consolidated, |
|
768 |
||
69234 | 769 |
writelnN, writeln, stateN, state, informationN, information, tracingN, tracing, |
770 |
warningN, warning, legacyN, legacy, errorN, error, reportN, report, no_reportN, no_report, |
|
771 |
||
772 |
intensifyN, intensify, |
|
773 |
Output, no_output) |
|
69225 | 774 |
where |
775 |
||
69248 | 776 |
import Prelude hiding (words, error, break) |
74182 | 777 |
import Data.Map.Strict (Map) |
778 |
import qualified Data.Map.Strict as Map |
|
69234 | 779 |
|
780 |
import Isabelle.Library |
|
69225 | 781 |
import qualified Isabelle.Properties as Properties |
69248 | 782 |
import qualified Isabelle.Value as Value |
74095 | 783 |
import qualified Isabelle.Bytes as Bytes |
784 |
import Isabelle.Bytes (Bytes) |
|
69225 | 785 |
|
786 |
||
69234 | 787 |
{- basic markup -} |
788 |
||
74095 | 789 |
type T = (Bytes, Properties.T) |
69225 | 790 |
|
791 |
empty :: T |
|
792 |
empty = ("", []) |
|
793 |
||
794 |
is_empty :: T -> Bool |
|
795 |
is_empty ("", _) = True |
|
796 |
is_empty _ = False |
|
797 |
||
69234 | 798 |
properties :: Properties.T -> T -> T |
799 |
properties more_props (elem, props) = |
|
800 |
(elem, fold_rev Properties.put more_props props) |
|
801 |
||
74095 | 802 |
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:
73178
diff
changeset
|
803 |
markup_elem name = (name, []) |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
804 |
|
74095 | 805 |
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:
73178
diff
changeset
|
806 |
markup_string name prop = \s -> (name, [(prop, s)]) |
69234 | 807 |
|
808 |
||
809 |
{- misc properties -} |
|
810 |
||
74095 | 811 |
nameN :: Bytes |
69234 | 812 |
nameN = \<open>Markup.nameN\<close> |
813 |
||
74095 | 814 |
name :: Bytes -> T -> T |
69234 | 815 |
name a = properties [(nameN, a)] |
816 |
||
74095 | 817 |
xnameN :: Bytes |
69234 | 818 |
xnameN = \<open>Markup.xnameN\<close> |
819 |
||
74095 | 820 |
xname :: Bytes -> T -> T |
69234 | 821 |
xname a = properties [(xnameN, a)] |
822 |
||
74095 | 823 |
kindN :: Bytes |
69234 | 824 |
kindN = \<open>Markup.kindN\<close> |
825 |
||
826 |
||
69315 | 827 |
{- formal entities -} |
828 |
||
74095 | 829 |
bindingN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
830 |
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:
73178
diff
changeset
|
831 |
binding :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
832 |
binding = markup_elem bindingN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
833 |
|
74095 | 834 |
entityN :: Bytes |
69315 | 835 |
entityN = \<open>Markup.entityN\<close> |
74095 | 836 |
entity :: Bytes -> Bytes -> T |
69315 | 837 |
entity kind name = |
838 |
(entityN, |
|
74095 | 839 |
(if Bytes.null name then [] else [(nameN, name)]) <> |
840 |
(if Bytes.null kind then [] else [(kindN, kind)])) |
|
841 |
||
842 |
defN :: Bytes |
|
69315 | 843 |
defN = \<open>Markup.defN\<close> |
844 |
||
74095 | 845 |
refN :: Bytes |
69315 | 846 |
refN = \<open>Markup.refN\<close> |
847 |
||
848 |
||
69288 | 849 |
{- completion -} |
850 |
||
74095 | 851 |
completionN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
852 |
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:
73178
diff
changeset
|
853 |
completion :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
854 |
completion = markup_elem completionN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
855 |
|
74095 | 856 |
no_completionN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
857 |
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:
73178
diff
changeset
|
858 |
no_completion :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
859 |
no_completion = markup_elem no_completionN |
69288 | 860 |
|
861 |
||
69234 | 862 |
{- position -} |
863 |
||
74095 | 864 |
lineN, end_lineN :: Bytes |
69234 | 865 |
lineN = \<open>Markup.lineN\<close> |
866 |
end_lineN = \<open>Markup.end_lineN\<close> |
|
867 |
||
74095 | 868 |
offsetN, end_offsetN :: Bytes |
69234 | 869 |
offsetN = \<open>Markup.offsetN\<close> |
870 |
end_offsetN = \<open>Markup.end_offsetN\<close> |
|
871 |
||
78021 | 872 |
labelN, fileN, idN :: Bytes |
873 |
labelN = \<open>Markup.labelN\<close> |
|
69234 | 874 |
fileN = \<open>Markup.fileN\<close> |
875 |
idN = \<open>Markup.idN\<close> |
|
876 |
||
74095 | 877 |
positionN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
878 |
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:
73178
diff
changeset
|
879 |
position :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
880 |
position = markup_elem positionN |
69234 | 881 |
|
74182 | 882 |
position_properties :: [Bytes] |
78021 | 883 |
position_properties = [lineN, offsetN, end_offsetN, labelN, fileN, idN] |
74182 | 884 |
|
885 |
||
886 |
{- position "def" names -} |
|
887 |
||
888 |
make_def :: Bytes -> Bytes |
|
889 |
make_def a = "def_" <> a |
|
890 |
||
891 |
def_names :: Map Bytes Bytes |
|
892 |
def_names = Map.fromList $ map (\a -> (a, make_def a)) position_properties |
|
893 |
||
894 |
def_name :: Bytes -> Bytes |
|
895 |
def_name a = |
|
896 |
case Map.lookup a def_names of |
|
897 |
Just b -> b |
|
898 |
Nothing -> make_def a |
|
899 |
||
69234 | 900 |
|
69291 | 901 |
{- expression -} |
902 |
||
74095 | 903 |
expressionN :: Bytes |
69291 | 904 |
expressionN = \<open>Markup.expressionN\<close> |
905 |
||
74095 | 906 |
expression :: Bytes -> T |
69291 | 907 |
expression kind = (expressionN, if kind == "" then [] else [(kindN, kind)]) |
908 |
||
909 |
||
910 |
{- external resources -} |
|
911 |
||
74095 | 912 |
pathN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
913 |
pathN = \<open>Markup.pathN\<close> |
74095 | 914 |
path :: Bytes -> T |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
915 |
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:
73178
diff
changeset
|
916 |
|
74095 | 917 |
urlN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
918 |
urlN = \<open>Markup.urlN\<close> |
74095 | 919 |
url :: Bytes -> T |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
920 |
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:
73178
diff
changeset
|
921 |
|
74095 | 922 |
docN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
923 |
docN = \<open>Markup.docN\<close> |
74095 | 924 |
doc :: Bytes -> T |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
925 |
doc = markup_string docN nameN |
69291 | 926 |
|
927 |
||
69248 | 928 |
{- pretty printing -} |
929 |
||
74095 | 930 |
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:
73178
diff
changeset
|
931 |
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:
73178
diff
changeset
|
932 |
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:
73178
diff
changeset
|
933 |
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:
73178
diff
changeset
|
934 |
indentN = \<open>Markup.indentN\<close> |
69248 | 935 |
|
74095 | 936 |
widthN :: Bytes |
69248 | 937 |
widthN = \<open>Markup.widthN\<close> |
938 |
||
74095 | 939 |
blockN :: Bytes |
69248 | 940 |
blockN = \<open>Markup.blockN\<close> |
941 |
block :: Bool -> Int -> T |
|
942 |
block c i = |
|
943 |
(blockN, |
|
74081 | 944 |
(if c then [(consistentN, Value.print_bool c)] else []) <> |
69248 | 945 |
(if i /= 0 then [(indentN, Value.print_int i)] else [])) |
946 |
||
74095 | 947 |
breakN :: Bytes |
69248 | 948 |
breakN = \<open>Markup.breakN\<close> |
949 |
break :: Int -> Int -> T |
|
950 |
break w i = |
|
951 |
(breakN, |
|
74081 | 952 |
(if w /= 0 then [(widthN, Value.print_int w)] else []) <> |
69248 | 953 |
(if i /= 0 then [(indentN, Value.print_int i)] else [])) |
954 |
||
74095 | 955 |
fbreakN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
956 |
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:
73178
diff
changeset
|
957 |
fbreak :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
958 |
fbreak = markup_elem fbreakN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
959 |
|
74095 | 960 |
itemN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
961 |
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:
73178
diff
changeset
|
962 |
item :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
963 |
item = markup_elem itemN |
69248 | 964 |
|
965 |
||
69234 | 966 |
{- text properties -} |
967 |
||
74095 | 968 |
wordsN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
969 |
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:
73178
diff
changeset
|
970 |
words :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
971 |
words = markup_elem wordsN |
69234 | 972 |
|
973 |
||
974 |
{- inner syntax -} |
|
975 |
||
74095 | 976 |
tfreeN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
977 |
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:
73178
diff
changeset
|
978 |
tfree :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
979 |
tfree = markup_elem tfreeN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
980 |
|
74095 | 981 |
tvarN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
982 |
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:
73178
diff
changeset
|
983 |
tvar :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
984 |
tvar = markup_elem tvarN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
985 |
|
74095 | 986 |
freeN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
987 |
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:
73178
diff
changeset
|
988 |
free :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
989 |
free = markup_elem freeN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
990 |
|
74095 | 991 |
skolemN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
992 |
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:
73178
diff
changeset
|
993 |
skolem :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
994 |
skolem = markup_elem skolemN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
995 |
|
74095 | 996 |
boundN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
997 |
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:
73178
diff
changeset
|
998 |
bound :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
999 |
bound = markup_elem boundN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1000 |
|
74095 | 1001 |
varN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1002 |
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:
73178
diff
changeset
|
1003 |
var :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1004 |
var = markup_elem varN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1005 |
|
74095 | 1006 |
numeralN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1007 |
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:
73178
diff
changeset
|
1008 |
numeral :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1009 |
numeral = markup_elem numeralN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1010 |
|
74095 | 1011 |
literalN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1012 |
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:
73178
diff
changeset
|
1013 |
literal :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1014 |
literal = markup_elem literalN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1015 |
|
74095 | 1016 |
delimiterN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1017 |
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:
73178
diff
changeset
|
1018 |
delimiter :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1019 |
delimiter = markup_elem delimiterN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1020 |
|
74095 | 1021 |
inner_stringN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1022 |
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:
73178
diff
changeset
|
1023 |
inner_string :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1024 |
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:
73178
diff
changeset
|
1025 |
|
74095 | 1026 |
inner_cartoucheN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1027 |
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:
73178
diff
changeset
|
1028 |
inner_cartouche :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1029 |
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:
73178
diff
changeset
|
1030 |
|
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1031 |
|
74095 | 1032 |
token_rangeN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1033 |
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:
73178
diff
changeset
|
1034 |
token_range :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1035 |
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:
73178
diff
changeset
|
1036 |
|
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1037 |
|
74095 | 1038 |
sortingN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1039 |
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:
73178
diff
changeset
|
1040 |
sorting :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1041 |
sorting = markup_elem sortingN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1042 |
|
74095 | 1043 |
typingN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1044 |
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:
73178
diff
changeset
|
1045 |
typing :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1046 |
typing = markup_elem typingN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1047 |
|
74095 | 1048 |
class_parameterN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1049 |
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:
73178
diff
changeset
|
1050 |
class_parameter :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1051 |
class_parameter = markup_elem class_parameterN |
69234 | 1052 |
|
1053 |
||
1054 |
{- antiquotations -} |
|
1055 |
||
74095 | 1056 |
antiquotedN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1057 |
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:
73178
diff
changeset
|
1058 |
antiquoted :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1059 |
antiquoted = markup_elem antiquotedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1060 |
|
74095 | 1061 |
antiquoteN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1062 |
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:
73178
diff
changeset
|
1063 |
antiquote :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1064 |
antiquote = markup_elem antiquoteN |
69234 | 1065 |
|
1066 |
||
1067 |
{- text structure -} |
|
1068 |
||
74095 | 1069 |
paragraphN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1070 |
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:
73178
diff
changeset
|
1071 |
paragraph :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1072 |
paragraph = markup_elem paragraphN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1073 |
|
74095 | 1074 |
text_foldN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1075 |
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:
73178
diff
changeset
|
1076 |
text_fold :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1077 |
text_fold = markup_elem text_foldN |
69234 | 1078 |
|
1079 |
||
1080 |
{- outer syntax -} |
|
1081 |
||
74095 | 1082 |
keyword1N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1083 |
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:
73178
diff
changeset
|
1084 |
keyword1 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1085 |
keyword1 = markup_elem keyword1N |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1086 |
|
74095 | 1087 |
keyword2N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1088 |
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:
73178
diff
changeset
|
1089 |
keyword2 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1090 |
keyword2 = markup_elem keyword2N |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1091 |
|
74095 | 1092 |
keyword3N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1093 |
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:
73178
diff
changeset
|
1094 |
keyword3 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1095 |
keyword3 = markup_elem keyword3N |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1096 |
|
74095 | 1097 |
quasi_keywordN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1098 |
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:
73178
diff
changeset
|
1099 |
quasi_keyword :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1100 |
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:
73178
diff
changeset
|
1101 |
|
74095 | 1102 |
improperN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1103 |
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:
73178
diff
changeset
|
1104 |
improper :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1105 |
improper = markup_elem improperN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1106 |
|
74095 | 1107 |
operatorN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1108 |
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:
73178
diff
changeset
|
1109 |
operator :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1110 |
operator = markup_elem operatorN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1111 |
|
74095 | 1112 |
stringN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1113 |
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:
73178
diff
changeset
|
1114 |
string :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1115 |
string = markup_elem stringN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1116 |
|
74095 | 1117 |
alt_stringN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1118 |
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:
73178
diff
changeset
|
1119 |
alt_string :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1120 |
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:
73178
diff
changeset
|
1121 |
|
74095 | 1122 |
verbatimN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1123 |
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:
73178
diff
changeset
|
1124 |
verbatim :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1125 |
verbatim = markup_elem verbatimN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1126 |
|
74095 | 1127 |
cartoucheN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1128 |
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:
73178
diff
changeset
|
1129 |
cartouche :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1130 |
cartouche = markup_elem cartoucheN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1131 |
|
74095 | 1132 |
commentN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1133 |
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:
73178
diff
changeset
|
1134 |
comment :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1135 |
comment = markup_elem commentN |
69234 | 1136 |
|
1137 |
||
69320 | 1138 |
{- comments -} |
1139 |
||
74095 | 1140 |
comment1N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1141 |
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:
73178
diff
changeset
|
1142 |
comment1 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1143 |
comment1 = markup_elem comment1N |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1144 |
|
74095 | 1145 |
comment2N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1146 |
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:
73178
diff
changeset
|
1147 |
comment2 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1148 |
comment2 = markup_elem comment2N |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1149 |
|
74095 | 1150 |
comment3N :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1151 |
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:
73178
diff
changeset
|
1152 |
comment3 :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1153 |
comment3 = markup_elem comment3N |
69320 | 1154 |
|
1155 |
||
69793 | 1156 |
{- command status -} |
1157 |
||
70667 | 1158 |
forkedN, joinedN, runningN, finishedN, failedN, canceledN, |
74095 | 1159 |
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:
73178
diff
changeset
|
1160 |
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:
73178
diff
changeset
|
1161 |
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:
73178
diff
changeset
|
1162 |
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:
73178
diff
changeset
|
1163 |
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:
73178
diff
changeset
|
1164 |
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:
73178
diff
changeset
|
1165 |
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:
73178
diff
changeset
|
1166 |
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:
73178
diff
changeset
|
1167 |
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:
73178
diff
changeset
|
1168 |
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:
73178
diff
changeset
|
1169 |
|
70667 | 1170 |
forked, joined, running, finished, failed, canceled, |
69793 | 1171 |
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:
73178
diff
changeset
|
1172 |
forked = markup_elem forkedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1173 |
joined = markup_elem joinedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1174 |
running = markup_elem runningN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1175 |
finished = markup_elem finishedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1176 |
failed = markup_elem failedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1177 |
canceled = markup_elem canceledN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1178 |
initialized = markup_elem initializedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1179 |
finalized = markup_elem finalizedN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1180 |
consolidated = markup_elem consolidatedN |
69793 | 1181 |
|
1182 |
||
69234 | 1183 |
{- messages -} |
1184 |
||
74095 | 1185 |
writelnN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1186 |
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:
73178
diff
changeset
|
1187 |
writeln :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1188 |
writeln = markup_elem writelnN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1189 |
|
74095 | 1190 |
stateN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1191 |
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:
73178
diff
changeset
|
1192 |
state :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1193 |
state = markup_elem stateN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1194 |
|
74095 | 1195 |
informationN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1196 |
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:
73178
diff
changeset
|
1197 |
information :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1198 |
information = markup_elem informationN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1199 |
|
74095 | 1200 |
tracingN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1201 |
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:
73178
diff
changeset
|
1202 |
tracing :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1203 |
tracing = markup_elem tracingN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1204 |
|
74095 | 1205 |
warningN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1206 |
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:
73178
diff
changeset
|
1207 |
warning :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1208 |
warning = markup_elem warningN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1209 |
|
74095 | 1210 |
legacyN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1211 |
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:
73178
diff
changeset
|
1212 |
legacy :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1213 |
legacy = markup_elem legacyN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1214 |
|
74095 | 1215 |
errorN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1216 |
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:
73178
diff
changeset
|
1217 |
error :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1218 |
error = markup_elem errorN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1219 |
|
74095 | 1220 |
reportN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1221 |
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:
73178
diff
changeset
|
1222 |
report :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1223 |
report = markup_elem reportN |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1224 |
|
74095 | 1225 |
no_reportN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1226 |
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:
73178
diff
changeset
|
1227 |
no_report :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1228 |
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:
73178
diff
changeset
|
1229 |
|
74095 | 1230 |
intensifyN :: Bytes |
73199
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1231 |
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:
73178
diff
changeset
|
1232 |
intensify :: T |
d300574cee4e
more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents:
73178
diff
changeset
|
1233 |
intensify = markup_elem intensifyN |
69234 | 1234 |
|
1235 |
||
1236 |
{- output -} |
|
69225 | 1237 |
|
74095 | 1238 |
type Output = (Bytes, Bytes) |
69225 | 1239 |
|
1240 |
no_output :: Output |
|
1241 |
no_output = ("", "") |
|
69222 | 1242 |
\<close> |
1243 |
||
74167 | 1244 |
generate_file "Isabelle/Position.hs" = \<open> |
1245 |
{- Title: Isabelle/Position.hs |
|
1246 |
Author: Makarius |
|
1247 |
LICENSE: BSD 3-clause (Isabelle) |
|
1248 |
||
74173 | 1249 |
Source positions starting from 1; values <= 0 mean "absent". Count Isabelle |
1250 |
symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a |
|
1251 |
right-open interval offset .. end_offset (exclusive). |
|
74167 | 1252 |
|
74178 | 1253 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/position.ML\<close>. |
74167 | 1254 |
-} |
1255 |
||
1256 |
{-# LANGUAGE OverloadedStrings #-} |
|
1257 |
||
1258 |
module Isabelle.Position ( |
|
78307 | 1259 |
T, line_of, column_of, offset_of, end_offset_of, label_of, file_of, id_of, |
1260 |
start, none, label, put_file, file, file_only, put_id, id, id_only, |
|
74174 | 1261 |
symbol, symbol_explode, symbol_explode_string, shift_offsets, |
74184 | 1262 |
of_properties, properties_of, def_properties_of, entity_markup, make_entity_markup, |
74187 | 1263 |
Report, Report_Text, is_reported, is_reported_range, here, |
74167 | 1264 |
Range, no_range, no_range_position, range_position, range |
74184 | 1265 |
) |
1266 |
where |
|
74167 | 1267 |
|
74185 | 1268 |
import Prelude hiding (id) |
74168
f0b2136e2204
tuned signature: prefer existing Haskell operations;
wenzelm
parents:
74167
diff
changeset
|
1269 |
import Data.Maybe (isJust, fromMaybe) |
74182 | 1270 |
import Data.Bifunctor (first) |
74167 | 1271 |
import qualified Isabelle.Properties as Properties |
1272 |
import qualified Isabelle.Bytes as Bytes |
|
1273 |
import qualified Isabelle.Value as Value |
|
1274 |
import Isabelle.Bytes (Bytes) |
|
1275 |
import qualified Isabelle.Markup as Markup |
|
1276 |
import qualified Isabelle.YXML as YXML |
|
1277 |
import Isabelle.Library |
|
74173 | 1278 |
import qualified Isabelle.Symbol as Symbol |
1279 |
import Isabelle.Symbol (Symbol) |
|
74167 | 1280 |
|
1281 |
||
1282 |
{- position -} |
|
1283 |
||
1284 |
data T = |
|
1285 |
Position { |
|
1286 |
_line :: Int, |
|
1287 |
_column :: Int, |
|
1288 |
_offset :: Int, |
|
1289 |
_end_offset :: Int, |
|
78021 | 1290 |
_label :: Bytes, |
74167 | 1291 |
_file :: Bytes, |
1292 |
_id :: Bytes } |
|
1293 |
deriving (Eq, Ord) |
|
1294 |
||
74173 | 1295 |
valid, invalid :: Int -> Bool |
1296 |
valid i = i > 0 |
|
1297 |
invalid = not . valid |
|
74167 | 1298 |
|
1299 |
maybe_valid :: Int -> Maybe Int |
|
74173 | 1300 |
maybe_valid i = if valid i then Just i else Nothing |
1301 |
||
1302 |
if_valid :: Int -> Int -> Int |
|
1303 |
if_valid i i' = if valid i then i' else i |
|
74167 | 1304 |
|
1305 |
||
1306 |
{- fields -} |
|
1307 |
||
1308 |
line_of, column_of, offset_of, end_offset_of :: T -> Maybe Int |
|
1309 |
line_of = maybe_valid . _line |
|
1310 |
column_of = maybe_valid . _column |
|
1311 |
offset_of = maybe_valid . _offset |
|
1312 |
end_offset_of = maybe_valid . _end_offset |
|
1313 |
||
78021 | 1314 |
label_of :: T -> Maybe Bytes |
1315 |
label_of = proper_string . _label |
|
1316 |
||
74167 | 1317 |
file_of :: T -> Maybe Bytes |
1318 |
file_of = proper_string . _file |
|
1319 |
||
1320 |
id_of :: T -> Maybe Bytes |
|
1321 |
id_of = proper_string . _id |
|
1322 |
||
1323 |
||
1324 |
{- make position -} |
|
1325 |
||
1326 |
start :: T |
|
78021 | 1327 |
start = Position 1 1 1 0 Bytes.empty Bytes.empty Bytes.empty |
74167 | 1328 |
|
1329 |
none :: T |
|
78021 | 1330 |
none = Position 0 0 0 0 Bytes.empty Bytes.empty Bytes.empty |
1331 |
||
1332 |
label :: Bytes -> T -> T |
|
1333 |
label label pos = pos { _label = label } |
|
74167 | 1334 |
|
1335 |
put_file :: Bytes -> T -> T |
|
1336 |
put_file file pos = pos { _file = file } |
|
1337 |
||
1338 |
file :: Bytes -> T |
|
1339 |
file file = put_file file start |
|
1340 |
||
1341 |
file_only :: Bytes -> T |
|
1342 |
file_only file = put_file file none |
|
1343 |
||
1344 |
put_id :: Bytes -> T -> T |
|
1345 |
put_id id pos = pos { _id = id } |
|
1346 |
||
74185 | 1347 |
id :: Bytes -> T |
1348 |
id id = put_id id start |
|
1349 |
||
1350 |
id_only :: Bytes -> T |
|
1351 |
id_only id = put_id id none |
|
1352 |
||
74167 | 1353 |
|
74176 | 1354 |
{- count position -} |
1355 |
||
1356 |
count_line :: Symbol -> Int -> Int |
|
1357 |
count_line "\n" line = if_valid line (line + 1) |
|
1358 |
count_line _ line = line |
|
1359 |
||
1360 |
count_column :: Symbol -> Int -> Int |
|
1361 |
count_column "\n" column = if_valid column 1 |
|
74177
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
1362 |
count_column s column = if Symbol.not_eof s then if_valid column (column + 1) else column |
74176 | 1363 |
|
1364 |
count_offset :: Symbol -> Int -> Int |
|
74177
a8b032dede5c
treat Symbol.eof as in ML (but: presently unused);
wenzelm
parents:
74176
diff
changeset
|
1365 |
count_offset s offset = if Symbol.not_eof s then if_valid offset (offset + 1) else offset |
74173 | 1366 |
|
74174 | 1367 |
symbol :: Symbol -> T -> T |
1368 |
symbol s pos = |
|
74167 | 1369 |
pos { |
74176 | 1370 |
_line = count_line s (_line pos), |
1371 |
_column = count_column s (_column pos), |
|
1372 |
_offset = count_offset s (_offset pos) } |
|
74173 | 1373 |
|
74174 | 1374 |
symbol_explode :: BYTES a => a -> T -> T |
1375 |
symbol_explode = fold symbol . Symbol.explode . make_bytes |
|
1376 |
||
1377 |
symbol_explode_string :: String -> T -> T |
|
1378 |
symbol_explode_string = symbol_explode |
|
74167 | 1379 |
|
1380 |
||
1381 |
{- shift offsets -} |
|
1382 |
||
1383 |
shift_offsets :: Int -> T -> T |
|
1384 |
shift_offsets shift pos = pos { _offset = offset', _end_offset = end_offset' } |
|
1385 |
where |
|
1386 |
offset = _offset pos |
|
1387 |
end_offset = _end_offset pos |
|
1388 |
offset' = if invalid offset || invalid shift then offset else offset + shift |
|
1389 |
end_offset' = if invalid end_offset || invalid shift then end_offset else end_offset + shift |
|
1390 |
||
1391 |
||
1392 |
{- markup properties -} |
|
1393 |
||
1394 |
get_string :: Properties.T -> Bytes -> Bytes |
|
74168
f0b2136e2204
tuned signature: prefer existing Haskell operations;
wenzelm
parents:
74167
diff
changeset
|
1395 |
get_string props name = fromMaybe "" (Properties.get_value Just props name) |
74167 | 1396 |
|
1397 |
get_int :: Properties.T -> Bytes -> Int |
|
74168
f0b2136e2204
tuned signature: prefer existing Haskell operations;
wenzelm
parents:
74167
diff
changeset
|
1398 |
get_int props name = fromMaybe 0 (Properties.get_value Value.parse_int props name) |
74167 | 1399 |
|
1400 |
of_properties :: Properties.T -> T |
|
1401 |
of_properties props = |
|
1402 |
none { |
|
1403 |
_line = get_int props Markup.lineN, |
|
1404 |
_offset = get_int props Markup.offsetN, |
|
1405 |
_end_offset = get_int props Markup.end_offsetN, |
|
78021 | 1406 |
_label = get_string props Markup.labelN, |
74167 | 1407 |
_file = get_string props Markup.fileN, |
1408 |
_id = get_string props Markup.idN } |
|
1409 |
||
1410 |
string_entry :: Bytes -> Bytes -> Properties.T |
|
1411 |
string_entry k s = if Bytes.null s then [] else [(k, s)] |
|
1412 |
||
1413 |
int_entry :: Bytes -> Int -> Properties.T |
|
1414 |
int_entry k i = if invalid i then [] else [(k, Value.print_int i)] |
|
1415 |
||
1416 |
properties_of :: T -> Properties.T |
|
1417 |
properties_of pos = |
|
1418 |
int_entry Markup.lineN (_line pos) ++ |
|
1419 |
int_entry Markup.offsetN (_offset pos) ++ |
|
1420 |
int_entry Markup.end_offsetN (_end_offset pos) ++ |
|
78021 | 1421 |
string_entry Markup.labelN (_label pos) ++ |
74167 | 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:
69381
diff
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 = |
|
80910 | 1570 |
implode_space (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:
69381
diff
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:
74167
diff
changeset
|
1595 |
import Data.Maybe (fromJust) |
f0b2136e2204
tuned signature: prefer existing Haskell operations;
wenzelm
parents:
74167
diff
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:
74167
diff
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:
69381
diff
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:
69381
diff
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:
74105
diff
changeset
|
1941 |
import qualified Isabelle.Name as Name |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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:
74105
diff
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:
74105
diff
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:
74105
diff
changeset
|
1952 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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:
74105
diff
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:
74105
diff
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:
74105
diff
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:
74105
diff
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:
69381
diff
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:
74105
diff
changeset
|
2166 |
generate_file "Isabelle/Name.hs" = \<open> |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2167 |
{- Title: Isabelle/Name.hs |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2168 |
Author: Makarius |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2169 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2170 |
Names of basic logical entities (variables etc.). |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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:
74105
diff
changeset
|
2173 |
-} |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2174 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2175 |
{-# LANGUAGE OverloadedStrings #-} |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2176 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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, |
80601 | 2181 |
Context, declare, declare_renamed, is_declared, declared, context, make_context, |
74327 | 2182 |
variant, variant_list |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2183 |
) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2184 |
where |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2185 |
|
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2186 |
import Data.Maybe (fromMaybe) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2187 |
import Data.Map.Strict (Map) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2188 |
import qualified Data.Map.Strict as Map |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2189 |
import Data.Word (Word8) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2190 |
import qualified Isabelle.Bytes as Bytes |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2191 |
import Isabelle.Bytes (Bytes) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2192 |
import qualified Isabelle.Symbol as Symbol |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2193 |
import Isabelle.Library |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2194 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2195 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2196 |
type Name = Bytes |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2197 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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:
74105
diff
changeset
|
2208 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2209 |
underscore :: Word8 |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2210 |
underscore = Bytes.byte '_' |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
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:
74310
diff
changeset
|
2224 |
clean_index :: (Name, Int) -> (Name, Int) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2225 |
clean_index (x, i) = |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2226 |
case dest_internal x of |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2227 |
Nothing -> (x, i) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2228 |
Just x' -> clean_index (x', i + 1) |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2229 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2230 |
clean :: Name -> Name |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2231 |
clean x = fst (clean_index (x, 0)) |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2232 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2233 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2234 |
{- context for used names -} |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2235 |
|
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
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:
74105
diff
changeset
|
2237 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2238 |
declare :: Name -> Context -> Context |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2239 |
declare x (Context names) = |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2240 |
Context ( |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2241 |
let a = clean x |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
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:
74310
diff
changeset
|
2243 |
|
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2244 |
declare_renaming :: (Name, Name) -> Context -> Context |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2245 |
declare_renaming (x, x') (Context names) = |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
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:
74105
diff
changeset
|
2247 |
|
80601 | 2248 |
declare_renamed :: (Name, Name) -> Context -> Context |
2249 |
declare_renamed (x, x') = |
|
2250 |
(if clean x /= clean x' then declare_renaming (x, x') else id) #> declare x' |
|
2251 |
||
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2252 |
is_declared :: Context -> Name -> Bool |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2253 |
is_declared (Context names) x = Map.member x names |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2254 |
|
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2255 |
declared :: Context -> Name -> Maybe (Maybe Name) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2256 |
declared (Context names) a = Map.lookup a names |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2257 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2258 |
context :: Context |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2259 |
context = Context Map.empty |> fold declare ["", "'"] |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2260 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2261 |
make_context :: [Name] -> Context |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2262 |
make_context used = fold declare used context |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2263 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2264 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2265 |
{- generating fresh names -} |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2266 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2267 |
bump_init :: Name -> Name |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2268 |
bump_init str = str <> "a" |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2269 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2270 |
bump_string :: Name -> Name |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2271 |
bump_string str = |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2272 |
let |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2273 |
a = Bytes.byte 'a' |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2274 |
z = Bytes.byte 'z' |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2275 |
bump (b : bs) | b == z = a : bump bs |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2276 |
bump (b : bs) | a <= b && b < z = b + 1 : bs |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2277 |
bump bs = a : bs |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2278 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2279 |
rev = reverse (Bytes.unpack str) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2280 |
part2 = reverse (takeWhile (Symbol.is_ascii_quasi . Bytes.char) rev) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2281 |
part1 = reverse (bump (drop (length part2) rev)) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2282 |
in Bytes.pack (part1 <> part2) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2283 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2284 |
variant :: Name -> Context -> (Name, Context) |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2285 |
variant name ctxt = |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2286 |
let |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2287 |
vary x = |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2288 |
case declared ctxt x of |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2289 |
Nothing -> x |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2290 |
Just x' -> vary (bump_string (fromMaybe x x')) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2291 |
|
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2292 |
(x, n) = clean_index (name, 0) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2293 |
(x', ctxt') = |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2294 |
if not (is_declared ctxt x) then (x, declare x ctxt) |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2295 |
else |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2296 |
let |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2297 |
x0 = bump_init x |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2298 |
x' = vary x0 |
80601 | 2299 |
ctxt' = ctxt |> declare_renamed (x0, x') |
74315
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2300 |
in (x', ctxt') |
30a0f5879d90
clarified operations: follow Isabelle/ML more closely;
wenzelm
parents:
74310
diff
changeset
|
2301 |
in (x' <> Bytes.pack (replicate n underscore), ctxt') |
74327 | 2302 |
|
2303 |
variant_list :: [Name] -> [Name] -> [Name] |
|
2304 |
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:
74105
diff
changeset
|
2305 |
\<close> |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2306 |
|
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
changeset
|
2307 |
generate_file "Isabelle/Term.hs" = \<open> |
69445 | 2308 |
{- Title: Isabelle/Term.hs |
69240 | 2309 |
Author: Makarius |
2310 |
LICENSE: BSD 3-clause (Isabelle) |
|
2311 |
||
2312 |
Lambda terms, types, sorts. |
|
69280 | 2313 |
|
74178 | 2314 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/term.scala\<close>. |
69240 | 2315 |
-} |
2316 |
||
74095 | 2317 |
{-# LANGUAGE OverloadedStrings #-} |
2318 |
||
69240 | 2319 |
module Isabelle.Term ( |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2320 |
Indexname, Sort, Typ(..), Term(..), |
74214 | 2321 |
Free, lambda, declare_frees, incr_boundvars, subst_bound, dest_lambda, strip_lambda, |
74197 | 2322 |
type_op0, type_op1, op0, op1, op2, typed_op0, typed_op1, typed_op2, binder, |
74105 | 2323 |
dummyS, dummyT, is_dummyT, propT, is_propT, (-->), dest_funT, (--->), |
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2324 |
aconv, list_comb, strip_comb, head_of |
74105 | 2325 |
) |
69240 | 2326 |
where |
2327 |
||
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2328 |
import Isabelle.Library |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2329 |
import qualified Isabelle.Name as Name |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2330 |
import Isabelle.Name (Name) |
74095 | 2331 |
|
74105 | 2332 |
infixr 5 --> |
2333 |
infixr ---> |
|
2334 |
||
2335 |
||
2336 |
{- types and terms -} |
|
2337 |
||
2338 |
type Indexname = (Name, Int) |
|
2339 |
||
2340 |
type Sort = [Name] |
|
2341 |
||
2342 |
data Typ = |
|
2343 |
Type (Name, [Typ]) |
|
2344 |
| TFree (Name, Sort) |
|
2345 |
| TVar (Indexname, Sort) |
|
2346 |
deriving (Show, Eq, Ord) |
|
2347 |
||
2348 |
data Term = |
|
2349 |
Const (Name, [Typ]) |
|
2350 |
| Free (Name, Typ) |
|
2351 |
| Var (Indexname, Typ) |
|
2352 |
| Bound Int |
|
2353 |
| Abs (Name, Typ, Term) |
|
2354 |
| App (Term, Term) |
|
80568
fbb655bf62d4
clarified Isabelle/Haskell type Term, following Isabelle/Scala (see 446b887e23c7);
wenzelm
parents:
79741
diff
changeset
|
2355 |
| OFCLASS (Typ, Name) |
74105 | 2356 |
deriving (Show, Eq, Ord) |
2357 |
||
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2358 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2359 |
{- free and bound variables -} |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2360 |
|
74105 | 2361 |
type Free = (Name, Typ) |
2362 |
||
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2363 |
lambda :: Free -> Term -> Term |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2364 |
lambda (name, typ) body = Abs (name, typ, abstract 0 body) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2365 |
where |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2366 |
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:
74105
diff
changeset
|
2367 |
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:
74105
diff
changeset
|
2368 |
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:
74105
diff
changeset
|
2369 |
abstract _ t = t |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2370 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2371 |
declare_frees :: Term -> Name.Context -> Name.Context |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2372 |
declare_frees (Free (x, _)) = Name.declare x |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2373 |
declare_frees (Abs (_, _, b)) = declare_frees b |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2374 |
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:
74105
diff
changeset
|
2375 |
declare_frees _ = id |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2376 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2377 |
incr_boundvars :: Int -> Term -> Term |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2378 |
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:
74105
diff
changeset
|
2379 |
where |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2380 |
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:
74105
diff
changeset
|
2381 |
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:
74105
diff
changeset
|
2382 |
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:
74105
diff
changeset
|
2383 |
incr _ t = t |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2384 |
|
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2385 |
subst_bound :: Term -> Term -> Term |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2386 |
subst_bound arg = subst 0 |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2387 |
where |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2388 |
subst lev (Bound i) = |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2389 |
if i < lev then Bound i |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2390 |
else if i == lev then incr_boundvars lev arg |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2391 |
else Bound (i - 1) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2392 |
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:
74105
diff
changeset
|
2393 |
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:
74105
diff
changeset
|
2394 |
subst _ t = t |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2395 |
|
74214 | 2396 |
dest_lambda :: Name.Context -> Term -> Maybe (Free, Term) |
2397 |
dest_lambda names (Abs (x, ty, b)) = |
|
74106
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2398 |
let |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2399 |
(x', _) = Name.variant x (declare_frees b names) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2400 |
v = (x', ty) |
4984fad0e91d
more operations, notably free and bound variables as in Isabelle/Pure;
wenzelm
parents:
74105
diff
changeset
|
2401 |
in Just (v, subst_bound (Free v) b) |
74214 | 2402 |
dest_lambda _ _ = Nothing |
2403 |
||
2404 |
strip_lambda :: Name.Context -> Term -> ([Free], Term) |
|
2405 |
strip_lambda names tm = |
|
2406 |
case dest_lambda names tm of |
|
74124 | 2407 |
Just (v, t) -> |
74214 | 2408 |
let (vs, t') = strip_lambda names t' |
74124 | 2409 |
in (v : vs, t') |
2410 |
Nothing -> ([], tm) |
|
2411 |
||
74105 | 2412 |
|
2413 |
{- type and term operators -} |
|
2414 |
||
2415 |
type_op0 :: Name -> (Typ, Typ -> Bool) |
|
2416 |
type_op0 name = (mk, is) |
|
2417 |
where |
|
2418 |
mk = Type (name, []) |
|
74205 | 2419 |
is (Type (c, _)) = c == name |
74105 | 2420 |
is _ = False |
2421 |
||
2422 |
type_op1 :: Name -> (Typ -> Typ, Typ -> Maybe Typ) |
|
2423 |
type_op1 name = (mk, dest) |
|
2424 |
where |
|
2425 |
mk ty = Type (name, [ty]) |
|
74205 | 2426 |
dest (Type (c, [ty])) | c == name = Just ty |
74105 | 2427 |
dest _ = Nothing |
2428 |
||
2429 |
type_op2 :: Name -> (Typ -> Typ -> Typ, Typ -> Maybe (Typ, Typ)) |
|
2430 |
type_op2 name = (mk, dest) |
|
2431 |
where |
|
2432 |
mk ty1 ty2 = Type (name, [ty1, ty2]) |
|
74205 | 2433 |
dest (Type (c, [ty1, ty2])) | c == name = Just (ty1, ty2) |
74105 | 2434 |
dest _ = Nothing |
2435 |
||
2436 |
op0 :: Name -> (Term, Term -> Bool) |
|
2437 |
op0 name = (mk, is) |
|
2438 |
where |
|
2439 |
mk = Const (name, []) |
|
2440 |
is (Const (c, _)) = c == name |
|
2441 |
is _ = False |
|
2442 |
||
2443 |
op1 :: Name -> (Term -> Term, Term -> Maybe Term) |
|
2444 |
op1 name = (mk, dest) |
|
2445 |
where |
|
2446 |
mk t = App (Const (name, []), t) |
|
2447 |
dest (App (Const (c, _), t)) | c == name = Just t |
|
2448 |
dest _ = Nothing |
|
2449 |
||
2450 |
op2 :: Name -> (Term -> Term -> Term, Term -> Maybe (Term, Term)) |
|
2451 |
op2 name = (mk, dest) |
|
2452 |
where |
|
2453 |
mk t u = App (App (Const (name, []), t), u) |
|
2454 |
dest (App (App (Const (c, _), t), u)) | c == name = Just (t, u) |
|
2455 |
dest _ = Nothing |
|
2456 |
||
74197 | 2457 |
typed_op0 :: Name -> (Typ -> Term, Term -> Maybe Typ) |
2458 |
typed_op0 name = (mk, dest) |
|
2459 |
where |
|
2460 |
mk ty = Const (name, [ty]) |
|
2461 |
dest (Const (c, [ty])) | c == name = Just ty |
|
2462 |
dest _ = Nothing |
|
2463 |
||
74188 | 2464 |
typed_op1 :: Name -> (Typ -> Term -> Term, Term -> Maybe (Typ, Term)) |
2465 |
typed_op1 name = (mk, dest) |
|
2466 |
where |
|
2467 |
mk ty t = App (Const (name, [ty]), t) |
|
2468 |
dest (App (Const (c, [ty]), t)) | c == name = Just (ty, t) |
|
2469 |
dest _ = Nothing |
|
2470 |
||
74105 | 2471 |
typed_op2 :: Name -> (Typ -> Term -> Term -> Term, Term -> Maybe (Typ, Term, Term)) |
2472 |
typed_op2 name = (mk, dest) |
|
2473 |
where |
|
2474 |
mk ty t u = App (App (Const (name, [ty]), t), u) |
|
2475 |
dest (App (App (Const (c, [ty]), t), u)) | c == name = Just (ty, t, u) |
|
2476 |
dest _ = Nothing |
|
2477 |
||
74124 | 2478 |
binder :: Name -> (Free -> Term -> Term, Name.Context -> Term -> Maybe (Free, Term)) |
2479 |
binder name = (mk, dest) |
|
2480 |
where |
|
2481 |
mk (a, ty) b = App (Const (name, [ty]), lambda (a, ty) b) |
|
74214 | 2482 |
dest names (App (Const (c, _), t)) | c == name = dest_lambda names t |
74124 | 2483 |
dest _ _ = Nothing |
74105 | 2484 |
|
2485 |
||
2486 |
{- type operations -} |
|
69240 | 2487 |
|
2488 |
dummyS :: Sort |
|
2489 |
dummyS = [""] |
|
2490 |
||
74105 | 2491 |
dummyT :: Typ; is_dummyT :: Typ -> Bool |
2492 |
(dummyT, is_dummyT) = type_op0 \<open>\<^type_name>\<open>dummy\<close>\<close> |
|
2493 |
||
2494 |
propT :: Typ; is_propT :: Typ -> Bool |
|
2495 |
(propT, is_propT) = type_op0 \<open>\<^type_name>\<open>prop\<close>\<close> |
|
2496 |
||
2497 |
(-->) :: Typ -> Typ -> Typ; dest_funT :: Typ -> Maybe (Typ, Typ) |
|
2498 |
((-->), dest_funT) = type_op2 \<open>\<^type_name>\<open>fun\<close>\<close> |
|
2499 |
||
2500 |
(--->) :: [Typ] -> Typ -> Typ |
|
2501 |
[] ---> b = b |
|
2502 |
(a : as) ---> b = a --> (as ---> b) |
|
2503 |
||
2504 |
||
2505 |
{- term operations -} |
|
74100 | 2506 |
|
2507 |
aconv :: Term -> Term -> Bool |
|
2508 |
aconv (App (t1, u1)) (App (t2, u2)) = aconv t1 t2 && aconv u1 u2 |
|
2509 |
aconv (Abs (_, ty1, t1)) (Abs (_, ty2, t2)) = aconv t1 t2 && ty1 == ty2 |
|
2510 |
aconv a1 a2 = a1 == a2 |
|
2511 |
||
2512 |
list_comb :: Term -> [Term] -> Term |
|
2513 |
list_comb f [] = f |
|
2514 |
list_comb f (t : ts) = list_comb (App (f, t)) ts |
|
2515 |
||
2516 |
strip_comb :: Term -> (Term, [Term]) |
|
2517 |
strip_comb tm = strip (tm, []) |
|
2518 |
where |
|
2519 |
strip (App (f, t), ts) = strip (f, t : ts) |
|
2520 |
strip x = x |
|
2521 |
||
2522 |
head_of :: Term -> Term |
|
2523 |
head_of (App (f, _)) = head_of f |
|
2524 |
head_of u = u |
|
69240 | 2525 |
\<close> |
2526 |
||
74105 | 2527 |
generate_file "Isabelle/Pure.hs" = \<open> |
2528 |
{- Title: Isabelle/Term.hs |
|
2529 |
Author: Makarius |
|
2530 |
LICENSE: BSD 3-clause (Isabelle) |
|
2531 |
||
2532 |
Support for Isabelle/Pure logic. |
|
2533 |
||
74178 | 2534 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/logic.ML\<close>. |
74105 | 2535 |
-} |
2536 |
||
2537 |
{-# LANGUAGE OverloadedStrings #-} |
|
2538 |
||
2539 |
module Isabelle.Pure ( |
|
74188 | 2540 |
mk_forall_op, dest_forall_op, mk_forall, dest_forall, |
2541 |
mk_equals, dest_equals, mk_implies, dest_implies |
|
74105 | 2542 |
) |
2543 |
where |
|
2544 |
||
74124 | 2545 |
import qualified Isabelle.Name as Name |
74105 | 2546 |
import Isabelle.Term |
2547 |
||
74188 | 2548 |
mk_forall_op :: Typ -> Term -> Term; dest_forall_op :: Term -> Maybe (Typ, Term) |
2549 |
(mk_forall_op, dest_forall_op) = typed_op1 \<open>\<^const_name>\<open>Pure.all\<close>\<close> |
|
2550 |
||
74124 | 2551 |
mk_forall :: Free -> Term -> Term; dest_forall :: Name.Context -> Term -> Maybe (Free, Term) |
2552 |
(mk_forall, dest_forall) = binder \<open>\<^const_name>\<open>Pure.all\<close>\<close> |
|
74105 | 2553 |
|
2554 |
mk_equals :: Typ -> Term -> Term -> Term; dest_equals :: Term -> Maybe (Typ, Term, Term) |
|
2555 |
(mk_equals, dest_equals) = typed_op2 \<open>\<^const_name>\<open>Pure.eq\<close>\<close> |
|
2556 |
||
2557 |
mk_implies :: Term -> Term -> Term; dest_implies :: Term -> Maybe (Term, Term) |
|
2558 |
(mk_implies, dest_implies) = op2 \<open>\<^const_name>\<open>Pure.imp\<close>\<close> |
|
2559 |
\<close> |
|
2560 |
||
2561 |
generate_file "Isabelle/HOL.hs" = \<open> |
|
2562 |
{- Title: Isabelle/Term.hs |
|
2563 |
Author: Makarius |
|
2564 |
LICENSE: BSD 3-clause (Isabelle) |
|
2565 |
||
2566 |
Support for Isabelle/HOL logic. |
|
2567 |
||
74178 | 2568 |
See \<^file>\<open>$ISABELLE_HOME/src/HOL/Tools/hologic.ML\<close>. |
74105 | 2569 |
-} |
2570 |
||
2571 |
{-# LANGUAGE OverloadedStrings #-} |
|
2572 |
||
2573 |
module Isabelle.HOL ( |
|
2574 |
boolT, is_boolT, mk_trueprop, dest_trueprop, |
|
2575 |
mk_setT, dest_setT, mk_mem, dest_mem, |
|
2576 |
mk_eq, dest_eq, true, is_true, false, is_false, |
|
2577 |
mk_not, dest_not, mk_conj, dest_conj, mk_disj, dest_disj, |
|
2578 |
mk_imp, dest_imp, mk_iff, dest_iff, |
|
74188 | 2579 |
mk_all_op, dest_all_op, mk_ex_op, dest_ex_op, |
74197 | 2580 |
mk_all, dest_all, mk_ex, dest_ex, |
2581 |
mk_undefined, dest_undefined |
|
74105 | 2582 |
) |
2583 |
where |
|
2584 |
||
74124 | 2585 |
import qualified Isabelle.Name as Name |
74105 | 2586 |
import Isabelle.Term |
2587 |
||
2588 |
||
2589 |
boolT :: Typ; is_boolT :: Typ -> Bool |
|
2590 |
(boolT, is_boolT) = type_op0 \<open>\<^type_name>\<open>bool\<close>\<close> |
|
2591 |
||
2592 |
mk_trueprop :: Term -> Term; dest_trueprop :: Term -> Maybe Term |
|
2593 |
(mk_trueprop, dest_trueprop) = op1 \<open>\<^const_name>\<open>Trueprop\<close>\<close> |
|
2594 |
||
2595 |
mk_setT :: Typ -> Typ; dest_setT :: Typ -> Maybe Typ |
|
2596 |
(mk_setT, dest_setT) = type_op1 \<open>\<^type_name>\<open>set\<close>\<close> |
|
2597 |
||
2598 |
mk_mem :: Typ -> Term -> Term -> Term; dest_mem :: Term -> Maybe (Typ, Term, Term) |
|
2599 |
(mk_mem, dest_mem) = typed_op2 \<open>\<^const_name>\<open>Set.member\<close>\<close> |
|
2600 |
||
2601 |
mk_eq :: Typ -> Term -> Term -> Term; dest_eq :: Term -> Maybe (Typ, Term, Term) |
|
2602 |
(mk_eq, dest_eq) = typed_op2 \<open>\<^const_name>\<open>HOL.eq\<close>\<close> |
|
2603 |
||
2604 |
true :: Term; is_true :: Term -> Bool |
|
2605 |
(true, is_true) = op0 \<open>\<^const_name>\<open>True\<close>\<close> |
|
2606 |
||
2607 |
false :: Term; is_false :: Term -> Bool |
|
2608 |
(false, is_false) = op0 \<open>\<^const_name>\<open>False\<close>\<close> |
|
2609 |
||
2610 |
mk_not :: Term -> Term; dest_not :: Term -> Maybe Term |
|
2611 |
(mk_not, dest_not) = op1 \<open>\<^const_name>\<open>Not\<close>\<close> |
|
2612 |
||
2613 |
mk_conj :: Term -> Term -> Term; dest_conj :: Term -> Maybe (Term, Term) |
|
2614 |
(mk_conj, dest_conj) = op2 \<open>\<^const_name>\<open>conj\<close>\<close> |
|
2615 |
||
2616 |
mk_disj :: Term -> Term -> Term; dest_disj :: Term -> Maybe (Term, Term) |
|
2617 |
(mk_disj, dest_disj) = op2 \<open>\<^const_name>\<open>disj\<close>\<close> |
|
2618 |
||
2619 |
mk_imp :: Term -> Term -> Term; dest_imp :: Term -> Maybe (Term, Term) |
|
2620 |
(mk_imp, dest_imp) = op2 \<open>\<^const_name>\<open>implies\<close>\<close> |
|
2621 |
||
2622 |
mk_iff :: Term -> Term -> Term |
|
2623 |
mk_iff = mk_eq boolT |
|
2624 |
||
2625 |
dest_iff :: Term -> Maybe (Term, Term) |
|
2626 |
dest_iff tm = |
|
2627 |
case dest_eq tm of |
|
2628 |
Just (ty, t, u) | ty == boolT -> Just (t, u) |
|
2629 |
_ -> Nothing |
|
2630 |
||
74188 | 2631 |
mk_all_op :: Typ -> Term -> Term; dest_all_op :: Term -> Maybe (Typ, Term) |
2632 |
(mk_all_op, dest_all_op) = typed_op1 \<open>\<^const_name>\<open>All\<close>\<close> |
|
2633 |
||
2634 |
mk_ex_op :: Typ -> Term -> Term; dest_ex_op :: Term -> Maybe (Typ, Term) |
|
2635 |
(mk_ex_op, dest_ex_op) = typed_op1 \<open>\<^const_name>\<open>Ex\<close>\<close> |
|
2636 |
||
74124 | 2637 |
mk_all :: Free -> Term -> Term; dest_all :: Name.Context -> Term -> Maybe (Free, Term) |
2638 |
(mk_all, dest_all) = binder \<open>\<^const_name>\<open>All\<close>\<close> |
|
2639 |
||
2640 |
mk_ex :: Free -> Term -> Term; dest_ex :: Name.Context -> Term -> Maybe (Free, Term) |
|
2641 |
(mk_ex, dest_ex) = binder \<open>\<^const_name>\<open>Ex\<close>\<close> |
|
74197 | 2642 |
|
2643 |
mk_undefined :: Typ -> Term; dest_undefined :: Term -> Maybe Typ |
|
2644 |
(mk_undefined, dest_undefined) = typed_op0 \<open>\<^const_name>\<open>undefined\<close>\<close> |
|
74105 | 2645 |
\<close> |
2646 |
||
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
changeset
|
2647 |
generate_file "Isabelle/Term_XML/Encode.hs" = \<open> |
69445 | 2648 |
{- Title: Isabelle/Term_XML/Encode.hs |
69240 | 2649 |
Author: Makarius |
2650 |
LICENSE: BSD 3-clause (Isabelle) |
|
2651 |
||
2652 |
XML data representation of lambda terms. |
|
69280 | 2653 |
|
74178 | 2654 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>. |
69240 | 2655 |
-} |
2656 |
||
2657 |
{-# LANGUAGE LambdaCase #-} |
|
2658 |
||
80589 | 2659 |
module Isabelle.Term_XML.Encode (indexname, sort, typ, term) |
69240 | 2660 |
where |
2661 |
||
2662 |
import Isabelle.Library |
|
2663 |
import Isabelle.XML.Encode |
|
2664 |
import Isabelle.Term |
|
2665 |
||
70845 | 2666 |
indexname :: P Indexname |
2667 |
indexname (a, b) = if b == 0 then [a] else [a, int_atom b] |
|
69240 | 2668 |
|
2669 |
sort :: T Sort |
|
2670 |
sort = list string |
|
2671 |
||
2672 |
typ :: T Typ |
|
2673 |
typ ty = |
|
2674 |
ty |> variant |
|
2675 |
[\case { Type (a, b) -> Just ([a], list typ b); _ -> Nothing }, |
|
2676 |
\case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing }, |
|
70845 | 2677 |
\case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }] |
2678 |
||
80589 | 2679 |
var_type :: T Typ |
2680 |
var_type ty = if is_dummyT ty then [] else typ ty |
|
69240 | 2681 |
|
2682 |
term :: T Term |
|
2683 |
term t = |
|
2684 |
t |> variant |
|
70784
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents:
70667
diff
changeset
|
2685 |
[\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing }, |
80589 | 2686 |
\case { Free (a, b) -> Just ([a], var_type b); _ -> Nothing }, |
2687 |
\case { Var (a, b) -> Just (indexname a, var_type b); _ -> Nothing }, |
|
70845 | 2688 |
\case { Bound a -> Just ([], int a); _ -> Nothing }, |
69240 | 2689 |
\case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing }, |
80568
fbb655bf62d4
clarified Isabelle/Haskell type Term, following Isabelle/Scala (see 446b887e23c7);
wenzelm
parents:
79741
diff
changeset
|
2690 |
\case { App a -> Just ([], pair term term a); _ -> Nothing }, |
fbb655bf62d4
clarified Isabelle/Haskell type Term, following Isabelle/Scala (see 446b887e23c7);
wenzelm
parents:
79741
diff
changeset
|
2691 |
\case { OFCLASS (a, b) -> Just ([b], typ a); _ -> Nothing }] |
69240 | 2692 |
\<close> |
2693 |
||
69444
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents:
69381
diff
changeset
|
2694 |
generate_file "Isabelle/Term_XML/Decode.hs" = \<open> |
69445 | 2695 |
{- Title: Isabelle/Term_XML/Decode.hs |
69240 | 2696 |
Author: Makarius |
2697 |
LICENSE: BSD 3-clause (Isabelle) |
|
2698 |
||
2699 |
XML data representation of lambda terms. |
|
69280 | 2700 |
|
74178 | 2701 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>. |
69240 | 2702 |
-} |
2703 |
||
73177 | 2704 |
{-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} |
2705 |
||
80589 | 2706 |
module Isabelle.Term_XML.Decode (indexname, sort, typ, term) |
69240 | 2707 |
where |
2708 |
||
2709 |
import Isabelle.Library |
|
2710 |
import Isabelle.XML.Decode |
|
2711 |
import Isabelle.Term |
|
2712 |
||
2713 |
||
70845 | 2714 |
indexname :: P Indexname |
2715 |
indexname [a] = (a, 0) |
|
2716 |
indexname [a, b] = (a, int_atom b) |
|
2717 |
||
69240 | 2718 |
sort :: T Sort |
2719 |
sort = list string |
|
2720 |
||
2721 |
typ :: T Typ |
|
2722 |
typ ty = |
|
2723 |
ty |> variant |
|
2724 |
[\([a], b) -> Type (a, list typ b), |
|
2725 |
\([a], b) -> TFree (a, sort b), |
|
70845 | 2726 |
\(a, b) -> TVar (indexname a, sort b)] |
2727 |
||
80589 | 2728 |
var_type :: T Typ |
2729 |
var_type [] = dummyT |
|
2730 |
var_type body = typ body |
|
69240 | 2731 |
|
2732 |
term :: T Term |
|
2733 |
term t = |
|
2734 |
t |> variant |
|
70784
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents:
70667
diff
changeset
|
2735 |
[\([a], b) -> Const (a, list typ b), |
80589 | 2736 |
\([a], b) -> Free (a, var_type b), |
2737 |
\(a, b) -> Var (indexname a, var_type b), |
|
70845 | 2738 |
\([], a) -> Bound (int a), |
69240 | 2739 |
\([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d), |
80568
fbb655bf62d4
clarified Isabelle/Haskell type Term, following Isabelle/Scala (see 446b887e23c7);
wenzelm
parents:
79741
diff
changeset
|
2740 |
\([], a) -> App (pair term term a), |
fbb655bf62d4
clarified Isabelle/Haskell type Term, following Isabelle/Scala (see 446b887e23c7);
wenzelm
parents:
79741
diff
changeset
|
2741 |
\([a], b) -> OFCLASS (typ b, a)] |
69240 | 2742 |
\<close> |
2743 |
||
74129 | 2744 |
generate_file "Isabelle/XML/Classes.hs" = \<open> |
2745 |
{- generated by Isabelle -} |
|
2746 |
||
2747 |
{- Title: Isabelle/XML/Classes.hs |
|
2748 |
Author: Makarius |
|
2749 |
LICENSE: BSD 3-clause (Isabelle) |
|
2750 |
||
2751 |
Type classes for XML data representation. |
|
2752 |
-} |
|
2753 |
||
2754 |
{-# LANGUAGE FlexibleInstances #-} |
|
2755 |
||
2756 |
module Isabelle.XML.Classes |
|
2757 |
(Encode_Atom(..), Decode_Atom(..), Encode (..), Decode (..)) |
|
2758 |
where |
|
2759 |
||
2760 |
import qualified Isabelle.XML as XML |
|
2761 |
import qualified Isabelle.XML.Encode as Encode |
|
2762 |
import qualified Isabelle.XML.Decode as Decode |
|
2763 |
import qualified Isabelle.Term_XML.Encode as Encode |
|
2764 |
import qualified Isabelle.Term_XML.Decode as Decode |
|
2765 |
import qualified Isabelle.Properties as Properties |
|
2766 |
import Isabelle.Bytes (Bytes) |
|
2767 |
import Isabelle.Term (Typ, Term) |
|
2768 |
||
2769 |
||
2770 |
class Encode_Atom a where encode_atom :: Encode.A a |
|
2771 |
class Decode_Atom a where decode_atom :: Decode.A a |
|
2772 |
||
2773 |
instance Encode_Atom Int where encode_atom = Encode.int_atom |
|
2774 |
instance Decode_Atom Int where decode_atom = Decode.int_atom |
|
2775 |
||
2776 |
instance Encode_Atom Bool where encode_atom = Encode.bool_atom |
|
2777 |
instance Decode_Atom Bool where decode_atom = Decode.bool_atom |
|
2778 |
||
2779 |
instance Encode_Atom () where encode_atom = Encode.unit_atom |
|
2780 |
instance Decode_Atom () where decode_atom = Decode.unit_atom |
|
2781 |
||
2782 |
||
2783 |
class Encode a where encode :: Encode.T a |
|
2784 |
class Decode a where decode :: Decode.T a |
|
2785 |
||
2786 |
instance Encode Bytes where encode = Encode.string |
|
2787 |
instance Decode Bytes where decode = Decode.string |
|
2788 |
||
2789 |
instance Encode Int where encode = Encode.int |
|
2790 |
instance Decode Int where decode = Decode.int |
|
2791 |
||
2792 |
instance Encode Bool where encode = Encode.bool |
|
2793 |
instance Decode Bool where decode = Decode.bool |
|
2794 |
||
2795 |
instance Encode () where encode = Encode.unit |
|
2796 |
instance Decode () where decode = Decode.unit |
|
2797 |
||
2798 |
instance (Encode a, Encode b) => Encode (a, b) |
|
2799 |
where encode = Encode.pair encode encode |
|
2800 |
instance (Decode a, Decode b) => Decode (a, b) |
|
2801 |
where decode = Decode.pair decode decode |
|
2802 |
||
2803 |
instance (Encode a, Encode b, Encode c) => Encode (a, b, c) |
|
2804 |
where encode = Encode.triple encode encode encode |
|
2805 |
instance (Decode a, Decode b, Decode c) => Decode (a, b, c) |
|
2806 |
where decode = Decode.triple decode decode decode |
|
2807 |
||
2808 |
instance Encode a => Encode [a] where encode = Encode.list encode |
|
2809 |
instance Decode a => Decode [a] where decode = Decode.list decode |
|
2810 |
||
2811 |
instance Encode a => Encode (Maybe a) where encode = Encode.option encode |
|
2812 |
instance Decode a => Decode (Maybe a) where decode = Decode.option decode |
|
2813 |
||
2814 |
instance Encode XML.Tree where encode = Encode.tree |
|
2815 |
instance Decode XML.Tree where decode = Decode.tree |
|
2816 |
||
2817 |
instance Encode Properties.T where encode = Encode.properties |
|
2818 |
instance Decode Properties.T where decode = Decode.properties |
|
2819 |
||
2820 |
instance Encode Typ where encode = Encode.typ |
|
2821 |
instance Decode Typ where decode = Decode.typ |
|
2822 |
||
2823 |
instance Encode Term where encode = Encode.term |
|
2824 |
instance Decode Term where decode = Decode.term |
|
2825 |
\<close> |
|
2826 |
||
69459 | 2827 |
generate_file "Isabelle/UUID.hs" = \<open> |
2828 |
{- Title: Isabelle/UUID.hs |
|
2829 |
Author: Makarius |
|
2830 |
LICENSE: BSD 3-clause (Isabelle) |
|
2831 |
||
2832 |
Universally unique identifiers. |
|
2833 |
||
2834 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/uuid.scala\<close>. |
|
2835 |
-} |
|
2836 |
||
74095 | 2837 |
module Isabelle.UUID (T, random, print, parse) |
69459 | 2838 |
where |
2839 |
||
74095 | 2840 |
import Prelude hiding (print) |
2841 |
||
69459 | 2842 |
import Data.UUID (UUID) |
2843 |
import qualified Data.UUID as UUID |
|
2844 |
import Data.UUID.V4 (nextRandom) |
|
2845 |
||
74084 | 2846 |
import qualified Isabelle.Bytes as Bytes |
2847 |
import Isabelle.Bytes (Bytes) |
|
2848 |
||
69459 | 2849 |
|
69462 | 2850 |
type T = UUID |
2851 |
||
2852 |
random :: IO T |
|
69459 | 2853 |
random = nextRandom |
2854 |
||
74095 | 2855 |
print :: T -> Bytes |
2856 |
print = Bytes.make . UUID.toASCIIBytes |
|
2857 |
||
2858 |
parse :: Bytes -> Maybe T |
|
2859 |
parse = UUID.fromASCIIBytes . Bytes.unmake |
|
69459 | 2860 |
\<close> |
2861 |
||
69448 | 2862 |
generate_file "Isabelle/Byte_Message.hs" = \<open> |
2863 |
{- Title: Isabelle/Byte_Message.hs |
|
69446 | 2864 |
Author: Makarius |
2865 |
LICENSE: BSD 3-clause (Isabelle) |
|
2866 |
||
69448 | 2867 |
Byte-oriented messages. |
2868 |
||
2869 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.ML\<close> |
|
2870 |
and \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.scala\<close>. |
|
69446 | 2871 |
-} |
2872 |
||
74083 | 2873 |
{-# LANGUAGE OverloadedStrings #-} |
73177 | 2874 |
{-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} |
2875 |
||
69452 | 2876 |
module Isabelle.Byte_Message ( |
69467 | 2877 |
write, write_line, |
74088 | 2878 |
read, read_block, read_line, |
74187 | 2879 |
make_message, write_message, read_message, |
2880 |
exchange_message, exchange_message0, |
|
73178 | 2881 |
make_line_message, write_line_message, read_line_message, |
2882 |
read_yxml, write_yxml |
|
69452 | 2883 |
) |
69446 | 2884 |
where |
2885 |
||
69449 | 2886 |
import Prelude hiding (read) |
69454 | 2887 |
import Data.Maybe |
69446 | 2888 |
import qualified Data.ByteString as ByteString |
74084 | 2889 |
import qualified Isabelle.Bytes as Bytes |
2890 |
import Isabelle.Bytes (Bytes) |
|
74135 | 2891 |
import qualified Isabelle.Symbol as Symbol |
74080
5b68a5cd7061
prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents:
73246
diff
changeset
|
2892 |
import qualified Isabelle.UTF8 as UTF8 |
73178 | 2893 |
import qualified Isabelle.XML as XML |
2894 |
import qualified Isabelle.YXML as YXML |
|
73177 | 2895 |
|
69446 | 2896 |
import Network.Socket (Socket) |
74082 | 2897 |
import qualified Network.Socket.ByteString as Socket |
69446 | 2898 |
|
74088 | 2899 |
import Isabelle.Library |
69446 | 2900 |
import qualified Isabelle.Value as Value |
2901 |
||
2902 |
||
69452 | 2903 |
{- output operations -} |
2904 |
||
74084 | 2905 |
write :: Socket -> [Bytes] -> IO () |
2906 |
write socket = Socket.sendMany socket . map Bytes.unmake |
|
2907 |
||
2908 |
write_line :: Socket -> Bytes -> IO () |
|
74083 | 2909 |
write_line socket s = write socket [s, "\n"] |
69467 | 2910 |
|
69452 | 2911 |
|
2912 |
{- input operations -} |
|
2913 |
||
74084 | 2914 |
read :: Socket -> Int -> IO Bytes |
69452 | 2915 |
read socket n = read_body 0 [] |
69449 | 2916 |
where |
74084 | 2917 |
result = Bytes.concat . reverse |
69452 | 2918 |
read_body len ss = |
69449 | 2919 |
if len >= n then return (result ss) |
2920 |
else |
|
2921 |
(do |
|
74082 | 2922 |
s <- Socket.recv socket (min (n - len) 8192) |
69449 | 2923 |
case ByteString.length s of |
2924 |
0 -> return (result ss) |
|
74084 | 2925 |
m -> read_body (len + m) (Bytes.make s : ss)) |
2926 |
||
2927 |
read_block :: Socket -> Int -> IO (Maybe Bytes, Int) |
|
69452 | 2928 |
read_block socket n = do |
2929 |
msg <- read socket n |
|
74084 | 2930 |
let len = Bytes.length msg |
69452 | 2931 |
return (if len == n then Just msg else Nothing, len) |
69449 | 2932 |
|
74084 | 2933 |
read_line :: Socket -> IO (Maybe Bytes) |
69452 | 2934 |
read_line socket = read_body [] |
69446 | 2935 |
where |
74134 | 2936 |
result = trim_line . Bytes.pack . reverse |
69452 | 2937 |
read_body bs = do |
74082 | 2938 |
s <- Socket.recv socket 1 |
69446 | 2939 |
case ByteString.length s of |
2940 |
0 -> return (if null bs then Nothing else Just (result bs)) |
|
2941 |
1 -> |
|
2942 |
case ByteString.head s of |
|
2943 |
10 -> return (Just (result bs)) |
|
69452 | 2944 |
b -> read_body (b : bs) |
69446 | 2945 |
|
69448 | 2946 |
|
69454 | 2947 |
{- messages with multiple chunks (arbitrary content) -} |
2948 |
||
74084 | 2949 |
make_header :: [Int] -> [Bytes] |
74095 | 2950 |
make_header ns = [space_implode "," (map Value.print_int ns), "\n"] |
69454 | 2951 |
|
74084 | 2952 |
make_message :: [Bytes] -> [Bytes] |
2953 |
make_message chunks = make_header (map Bytes.length chunks) <> chunks |
|
2954 |
||
2955 |
write_message :: Socket -> [Bytes] -> IO () |
|
69476 | 2956 |
write_message socket = write socket . make_message |
69454 | 2957 |
|
74084 | 2958 |
parse_header :: Bytes -> [Int] |
69454 | 2959 |
parse_header line = |
2960 |
let |
|
74132 | 2961 |
res = map Value.parse_nat (space_explode ',' line) |
69454 | 2962 |
in |
2963 |
if all isJust res then map fromJust res |
|
74081 | 2964 |
else error ("Malformed message header: " <> quote (UTF8.decode line)) |
69454 | 2965 |
|
74084 | 2966 |
read_chunk :: Socket -> Int -> IO Bytes |
69454 | 2967 |
read_chunk socket n = do |
2968 |
res <- read_block socket n |
|
2969 |
return $ |
|
2970 |
case res of |
|
2971 |
(Just chunk, _) -> chunk |
|
2972 |
(Nothing, len) -> |
|
74081 | 2973 |
error ("Malformed message chunk: unexpected EOF after " <> |
2974 |
show len <> " of " <> show n <> " bytes") |
|
69454 | 2975 |
|
74084 | 2976 |
read_message :: Socket -> IO (Maybe [Bytes]) |
69454 | 2977 |
read_message socket = do |
2978 |
res <- read_line socket |
|
2979 |
case res of |
|
2980 |
Just line -> Just <$> mapM (read_chunk socket) (parse_header line) |
|
2981 |
Nothing -> return Nothing |
|
2982 |
||
74186 | 2983 |
exchange_message :: Socket -> [Bytes] -> IO (Maybe [Bytes]) |
2984 |
exchange_message socket msg = do |
|
2985 |
write_message socket msg |
|
2986 |
read_message socket |
|
2987 |
||
74187 | 2988 |
exchange_message0 :: Socket -> [Bytes] -> IO () |
2989 |
exchange_message0 socket msg = do |
|
2990 |
_ <- exchange_message socket msg |
|
2991 |
return () |
|
2992 |
||
69454 | 2993 |
|
69448 | 2994 |
-- hybrid messages: line or length+block (with content restriction) |
2995 |
||
74084 | 2996 |
is_length :: Bytes -> Bool |
69452 | 2997 |
is_length msg = |
74135 | 2998 |
not (Bytes.null msg) && Bytes.all_char (\c -> '0' <= c && c <= '9') msg |
74084 | 2999 |
|
3000 |
is_terminated :: Bytes -> Bool |
|
69452 | 3001 |
is_terminated msg = |
74135 | 3002 |
not (Bytes.null msg) && Symbol.is_ascii_line_terminator (Bytes.char $ Bytes.last msg) |
74084 | 3003 |
|
3004 |
make_line_message :: Bytes -> [Bytes] |
|
69476 | 3005 |
make_line_message msg = |
74084 | 3006 |
let n = Bytes.length msg in |
69476 | 3007 |
if is_length msg || is_terminated msg then |
74081 | 3008 |
error ("Bad content for line message:\n" <> take 100 (UTF8.decode msg)) |
69476 | 3009 |
else |
74135 | 3010 |
(if n > 100 || Bytes.any_char (== '\n') msg then make_header [n + 1] else []) <> [msg, "\n"] |
74084 | 3011 |
|
3012 |
write_line_message :: Socket -> Bytes -> IO () |
|
69476 | 3013 |
write_line_message socket = write socket . make_line_message |
69448 | 3014 |
|
74084 | 3015 |
read_line_message :: Socket -> IO (Maybe Bytes) |
69448 | 3016 |
read_line_message socket = do |
69446 | 3017 |
opt_line <- read_line socket |
3018 |
case opt_line of |
|
3019 |
Nothing -> return Nothing |
|
3020 |
Just line -> |
|
74095 | 3021 |
case Value.parse_nat line of |
69446 | 3022 |
Nothing -> return $ Just line |
74134 | 3023 |
Just n -> fmap trim_line . fst <$> read_block socket n |
73178 | 3024 |
|
3025 |
||
3026 |
read_yxml :: Socket -> IO (Maybe XML.Body) |
|
3027 |
read_yxml socket = do |
|
3028 |
res <- read_line_message socket |
|
74095 | 3029 |
return (YXML.parse_body <$> res) |
73178 | 3030 |
|
3031 |
write_yxml :: Socket -> XML.Body -> IO () |
|
3032 |
write_yxml socket body = |
|
74095 | 3033 |
write_line_message socket (YXML.string_of_body body) |
69446 | 3034 |
\<close> |
3035 |
||
71692 | 3036 |
generate_file "Isabelle/Isabelle_Thread.hs" = \<open> |
3037 |
{- Title: Isabelle/Isabelle_Thread.hs |
|
69473 | 3038 |
Author: Makarius |
3039 |
LICENSE: BSD 3-clause (Isabelle) |
|
3040 |
||
71692 | 3041 |
Isabelle-specific thread management. |
3042 |
||
3043 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\<close> |
|
3044 |
and \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\<close>. |
|
69473 | 3045 |
-} |
3046 |
||
3047 |
{-# LANGUAGE NamedFieldPuns #-} |
|
3048 |
||
71692 | 3049 |
module Isabelle.Isabelle_Thread ( |
69473 | 3050 |
ThreadId, Result, |
69494 | 3051 |
find_id, |
69473 | 3052 |
properties, change_properties, |
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3053 |
add_resource, del_resource, bracket_resource, |
69497 | 3054 |
is_stopped, expose_stopped, stop, |
69496 | 3055 |
my_uuid, stop_uuid, |
69494 | 3056 |
Fork, fork_finally, fork) |
69473 | 3057 |
where |
3058 |
||
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3059 |
import Data.Unique |
69473 | 3060 |
import Data.IORef |
3061 |
import System.IO.Unsafe |
|
3062 |
||
69494 | 3063 |
import qualified Data.List as List |
69497 | 3064 |
import Control.Monad (when, forM_) |
69473 | 3065 |
import Data.Map.Strict (Map) |
3066 |
import qualified Data.Map.Strict as Map |
|
3067 |
import Control.Exception as Exception |
|
3068 |
import Control.Concurrent (ThreadId) |
|
3069 |
import qualified Control.Concurrent as Concurrent |
|
3070 |
import Control.Concurrent.Thread (Result) |
|
3071 |
import qualified Control.Concurrent.Thread as Thread |
|
69494 | 3072 |
import qualified Isabelle.UUID as UUID |
69473 | 3073 |
import qualified Isabelle.Properties as Properties |
3074 |
||
3075 |
||
3076 |
{- thread info -} |
|
3077 |
||
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3078 |
type Resources = Map Unique (IO ()) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3079 |
data Info = Info {uuid :: UUID.T, props :: Properties.T, stopped :: Bool, resources :: Resources} |
69473 | 3080 |
type Infos = Map ThreadId Info |
3081 |
||
3082 |
lookup_info :: Infos -> ThreadId -> Maybe Info |
|
69494 | 3083 |
lookup_info infos id = Map.lookup id infos |
3084 |
||
3085 |
init_info :: ThreadId -> UUID.T -> Infos -> (Infos, ()) |
|
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3086 |
init_info id uuid infos = (Map.insert id (Info uuid [] False Map.empty) infos, ()) |
69473 | 3087 |
|
3088 |
||
3089 |
{- global state -} |
|
3090 |
||
3091 |
{-# NOINLINE global_state #-} |
|
3092 |
global_state :: IORef Infos |
|
3093 |
global_state = unsafePerformIO (newIORef Map.empty) |
|
3094 |
||
69494 | 3095 |
find_id :: UUID.T -> IO (Maybe ThreadId) |
3096 |
find_id uuid = do |
|
3097 |
state <- readIORef global_state |
|
3098 |
return $ fst <$> List.find (\(_, Info{uuid = uuid'}) -> uuid == uuid') (Map.assocs state) |
|
3099 |
||
69473 | 3100 |
get_info :: ThreadId -> IO (Maybe Info) |
3101 |
get_info id = do |
|
3102 |
state <- readIORef global_state |
|
3103 |
return $ lookup_info state id |
|
3104 |
||
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3105 |
map_info :: ThreadId -> (Info -> Info) -> IO (Maybe Info) |
69473 | 3106 |
map_info id f = |
3107 |
atomicModifyIORef' global_state |
|
3108 |
(\infos -> |
|
3109 |
case lookup_info infos id of |
|
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3110 |
Nothing -> (infos, Nothing) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3111 |
Just info -> |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3112 |
let info' = f info |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3113 |
in (Map.insert id info' infos, Just info')) |
69473 | 3114 |
|
69495 | 3115 |
delete_info :: ThreadId -> IO () |
3116 |
delete_info id = |
|
3117 |
atomicModifyIORef' global_state (\infos -> (Map.delete id infos, ())) |
|
3118 |
||
69473 | 3119 |
|
3120 |
{- thread properties -} |
|
3121 |
||
69498
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3122 |
my_info :: IO (Maybe Info) |
69473 | 3123 |
my_info = do |
3124 |
id <- Concurrent.myThreadId |
|
69498
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3125 |
get_info id |
69473 | 3126 |
|
3127 |
properties :: IO Properties.T |
|
69498
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3128 |
properties = maybe [] props <$> my_info |
69473 | 3129 |
|
3130 |
change_properties :: (Properties.T -> Properties.T) -> IO () |
|
3131 |
change_properties f = do |
|
3132 |
id <- Concurrent.myThreadId |
|
3133 |
map_info id (\info -> info {props = f (props info)}) |
|
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3134 |
return () |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3135 |
|
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3136 |
|
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3137 |
{- managed resources -} |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3138 |
|
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3139 |
add_resource :: IO () -> IO Unique |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3140 |
add_resource resource = do |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3141 |
id <- Concurrent.myThreadId |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3142 |
u <- newUnique |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3143 |
map_info id (\info -> info {resources = Map.insert u resource (resources info)}) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3144 |
return u |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3145 |
|
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3146 |
del_resource :: Unique -> IO () |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3147 |
del_resource u = do |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3148 |
id <- Concurrent.myThreadId |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3149 |
map_info id (\info -> info {resources = Map.delete u (resources info)}) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3150 |
return () |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3151 |
|
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3152 |
bracket_resource :: IO () -> IO a -> IO a |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3153 |
bracket_resource resource body = |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3154 |
Exception.bracket (add_resource resource) del_resource (const body) |
69473 | 3155 |
|
3156 |
||
3157 |
{- stop -} |
|
3158 |
||
3159 |
is_stopped :: IO Bool |
|
69498
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3160 |
is_stopped = maybe False stopped <$> my_info |
69473 | 3161 |
|
69497 | 3162 |
expose_stopped :: IO () |
3163 |
expose_stopped = do |
|
3164 |
stopped <- is_stopped |
|
3165 |
when stopped $ throw ThreadKilled |
|
3166 |
||
69473 | 3167 |
stop :: ThreadId -> IO () |
69499
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3168 |
stop id = do |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3169 |
info <- map_info id (\info -> info {stopped = True}) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3170 |
let ops = case info of Nothing -> []; Just Info{resources} -> map snd (Map.toDescList resources) |
638fdbbc7d1f
more Haskell operations: managed resources for threads;
wenzelm
parents:
69498
diff
changeset
|
3171 |
sequence_ ops |
69473 | 3172 |
|
3173 |
||
69496 | 3174 |
{- UUID -} |
3175 |
||
69498
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3176 |
my_uuid :: IO (Maybe UUID.T) |
22e958b76bf6
more robust: suitable defaults for unmanaged threads;
wenzelm
parents:
69497
diff
changeset
|
3177 |
my_uuid = fmap uuid <$> my_info |
69496 | 3178 |
|
3179 |
stop_uuid :: UUID.T -> IO () |
|
3180 |
stop_uuid uuid = do |
|
3181 |
id <- find_id uuid |
|
3182 |
forM_ id stop |
|
3183 |
||
3184 |
||
69473 | 3185 |
{- fork -} |
3186 |
||
69494 | 3187 |
type Fork a = (ThreadId, UUID.T, IO (Result a)) |
3188 |
||
3189 |
fork_finally :: IO a -> (Either SomeException a -> IO b) -> IO (Fork b) |
|
69473 | 3190 |
fork_finally body finally = do |
69494 | 3191 |
uuid <- UUID.random |
69473 | 3192 |
(id, result) <- |
3193 |
Exception.mask (\restore -> |
|
3194 |
Thread.forkIO |
|
3195 |
(Exception.try |
|
3196 |
(do |
|
3197 |
id <- Concurrent.myThreadId |
|
69494 | 3198 |
atomicModifyIORef' global_state (init_info id uuid) |
69495 | 3199 |
restore body) |
3200 |
>>= (\res -> do id <- Concurrent.myThreadId; delete_info id; finally res))) |
|
69494 | 3201 |
return (id, uuid, result) |
3202 |
||
3203 |
fork :: IO a -> IO (Fork a) |
|
69473 | 3204 |
fork body = fork_finally body Thread.result |
3205 |
\<close> |
|
3206 |
||
69455 | 3207 |
generate_file "Isabelle/Server.hs" = \<open> |
3208 |
{- Title: Isabelle/Server.hs |
|
3209 |
Author: Makarius |
|
3210 |
LICENSE: BSD 3-clause (Isabelle) |
|
3211 |
||
3212 |
TCP server on localhost. |
|
3213 |
-} |
|
3214 |
||
74127 | 3215 |
{-# LANGUAGE OverloadedStrings #-} |
3216 |
||
69465 | 3217 |
module Isabelle.Server ( |
74216 | 3218 |
localhost_name, localhost_prefix, localhost, publish_text, publish_stdout, |
73178 | 3219 |
server, connection |
69465 | 3220 |
) |
3221 |
where |
|
69455 | 3222 |
|
69465 | 3223 |
import Control.Monad (forever, when) |
69455 | 3224 |
import qualified Control.Exception as Exception |
3225 |
import Network.Socket (Socket) |
|
3226 |
import qualified Network.Socket as Socket |
|
69472 | 3227 |
import qualified System.IO as IO |
74127 | 3228 |
import qualified Data.ByteString.Char8 as Char8 |
69455 | 3229 |
|
69462 | 3230 |
import Isabelle.Library |
74127 | 3231 |
import qualified Isabelle.Bytes as Bytes |
74084 | 3232 |
import Isabelle.Bytes (Bytes) |
69462 | 3233 |
import qualified Isabelle.UUID as UUID |
69465 | 3234 |
import qualified Isabelle.Byte_Message as Byte_Message |
71692 | 3235 |
import qualified Isabelle.Isabelle_Thread as Isabelle_Thread |
69462 | 3236 |
|
3237 |
||
3238 |
{- server address -} |
|
69455 | 3239 |
|
74127 | 3240 |
localhost_name :: Bytes |
69463 | 3241 |
localhost_name = "127.0.0.1" |
3242 |
||
74211 | 3243 |
localhost_prefix :: Bytes |
3244 |
localhost_prefix = localhost_name <> ":" |
|
3245 |
||
69455 | 3246 |
localhost :: Socket.HostAddress |
3247 |
localhost = Socket.tupleToHostAddress (127, 0, 0, 1) |
|
3248 |
||
74127 | 3249 |
publish_text :: Bytes -> Bytes -> UUID.T -> Bytes |
69462 | 3250 |
publish_text name address password = |
74127 | 3251 |
"server " <> quote name <> " = " <> address <> |
74128 | 3252 |
" (password " <> quote (show_bytes password) <> ")" |
74127 | 3253 |
|
3254 |
publish_stdout :: Bytes -> Bytes -> UUID.T -> IO () |
|
74095 | 3255 |
publish_stdout name address password = |
74127 | 3256 |
Char8.putStrLn (Bytes.unmake $ publish_text name address password) |
69462 | 3257 |
|
3258 |
||
3259 |
{- server -} |
|
3260 |
||
74127 | 3261 |
server :: (Bytes -> UUID.T -> IO ()) -> (Socket -> IO ()) -> IO () |
69455 | 3262 |
server publish handle = |
69462 | 3263 |
Socket.withSocketsDo $ Exception.bracket open (Socket.close . fst) (uncurry loop) |
69455 | 3264 |
where |
74084 | 3265 |
open :: IO (Socket, Bytes) |
69455 | 3266 |
open = do |
69466 | 3267 |
server_socket <- Socket.socket Socket.AF_INET Socket.Stream Socket.defaultProtocol |
3268 |
Socket.bind server_socket (Socket.SockAddrInet 0 localhost) |
|
3269 |
Socket.listen server_socket 50 |
|
69462 | 3270 |
|
69466 | 3271 |
port <- Socket.socketPort server_socket |
74128 | 3272 |
let address = localhost_name <> ":" <> show_bytes port |
69462 | 3273 |
password <- UUID.random |
3274 |
publish address password |
|
69455 | 3275 |
|
74095 | 3276 |
return (server_socket, UUID.print password) |
69462 | 3277 |
|
74084 | 3278 |
loop :: Socket -> Bytes -> IO () |
69466 | 3279 |
loop server_socket password = forever $ do |
69480 | 3280 |
(connection, _) <- Socket.accept server_socket |
71692 | 3281 |
Isabelle_Thread.fork_finally |
69465 | 3282 |
(do |
3283 |
line <- Byte_Message.read_line connection |
|
3284 |
when (line == Just password) $ handle connection) |
|
69473 | 3285 |
(\finally -> do |
69472 | 3286 |
Socket.close connection |
69473 | 3287 |
case finally of |
69472 | 3288 |
Left exn -> IO.hPutStrLn IO.stderr $ Exception.displayException exn |
3289 |
Right () -> return ()) |
|
69455 | 3290 |
return () |
73178 | 3291 |
|
3292 |
||
3293 |
{- client connection -} |
|
3294 |
||
74095 | 3295 |
connection :: String -> Bytes -> (Socket -> IO a) -> IO a |
73178 | 3296 |
connection port password client = |
3297 |
Socket.withSocketsDo $ do |
|
3298 |
addr <- resolve |
|
3299 |
Exception.bracket (open addr) Socket.close body |
|
3300 |
where |
|
3301 |
resolve = do |
|
3302 |
let hints = |
|
3303 |
Socket.defaultHints { |
|
3304 |
Socket.addrFlags = [Socket.AI_NUMERICHOST, Socket.AI_NUMERICSERV], |
|
3305 |
Socket.addrSocketType = Socket.Stream } |
|
74211 | 3306 |
head <$> Socket.getAddrInfo (Just hints) (Just $ make_string localhost_name) (Just port) |
73178 | 3307 |
|
3308 |
open addr = do |
|
3309 |
socket <- Socket.socket (Socket.addrFamily addr) (Socket.addrSocketType addr) |
|
3310 |
(Socket.addrProtocol addr) |
|
3311 |
Socket.connect socket $ Socket.addrAddress addr |
|
3312 |
return socket |
|
3313 |
||
3314 |
body socket = do |
|
74095 | 3315 |
Byte_Message.write_line socket password |
73178 | 3316 |
client socket |
69455 | 3317 |
\<close> |
3318 |
||
74161 | 3319 |
generate_file "Isabelle/Time.hs" = \<open> |
3320 |
{- Title: Isabelle/Time.hs |
|
3321 |
Author: Makarius |
|
3322 |
LICENSE: BSD 3-clause (Isabelle) |
|
3323 |
||
3324 |
Time based on milliseconds. |
|
3325 |
||
3326 |
See \<^file>\<open>~~/src/Pure/General/time.scala\<close> |
|
3327 |
-} |
|
3328 |
||
3329 |
{-# LANGUAGE OverloadedStrings #-} |
|
3330 |
||
3331 |
module Isabelle.Time ( |
|
3332 |
Time, seconds, minutes, ms, zero, is_zero, is_relevant, |
|
74209 | 3333 |
get_seconds, get_minutes, get_ms, message, now |
74161 | 3334 |
) |
3335 |
where |
|
3336 |
||
3337 |
import Text.Printf (printf) |
|
74209 | 3338 |
import Data.Time.Clock.POSIX (getPOSIXTime) |
74161 | 3339 |
import Isabelle.Bytes (Bytes) |
3340 |
import Isabelle.Library |
|
3341 |
||
3342 |
||
74218 | 3343 |
newtype Time = Time Int |
74161 | 3344 |
|
75066 | 3345 |
instance Eq Time where Time a == Time b = a == b |
3346 |
instance Ord Time where compare (Time a) (Time b) = compare a b |
|
3347 |
instance Num Time where |
|
3348 |
fromInteger = Time . fromInteger |
|
3349 |
Time a + Time b = Time (a + b) |
|
3350 |
Time a - Time b = Time (a - b) |
|
3351 |
Time a * Time b = Time (a * b) |
|
3352 |
abs (Time a) = Time (abs a) |
|
3353 |
signum (Time a) = Time (signum a) |
|
74161 | 3354 |
|
3355 |
seconds :: Double -> Time |
|
3356 |
seconds s = Time (round (s * 1000.0)) |
|
3357 |
||
3358 |
minutes :: Double -> Time |
|
3359 |
minutes m = Time (round (m * 60000.0)) |
|
3360 |
||
3361 |
ms :: Int -> Time |
|
3362 |
ms = Time |
|
3363 |
||
3364 |
zero :: Time |
|
3365 |
zero = ms 0 |
|
3366 |
||
3367 |
is_zero :: Time -> Bool |
|
3368 |
is_zero (Time ms) = ms == 0 |
|
3369 |
||
3370 |
is_relevant :: Time -> Bool |
|
3371 |
is_relevant (Time ms) = ms >= 1 |
|
3372 |
||
3373 |
get_seconds :: Time -> Double |
|
3374 |
get_seconds (Time ms) = fromIntegral ms / 1000.0 |
|
3375 |
||
3376 |
get_minutes :: Time -> Double |
|
3377 |
get_minutes (Time ms) = fromIntegral ms / 60000.0 |
|
3378 |
||
3379 |
get_ms :: Time -> Int |
|
3380 |
get_ms (Time ms) = ms |
|
3381 |
||
3382 |
instance Show Time where |
|
3383 |
show t = printf "%.3f" (get_seconds t) |
|
3384 |
||
3385 |
message :: Time -> Bytes |
|
3386 |
message t = make_bytes (show t) <> "s" |
|
74209 | 3387 |
|
3388 |
now :: IO Time |
|
3389 |
now = do |
|
3390 |
t <- getPOSIXTime |
|
3391 |
return $ Time (round (realToFrac t * 1000.0 :: Double)) |
|
74161 | 3392 |
\<close> |
3393 |
||
3394 |
generate_file "Isabelle/Timing.hs" = \<open> |
|
3395 |
{- Title: Isabelle/Timing.hs |
|
3396 |
Author: Makarius |
|
3397 |
LICENSE: BSD 3-clause (Isabelle) |
|
3398 |
||
3399 |
Support for time measurement. |
|
3400 |
||
3401 |
See \<^file>\<open>~~/src/Pure/General/timing.ML\<close> |
|
3402 |
and \<^file>\<open>~~/src/Pure/General/timing.scala\<close> |
|
3403 |
-} |
|
3404 |
||
3405 |
module Isabelle.Timing ( |
|
3406 |
Timing (..), zero, is_zero, is_relevant |
|
3407 |
) |
|
3408 |
where |
|
3409 |
||
3410 |
import qualified Isabelle.Time as Time |
|
3411 |
import Isabelle.Time (Time) |
|
3412 |
||
3413 |
data Timing = Timing {elapsed :: Time, cpu :: Time, gc :: Time} |
|
3414 |
deriving (Show, Eq) |
|
3415 |
||
3416 |
zero :: Timing |
|
3417 |
zero = Timing Time.zero Time.zero Time.zero |
|
3418 |
||
3419 |
is_zero :: Timing -> Bool |
|
3420 |
is_zero t = Time.is_zero (elapsed t) && Time.is_zero (cpu t) && Time.is_zero (gc t) |
|
3421 |
||
3422 |
is_relevant :: Timing -> Bool |
|
3423 |
is_relevant t = Time.is_relevant (elapsed t) || Time.is_relevant (cpu t) || Time.is_relevant (gc t) |
|
3424 |
\<close> |
|
3425 |
||
3426 |
generate_file "Isabelle/Bash.hs" = \<open> |
|
3427 |
{- Title: Isabelle/Bash.hs |
|
3428 |
Author: Makarius |
|
3429 |
LICENSE: BSD 3-clause (Isabelle) |
|
3430 |
||
3431 |
Support for GNU bash. |
|
3432 |
||
3433 |
See \<^file>\<open>$ISABELLE_HOME/src/Pure/System/bash.ML\<close> |
|
3434 |
-} |
|
3435 |
||
3436 |
{-# LANGUAGE OverloadedStrings #-} |
|
3437 |
||
3438 |
module Isabelle.Bash ( |
|
3439 |
string, strings, |
|
3440 |
||
3441 |
Params, |
|
3442 |
get_script, get_input, get_cwd, get_putenv, get_redirect, |
|
3443 |
get_timeout, get_description, |
|
3444 |
script, input, cwd, putenv, redirect, timeout, description, |
|
74211 | 3445 |
server_run, server_kill, |
3446 |
server_uuid, server_interrupt, server_failure, server_result |
|
74161 | 3447 |
) |
3448 |
where |
|
3449 |
||
3450 |
import Text.Printf (printf) |
|
3451 |
import qualified Isabelle.Symbol as Symbol |
|
3452 |
import qualified Isabelle.Bytes as Bytes |
|
3453 |
import Isabelle.Bytes (Bytes) |
|
3454 |
import qualified Isabelle.Time as Time |
|
3455 |
import Isabelle.Time (Time) |
|
3456 |
import Isabelle.Library |
|
3457 |
||
3458 |
||
3459 |
{- concrete syntax -} |
|
3460 |
||
3461 |
string :: Bytes -> Bytes |
|
3462 |
string str = |
|
3463 |
if Bytes.null str then "\"\"" |
|
3464 |
else str |> Bytes.unpack |> map trans |> Bytes.concat |
|
3465 |
where |
|
3466 |
trans b = |
|
3467 |
case Bytes.char b of |
|
3468 |
'\t' -> "$'\\t'" |
|
3469 |
'\n' -> "$'\\n'" |
|
3470 |
'\f' -> "$'\\f'" |
|
3471 |
'\r' -> "$'\\r'" |
|
3472 |
c -> |
|
74169 | 3473 |
if Symbol.is_ascii_letter c || Symbol.is_ascii_digit c || c `elem` ("+,-./:_" :: String) |
74161 | 3474 |
then Bytes.singleton b |
3475 |
else if b < 32 || b >= 127 then make_bytes (printf "$'\\x%02x'" b :: String) |
|
3476 |
else "\\" <> Bytes.singleton b |
|
3477 |
||
3478 |
strings :: [Bytes] -> Bytes |
|
80910 | 3479 |
strings = implode_space . map string |
74161 | 3480 |
|
3481 |
||
3482 |
{- server parameters -} |
|
3483 |
||
3484 |
data Params = Params { |
|
3485 |
_script :: Bytes, |
|
3486 |
_input :: Bytes, |
|
3487 |
_cwd :: Maybe Bytes, |
|
3488 |
_putenv :: [(Bytes, Bytes)], |
|
3489 |
_redirect :: Bool, |
|
3490 |
_timeout :: Time, |
|
3491 |
_description :: Bytes} |
|
3492 |
deriving (Show, Eq) |
|
3493 |
||
3494 |
get_script :: Params -> Bytes |
|
3495 |
get_script = _script |
|
3496 |
||
3497 |
get_input :: Params -> Bytes |
|
3498 |
get_input = _input |
|
3499 |
||
3500 |
get_cwd :: Params -> Maybe Bytes |
|
3501 |
get_cwd = _cwd |
|
3502 |
||
3503 |
get_putenv :: Params -> [(Bytes, Bytes)] |
|
3504 |
get_putenv = _putenv |
|
3505 |
||
3506 |
get_redirect :: Params -> Bool |
|
3507 |
get_redirect = _redirect |
|
3508 |
||
3509 |
get_timeout :: Params -> Time |
|
3510 |
get_timeout = _timeout |
|
3511 |
||
3512 |
get_description :: Params -> Bytes |
|
3513 |
get_description = _description |
|
3514 |
||
3515 |
script :: Bytes -> Params |
|
3516 |
script script = Params script "" Nothing [] False Time.zero "" |
|
3517 |
||
3518 |
input :: Bytes -> Params -> Params |
|
3519 |
input input params = params { _input = input } |
|
3520 |
||
3521 |
cwd :: Bytes -> Params -> Params |
|
3522 |
cwd cwd params = params { _cwd = Just cwd } |
|
3523 |
||
3524 |
putenv :: [(Bytes, Bytes)] -> Params -> Params |
|
3525 |
putenv putenv params = params { _putenv = putenv } |
|
3526 |
||
3527 |
redirect :: Params -> Params |
|
3528 |
redirect params = params { _redirect = True } |
|
3529 |
||
3530 |
timeout :: Time -> Params -> Params |
|
3531 |
timeout timeout params = params { _timeout = timeout } |
|
3532 |
||
3533 |
description :: Bytes -> Params -> Params |
|
3534 |
description description params = params { _description = description } |
|
74211 | 3535 |
|
3536 |
||
3537 |
{- server messages -} |
|
3538 |
||
3539 |
server_run, server_kill :: Bytes |
|
3540 |
server_run = \<open>Bash.server_run\<close>; |
|
3541 |
server_kill = \<open>Bash.server_kill\<close>; |
|
3542 |
||
3543 |
server_uuid, server_interrupt, server_failure, server_result :: Bytes |
|
3544 |
server_uuid = \<open>Bash.server_uuid\<close>; |
|
3545 |
server_interrupt = \<open>Bash.server_interrupt\<close>; |
|
3546 |
server_failure = \<open>Bash.server_failure\<close>; |
|
3547 |
server_result = \<open>Bash.server_result\<close>; |
|
74161 | 3548 |
\<close> |
3549 |
||
3550 |
generate_file "Isabelle/Process_Result.hs" = \<open> |
|
3551 |
{- Title: Isabelle/Process_Result.hs |
|
3552 |
Author: Makarius |
|
3553 |
LICENSE: BSD 3-clause (Isabelle) |
|
3554 |
||
3555 |
Result of system process. |
|
3556 |
||
3557 |
See \<^file>\<open>~~/src/Pure/System/process_result.ML\<close> |
|
3558 |
and \<^file>\<open>~~/src/Pure/System/process_result.scala\<close> |
|
3559 |
-} |
|
3560 |
||
3561 |
{-# LANGUAGE OverloadedStrings #-} |
|
3562 |
||
3563 |
module Isabelle.Process_Result ( |
|
74310 | 3564 |
ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc, |
74161 | 3565 |
|
3566 |
T, make, rc, out_lines, err_lines, timing, timing_elapsed, out, err, ok, check |
|
3567 |
) |
|
3568 |
where |
|
3569 |
||
3570 |
import Isabelle.Time (Time) |
|
3571 |
import qualified Isabelle.Timing as Timing |
|
3572 |
import Isabelle.Timing (Timing) |
|
3573 |
import Isabelle.Bytes (Bytes) |
|
3574 |
import Isabelle.Library |
|
3575 |
||
3576 |
||
74310 | 3577 |
ok_rc, error_rc, failure_rc, interrupt_rc , timeout_rc :: Int |
3578 |
ok_rc = 0 |
|
3579 |
error_rc = 1 |
|
3580 |
failure_rc = 2 |
|
74161 | 3581 |
interrupt_rc = 130 |
3582 |
timeout_rc = 142 |
|
3583 |
||
3584 |
data T = |
|
3585 |
Process_Result { |
|
3586 |
_rc :: Int, |
|
3587 |
_out_lines :: [Bytes], |
|
3588 |
_err_lines :: [Bytes], |
|
3589 |
_timing :: Timing} |
|
3590 |
deriving (Show, Eq) |
|
3591 |
||
3592 |
make :: Int -> [Bytes] -> [Bytes] -> Timing -> T |
|
3593 |
make = Process_Result |
|
3594 |
||
3595 |
rc :: T -> Int |
|
3596 |
rc = _rc |
|
3597 |
||
3598 |
out_lines :: T -> [Bytes] |
|
3599 |
out_lines = _out_lines |
|
3600 |
||
3601 |
err_lines :: T -> [Bytes] |
|
3602 |
err_lines = _err_lines |
|
3603 |
||
3604 |
timing :: T -> Timing |
|
3605 |
timing = _timing |
|
3606 |
||
3607 |
timing_elapsed :: T -> Time |
|
3608 |
timing_elapsed = Timing.elapsed . timing |
|
3609 |
||
3610 |
out :: T -> Bytes |
|
3611 |
out = trim_line . cat_lines . out_lines |
|
3612 |
||
3613 |
err :: T -> Bytes |
|
3614 |
err = trim_line . cat_lines . err_lines |
|
3615 |
||
3616 |
ok :: T -> Bool |
|
74310 | 3617 |
ok result = rc result == ok_rc |
74161 | 3618 |
|
3619 |
check :: T -> T |
|
3620 |
check result = if ok result then result else error (make_string $ err result) |
|
3621 |
\<close> |
|
3622 |
||
3623 |
generate_file "Isabelle/Options.hs" = \<open> |
|
3624 |
{- Title: Isabelle/Options.hs |
|
3625 |
Author: Makarius |
|
3626 |
LICENSE: BSD 3-clause (Isabelle) |
|
3627 |
||
3628 |
System options with external string representation. |
|
3629 |
||
3630 |
See \<^file>\<open>~~/src/Pure/System/options.ML\<close> |
|
3631 |
and \<^file>\<open>~~/src/Pure/System/options.scala\<close> |
|
3632 |
-} |
|
3633 |
||
3634 |
{-# LANGUAGE OverloadedStrings #-} |
|
3635 |
{-# LANGUAGE InstanceSigs #-} |
|
3636 |
||
3637 |
module Isabelle.Options ( |
|
3638 |
boolT, intT, realT, stringT, unknownT, |
|
3639 |
||
3640 |
T, typ, bool, int, real, seconds, string, |
|
3641 |
decode |
|
3642 |
) |
|
3643 |
where |
|
3644 |
||
3645 |
import qualified Data.Map.Strict as Map |
|
3646 |
import Data.Map.Strict (Map) |
|
3647 |
import qualified Isabelle.Properties as Properties |
|
3648 |
import Isabelle.Bytes (Bytes) |
|
3649 |
import qualified Isabelle.Value as Value |
|
3650 |
import qualified Isabelle.Time as Time |
|
3651 |
import Isabelle.Time (Time) |
|
3652 |
import Isabelle.Library |
|
3653 |
import qualified Isabelle.XML.Decode as Decode |
|
3654 |
import Isabelle.XML.Classes (Decode (..)) |
|
3655 |
||
3656 |
||
3657 |
{- representation -} |
|
3658 |
||
3659 |
boolT :: Bytes |
|
3660 |
boolT = "bool" |
|
3661 |
||
3662 |
intT :: Bytes |
|
3663 |
intT = "int" |
|
3664 |
||
3665 |
realT :: Bytes |
|
3666 |
realT = "real" |
|
3667 |
||
3668 |
stringT :: Bytes |
|
3669 |
stringT = "string" |
|
3670 |
||
3671 |
unknownT :: Bytes |
|
3672 |
unknownT = "unknown" |
|
3673 |
||
3674 |
data Opt = Opt { |
|
3675 |
_pos :: Properties.T, |
|
3676 |
_name :: Bytes, |
|
3677 |
_typ :: Bytes, |
|
3678 |
_value :: Bytes } |
|
3679 |
||
75015 | 3680 |
newtype T = Options (Map Bytes Opt) |
74161 | 3681 |
|
3682 |
||
3683 |
{- check -} |
|
3684 |
||
3685 |
check_name :: T -> Bytes -> Opt |
|
3686 |
check_name (Options map) name = |
|
3687 |
case Map.lookup name map of |
|
3688 |
Just opt | _typ opt /= unknownT -> opt |
|
3689 |
_ -> error (make_string ("Unknown system option " <> quote name)) |
|
3690 |
||
3691 |
check_type :: T -> Bytes -> Bytes -> Opt |
|
3692 |
check_type options name typ = |
|
3693 |
let |
|
3694 |
opt = check_name options name |
|
3695 |
t = _typ opt |
|
3696 |
in |
|
3697 |
if t == typ then opt |
|
3698 |
else error (make_string ("Ill-typed system option " <> quote name <> " : " <> t <> " vs. " <> typ)) |
|
3699 |
||
3700 |
||
3701 |
{- get typ -} |
|
3702 |
||
3703 |
typ :: T -> Bytes -> Bytes |
|
3704 |
typ options name = _typ (check_name options name) |
|
3705 |
||
3706 |
||
3707 |
{- get value -} |
|
3708 |
||
3709 |
get :: Bytes -> (Bytes -> Maybe a) -> T -> Bytes -> a |
|
3710 |
get typ parse options name = |
|
3711 |
let opt = check_type options name typ in |
|
3712 |
case parse (_value opt) of |
|
3713 |
Just x -> x |
|
3714 |
Nothing -> |
|
3715 |
error (make_string ("Malformed value for system option " <> quote name <> |
|
3716 |
" : " <> typ <> " =\n" <> quote (_value opt))) |
|
3717 |
||
3718 |
bool :: T -> Bytes -> Bool |
|
3719 |
bool = get boolT Value.parse_bool |
|
3720 |
||
3721 |
int :: T -> Bytes -> Int |
|
3722 |
int = get intT Value.parse_int |
|
3723 |
||
3724 |
real :: T -> Bytes -> Double |
|
3725 |
real = get realT Value.parse_real |
|
3726 |
||
3727 |
seconds :: T -> Bytes -> Time |
|
3728 |
seconds options = Time.seconds . real options |
|
3729 |
||
3730 |
string :: T -> Bytes -> Bytes |
|
3731 |
string = get stringT Just |
|
3732 |
||
3733 |
||
3734 |
{- decode -} |
|
3735 |
||
3736 |
instance Decode T where |
|
3737 |
decode :: Decode.T T |
|
3738 |
decode = |
|
3739 |
let |
|
3740 |
decode_entry :: Decode.T (Bytes, Opt) |
|
3741 |
decode_entry body = |
|
3742 |
let |
|
3743 |
(pos, (name, (typ, value))) = |
|
3744 |
Decode.pair Decode.properties (Decode.pair Decode.string (Decode.pair Decode.string Decode.string)) body |
|
3745 |
in (name, Opt { _pos = pos, _name = name, _typ = typ, _value = value }) |
|
3746 |
in Options . Map.fromList . Decode.list decode_entry |
|
3747 |
\<close> |
|
3748 |
||
74211 | 3749 |
generate_file "Isabelle/Isabelle_System.hs" = \<open> |
3750 |
{- Title: Isabelle/Isabelle_System.hs |
|
3751 |
Author: Makarius |
|
3752 |
LICENSE: BSD 3-clause (Isabelle) |
|
3753 |
||
3754 |
Isabelle system support. |
|
3755 |
||
3756 |
See \<^file>\<open>~~/src/Pure/System/isabelle_system.ML\<close> |
|
3757 |
and \<^file>\<open>~~/src/Pure/System/isabelle_system.scala\<close> |
|
3758 |
-} |
|
3759 |
||
3760 |
{-# LANGUAGE OverloadedStrings #-} |
|
3761 |
||
3762 |
module Isabelle.Isabelle_System ( |
|
3763 |
bash_process, bash_process0 |
|
3764 |
) |
|
3765 |
where |
|
3766 |
||
3767 |
import Data.Maybe (fromMaybe) |
|
3768 |
import Control.Exception (throw, AsyncException (UserInterrupt)) |
|
3769 |
import Network.Socket (Socket) |
|
74216 | 3770 |
import qualified Isabelle.Bytes as Bytes |
74211 | 3771 |
import Isabelle.Bytes (Bytes) |
3772 |
import qualified Isabelle.Byte_Message as Byte_Message |
|
3773 |
import qualified Isabelle.Time as Time |
|
3774 |
import Isabelle.Timing (Timing (..)) |
|
3775 |
import qualified Isabelle.Options as Options |
|
3776 |
import qualified Isabelle.Bash as Bash |
|
3777 |
import qualified Isabelle.Process_Result as Process_Result |
|
3778 |
import qualified Isabelle.XML.Encode as Encode |
|
3779 |
import qualified Isabelle.YXML as YXML |
|
3780 |
import qualified Isabelle.Value as Value |
|
3781 |
import qualified Isabelle.Server as Server |
|
3782 |
import qualified Isabelle.Isabelle_Thread as Isabelle_Thread |
|
3783 |
import Isabelle.Library |
|
3784 |
||
3785 |
||
3786 |
{- bash_process -} |
|
3787 |
||
3788 |
absolute_path :: Bytes -> Bytes -- FIXME dummy |
|
3789 |
absolute_path = id |
|
3790 |
||
3791 |
bash_process :: Options.T -> Bash.Params -> IO Process_Result.T |
|
3792 |
bash_process options = bash_process0 address password |
|
3793 |
where |
|
3794 |
address = Options.string options "bash_process_address" |
|
3795 |
password = Options.string options "bash_process_password" |
|
3796 |
||
3797 |
bash_process0 :: Bytes -> Bytes -> Bash.Params -> IO Process_Result.T |
|
3798 |
bash_process0 address password params = do |
|
3799 |
Server.connection port password |
|
3800 |
(\socket -> do |
|
3801 |
isabelle_tmp <- getenv "ISABELLE_TMP" |
|
3802 |
Byte_Message.write_message socket (run isabelle_tmp) |
|
3803 |
loop Nothing socket) |
|
3804 |
where |
|
3805 |
port = |
|
74216 | 3806 |
case Bytes.try_unprefix Server.localhost_prefix address of |
74211 | 3807 |
Just port -> make_string port |
3808 |
Nothing -> errorWithoutStackTrace "Bad bash_process server address" |
|
3809 |
||
3810 |
script = Bash.get_script params |
|
3811 |
input = Bash.get_input params |
|
3812 |
cwd = Bash.get_cwd params |
|
3813 |
putenv = Bash.get_putenv params |
|
3814 |
redirect = Bash.get_redirect params |
|
3815 |
timeout = Bash.get_timeout params |
|
3816 |
description = Bash.get_description params |
|
3817 |
||
3818 |
run :: Bytes -> [Bytes] |
|
3819 |
run isabelle_tmp = |
|
3820 |
[Bash.server_run, script, input, |
|
3821 |
YXML.string_of_body (Encode.option (Encode.string . absolute_path) cwd), |
|
3822 |
YXML.string_of_body |
|
3823 |
(Encode.list (Encode.pair Encode.string Encode.string) |
|
3824 |
(("ISABELLE_TMP", isabelle_tmp) : putenv)), |
|
3825 |
Value.print_bool redirect, |
|
3826 |
Value.print_real (Time.get_seconds timeout), |
|
3827 |
description] |
|
3828 |
||
3829 |
kill :: Maybe Bytes -> IO () |
|
3830 |
kill maybe_uuid = do |
|
3831 |
case maybe_uuid of |
|
3832 |
Just uuid -> |
|
3833 |
Server.connection port password (\socket -> |
|
3834 |
Byte_Message.write_message socket [Bash.server_kill, uuid]) |
|
3835 |
Nothing -> return () |
|
3836 |
||
3837 |
err = errorWithoutStackTrace "Malformed result from bash_process server" |
|
3838 |
the = fromMaybe err |
|
3839 |
||
3840 |
loop :: Maybe Bytes -> Socket -> IO Process_Result.T |
|
3841 |
loop maybe_uuid socket = do |
|
3842 |
result <- Isabelle_Thread.bracket_resource (kill maybe_uuid) (Byte_Message.read_message socket) |
|
3843 |
case result of |
|
3844 |
Just [head, uuid] | head == Bash.server_uuid -> loop (Just uuid) socket |
|
3845 |
Just [head] | head == Bash.server_interrupt -> throw UserInterrupt |
|
3846 |
Just [head, msg] | head == Bash.server_failure -> errorWithoutStackTrace $ make_string msg |
|
3847 |
Just (head : a : b : c : d : lines) | head == Bash.server_result -> |
|
3848 |
let |
|
3849 |
rc = the $ Value.parse_int a |
|
3850 |
elapsed = Time.ms $ the $ Value.parse_int b |
|
3851 |
cpu = Time.ms $ the $ Value.parse_int c |
|
3852 |
timing = Timing elapsed cpu Time.zero |
|
3853 |
n = the $ Value.parse_int d |
|
3854 |
out_lines = take n lines |
|
3855 |
err_lines = drop n lines |
|
3856 |
in return $ Process_Result.make rc out_lines err_lines timing |
|
3857 |
_ -> err |
|
3858 |
\<close> |
|
3859 |
||
75068 | 3860 |
generate_file "Isabelle/Cache.hs" = \<open> |
3861 |
{- Title: Isabelle/Cache.hs |
|
3862 |
Author: Makarius |
|
3863 |
LICENSE: BSD 3-clause (Isabelle) |
|
3864 |
||
3865 |
Cache for slow computations. |
|
3866 |
-} |
|
3867 |
||
3868 |
module Isabelle.Cache ( |
|
3869 |
T, init, apply, prune |
|
3870 |
) |
|
3871 |
where |
|
3872 |
||
3873 |
import Prelude hiding (init) |
|
3874 |
import Data.IORef |
|
3875 |
import Data.Map.Strict (Map) |
|
3876 |
import qualified Data.Map.Strict as Map |
|
3877 |
import qualified Data.List as List |
|
3878 |
||
3879 |
import Isabelle.Time (Time) |
|
3880 |
import qualified Isabelle.Time as Time |
|
3881 |
||
3882 |
||
3883 |
data Entry v = Entry {_value :: v, _access :: Time, _timing :: Time} |
|
3884 |
||
3885 |
newtype T k v = Cache (IORef (Map k (Entry v))) |
|
3886 |
||
3887 |
init :: IO (T k v) |
|
3888 |
init = Cache <$> newIORef Map.empty |
|
3889 |
||
3890 |
commit :: Ord k => T k v -> k -> Entry v -> IO v |
|
3891 |
commit (Cache ref) x e = do |
|
3892 |
atomicModifyIORef' ref (\entries -> |
|
3893 |
let |
|
3894 |
entry = |
|
3895 |
case Map.lookup x entries of |
|
3896 |
Just e' | _access e' > _access e -> e' |
|
3897 |
_ -> e |
|
3898 |
in (Map.insert x entry entries, _value entry)) |
|
3899 |
||
75072 | 3900 |
apply :: Ord k => T k v -> k -> IO v -> IO v |
3901 |
apply cache@(Cache ref) x body = do |
|
75068 | 3902 |
start <- Time.now |
3903 |
entries <- readIORef ref |
|
3904 |
case Map.lookup x entries of |
|
3905 |
Just entry -> do |
|
3906 |
commit cache x (entry {_access = start}) |
|
3907 |
Nothing -> do |
|
75072 | 3908 |
y <- body |
75068 | 3909 |
stop <- Time.now |
3910 |
commit cache x (Entry y start (stop - start)) |
|
3911 |
||
3912 |
prune :: Ord k => T k v -> Int -> Time -> IO () |
|
3913 |
prune (Cache ref) max_size min_timing = do |
|
3914 |
atomicModifyIORef' ref (\entries -> |
|
3915 |
let |
|
3916 |
sort = List.sortBy (\(_, e1) (_, e2) -> compare (_access e2) (_access e1)) |
|
79741
513829904beb
more scalable: avoid potentially expensive ordering of underlying key data type, e.g. in MESON.Cache of Naproche;
wenzelm
parents:
78307
diff
changeset
|
3917 |
entries1 = Map.filter (\e -> _timing e >= min_timing) entries |
75068 | 3918 |
entries2 = |
3919 |
if Map.size entries1 <= max_size then entries1 |
|
3920 |
else Map.fromList $ List.take max_size $ sort $ Map.toList entries1 |
|
3921 |
in (entries2, ())) |
|
3922 |
\<close> |
|
3923 |
||
70047
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
wenzelm
parents:
69968
diff
changeset
|
3924 |
export_generated_files _ |
69628 | 3925 |
|
69222 | 3926 |
end |