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