src/Tools/Haskell/Haskell.thy
author wenzelm
Sat, 31 Jul 2021 12:24:13 +0200
changeset 74092 1d26f1a49480
parent 74091 5721f1843e93
child 74093 dc962d4248ca
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69222
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     1
(*  Title:      Tools/Haskell/Haskell.thy
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
     3
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
     4
Support for Isabelle tools in Haskell.
69222
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     5
*)
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     6
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     7
theory Haskell
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     8
  imports Pure
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
     9
begin
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
    10
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    11
generate_file "Isabelle/Bytes.hs" = \<open>
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    12
{-  Title:      Isabelle/Bytes.hs
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    13
    Author:     Makarius
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    14
    LICENSE:    BSD 3-clause (Isabelle)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    15
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    16
Compact byte strings.
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    17
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    18
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/bytes.ML\<close>
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    19
and \<^file>\<open>$ISABELLE_HOME/src/Pure/General/bytes.scala\<close>.
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    20
-}
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    21
74089
wenzelm
parents: 74088
diff changeset
    22
{-# LANGUAGE ScopedTypeVariables #-}
wenzelm
parents: 74088
diff changeset
    23
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    24
module Isabelle.Bytes (
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    25
  Bytes,
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    26
  make, unmake, pack, unpack,
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    27
  empty, null, length, index, all, any,
74090
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    28
  head, last, take, drop, concat, trim_line,
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    29
  singleton
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    30
)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    31
where
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    32
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    33
import Prelude hiding (null, length, all, any, head, last, take, drop, concat)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    34
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    35
import qualified Data.ByteString.Short as ShortByteString
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    36
import Data.ByteString.Short (ShortByteString)
74090
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    37
import qualified Data.ByteString as ByteString
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    38
import Data.ByteString (ByteString)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    39
import qualified Data.List as List
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    40
import Data.Word (Word8)
74090
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    41
import Data.Array (Array, array, (!))
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    42
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    43
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    44
type Bytes = ShortByteString
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    45
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    46
make :: ByteString -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    47
make = ShortByteString.toShort
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    48
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    49
unmake :: Bytes -> ByteString
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    50
unmake = ShortByteString.fromShort
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    51
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    52
pack :: [Word8] -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    53
pack = ShortByteString.pack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    54
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    55
unpack :: Bytes -> [Word8]
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    56
unpack = ShortByteString.unpack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    57
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    58
empty :: Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    59
empty = ShortByteString.empty
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    60
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    61
null :: Bytes -> Bool
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    62
null = ShortByteString.null
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    63
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    64
length :: Bytes -> Int
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    65
length = ShortByteString.length
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    66
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    67
index :: Bytes -> Int -> Word8
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    68
index = ShortByteString.index
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    69
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    70
all :: (Word8 -> Bool) -> Bytes -> Bool
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    71
all p = List.all p . unpack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    72
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    73
any :: (Word8 -> Bool) -> Bytes -> Bool
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    74
any p = List.any p . unpack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    75
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    76
head :: Bytes -> Word8
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    77
head bytes = index bytes 0
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    78
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    79
last :: Bytes -> Word8
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    80
last bytes = index bytes (length bytes - 1)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    81
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    82
take :: Int -> Bytes -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    83
take n = pack . List.take n . unpack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    84
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    85
drop :: Int -> Bytes -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    86
drop n = pack . List.drop n . unpack
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    87
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    88
concat :: [Bytes] -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
    89
concat = mconcat
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
    90
74090
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    91
singletons :: Array Word8 Bytes
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    92
singletons =
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    93
  array (minBound, maxBound) $!
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    94
  [(i, make (ByteString.singleton i)) | i <- [minBound .. maxBound]]
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    95
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    96
singleton :: Word8 -> Bytes
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    97
singleton b = singletons ! b
c26f4ec59835 more operations;
wenzelm
parents: 74089
diff changeset
    98
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
    99
trim_line :: Bytes -> Bytes
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
   100
trim_line s =
74089
wenzelm
parents: 74088
diff changeset
   101
  if n >= 2 && at (n - 2) == '\r' && at (n - 1) == '\n' then take (n - 2) s
wenzelm
parents: 74088
diff changeset
   102
  else if n >= 1 && (at (n - 1) == '\r' || at (n - 1) == '\n') then take (n - 1) s
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
   103
  else s
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
   104
  where
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
   105
    n = length s
74089
wenzelm
parents: 74088
diff changeset
   106
    at :: Int -> Char = toEnum . fromEnum . index s
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   107
\<close>
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   108
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   109
generate_file "Isabelle/UTF8.hs" = \<open>
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   110
{-  Title:      Isabelle/UTF8.hs
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   111
    Author:     Makarius
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   112
    LICENSE:    BSD 3-clause (Isabelle)
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   113
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   114
Variations on UTF-8.
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   115
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   116
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
   117
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
   118
-}
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   119
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   120
{-# LANGUAGE MultiParamTypeClasses #-}
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   121
{-# LANGUAGE TypeSynonymInstances #-}
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   122
{-# LANGUAGE FlexibleInstances #-}
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   123
{-# LANGUAGE InstanceSigs #-}
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   124
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   125
module Isabelle.UTF8 (
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   126
  Recode (..)
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   127
)
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   128
where
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   129
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   130
import Data.Text (Text)
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   131
import qualified Data.Text as Text
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   132
import qualified Data.Text.Encoding as Encoding
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   133
import qualified Data.Text.Encoding.Error as Error
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   134
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   135
import Data.ByteString (ByteString)
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   136
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   137
import qualified Isabelle.Bytes as Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   138
import Isabelle.Bytes (Bytes)
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   139
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   140
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   141
class Recode a b where
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   142
  encode :: a -> b
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   143
  decode :: b -> a
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   144
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   145
instance Recode Text ByteString where
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   146
  encode :: Text -> ByteString
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   147
  encode = Encoding.encodeUtf8
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   148
  decode :: ByteString -> Text
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   149
  decode = Encoding.decodeUtf8With Error.lenientDecode
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   150
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   151
instance Recode Text Bytes where
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   152
  encode :: Text -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   153
  encode = Bytes.make . encode
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   154
  decode :: Bytes -> Text
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   155
  decode = decode . Bytes.unmake
74080
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
instance Recode String ByteString where
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   158
  encode :: String -> ByteString
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   159
  encode = encode . Text.pack
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   160
  decode :: ByteString -> String
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   161
  decode = Text.unpack . decode
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   162
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   163
instance Recode String Bytes where
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   164
  encode :: String -> Bytes
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   165
  encode = encode . Text.pack
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
   166
  decode :: Bytes -> String
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   167
  decode = Text.unpack . decode
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   168
\<close>
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
   169
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   170
generate_file "Isabelle/Library.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
   171
{-  Title:      Isabelle/Library.hs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   172
    Author:     Makarius
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   173
    LICENSE:    BSD 3-clause (Isabelle)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   174
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   175
Basic library of Isabelle idioms.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   176
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   177
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/basics.ML\<close>, \<^file>\<open>$ISABELLE_HOME/src/Pure/library.ML\<close>.
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   178
-}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   179
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   180
{-# LANGUAGE OverloadedStrings #-}
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   181
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   182
module Isabelle.Library (
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   183
  (|>), (|->), (#>), (#->),
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   184
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   185
  the, the_default,
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   186
74087
12c984b7d391 tuned signature;
wenzelm
parents: 74086
diff changeset
   187
  fold, fold_rev, single, map_index, get_index, separate,
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   188
69477
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   189
  proper_string, quote, space_implode, commas, commas_quote, cat_lines,
69453
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   190
  space_explode, split_lines, trim_line, clean_name)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   191
where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   192
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   193
import Data.String (IsString)
69453
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   194
import qualified Data.List.Split as Split
69491
wenzelm
parents: 69490
diff changeset
   195
import qualified Isabelle.Symbol as Symbol
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   196
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   197
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   198
{- functions -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   199
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   200
(|>) :: a -> (a -> b) -> b
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   201
x |> f = f x
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   202
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   203
(|->) :: (a, b) -> (a -> b -> c) -> c
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   204
(x, y) |-> f = f x y
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   205
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   206
(#>) :: (a -> b) -> (b -> c) -> a -> c
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   207
(f #> g) x = x |> f |> g
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   208
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   209
(#->) :: (a -> (c, b)) -> (c -> b -> d) -> a -> d
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   210
(f #-> g) x  = x |> f |-> g
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   211
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   212
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   213
{- options -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   214
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   215
the :: Maybe a -> a
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   216
the (Just x) = x
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   217
the Nothing = error "the Nothing"
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   218
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   219
the_default :: a -> Maybe a -> a
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   220
the_default x Nothing = x
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   221
the_default _ (Just y) = y
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   222
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   223
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   224
{- lists -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   226
fold :: (a -> b -> b) -> [a] -> b -> b
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   227
fold _ [] y = y
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   228
fold f (x : xs) y = fold f xs (f x y)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   229
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   230
fold_rev :: (a -> b -> b) -> [a] -> b -> b
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   231
fold_rev _ [] y = y
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   232
fold_rev f (x : xs) y = f x (fold_rev f xs y)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   233
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   234
single :: a -> [a]
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   235
single x = [x]
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   236
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   237
map_index :: ((Int, a) -> b) -> [a] -> [b]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   238
map_index f = map_aux 0
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   239
  where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   240
    map_aux _ [] = []
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   241
    map_aux i (x : xs) = f (i, x) : map_aux (i + 1) xs
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   242
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   243
get_index :: (a -> Maybe b) -> [a] -> Maybe (Int, b)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   244
get_index f = get_aux 0
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   245
  where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   246
    get_aux _ [] = Nothing
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   247
    get_aux i (x : xs) =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   248
      case f x of
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   249
        Nothing -> get_aux (i + 1) xs
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   250
        Just y -> Just (i, y)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
   251
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   252
separate :: a -> [a] -> [a]
74092
wenzelm
parents: 74091
diff changeset
   253
separate s (x : xs @ (_ : _)) = x : s : separate s xs
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   254
separate _ xs = xs;
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   255
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   256
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   257
{- strings -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   258
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   259
proper_string :: (IsString a, Eq a) => a -> Maybe a
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   260
proper_string s = if s == "" then Nothing else Just s
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   261
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   262
quote :: (IsString a, Semigroup a) => a -> a
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
   263
quote s = "\"" <> s <> "\""
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   264
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   265
space_implode :: (IsString a, Monoid a) => a -> [a] -> a
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   266
space_implode s = mconcat . separate s
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   267
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   268
commas, commas_quote :: (IsString a, Monoid a) => [a] -> a
69453
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   269
commas = space_implode ", "
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   270
commas_quote = commas . map quote
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   271
74086
73487ebd7332 clarified signature;
wenzelm
parents: 74084
diff changeset
   272
cat_lines :: (IsString a, Monoid a) => [a] -> a
69453
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   273
cat_lines = space_implode "\n"
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   274
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   275
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   276
space_explode :: Char -> String -> [String]
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   277
space_explode c = Split.split (Split.dropDelims (Split.whenElt (== c)))
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   278
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   279
split_lines :: String -> [String]
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   280
split_lines = space_explode '\n'
dcea1fffbfe6 more Haskell operations;
wenzelm
parents: 69452
diff changeset
   281
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   282
trim_line :: String -> String
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   283
trim_line line =
69491
wenzelm
parents: 69490
diff changeset
   284
  if not (null line) && Symbol.is_ascii_line_terminator (last line) then
69486
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   285
    case reverse line of
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   286
      '\n' : '\r' : rest -> reverse rest
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   287
      '\r' : rest -> reverse rest
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   288
      '\n' : rest -> reverse rest
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   289
      _ -> line
59ada9f63cc5 proper trim_line according to ML/Scala versions;
wenzelm
parents: 69482
diff changeset
   290
  else line
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   291
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   292
clean_name :: String -> String
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   293
clean_name = reverse #> dropWhile (== '_') #> reverse
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   294
\<close>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   295
74091
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   296
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   297
generate_file "Isabelle/Symbol.hs" = \<open>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   298
{-  Title:      Isabelle/Symbols.hs
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   299
    Author:     Makarius
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   300
    LICENSE:    BSD 3-clause (Isabelle)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   301
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   302
Isabelle text symbols.
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   303
-}
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   304
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   305
module Isabelle.Symbol where
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   306
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   307
{- ASCII characters -}
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   308
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   309
is_ascii_letter :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   310
is_ascii_letter c = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z'
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   311
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   312
is_ascii_digit :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   313
is_ascii_digit c = '0' <= c && c <= '9'
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   314
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   315
is_ascii_hex :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   316
is_ascii_hex c = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f'
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   317
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   318
is_ascii_quasi :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   319
is_ascii_quasi c = c == '_' || c == '\''
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   320
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   321
is_ascii_blank :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   322
is_ascii_blank c = c `elem` " \t\n\11\f\r"
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   323
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   324
is_ascii_line_terminator :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   325
is_ascii_line_terminator c = c == '\r' || c == '\n'
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   326
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   327
is_ascii_letdig :: Char -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   328
is_ascii_letdig c = is_ascii_letter c || is_ascii_digit c || is_ascii_quasi c
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   329
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   330
is_ascii_identifier :: String -> Bool
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   331
is_ascii_identifier s =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   332
  not (null s) && is_ascii_letter (head s) && all is_ascii_letdig s
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   333
\<close>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
   334
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   335
generate_file "Isabelle/Value.hs" = \<open>
73246
b9c480878663 tuned comments;
wenzelm
parents: 73199
diff changeset
   336
{-  Title:      Isabelle/Value.hs
69233
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   337
    Author:     Makarius
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   338
    LICENSE:    BSD 3-clause (Isabelle)
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   339
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   340
Plain values, represented as string.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   341
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   342
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/value.ML\<close>.
69233
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   343
-}
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   344
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   345
module Isabelle.Value
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   346
  (print_bool, parse_bool, parse_nat, print_int, parse_int, print_real, parse_real)
69233
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   347
where
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   348
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   349
import qualified Data.List as List
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   350
import qualified Text.Read as Read
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   351
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   352
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   353
{- bool -}
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   354
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   355
print_bool :: Bool -> String
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   356
print_bool True = "true"
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   357
print_bool False = "false"
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   358
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   359
parse_bool :: String -> Maybe Bool
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   360
parse_bool "true" = Just True
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   361
parse_bool "false" = Just False
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   362
parse_bool _ = Nothing
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   363
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   364
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   365
{- nat -}
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   366
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   367
parse_nat :: String -> Maybe Int
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   368
parse_nat s =
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   369
  case Read.readMaybe s of
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   370
    Just n | n >= 0 -> Just n
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   371
    _ -> Nothing
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   372
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
   373
69233
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   374
{- int -}
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   375
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   376
print_int :: Int -> String
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   377
print_int = show
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   378
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   379
parse_int :: String -> Maybe Int
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   380
parse_int = Read.readMaybe
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   381
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   382
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   383
{- real -}
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   384
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   385
print_real :: Double -> String
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   386
print_real x =
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   387
  let s = show x in
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   388
    case span (/= '.') s of
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   389
      (a, '.' : b) | List.all (== '0') b -> a
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   390
      _ -> s
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   391
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   392
parse_real :: String -> Maybe Double
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   393
parse_real = Read.readMaybe
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   394
\<close>
560263485988 ported from src/Pure/General/value.ML;
wenzelm
parents: 69228
diff changeset
   395
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   396
generate_file "Isabelle/Buffer.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
   397
{-  Title:      Isabelle/Buffer.hs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   398
    Author:     Makarius
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   399
    LICENSE:    BSD 3-clause (Isabelle)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   400
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   401
Efficient text buffers.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   402
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   403
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/buffer.ML\<close>.
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   404
-}
69222
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
   405
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   406
module Isabelle.Buffer (T, empty, add, content)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   407
where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   408
69474
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   409
import Isabelle.Library
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   410
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   411
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   412
newtype T = Buffer [Char]
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   413
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   414
empty :: T
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   415
empty = Buffer []
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   416
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   417
add :: String -> T -> T
69474
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   418
add s (Buffer cs) = Buffer (fold (:) s cs)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   419
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   420
content :: T -> String
69474
2633cf136335 tuned -- more compact;
wenzelm
parents: 69473
diff changeset
   421
content (Buffer cs) = reverse cs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   422
\<close>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   423
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   424
generate_file "Isabelle/Properties.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
   425
{-  Title:      Isabelle/Properties.hs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   426
    Author:     Makarius
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   427
    LICENSE:    BSD 3-clause (Isabelle)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   428
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   429
Property lists.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   430
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   431
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/properties.ML\<close>.
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   432
-}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   433
69477
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   434
module Isabelle.Properties (Entry, T, defined, get, get_value, put, remove)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   435
where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   436
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   437
import qualified Data.List as List
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   438
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   439
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   440
type Entry = (String, String)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   441
type T = [Entry]
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   442
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   443
defined :: T -> String -> Bool
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   444
defined props name = any (\(a, _) -> a == name) props
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   445
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   446
get :: T -> String -> Maybe String
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   447
get props name = List.lookup name props
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   448
69477
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   449
get_value :: (String -> Maybe a) -> T -> String -> Maybe a
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   450
get_value parse props name =
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   451
  case get props name of
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   452
    Nothing -> Nothing
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   453
    Just s -> parse s
1690ba936016 more Isabelle/Haskell operations;
wenzelm
parents: 69476
diff changeset
   454
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   455
put :: Entry -> T -> T
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   456
put entry props = entry : remove (fst entry) props
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   457
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   458
remove :: String -> T -> T
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   459
remove name props =
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   460
  if defined props name then filter (\(a, _) -> a /= name) props
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   461
  else props
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   462
\<close>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   463
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   464
generate_file "Isabelle/Markup.hs" = \<open>
73246
b9c480878663 tuned comments;
wenzelm
parents: 73199
diff changeset
   465
{-  Title:      Isabelle/Markup.hs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   466
    Author:     Makarius
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   467
    LICENSE:    BSD 3-clause (Isabelle)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   468
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   469
Quasi-abstract markup elements.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   470
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   471
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/markup.ML\<close>.
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   472
-}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   473
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
   474
{-# OPTIONS_GHC -fno-warn-missing-signatures #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
   475
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   476
module Isabelle.Markup (
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   477
  T, empty, is_empty, properties,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   478
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   479
  nameN, name, xnameN, xname, kindN,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   480
69315
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   481
  bindingN, binding, entityN, entity, defN, refN,
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   482
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   483
  completionN, completion, no_completionN, no_completion,
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   484
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   485
  lineN, end_lineN, offsetN, end_offsetN, fileN, idN, positionN, position,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   486
69291
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   487
  expressionN, expression,
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   488
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   489
  citationN, citation,
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   490
71489
e8da4a8d364a more complete signature;
wenzelm
parents: 70845
diff changeset
   491
  pathN, path, urlN, url, docN, doc,
69291
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   492
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   493
  markupN, consistentN, unbreakableN, indentN, widthN,
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   494
  blockN, block, breakN, break, fbreakN, fbreak, itemN, item,
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   495
69968
1a400b14fd3a clarified spell-checking (see also 30233285270a);
wenzelm
parents: 69794
diff changeset
   496
  wordsN, words,
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   497
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   498
  tfreeN, tfree, tvarN, tvar, freeN, free, skolemN, skolem, boundN, bound, varN, var,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   499
  numeralN, numeral, literalN, literal, delimiterN, delimiter, inner_stringN, inner_string,
69320
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   500
  inner_cartoucheN, inner_cartouche,
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   501
  token_rangeN, token_range,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   502
  sortingN, sorting, typingN, typing, class_parameterN, class_parameter,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   503
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   504
  antiquotedN, antiquoted, antiquoteN, antiquote,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   505
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   506
  paragraphN, paragraph, text_foldN, text_fold,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   507
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   508
  keyword1N, keyword1, keyword2N, keyword2, keyword3N, keyword3, quasi_keywordN, quasi_keyword,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   509
  improperN, improper, operatorN, operator, stringN, string, alt_stringN, alt_string,
69320
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   510
  verbatimN, verbatim, cartoucheN, cartouche, commentN, comment, comment1N, comment1,
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   511
  comment2N, comment2, comment3N, comment3,
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   512
70667
3cab8dad5b40 obsolete (see 94442fce40a5);
wenzelm
parents: 70047
diff changeset
   513
  forkedN, forked, joinedN, joined, runningN, running, finishedN, finished,
69794
a19fdf64726c tuned signature: proper exports;
wenzelm
parents: 69793
diff changeset
   514
  failedN, failed, canceledN, canceled, initializedN, initialized, finalizedN, finalized,
a19fdf64726c tuned signature: proper exports;
wenzelm
parents: 69793
diff changeset
   515
  consolidatedN, consolidated,
a19fdf64726c tuned signature: proper exports;
wenzelm
parents: 69793
diff changeset
   516
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   517
  writelnN, writeln, stateN, state, informationN, information, tracingN, tracing,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   518
  warningN, warning, legacyN, legacy, errorN, error, reportN, report, no_reportN, no_report,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   519
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   520
  intensifyN, intensify,
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   521
  Output, no_output)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   522
where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   523
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   524
import Prelude hiding (words, error, break)
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   525
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   526
import Isabelle.Library
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   527
import qualified Isabelle.Properties as Properties
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   528
import qualified Isabelle.Value as Value
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   529
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   530
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   531
{- basic markup -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   532
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   533
type T = (String, Properties.T)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   534
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   535
empty :: T
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   536
empty = ("", [])
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   537
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   538
is_empty :: T -> Bool
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   539
is_empty ("", _) = True
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   540
is_empty _ = False
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   541
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   542
properties :: Properties.T -> T -> T
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   543
properties more_props (elem, props) =
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   544
  (elem, fold_rev Properties.put more_props props)
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   545
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   546
markup_elem :: 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
   547
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
   548
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   549
markup_string :: String -> String -> 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
   550
markup_string name prop = \s -> (name, [(prop, s)])
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   551
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   552
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   553
{- misc properties -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   554
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   555
nameN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   556
nameN = \<open>Markup.nameN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   557
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   558
name :: String -> T -> T
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   559
name a = properties [(nameN, a)]
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   560
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   561
xnameN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   562
xnameN = \<open>Markup.xnameN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   563
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   564
xname :: String -> T -> T
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   565
xname a = properties [(xnameN, a)]
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   566
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   567
kindN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   568
kindN = \<open>Markup.kindN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   569
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   570
69315
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   571
{- formal entities -}
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   572
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   573
bindingN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   574
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
   575
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
   576
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
   577
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   578
entityN :: String
69315
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   579
entityN = \<open>Markup.entityN\<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
   580
entity :: String -> String -> T
69315
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   581
entity kind name =
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   582
  (entityN,
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
   583
    (if null name then [] else [(nameN, name)]) <> (if null kind then [] else [(kindN, kind)]))
69315
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   584
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   585
defN :: String
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   586
defN = \<open>Markup.defN\<close>
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   587
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   588
refN :: String
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   589
refN = \<open>Markup.refN\<close>
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   590
fc1a8df3062d more Haskell operations;
wenzelm
parents: 69291
diff changeset
   591
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   592
{- completion -}
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   593
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   594
completionN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   595
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
   596
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
   597
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
   598
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   599
no_completionN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   600
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
   601
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
   602
no_completion = markup_elem no_completionN
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   603
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
   604
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   605
{- position -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   606
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   607
lineN, end_lineN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   608
lineN = \<open>Markup.lineN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   609
end_lineN = \<open>Markup.end_lineN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   610
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   611
offsetN, end_offsetN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   612
offsetN = \<open>Markup.offsetN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   613
end_offsetN = \<open>Markup.end_offsetN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   614
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   615
fileN, idN :: String
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   616
fileN = \<open>Markup.fileN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   617
idN = \<open>Markup.idN\<close>
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   618
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   619
positionN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   620
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
   621
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
   622
position = markup_elem positionN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   623
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   624
69291
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   625
{- expression -}
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   626
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   627
expressionN :: String
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   628
expressionN = \<open>Markup.expressionN\<close>
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   629
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   630
expression :: String -> T
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   631
expression kind = (expressionN, if kind == "" then [] else [(kindN, kind)])
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   632
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   633
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   634
{- citation -}
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   635
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   636
citationN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   637
citationN = \<open>Markup.citationN\<close>
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   638
citation :: 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
   639
citation = markup_string nameN citationN
69291
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   640
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   641
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   642
{- external resources -}
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   643
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   644
pathN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   645
pathN = \<open>Markup.pathN\<close>
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   646
path :: 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
   647
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
   648
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   649
urlN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   650
urlN = \<open>Markup.urlN\<close>
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   651
url :: 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
   652
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
   653
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   654
docN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   655
docN = \<open>Markup.docN\<close>
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   656
doc :: 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
   657
doc = markup_string docN nameN
69291
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   658
36d711008292 more Haskell operations;
wenzelm
parents: 69290
diff changeset
   659
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   660
{- pretty printing -}
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   661
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   662
markupN, consistentN, unbreakableN, indentN :: String
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   663
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
   664
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
   665
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
   666
indentN = \<open>Markup.indentN\<close>
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   667
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   668
widthN :: String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   669
widthN = \<open>Markup.widthN\<close>
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   670
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   671
blockN :: String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   672
blockN = \<open>Markup.blockN\<close>
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   673
block :: Bool -> Int -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   674
block c i =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   675
  (blockN,
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
   676
    (if c then [(consistentN, Value.print_bool c)] else []) <>
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   677
    (if i /= 0 then [(indentN, Value.print_int i)] else []))
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   678
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   679
breakN :: String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   680
breakN = \<open>Markup.breakN\<close>
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   681
break :: Int -> Int -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   682
break w i =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   683
  (breakN,
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
   684
    (if w /= 0 then [(widthN, Value.print_int w)] else []) <>
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   685
    (if i /= 0 then [(indentN, Value.print_int i)] else []))
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   686
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   687
fbreakN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   688
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
   689
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
   690
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
   691
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   692
itemN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   693
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
   694
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
   695
item = markup_elem itemN
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   696
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
   697
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   698
{- text properties -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   699
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   700
wordsN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   701
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
   702
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
   703
words = markup_elem wordsN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   704
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   705
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   706
{- inner syntax -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   707
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   708
tfreeN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   709
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
   710
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
   711
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
   712
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   713
tvarN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   714
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
   715
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
   716
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
   717
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   718
freeN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   719
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
   720
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
   721
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
   722
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   723
skolemN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   724
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
   725
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
   726
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
   727
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   728
boundN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   729
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
   730
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
   731
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
   732
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   733
varN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   734
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
   735
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
   736
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
   737
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   738
numeralN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   739
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
   740
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
   741
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
   742
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   743
literalN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   744
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
   745
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
   746
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
   747
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   748
delimiterN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   749
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
   750
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
   751
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
   752
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   753
inner_stringN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   754
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
   755
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
   756
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
   757
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   758
inner_cartoucheN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   759
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
   760
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
   761
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
   762
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   763
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   764
token_rangeN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   765
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
   766
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
   767
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
   768
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   769
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   770
sortingN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   771
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
   772
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
   773
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
   774
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   775
typingN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   776
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
   777
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
   778
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
   779
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   780
class_parameterN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   781
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
   782
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
   783
class_parameter = markup_elem class_parameterN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   784
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   785
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   786
{- antiquotations -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   787
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   788
antiquotedN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   789
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
   790
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
   791
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
   792
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   793
antiquoteN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   794
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
   795
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
   796
antiquote = markup_elem antiquoteN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   797
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   798
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   799
{- text structure -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   800
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   801
paragraphN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   802
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
   803
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
   804
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
   805
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   806
text_foldN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   807
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
   808
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
   809
text_fold = markup_elem text_foldN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   810
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   811
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   812
{- outer syntax -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   813
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   814
keyword1N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   815
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
   816
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
   817
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
   818
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   819
keyword2N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   820
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
   821
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
   822
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
   823
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   824
keyword3N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   825
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
   826
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
   827
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
   828
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   829
quasi_keywordN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   830
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
   831
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
   832
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
   833
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   834
improperN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   835
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
   836
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
   837
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
   838
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   839
operatorN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   840
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
   841
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
   842
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
   843
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   844
stringN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   845
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
   846
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
   847
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
   848
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   849
alt_stringN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   850
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
   851
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
   852
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
   853
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   854
verbatimN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   855
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
   856
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
   857
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
   858
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   859
cartoucheN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   860
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
   861
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
   862
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
   863
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   864
commentN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   865
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
   866
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
   867
comment = markup_elem commentN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   868
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   869
69320
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   870
{- comments -}
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   871
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   872
comment1N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   873
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
   874
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
   875
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
   876
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   877
comment2N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   878
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
   879
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
   880
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
   881
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   882
comment3N :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   883
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
   884
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
   885
comment3 = markup_elem comment3N
69320
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   886
fc221fa79741 more comment markup;
wenzelm
parents: 69315
diff changeset
   887
69793
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   888
{- command status -}
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   889
70667
3cab8dad5b40 obsolete (see 94442fce40a5);
wenzelm
parents: 70047
diff changeset
   890
forkedN, joinedN, runningN, finishedN, failedN, canceledN,
69793
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   891
  initializedN, finalizedN, consolidatedN :: String
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   892
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
   893
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
   894
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
   895
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
   896
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
   897
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
   898
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
   899
finalizedN = \<open>Markup.finalizedN\<close>
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   900
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
   901
70667
3cab8dad5b40 obsolete (see 94442fce40a5);
wenzelm
parents: 70047
diff changeset
   902
forked, joined, running, finished, failed, canceled,
69793
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   903
  initialized, finalized, consolidated :: T
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   904
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
   905
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
   906
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
   907
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
   908
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
   909
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
   910
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
   911
finalized = markup_elem finalizedN
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   912
consolidated = markup_elem consolidatedN
69793
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   913
0f2dc49250fb more Haskell operations;
wenzelm
parents: 69662
diff changeset
   914
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   915
{- messages -}
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   916
73199
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   917
writelnN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   918
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
   919
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
   920
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
   921
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   922
stateN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   923
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
   924
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
   925
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
   926
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   927
informationN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   928
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
   929
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
   930
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
   931
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   932
tracingN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   933
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
   934
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
   935
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
   936
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   937
warningN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   938
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
   939
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
   940
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
   941
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   942
legacyN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   943
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
   944
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
   945
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
   946
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   947
errorN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   948
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
   949
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
   950
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
   951
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   952
reportN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   953
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
   954
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
   955
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
   956
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   957
no_reportN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   958
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
   959
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
   960
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
   961
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   962
intensifyN :: String
d300574cee4e more robust type signatures, notably for the sake of haskell-stack-trace-plugin-0.1.1.1;
wenzelm
parents: 73178
diff changeset
   963
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
   964
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
   965
intensify = markup_elem intensifyN
69234
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   966
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   967
2dec32c7313f more Haskell operations;
wenzelm
parents: 69233
diff changeset
   968
{- output -}
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   969
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   970
type Output = (String, String)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   971
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   972
no_output :: Output
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   973
no_output = ("", "")
69222
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
   974
\<close>
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
   975
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
   976
generate_file "Isabelle/XML.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
   977
{-  Title:      Isabelle/XML.hs
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   978
    Author:     Makarius
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   979
    LICENSE:    BSD 3-clause (Isabelle)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   980
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   981
Untyped XML trees and representation of ML values.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   982
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
   983
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>.
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   984
-}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   985
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
   986
{-# OPTIONS_GHC -fno-warn-missing-signatures #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
   987
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   988
module Isabelle.XML (Attributes, Body, Tree(..), wrap_elem, unwrap_elem, content_of)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   989
where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   990
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   991
import Isabelle.Library
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   992
import qualified Isabelle.Properties as Properties
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   993
import qualified Isabelle.Markup as Markup
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   994
import qualified Isabelle.Buffer as Buffer
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   995
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   996
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   997
{- types -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   998
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
   999
type Attributes = Properties.T
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1000
type Body = [Tree]
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1001
data Tree = Elem (Markup.T, Body) | Text String
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1002
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1003
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1004
{- wrapped elements -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1005
69482
186b03abb764 tuned signature;
wenzelm
parents: 69480
diff changeset
  1006
wrap_elem :: ((Markup.T, Body), [Tree]) -> Tree
69236
wenzelm
parents: 69234
diff changeset
  1007
wrap_elem (((a, atts), body1), body2) =
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1008
  Elem ((\<open>XML.xml_elemN\<close>, (\<open>XML.xml_nameN\<close>, a) : atts), Elem ((\<open>XML.xml_bodyN\<close>, []), body1) : body2)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1009
69482
186b03abb764 tuned signature;
wenzelm
parents: 69480
diff changeset
  1010
unwrap_elem :: Tree -> Maybe ((Markup.T, Body), [Tree])
69236
wenzelm
parents: 69234
diff changeset
  1011
unwrap_elem
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1012
  (Elem ((\<open>XML.xml_elemN\<close>, (\<open>XML.xml_nameN\<close>, a) : atts), Elem ((\<open>XML.xml_bodyN\<close>, []), body1) : body2)) =
69236
wenzelm
parents: 69234
diff changeset
  1013
  Just (((a, atts), body1), body2)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1014
unwrap_elem _ = Nothing
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1015
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1016
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1017
{- text content -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1018
69482
186b03abb764 tuned signature;
wenzelm
parents: 69480
diff changeset
  1019
add_content :: Tree -> Buffer.T -> Buffer.T
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1020
add_content tree =
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1021
  case unwrap_elem tree of
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1022
    Just (_, ts) -> fold add_content ts
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1023
    Nothing ->
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1024
      case tree of
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1025
        Elem (_, ts) -> fold add_content ts
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1026
        Text s -> Buffer.add s
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1027
69482
186b03abb764 tuned signature;
wenzelm
parents: 69480
diff changeset
  1028
content_of :: Body -> String
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1029
content_of body = Buffer.empty |> fold add_content body |> Buffer.content
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1030
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1031
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1032
{- string representation -}
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1033
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1034
encode '<' = "&lt;"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1035
encode '>' = "&gt;"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1036
encode '&' = "&amp;"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1037
encode '\'' = "&apos;"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1038
encode '\"' = "&quot;"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1039
encode c = [c]
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1040
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1041
instance Show Tree where
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1042
  show tree =
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1043
    Buffer.empty |> show_tree tree |> Buffer.content
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1044
    where
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1045
      show_tree (Elem ((name, atts), [])) =
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1046
        Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add "/>"
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1047
      show_tree (Elem ((name, atts), ts)) =
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1048
        Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add ">" #>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1049
        fold show_tree ts #>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1050
        Buffer.add "</" #> Buffer.add name #> Buffer.add ">"
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1051
      show_tree (Text s) = Buffer.add (show_text s)
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1052
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1053
      show_elem name atts =
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1054
        unwords (name : map (\(a, x) -> a <> "=\"" <> show_text x <> "\"") atts)
69225
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1055
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1056
      show_text = concatMap encode
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1057
\<close>
bf2fecda8383 support for Isabelle tools in Haskell;
wenzelm
parents: 69222
diff changeset
  1058
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1059
generate_file "Isabelle/XML/Encode.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1060
{-  Title:      Isabelle/XML/Encode.hs
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1061
    Author:     Makarius
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1062
    LICENSE:    BSD 3-clause (Isabelle)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1063
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1064
XML as data representation language.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1065
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1066
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>.
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1067
-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1068
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1069
{-# OPTIONS_GHC -fno-warn-missing-signatures #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1070
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1071
module Isabelle.XML.Encode (
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1072
  A, T, V, P,
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1073
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1074
  int_atom, bool_atom, unit_atom,
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1075
71490
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1076
  tree, properties, string, int, bool, unit, pair, triple, list, option, variant
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1077
)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1078
where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1079
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1080
import Isabelle.Library
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1081
import qualified Isabelle.Value as Value
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1082
import qualified Isabelle.Properties as Properties
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1083
import qualified Isabelle.XML as XML
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1084
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1085
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1086
type A a = a -> String
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1087
type T a = a -> XML.Body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1088
type V a = a -> Maybe ([String], XML.Body)
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1089
type P a = a -> [String]
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1090
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1091
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1092
-- atomic values
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1093
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1094
int_atom :: A Int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1095
int_atom = Value.print_int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1096
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1097
bool_atom :: A Bool
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1098
bool_atom False = "0"
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1099
bool_atom True = "1"
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1100
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1101
unit_atom :: A ()
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1102
unit_atom () = ""
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1103
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1104
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1105
-- structural nodes
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1106
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1107
node ts = XML.Elem ((":", []), ts)
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1108
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1109
vector = map_index (\(i, x) -> (int_atom i, x))
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1110
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1111
tagged (tag, (xs, ts)) = XML.Elem ((int_atom tag, vector xs), ts)
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1112
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1113
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1114
-- representation of standard types
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1115
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1116
tree :: T XML.Tree
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1117
tree t = [t]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1118
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1119
properties :: T Properties.T
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1120
properties props = [XML.Elem ((":", props), [])]
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1121
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1122
string :: T String
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1123
string "" = []
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1124
string s = [XML.Text s]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1125
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1126
int :: T Int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1127
int = string . int_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1128
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1129
bool :: T Bool
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1130
bool = string . bool_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1131
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1132
unit :: T ()
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1133
unit = string . unit_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1134
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1135
pair :: T a -> T b -> T (a, b)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1136
pair f g (x, y) = [node (f x), node (g y)]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1137
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1138
triple :: T a -> T b -> T c -> T (a, b, c)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1139
triple f g h (x, y, z) = [node (f x), node (g y), node (h z)]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1140
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1141
list :: T a -> T [a]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1142
list f xs = map (node . f) xs
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1143
71490
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1144
option :: T a -> T (Maybe a)
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1145
option _ Nothing = []
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1146
option f (Just x) = [node (f x)]
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1147
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1148
variant :: [V a] -> T a
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1149
variant fs x = [tagged (the (get_index (\f -> f x) fs))]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1150
\<close>
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1151
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1152
generate_file "Isabelle/XML/Decode.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1153
{-  Title:      Isabelle/XML/Decode.hs
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1154
    Author:     Makarius
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1155
    LICENSE:    BSD 3-clause (Isabelle)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1156
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1157
XML as data representation language.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1158
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1159
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/xml.ML\<close>.
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1160
-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1161
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1162
{-# OPTIONS_GHC -fno-warn-missing-signatures #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1163
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1164
module Isabelle.XML.Decode (
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1165
  A, T, V, P,
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1166
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1167
  int_atom, bool_atom, unit_atom,
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1168
71490
3488c0eb4cc8 more complete signature;
wenzelm
parents: 71489
diff changeset
  1169
  tree, properties, string, int, bool, unit, pair, triple, list, option, variant
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1170
)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1171
where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1172
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1173
import Isabelle.Library
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1174
import qualified Isabelle.Value as Value
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1175
import qualified Isabelle.Properties as Properties
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1176
import qualified Isabelle.XML as XML
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1177
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1178
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1179
type A a = String -> a
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1180
type T a = XML.Body -> a
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1181
type V a = ([String], XML.Body) -> a
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1182
type P a = [String] -> a
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1183
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1184
err_atom = error "Malformed XML atom"
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1185
err_body = error "Malformed XML body"
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1186
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1187
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1188
{- atomic values -}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1189
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1190
int_atom :: A Int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1191
int_atom s =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1192
  case Value.parse_int s of
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1193
    Just i -> i
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1194
    Nothing -> err_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1195
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1196
bool_atom :: A Bool
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1197
bool_atom "0" = False
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1198
bool_atom "1" = True
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1199
bool_atom _ = err_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1200
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1201
unit_atom :: A ()
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1202
unit_atom "" = ()
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1203
unit_atom _ = err_atom
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1204
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1205
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1206
{- structural nodes -}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1207
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1208
node (XML.Elem ((":", []), ts)) = ts
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1209
node _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1210
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1211
vector atts =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1212
  map_index (\(i, (a, x)) -> if int_atom a == i then x else err_atom) atts
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1213
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1214
tagged (XML.Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts))
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1215
tagged _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1216
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1217
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1218
{- representation of standard types -}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1219
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1220
tree :: T XML.Tree
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1221
tree [t] = t
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1222
tree _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1223
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1224
properties :: T Properties.T
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1225
properties [XML.Elem ((":", props), [])] = props
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1226
properties _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1227
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1228
string :: T String
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1229
string [] = ""
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1230
string [XML.Text s] = s
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1231
string _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1232
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1233
int :: T Int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1234
int = int_atom . string
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1235
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1236
bool :: T Bool
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1237
bool = bool_atom . string
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1238
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1239
unit :: T ()
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1240
unit = unit_atom . string
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1241
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1242
pair :: T a -> T b -> T (a, b)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1243
pair f g [t1, t2] = (f (node t1), g (node t2))
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1244
pair _ _ _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1245
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1246
triple :: T a -> T b -> T c -> T (a, b, c)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1247
triple f g h [t1, t2, t3] = (f (node t1), g (node t2), h (node t3))
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1248
triple _ _ _ _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1249
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1250
list :: T a -> T [a]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1251
list f ts = map (f . node) ts
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1252
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1253
option :: T a -> T (Maybe a)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1254
option _ [] = Nothing
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1255
option f [t] = Just (f (node t))
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1256
option _ _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1257
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1258
variant :: [V a] -> T a
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1259
variant fs [t] = (fs !! tag) (xs, ts)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1260
  where (tag, (xs, ts)) = tagged t
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1261
variant _ _ = err_body
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1262
\<close>
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1263
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1264
generate_file "Isabelle/YXML.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1265
{-  Title:      Isabelle/YXML.hs
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1266
    Author:     Makarius
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1267
    LICENSE:    BSD 3-clause (Isabelle)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1268
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1269
Efficient text representation of XML trees.  Suitable for direct
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1270
inlining into plain text.
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1271
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1272
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/yxml.ML\<close>.
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1273
-}
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1274
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1275
{-# OPTIONS_GHC -fno-warn-missing-signatures -fno-warn-incomplete-patterns #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1276
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1277
module Isabelle.YXML (charX, charY, strX, strY, detect, output_markup,
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1278
  buffer_body, buffer, string_of_body, string_of, parse_body, parse)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1279
where
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1280
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1281
import qualified Data.Char as Char
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1282
import qualified Data.List as List
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1283
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1284
import Isabelle.Library
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1285
import qualified Isabelle.Markup as Markup
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1286
import qualified Isabelle.XML as XML
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1287
import qualified Isabelle.Buffer as Buffer
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1289
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1290
{- markers -}
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1291
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1292
charX, charY :: Char
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1293
charX = Char.chr 5
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1294
charY = Char.chr 6
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1295
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1296
strX, strY, strXY, strXYX :: String
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1297
strX = [charX]
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1298
strY = [charY]
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1299
strXY = strX <> strY
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1300
strXYX = strXY <> strX
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1301
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1302
detect :: String -> Bool
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1303
detect = any (\c -> c == charX || c == charY)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1304
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1305
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1306
{- output -}
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1307
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1308
output_markup :: Markup.T -> Markup.Output
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1309
output_markup markup@(name, atts) =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1310
  if Markup.is_empty markup then Markup.no_output
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1311
  else (strXY <> name <> concatMap (\(a, x) -> strY <> a <> "=" <> x) atts <> strX, strXYX)
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1312
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1313
buffer_attrib (a, x) =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1314
  Buffer.add strY #> Buffer.add a #> Buffer.add "=" #> Buffer.add x
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1315
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1316
buffer_body :: XML.Body -> Buffer.T -> Buffer.T
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1317
buffer_body = fold buffer
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1318
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1319
buffer :: XML.Tree -> Buffer.T -> Buffer.T
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1320
buffer (XML.Elem ((name, atts), ts)) =
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1321
  Buffer.add strXY #> Buffer.add name #> fold buffer_attrib atts #> Buffer.add strX #>
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1322
  buffer_body ts #>
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1323
  Buffer.add strXYX
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1324
buffer (XML.Text s) = Buffer.add s
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1325
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1326
string_of_body :: XML.Body -> String
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1327
string_of_body body = Buffer.empty |> buffer_body body |> Buffer.content
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1328
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1329
string_of :: XML.Tree -> String
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1330
string_of = string_of_body . single
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1331
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1332
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1333
{- parse -}
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1334
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1335
-- split: fields or non-empty tokens
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1336
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1337
split :: Bool -> Char -> String -> [String]
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1338
split _ _ [] = []
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1339
split fields sep str = splitting str
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1340
  where
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1341
    splitting rest =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1342
      case span (/= sep) rest of
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1343
        (_, []) -> cons rest []
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1344
        (prfx, _ : rest') -> cons prfx (splitting rest')
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1345
    cons item = if fields || not (null item) then (:) item else id
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1346
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1347
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1348
-- structural errors
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1349
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1350
err msg = error ("Malformed YXML: " <> msg)
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1351
err_attribute = err "bad attribute"
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1352
err_element = err "bad element"
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1353
err_unbalanced "" = err "unbalanced element"
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1354
err_unbalanced name = err ("unbalanced element " <> quote name)
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1355
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1356
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1357
-- stack operations
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1358
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1359
add x ((elem, body) : pending) = (elem, x : body) : pending
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1360
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1361
push "" _ _ = err_element
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1362
push name atts pending = ((name, atts), []) : pending
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1363
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1364
pop ((("", _), _) : _) = err_unbalanced ""
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1365
pop ((markup, body) : pending) = add (XML.Elem (markup, reverse body)) pending
69288
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1366
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1367
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1368
-- parsing
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1369
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1370
parse_attrib s =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1371
  case List.elemIndex '=' s of
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1372
    Just i | i > 0 -> (take i s, drop (i + 1) s)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1373
    _ -> err_attribute
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1374
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1375
parse_chunk ["", ""] = pop
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1376
parse_chunk ("" : name : atts) = push name (map parse_attrib atts)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1377
parse_chunk txts = fold (add . XML.Text) txts
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1378
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1379
parse_body :: String -> XML.Body
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1380
parse_body source =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1381
  case fold parse_chunk chunks [(("", []), [])] of
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1382
    [(("", _), result)] -> reverse result
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1383
    ((name, _), _) : _ -> err_unbalanced name
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1384
  where chunks = split False charX source |> map (split True charY)
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1385
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1386
parse :: String -> XML.Tree
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1387
parse source =
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1388
  case parse_body source of
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1389
    [result] -> result
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1390
    [] -> XML.Text ""
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1391
    _ -> err "multiple results"
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1392
\<close>
4c3704ecb0e6 more Haskell operations;
wenzelm
parents: 69287
diff changeset
  1393
74091
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1394
generate_file "Isabelle/Completion.hs" = \<open>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1395
{-  Title:      Isabelle/Completion.hs
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1396
    Author:     Makarius
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1397
    LICENSE:    BSD 3-clause (Isabelle)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1398
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1399
Completion of names.
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1400
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1401
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/completion.ML\<close>.
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1402
-}
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1403
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1404
module Isabelle.Completion (
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1405
    Name, T, names, none, make, markup_element, markup_report, make_report
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1406
  ) where
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1407
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1408
import qualified Data.List as List
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1409
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1410
import Isabelle.Library
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1411
import qualified Isabelle.Properties as Properties
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1412
import qualified Isabelle.Markup as Markup
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1413
import qualified Isabelle.XML.Encode as Encode
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1414
import qualified Isabelle.XML as XML
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1415
import qualified Isabelle.YXML as YXML
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1416
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1417
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1418
type Name = (String, (String, String))  -- external name, kind, internal name
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1419
data T = Completion Properties.T Int [Name]  -- position, total length, names
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1420
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1421
names :: Int -> Properties.T -> [Name] -> T
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1422
names limit props names = Completion props (length names) (take limit names)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1423
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1424
none :: T
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1425
none = names 0 [] []
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1426
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1427
make :: Int -> (String, Properties.T) -> ((String -> Bool) -> [Name]) -> T
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1428
make limit (name, props) make_names =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1429
  if name /= "" && name /= "_"
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1430
  then names limit props (make_names $ List.isPrefixOf $ clean_name name)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1431
  else none
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1432
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1433
markup_element :: T -> (Markup.T, XML.Body)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1434
markup_element (Completion props total names) =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1435
  if not (null names) then
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1436
    let
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1437
      markup = Markup.properties props Markup.completion
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1438
      body =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1439
        Encode.pair Encode.int
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1440
          (Encode.list (Encode.pair Encode.string (Encode.pair Encode.string Encode.string)))
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1441
          (total, names)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1442
    in (markup, body)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1443
  else (Markup.empty, [])
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1444
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1445
markup_report :: [T] -> String
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1446
markup_report [] = ""
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1447
markup_report elems =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1448
  YXML.string_of $ XML.Elem (Markup.report, map (XML.Elem . markup_element) elems)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1449
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1450
make_report :: Int -> (String, Properties.T) -> ((String -> Bool) -> [Name]) -> String
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1451
make_report limit name_props make_names =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1452
  markup_report [make limit name_props make_names]
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1453
\<close>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1454
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1455
generate_file "Isabelle/File.hs" = \<open>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1456
{-  Title:      Isabelle/File.hs
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1457
    Author:     Makarius
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1458
    LICENSE:    BSD 3-clause (Isabelle)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1459
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1460
File-system operations.
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1461
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1462
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/file.ML\<close>.
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1463
-}
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1464
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1465
module Isabelle.File (setup, read, write, append) where
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1466
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1467
import Prelude hiding (read)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1468
import qualified System.IO as IO
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1469
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1470
setup :: IO.Handle -> IO ()
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1471
setup h = do
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1472
  IO.hSetEncoding h IO.utf8
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1473
  IO.hSetNewlineMode h IO.noNewlineTranslation
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1474
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1475
read :: IO.FilePath -> IO String
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1476
read path =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1477
  IO.withFile path IO.ReadMode (\h ->
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1478
    do setup h; IO.hGetContents h >>= \s -> length s `seq` return s)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1479
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1480
write :: IO.FilePath -> String -> IO ()
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1481
write path s =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1482
  IO.withFile path IO.WriteMode (\h -> do setup h; IO.hPutStr h s)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1483
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1484
append :: IO.FilePath -> String -> IO ()
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1485
append path s =
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1486
  IO.withFile path IO.AppendMode (\h -> do setup h; IO.hPutStr h s)
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1487
\<close>
5721f1843e93 clarified order of modules;
wenzelm
parents: 74090
diff changeset
  1488
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1489
generate_file "Isabelle/Pretty.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1490
{-  Title:      Isabelle/Pretty.hs
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1491
    Author:     Makarius
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1492
    LICENSE:    BSD 3-clause (Isabelle)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1493
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1494
Generic pretty printing module.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1495
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1496
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/General/pretty.ML\<close>.
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1497
-}
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1498
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1499
{-# OPTIONS_GHC -fno-warn-missing-signatures #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1500
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1501
module Isabelle.Pretty (
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1502
  T, symbolic, formatted, unformatted,
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1503
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1504
  str, brk_indent, brk, fbrk, breaks, fbreaks, blk, block, strs, markup, mark, mark_str, marks_str,
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1505
  item, text_fold, keyword1, keyword2, text, paragraph, para, quote, cartouche, separate,
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1506
  commas, enclose, enum, list, str_list, big_list)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1507
where
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1508
74087
12c984b7d391 tuned signature;
wenzelm
parents: 74086
diff changeset
  1509
import Isabelle.Library hiding (quote, separate, commas)
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1510
import qualified Data.List as List
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1511
import qualified Isabelle.Buffer as Buffer
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1512
import qualified Isabelle.Markup as Markup
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1513
import qualified Isabelle.XML as XML
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1514
import qualified Isabelle.YXML as YXML
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1515
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1516
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1517
data T =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1518
    Block Markup.T Bool Int [T]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1519
  | Break Int Int
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1520
  | Str String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1521
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1522
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1523
{- output -}
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1524
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1525
output_spaces n = replicate n ' '
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1526
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1527
symbolic_text "" = []
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1528
symbolic_text s = [XML.Text s]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1529
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1530
symbolic_markup markup body =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1531
  if Markup.is_empty markup then body
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1532
  else [XML.Elem (markup, body)]
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1533
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1534
symbolic :: T -> XML.Body
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1535
symbolic (Block markup consistent indent prts) =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1536
  concatMap symbolic prts
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1537
  |> symbolic_markup block_markup
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1538
  |> symbolic_markup markup
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1539
  where block_markup = if null prts then Markup.empty else Markup.block consistent indent
69290
fb77612d11eb tuned signature;
wenzelm
parents: 69289
diff changeset
  1540
symbolic (Break wd ind) = [XML.Elem (Markup.break wd ind, symbolic_text (output_spaces wd))]
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1541
symbolic (Str s) = symbolic_text s
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1542
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1543
formatted :: T -> String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1544
formatted = YXML.string_of_body . symbolic
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1545
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1546
unformatted :: T -> String
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1547
unformatted prt = Buffer.empty |> out prt |> Buffer.content
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1548
  where
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1549
    out (Block markup _ _ prts) =
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1550
      let (bg, en) = YXML.output_markup markup
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1551
      in Buffer.add bg #> fold out prts #> Buffer.add en
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1552
    out (Break _ wd) = Buffer.add (output_spaces wd)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1553
    out (Str s) = Buffer.add s
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1554
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1555
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1556
{- derived operations to create formatting expressions -}
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1557
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1558
force_nat n | n < 0 = 0
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1559
force_nat n = n
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1560
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1561
str :: String -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1562
str = Str
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1563
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1564
brk_indent :: Int -> Int -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1565
brk_indent wd ind = Break (force_nat wd) ind
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1566
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1567
brk :: Int -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1568
brk wd = brk_indent wd 0
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1569
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1570
fbrk :: T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1571
fbrk = str "\n"
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1572
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1573
breaks, fbreaks :: [T] -> [T]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1574
breaks = List.intersperse (brk 1)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1575
fbreaks = List.intersperse fbrk
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1576
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1577
blk :: (Int, [T]) -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1578
blk (indent, es) = Block Markup.empty False (force_nat indent) es
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1579
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1580
block :: [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1581
block prts = blk (2, prts)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1582
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1583
strs :: [String] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1584
strs = block . breaks . map str
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1585
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1586
markup :: Markup.T -> [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1587
markup m = Block m False 0
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1588
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1589
mark :: Markup.T -> T -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1590
mark m prt = if m == Markup.empty then prt else markup m [prt]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1591
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1592
mark_str :: (Markup.T, String) -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1593
mark_str (m, s) = mark m (str s)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1594
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1595
marks_str :: ([Markup.T], String) -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1596
marks_str (ms, s) = fold_rev mark ms (str s)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1597
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1598
item :: [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1599
item = markup Markup.item
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1600
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1601
text_fold :: [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1602
text_fold = markup Markup.text_fold
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1603
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1604
keyword1, keyword2 :: String -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1605
keyword1 name = mark_str (Markup.keyword1, name)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1606
keyword2 name = mark_str (Markup.keyword2, name)
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1607
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1608
text :: String -> [T]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1609
text = breaks . map str . words
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1610
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1611
paragraph :: [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1612
paragraph = markup Markup.paragraph
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1613
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1614
para :: String -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1615
para = paragraph . text
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1616
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1617
quote :: T -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1618
quote prt = blk (1, [str "\"", prt, str "\""])
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1619
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1620
cartouche :: T -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1621
cartouche prt = blk (1, [str "\92<open>", prt, str "\92<close>"])
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1622
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1623
separate :: String -> [T] -> [T]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1624
separate sep = List.intercalate [str sep, brk 1] . map single
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1625
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1626
commas :: [T] -> [T]
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1627
commas = separate ","
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1628
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1629
enclose :: String -> String -> [T] -> T
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1630
enclose lpar rpar prts = block (str lpar : prts <> [str rpar])
69248
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1631
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1632
enum :: String -> String -> String -> [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1633
enum sep lpar rpar = enclose lpar rpar . separate sep
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1634
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1635
list :: String -> String -> [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1636
list = enum ","
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1637
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1638
str_list :: String -> String -> [String] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1639
str_list lpar rpar = list lpar rpar . map str
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1640
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1641
big_list :: String -> [T] -> T
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1642
big_list name prts = block (fbreaks (str name : prts))
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1643
\<close>
9f21381600e3 more Haskell operations;
wenzelm
parents: 69240
diff changeset
  1644
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1645
generate_file "Isabelle/Term.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1646
{-  Title:      Isabelle/Term.hs
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1647
    Author:     Makarius
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1648
    LICENSE:    BSD 3-clause (Isabelle)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1649
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1650
Lambda terms, types, sorts.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1651
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1652
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/term.scala\<close>.
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1653
-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1654
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1655
module Isabelle.Term (
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1656
  Indexname,
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1657
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1658
  Sort, dummyS,
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1659
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1660
  Typ(..), dummyT, is_dummyT, Term(..))
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1661
where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1662
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1663
type Indexname = (String, Int)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1664
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1665
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1666
type Sort = [String]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1667
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1668
dummyS :: Sort
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1669
dummyS = [""]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1670
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1671
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1672
data Typ =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1673
    Type (String, [Typ])
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1674
  | TFree (String, Sort)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1675
  | TVar (Indexname, Sort)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1676
  deriving Show
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1677
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1678
dummyT :: Typ
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1679
dummyT = Type (\<open>\<^type_name>\<open>dummy\<close>\<close>, [])
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1680
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1681
is_dummyT :: Typ -> Bool
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1682
is_dummyT (Type (\<open>\<^type_name>\<open>dummy\<close>\<close>, [])) = True
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1683
is_dummyT _ = False
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1684
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1685
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1686
data Term =
70784
799437173553 Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents: 70667
diff changeset
  1687
    Const (String, [Typ])
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1688
  | Free (String, Typ)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1689
  | Var (Indexname, Typ)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1690
  | Bound Int
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1691
  | Abs (String, Typ, Term)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1692
  | App (Term, Term)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1693
  deriving Show
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1694
\<close>
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1695
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1696
generate_file "Isabelle/Term_XML/Encode.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1697
{-  Title:      Isabelle/Term_XML/Encode.hs
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1698
    Author:     Makarius
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1699
    LICENSE:    BSD 3-clause (Isabelle)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1700
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1701
XML data representation of lambda terms.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1702
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1703
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>.
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1704
-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1705
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1706
{-# LANGUAGE LambdaCase #-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1707
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1708
module Isabelle.Term_XML.Encode (indexname, sort, typ, typ_body, term)
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1709
where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1710
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1711
import Isabelle.Library
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1712
import Isabelle.XML.Encode
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1713
import Isabelle.Term
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1714
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1715
indexname :: P Indexname
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1716
indexname (a, b) = if b == 0 then [a] else [a, int_atom b]
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1717
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1718
sort :: T Sort
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1719
sort = list string
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1720
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1721
typ :: T Typ
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1722
typ ty =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1723
  ty |> variant
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1724
   [\case { Type (a, b) -> Just ([a], list typ b); _ -> Nothing },
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1725
    \case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing },
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1726
    \case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }]
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1727
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1728
typ_body :: T Typ
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1729
typ_body ty = if is_dummyT ty then [] else typ ty
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1730
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1731
term :: T Term
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1732
term t =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1733
  t |> variant
70784
799437173553 Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents: 70667
diff changeset
  1734
   [\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing },
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1735
    \case { Free (a, b) -> Just ([a], typ_body b); _ -> Nothing },
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1736
    \case { Var (a, b) -> Just (indexname a, typ_body b); _ -> Nothing },
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1737
    \case { Bound a -> Just ([], int a); _ -> Nothing },
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1738
    \case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing },
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1739
    \case { App a -> Just ([], pair term term a); _ -> Nothing }]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1740
\<close>
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1741
69444
c3c9440cbf9b more formal Haskell project setup, with dependencies on packages from "stackage";
wenzelm
parents: 69381
diff changeset
  1742
generate_file "Isabelle/Term_XML/Decode.hs" = \<open>
69445
bff0011cdf42 tuned headers;
wenzelm
parents: 69444
diff changeset
  1743
{-  Title:      Isabelle/Term_XML/Decode.hs
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1744
    Author:     Makarius
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1745
    LICENSE:    BSD 3-clause (Isabelle)
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1746
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1747
XML data representation of lambda terms.
69280
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1748
e1d01b351724 more formal references;
wenzelm
parents: 69278
diff changeset
  1749
See also \<^file>\<open>$ISABELLE_HOME/src/Pure/term_xml.ML\<close>.
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1750
-}
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1751
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1752
{-# OPTIONS_GHC -fno-warn-incomplete-patterns #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1753
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1754
module Isabelle.Term_XML.Decode (indexname, sort, typ, typ_body, term)
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1755
where
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1756
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1757
import Isabelle.Library
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1758
import Isabelle.XML.Decode
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1759
import Isabelle.Term
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1760
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1761
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1762
indexname :: P Indexname
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1763
indexname [a] = (a, 0)
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1764
indexname [a, b] = (a, int_atom b)
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1765
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1766
sort :: T Sort
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1767
sort = list string
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1768
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1769
typ :: T Typ
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1770
typ ty =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1771
  ty |> variant
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1772
  [\([a], b) -> Type (a, list typ b),
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1773
   \([a], b) -> TFree (a, sort b),
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1774
   \(a, b) -> TVar (indexname a, sort b)]
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1775
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1776
typ_body :: T Typ
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1777
typ_body [] = dummyT
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1778
typ_body body = typ body
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1779
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1780
term :: T Term
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1781
term t =
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1782
  t |> variant
70784
799437173553 Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents: 70667
diff changeset
  1783
   [\([a], b) -> Const (a, list typ b),
70845
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1784
    \([a], b) -> Free (a, typ_body b),
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1785
    \(a, b) -> Var (indexname a, typ_body b),
8e51ea8d4609 adapted to ML version;
wenzelm
parents: 70784
diff changeset
  1786
    \([], a) -> Bound (int a),
69240
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1787
    \([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d),
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1788
    \([], a) -> App (pair term term a)]
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1789
\<close>
16ca270090b6 more Haskell operations;
wenzelm
parents: 69236
diff changeset
  1790
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1791
generate_file "Isabelle/UUID.hs" = \<open>
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1792
{-  Title:      Isabelle/UUID.hs
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1793
    Author:     Makarius
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1794
    LICENSE:    BSD 3-clause (Isabelle)
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1795
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1796
Universally unique identifiers.
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1797
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1798
See \<^file>\<open>$ISABELLE_HOME/src/Pure/General/uuid.scala\<close>.
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1799
-}
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1800
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1801
module Isabelle.UUID (
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1802
    T,
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1803
    parse_string, parse_bytes,
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1804
    string, bytes,
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1805
    random, random_string, random_bytes
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1806
  )
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1807
where
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1808
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1809
import Data.UUID (UUID)
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1810
import qualified Data.UUID as UUID
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1811
import Data.UUID.V4 (nextRandom)
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1812
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1813
import qualified Isabelle.Bytes as Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1814
import Isabelle.Bytes (Bytes)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1815
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1816
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  1817
type T = UUID
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  1818
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1819
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1820
{- parse -}
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1821
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1822
parse_string :: String -> Maybe T
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1823
parse_string = UUID.fromString
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1824
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1825
parse_bytes :: Bytes -> Maybe T
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1826
parse_bytes = UUID.fromASCIIBytes . Bytes.unmake
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1827
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1828
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1829
{- print -}
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1830
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1831
string :: T -> String
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1832
string = UUID.toString
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1833
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1834
bytes :: T -> Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1835
bytes = Bytes.make . UUID.toASCIIBytes
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1836
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1837
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1838
{- random id -}
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1839
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  1840
random :: IO T
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1841
random = nextRandom
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1842
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1843
random_string :: IO String
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1844
random_string = string <$> random
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1845
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1846
random_bytes :: IO Bytes
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  1847
random_bytes = bytes <$> random
69459
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1848
\<close>
bbb61a9cb99a more Haskell operations;
wenzelm
parents: 69455
diff changeset
  1849
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1850
generate_file "Isabelle/Byte_Message.hs" = \<open>
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1851
{-  Title:      Isabelle/Byte_Message.hs
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1852
    Author:     Makarius
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1853
    LICENSE:    BSD 3-clause (Isabelle)
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1854
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1855
Byte-oriented messages.
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1856
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1857
See \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.ML\<close>
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1858
and \<^file>\<open>$ISABELLE_HOME/src/Pure/PIDE/byte_message.scala\<close>.
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1859
-}
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1860
74083
wenzelm
parents: 74082
diff changeset
  1861
{-# LANGUAGE OverloadedStrings #-}
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1862
{-# OPTIONS_GHC -fno-warn-incomplete-patterns #-}
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1863
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1864
module Isabelle.Byte_Message (
69467
e8893c893241 tuned signature;
wenzelm
parents: 69466
diff changeset
  1865
    write, write_line,
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
  1866
    read, read_block, read_line,
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1867
    make_message, write_message, read_message,
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  1868
    make_line_message, write_line_message, read_line_message,
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  1869
    read_yxml, write_yxml
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1870
  )
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1871
where
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1872
69449
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1873
import Prelude hiding (read)
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1874
import Data.Maybe
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1875
import qualified Data.ByteString as ByteString
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1876
import qualified Isabelle.Bytes as Bytes
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1877
import Isabelle.Bytes (Bytes)
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  1878
import qualified Isabelle.UTF8 as UTF8
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  1879
import qualified Isabelle.XML as XML
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  1880
import qualified Isabelle.YXML as YXML
73177
9288ac2eda12 fewer warnings, notably in Naproche-SAD;
wenzelm
parents: 71692
diff changeset
  1881
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1882
import Network.Socket (Socket)
74082
f81d2a1cad69 tuned signature;
wenzelm
parents: 74081
diff changeset
  1883
import qualified Network.Socket.ByteString as Socket
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1884
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
  1885
import Isabelle.Library
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1886
import qualified Isabelle.Value as Value
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1887
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1888
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1889
{- output operations -}
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1890
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1891
write :: Socket -> [Bytes] -> IO ()
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1892
write socket = Socket.sendMany socket . map Bytes.unmake
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1893
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1894
write_line :: Socket -> Bytes -> IO ()
74083
wenzelm
parents: 74082
diff changeset
  1895
write_line socket s = write socket [s, "\n"]
69467
e8893c893241 tuned signature;
wenzelm
parents: 69466
diff changeset
  1896
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1897
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1898
{- input operations -}
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1899
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1900
read :: Socket -> Int -> IO Bytes
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1901
read socket n = read_body 0 []
69449
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1902
  where
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1903
    result = Bytes.concat . reverse
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1904
    read_body len ss =
69449
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1905
      if len >= n then return (result ss)
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1906
      else
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1907
        (do
74082
f81d2a1cad69 tuned signature;
wenzelm
parents: 74081
diff changeset
  1908
          s <- Socket.recv socket (min (n - len) 8192)
69449
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1909
          case ByteString.length s of
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1910
            0 -> return (result ss)
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1911
            m -> read_body (len + m) (Bytes.make s : ss))
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1912
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1913
read_block :: Socket -> Int -> IO (Maybe Bytes, Int)
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1914
read_block socket n = do
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1915
  msg <- read socket n
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1916
  let len = Bytes.length msg
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1917
  return (if len == n then Just msg else Nothing, len)
69449
b516fdf8005c more uniform multi-language operations;
wenzelm
parents: 69448
diff changeset
  1918
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1919
read_line :: Socket -> IO (Maybe Bytes)
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1920
read_line socket = read_body []
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1921
  where
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
  1922
    result = Bytes.trim_line . Bytes.pack . reverse
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1923
    read_body bs = do
74082
f81d2a1cad69 tuned signature;
wenzelm
parents: 74081
diff changeset
  1924
      s <- Socket.recv socket 1
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1925
      case ByteString.length s of
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1926
        0 -> return (if null bs then Nothing else Just (result bs))
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1927
        1 ->
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1928
          case ByteString.head s of
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1929
            10 -> return (Just (result bs))
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1930
            b -> read_body (b : bs)
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1931
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1932
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1933
{- messages with multiple chunks (arbitrary content) -}
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1934
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1935
make_header :: [Int] -> [Bytes]
74083
wenzelm
parents: 74082
diff changeset
  1936
make_header ns = [UTF8.encode (space_implode "," (map Value.print_int ns)), "\n"]
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1937
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1938
make_message :: [Bytes] -> [Bytes]
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1939
make_message chunks = make_header (map Bytes.length chunks) <> chunks
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1940
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1941
write_message :: Socket -> [Bytes] -> IO ()
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1942
write_message socket = write socket . make_message
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1943
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1944
parse_header :: Bytes -> [Int]
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1945
parse_header line =
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1946
  let
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  1947
    res = map Value.parse_nat (space_explode ',' (UTF8.decode line))
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1948
  in
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1949
    if all isJust res then map fromJust res
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1950
    else error ("Malformed message header: " <> quote (UTF8.decode line))
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1951
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1952
read_chunk :: Socket -> Int -> IO Bytes
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1953
read_chunk socket n = do
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1954
  res <- read_block socket n
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1955
  return $
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1956
    case res of
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1957
      (Just chunk, _) -> chunk
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1958
      (Nothing, len) ->
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1959
        error ("Malformed message chunk: unexpected EOF after " <>
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1960
          show len <> " of " <> show n <> " bytes")
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1961
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1962
read_message :: Socket -> IO (Maybe [Bytes])
69454
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1963
read_message socket = do
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1964
  res <- read_line socket
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1965
  case res of
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1966
    Just line -> Just <$> mapM (read_chunk socket) (parse_header line)
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1967
    Nothing -> return Nothing
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1968
ef051edd4d10 more uniform multi-language operations;
wenzelm
parents: 69453
diff changeset
  1969
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1970
-- hybrid messages: line or length+block (with content restriction)
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1971
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1972
is_length :: Bytes -> Bool
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1973
is_length msg =
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1974
  not (Bytes.null msg) && Bytes.all (\b -> 48 <= b && b <= 57) msg
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1975
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1976
is_terminated :: Bytes -> Bool
69452
704915cf59fa more uniform multi-language operations;
wenzelm
parents: 69449
diff changeset
  1977
is_terminated msg =
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1978
  not (Bytes.null msg) && (Bytes.last msg == 13 || Bytes.last msg == 10)
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1979
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1980
make_line_message :: Bytes -> [Bytes]
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1981
make_line_message msg =
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1982
  let n = Bytes.length msg in
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1983
    if is_length msg || is_terminated msg then
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  1984
      error ("Bad content for line message:\n" <> take 100 (UTF8.decode msg))
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1985
    else
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1986
      (if n > 100 || Bytes.any (== 10) msg then make_header [n + 1] else []) <> [msg, "\n"]
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1987
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1988
write_line_message :: Socket -> Bytes -> IO ()
69476
d93fe3557a98 clarified signature;
wenzelm
parents: 69474
diff changeset
  1989
write_line_message socket = write socket . make_line_message
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1990
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  1991
read_line_message :: Socket -> IO (Maybe Bytes)
69448
51e696887b81 more uniform multi-language operations;
wenzelm
parents: 69446
diff changeset
  1992
read_line_message socket = do
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1993
  opt_line <- read_line socket
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1994
  case opt_line of
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1995
    Nothing -> return Nothing
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1996
    Just line ->
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  1997
      case Value.parse_nat (UTF8.decode line) of
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  1998
        Nothing -> return $ Just line
74088
6d8674ffb962 clarified signature;
wenzelm
parents: 74087
diff changeset
  1999
        Just n -> fmap Bytes.trim_line . fst <$> read_block socket n
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2000
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2001
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2002
read_yxml :: Socket -> IO (Maybe XML.Body)
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2003
read_yxml socket = do
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2004
  res <- read_line_message socket
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  2005
  return (YXML.parse_body . UTF8.decode <$> res)
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2006
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2007
write_yxml :: Socket -> XML.Body -> IO ()
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2008
write_yxml socket body =
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  2009
  write_line_message socket (UTF8.encode (YXML.string_of_body body))
69446
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  2010
\<close>
9cf0b79dfb7f more Haskell operations;
wenzelm
parents: 69445
diff changeset
  2011
71692
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2012
generate_file "Isabelle/Isabelle_Thread.hs" = \<open>
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2013
{-  Title:      Isabelle/Isabelle_Thread.hs
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2014
    Author:     Makarius
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2015
    LICENSE:    BSD 3-clause (Isabelle)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2016
71692
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2017
Isabelle-specific thread management.
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2018
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2019
See \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\<close>
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2020
and \<^file>\<open>$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\<close>.
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2021
-}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2022
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2023
{-# LANGUAGE NamedFieldPuns #-}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2024
71692
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2025
module Isabelle.Isabelle_Thread (
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2026
  ThreadId, Result,
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2027
  find_id,
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2028
  properties, change_properties,
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2029
  add_resource, del_resource, bracket_resource,
69497
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2030
  is_stopped, expose_stopped, stop,
69496
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2031
  my_uuid, stop_uuid,
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2032
  Fork, fork_finally, fork)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2033
where
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2034
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2035
import Data.Unique
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2036
import Data.IORef
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2037
import System.IO.Unsafe
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2038
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2039
import qualified Data.List as List
69497
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2040
import Control.Monad (when, forM_)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2041
import Data.Map.Strict (Map)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2042
import qualified Data.Map.Strict as Map
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2043
import Control.Exception as Exception
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2044
import Control.Concurrent (ThreadId)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2045
import qualified Control.Concurrent as Concurrent
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2046
import Control.Concurrent.Thread (Result)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2047
import qualified Control.Concurrent.Thread as Thread
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2048
import qualified Isabelle.UUID as UUID
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2049
import qualified Isabelle.Properties as Properties
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2050
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2051
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2052
{- thread info -}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2053
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2054
type Resources = Map Unique (IO ())
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2055
data Info = Info {uuid :: UUID.T, props :: Properties.T, stopped :: Bool, resources :: Resources}
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2056
type Infos = Map ThreadId Info
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2057
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2058
lookup_info :: Infos -> ThreadId -> Maybe Info
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2059
lookup_info infos id = Map.lookup id infos
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2060
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2061
init_info :: ThreadId -> UUID.T -> Infos -> (Infos, ())
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2062
init_info id uuid infos = (Map.insert id (Info uuid [] False Map.empty) infos, ())
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2063
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2064
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2065
{- global state -}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2066
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2067
{-# NOINLINE global_state #-}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2068
global_state :: IORef Infos
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2069
global_state = unsafePerformIO (newIORef Map.empty)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2070
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2071
find_id :: UUID.T -> IO (Maybe ThreadId)
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2072
find_id uuid = do
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2073
  state <- readIORef global_state
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2074
  return $ fst <$> List.find (\(_, Info{uuid = uuid'}) -> uuid == uuid') (Map.assocs state)
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2075
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2076
get_info :: ThreadId -> IO (Maybe Info)
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2077
get_info id = do
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2078
  state <- readIORef global_state
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2079
  return $ lookup_info state id
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2080
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2081
map_info :: ThreadId -> (Info -> Info) -> IO (Maybe Info)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2082
map_info id f =
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2083
  atomicModifyIORef' global_state
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2084
    (\infos ->
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2085
      case lookup_info infos id of
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2086
        Nothing -> (infos, Nothing)
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2087
        Just info ->
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2088
          let info' = f info
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2089
          in (Map.insert id info' infos, Just info'))
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2090
69495
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2091
delete_info :: ThreadId -> IO ()
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2092
delete_info id =
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2093
  atomicModifyIORef' global_state (\infos -> (Map.delete id infos, ()))
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2094
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2095
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2096
{- thread properties -}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2097
69498
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2098
my_info :: IO (Maybe Info)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2099
my_info = do
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2100
  id <- Concurrent.myThreadId
69498
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2101
  get_info id
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2102
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2103
properties :: IO Properties.T
69498
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2104
properties = maybe [] props <$> my_info
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2105
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2106
change_properties :: (Properties.T -> Properties.T) -> IO ()
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2107
change_properties f = do
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2108
  id <- Concurrent.myThreadId
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2109
  map_info id (\info -> info {props = f (props info)})
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2110
  return ()
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2111
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2112
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2113
{- managed resources -}
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2114
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2115
add_resource :: IO () -> IO Unique
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2116
add_resource resource = do
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2117
  id <- Concurrent.myThreadId
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2118
  u <- newUnique
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2119
  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
  2120
  return u
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2121
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2122
del_resource :: Unique -> IO ()
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2123
del_resource u = do
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2124
  id <- Concurrent.myThreadId
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2125
  map_info id (\info -> info {resources = Map.delete u (resources info)})
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2126
  return ()
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2127
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2128
bracket_resource :: IO () -> IO a -> IO a
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2129
bracket_resource resource body =
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2130
  Exception.bracket (add_resource resource) del_resource (const body)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2131
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2132
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2133
{- stop -}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2134
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2135
is_stopped :: IO Bool
69498
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2136
is_stopped = maybe False stopped <$> my_info
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2137
69497
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2138
expose_stopped :: IO ()
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2139
expose_stopped = do
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2140
  stopped <- is_stopped
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2141
  when stopped $ throw ThreadKilled
c434ca819aea more Haskell operations;
wenzelm
parents: 69496
diff changeset
  2142
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2143
stop :: ThreadId -> IO ()
69499
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2144
stop id = do
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2145
  info <- map_info id (\info -> info {stopped = True})
638fdbbc7d1f more Haskell operations: managed resources for threads;
wenzelm
parents: 69498
diff changeset
  2146
  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
  2147
  sequence_ ops
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2148
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2149
69496
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2150
{- UUID -}
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2151
69498
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2152
my_uuid :: IO (Maybe UUID.T)
22e958b76bf6 more robust: suitable defaults for unmanaged threads;
wenzelm
parents: 69497
diff changeset
  2153
my_uuid = fmap uuid <$> my_info
69496
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2154
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2155
stop_uuid :: UUID.T -> IO ()
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2156
stop_uuid uuid = do
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2157
  id <- find_id uuid
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2158
  forM_ id stop
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2159
5256e7f26640 more Haskell operations;
wenzelm
parents: 69495
diff changeset
  2160
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2161
{- fork -}
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2162
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2163
type Fork a = (ThreadId, UUID.T, IO (Result a))
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2164
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2165
fork_finally :: IO a -> (Either SomeException a -> IO b) -> IO (Fork b)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2166
fork_finally body finally = do
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2167
  uuid <- UUID.random
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2168
  (id, result) <-
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2169
    Exception.mask (\restore ->
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2170
      Thread.forkIO
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2171
        (Exception.try
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2172
          (do
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2173
            id <- Concurrent.myThreadId
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2174
            atomicModifyIORef' global_state (init_info id uuid)
69495
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2175
            restore body)
c34dfa431b89 proper cleanup;
wenzelm
parents: 69494
diff changeset
  2176
         >>= (\res -> do id <- Concurrent.myThreadId; delete_info id; finally res)))
69494
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2177
  return (id, uuid, result)
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2178
7e44f8e2cc49 more Haskell operations;
wenzelm
parents: 69491
diff changeset
  2179
fork :: IO a -> IO (Fork a)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2180
fork body = fork_finally body Thread.result
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2181
\<close>
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2182
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2183
generate_file "Isabelle/Server.hs" = \<open>
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2184
{-  Title:      Isabelle/Server.hs
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2185
    Author:     Makarius
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2186
    LICENSE:    BSD 3-clause (Isabelle)
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2187
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2188
TCP server on localhost.
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2189
-}
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2190
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2191
module Isabelle.Server (
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2192
  localhost_name, localhost, publish_text, publish_stdout,
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2193
  server, connection
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2194
)
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2195
where
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2196
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2197
import Control.Monad (forever, when)
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2198
import qualified Control.Exception as Exception
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2199
import Network.Socket (Socket)
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2200
import qualified Network.Socket as Socket
69472
d016ef70c069 tuned messages;
wenzelm
parents: 69467
diff changeset
  2201
import qualified System.IO as IO
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2202
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2203
import Isabelle.Library
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  2204
import Isabelle.Bytes (Bytes)
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2205
import qualified Isabelle.UUID as UUID
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2206
import qualified Isabelle.Byte_Message as Byte_Message
71692
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2207
import qualified Isabelle.Isabelle_Thread as Isabelle_Thread
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  2208
import qualified Isabelle.UTF8 as UTF8
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2209
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2210
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2211
{- server address -}
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2212
69463
6439c9024dcc clarified signature;
wenzelm
parents: 69462
diff changeset
  2213
localhost_name :: String
6439c9024dcc clarified signature;
wenzelm
parents: 69462
diff changeset
  2214
localhost_name = "127.0.0.1"
6439c9024dcc clarified signature;
wenzelm
parents: 69462
diff changeset
  2215
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2216
localhost :: Socket.HostAddress
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2217
localhost = Socket.tupleToHostAddress (127, 0, 0, 1)
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2218
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2219
publish_text :: String -> String -> UUID.T -> String
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2220
publish_text name address password =
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  2221
  "server " <> quote name <> " = " <> address <> " (password " <> quote (show password) <> ")"
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2222
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2223
publish_stdout :: String -> String -> UUID.T -> IO ()
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2224
publish_stdout name address password = putStrLn (publish_text name address password)
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2225
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2226
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2227
{- server -}
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2228
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2229
server :: (String -> UUID.T -> IO ()) -> (Socket -> IO ()) -> IO ()
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2230
server publish handle =
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2231
  Socket.withSocketsDo $ Exception.bracket open (Socket.close . fst) (uncurry loop)
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2232
  where
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  2233
    open :: IO (Socket, Bytes)
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2234
    open = do
69466
wenzelm
parents: 69465
diff changeset
  2235
      server_socket <- Socket.socket Socket.AF_INET Socket.Stream Socket.defaultProtocol
wenzelm
parents: 69465
diff changeset
  2236
      Socket.bind server_socket (Socket.SockAddrInet 0 localhost)
wenzelm
parents: 69465
diff changeset
  2237
      Socket.listen server_socket 50
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2238
69466
wenzelm
parents: 69465
diff changeset
  2239
      port <- Socket.socketPort server_socket
74081
adaa2e9a4111 tuned signature: more generic operations;
wenzelm
parents: 74080
diff changeset
  2240
      let address = localhost_name <> ":" <> show port
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2241
      password <- UUID.random
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2242
      publish address password
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2243
69466
wenzelm
parents: 69465
diff changeset
  2244
      return (server_socket, UUID.bytes password)
69462
fe125722f7a9 more Haskell operations;
wenzelm
parents: 69459
diff changeset
  2245
74084
a8bbeb266651 prefer Isabelle.Bytes, based on ShortByteString;
wenzelm
parents: 74083
diff changeset
  2246
    loop :: Socket -> Bytes -> IO ()
69466
wenzelm
parents: 69465
diff changeset
  2247
    loop server_socket password = forever $ do
69480
wenzelm
parents: 69477
diff changeset
  2248
      (connection, _) <- Socket.accept server_socket
71692
f8e52c0152fe clarified names;
wenzelm
parents: 71490
diff changeset
  2249
      Isabelle_Thread.fork_finally
69465
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2250
        (do
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2251
          line <- Byte_Message.read_line connection
16fa609a62b1 clarified UUID operations;
wenzelm
parents: 69463
diff changeset
  2252
          when (line == Just password) $ handle connection)
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2253
        (\finally -> do
69472
d016ef70c069 tuned messages;
wenzelm
parents: 69467
diff changeset
  2254
          Socket.close connection
69473
f71598c11fae more Haskell operations;
wenzelm
parents: 69472
diff changeset
  2255
          case finally of
69472
d016ef70c069 tuned messages;
wenzelm
parents: 69467
diff changeset
  2256
            Left exn -> IO.hPutStrLn IO.stderr $ Exception.displayException exn
d016ef70c069 tuned messages;
wenzelm
parents: 69467
diff changeset
  2257
            Right () -> return ())
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2258
      return ()
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2259
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2260
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2261
{- client connection -}
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2262
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2263
connection :: String -> String -> (Socket -> IO a) -> IO a
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2264
connection port password client =
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2265
  Socket.withSocketsDo $ do
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2266
    addr <- resolve
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2267
    Exception.bracket (open addr) Socket.close body
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2268
  where
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2269
    resolve = do
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2270
      let hints =
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2271
            Socket.defaultHints {
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2272
              Socket.addrFlags = [Socket.AI_NUMERICHOST, Socket.AI_NUMERICSERV],
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2273
              Socket.addrSocketType = Socket.Stream }
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2274
      head <$> Socket.getAddrInfo (Just hints) (Just "127.0.0.1") (Just port)
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2275
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2276
    open addr = do
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2277
      socket <- Socket.socket (Socket.addrFamily addr) (Socket.addrSocketType addr)
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2278
                  (Socket.addrProtocol addr)
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2279
      Socket.connect socket $ Socket.addrAddress addr
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2280
      return socket
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2281
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2282
    body socket = do
74080
5b68a5cd7061 prefer UTF8 implementation from Data.Text.Encoding (foreign C);
wenzelm
parents: 73246
diff changeset
  2283
      Byte_Message.write_line socket (UTF8.encode password)
73178
7e70d7dd1baa more operations for client connection;
wenzelm
parents: 73177
diff changeset
  2284
      client socket
69455
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2285
\<close>
6a901078a294 more Haskell operations;
wenzelm
parents: 69454
diff changeset
  2286
70047
96fe857a7a6f clarified signature: more explicit operations for corresponding Isar commands;
wenzelm
parents: 69968
diff changeset
  2287
export_generated_files _
69628
a2fbfdc5e62d export generated files;
wenzelm
parents: 69499
diff changeset
  2288
69222
8365124a86ae support for Isabelle tool development in Haskell;
wenzelm
parents:
diff changeset
  2289
end