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