src/HOL/Tools/SMT2/smtlib2.ML
changeset 56078 624faeda77b5
child 57229 489083abce44
equal deleted inserted replaced
56077:d397030fb27e 56078:624faeda77b5
       
     1 (*  Title:      HOL/Tools/SMT2/smtlib2.ML
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 
       
     4 Parsing and generating SMT-LIB 2.
       
     5 *)
       
     6 
       
     7 signature SMTLIB2 =
       
     8 sig
       
     9   exception PARSE of int * string
       
    10   datatype tree = 
       
    11     Num of int |
       
    12     Dec of int * int |
       
    13     Str of string |
       
    14     Sym of string |
       
    15     Key of string |
       
    16     S of tree list
       
    17   val parse: string list -> tree
       
    18   val pretty_tree: tree -> Pretty.T
       
    19   val str_of: tree -> string
       
    20 end
       
    21 
       
    22 structure SMTLIB2: SMTLIB2 =
       
    23 struct
       
    24 
       
    25 (* data structures *)
       
    26 
       
    27 exception PARSE of int * string
       
    28 
       
    29 datatype tree = 
       
    30   Num of int |
       
    31   Dec of int * int |
       
    32   Str of string |
       
    33   Sym of string |
       
    34   Key of string |
       
    35   S of tree list
       
    36 
       
    37 datatype unfinished = None | String of string | Symbol of string
       
    38 
       
    39 
       
    40 
       
    41 (* utilities *)
       
    42 
       
    43 fun read_raw pred l cs =
       
    44   (case take_prefix pred cs of
       
    45     ([], []) => raise PARSE (l, "empty token")
       
    46   | ([], c :: _) => raise PARSE (l, "unexpected character " ^ quote c)
       
    47   | x => x)
       
    48 
       
    49 
       
    50 
       
    51 (* numerals and decimals *)
       
    52 
       
    53 fun int_of cs = fst (read_int cs)
       
    54 
       
    55 fun read_num l cs =
       
    56   (case read_raw Symbol.is_ascii_digit l cs of
       
    57     (cs1, "." :: cs') =>
       
    58       let val (cs2, cs'') = read_raw Symbol.is_ascii_digit l cs'
       
    59       in (Dec (int_of cs1, int_of cs2), cs'') end
       
    60   | (cs1, cs2) => (Num (int_of cs1), cs2))
       
    61 
       
    62 
       
    63 
       
    64 (* binary numbers *)
       
    65 
       
    66 fun is_bin c = (c = "0" orelse c = "1")
       
    67 
       
    68 fun read_bin l cs = read_raw is_bin l cs |>> Num o fst o read_radix_int 2
       
    69 
       
    70 
       
    71 
       
    72 (* hex numbers *)
       
    73 
       
    74 val is_hex = member (op =) (raw_explode "0123456789abcdefABCDEF")
       
    75 
       
    76 fun within c1 c2 c = (ord c1 <= ord c andalso ord c <= ord c2)
       
    77 
       
    78 fun unhex i [] = i
       
    79   | unhex i (c :: cs) =
       
    80       if within "0" "9" c then unhex (i * 16 + (ord c - ord "0")) cs
       
    81       else if within "a" "f" c then unhex (i * 16 + (ord c - ord "a" + 10)) cs
       
    82       else if within "A" "F" c then unhex (i * 16 + (ord c - ord "A" + 10)) cs
       
    83       else raise Fail ("bad hex character " ^ quote c)
       
    84 
       
    85 fun read_hex l cs = read_raw is_hex l cs |>> Num o unhex 0
       
    86 
       
    87 
       
    88 
       
    89 (* symbols *)
       
    90 
       
    91 val symbol_chars = raw_explode "~!@$%^&*_+=<>.?/-" 
       
    92 
       
    93 fun is_sym c =
       
    94   Symbol.is_ascii_letter c orelse
       
    95   Symbol.is_ascii_digit c orelse
       
    96   member (op =) symbol_chars c
       
    97 
       
    98 fun read_sym f l cs = read_raw is_sym l cs |>> f o implode
       
    99 
       
   100 
       
   101 
       
   102 (* quoted tokens *)
       
   103 
       
   104 fun read_quoted stop (escape, replacement) cs =
       
   105   let
       
   106     fun read _ [] = NONE
       
   107       | read rs (cs as (c :: cs')) =
       
   108           if is_prefix (op =) stop cs then
       
   109             SOME (implode (rev rs), drop (length stop) cs)
       
   110           else if not (null escape) andalso is_prefix (op =) escape cs then
       
   111             read (replacement :: rs) (drop (length escape) cs)
       
   112           else read (c :: rs) cs'
       
   113   in read [] cs end
       
   114 
       
   115 fun read_string cs = read_quoted ["\\", "\""] (["\\", "\\"], "\\") cs
       
   116 fun read_symbol cs = read_quoted ["|"] ([], "") cs
       
   117 
       
   118 
       
   119 
       
   120 (* core parser *)
       
   121 
       
   122 fun read _ [] rest tss = (rest, tss)
       
   123   | read l ("(" :: cs) None tss = read l cs None ([] :: tss)
       
   124   | read l (")" :: cs) None (ts1 :: ts2 :: tss) =
       
   125       read l cs None ((S (rev ts1) :: ts2) :: tss)
       
   126   | read l ("#" :: "x" :: cs) None (ts :: tss) =
       
   127       token read_hex l cs ts tss
       
   128   | read l ("#" :: cs) None (ts :: tss) =
       
   129       token read_bin l cs ts tss
       
   130   | read l (":" :: cs) None (ts :: tss) =
       
   131       token (read_sym Key) l cs ts tss
       
   132   | read l ("\"" :: cs) None (ts :: tss) =
       
   133       quoted read_string String Str l "" cs ts tss
       
   134   | read l ("|" :: cs) None (ts :: tss) =
       
   135       quoted read_symbol Symbol Sym l "" cs ts tss
       
   136   | read l ((c as "!") :: cs) None (ts :: tss) =
       
   137       token (fn _ => pair (Sym c)) l cs ts tss
       
   138   | read l (c :: cs) None (ts :: tss) =
       
   139       if Symbol.is_ascii_blank c then read l cs None (ts :: tss)
       
   140       else if Symbol.is_digit c then token read_num l (c :: cs) ts tss
       
   141       else token (read_sym Sym) l (c :: cs) ts tss
       
   142   | read l cs (String s) (ts :: tss) =
       
   143       quoted read_string String Str l s cs ts tss
       
   144   | read l cs (Symbol s) (ts :: tss) =
       
   145       quoted read_symbol Symbol Sym l s cs ts tss
       
   146   | read l _ _ [] = raise PARSE (l, "bad parser state")
       
   147 
       
   148 and token f l cs ts tss =
       
   149   let val (t, cs') = f l cs
       
   150   in read l cs' None ((t :: ts) :: tss) end
       
   151 
       
   152 and quoted r f g l s cs ts tss =
       
   153   (case r cs of
       
   154     NONE => (f (s ^ implode cs), ts :: tss)
       
   155   | SOME (s', cs') => read l cs' None ((g (s ^ s') :: ts) :: tss))
       
   156   
       
   157 
       
   158 
       
   159 (* overall parser *)
       
   160 
       
   161 fun read_line l line = read l (raw_explode line)
       
   162 
       
   163 fun add_line line (l, (None, tss)) =
       
   164       if size line = 0 orelse nth_string line 0 = ";" then (l + 1, (None, tss))
       
   165       else (l + 1, read_line l line None tss)
       
   166   | add_line line (l, (unfinished, tss)) =
       
   167       (l + 1, read_line l line unfinished tss)
       
   168 
       
   169 fun finish (_, (None, [[t]])) = t
       
   170   | finish (l, _) = raise PARSE (l, "bad nesting")
       
   171 
       
   172 fun parse lines = finish (fold add_line lines (1, (None, [[]])))
       
   173 
       
   174 
       
   175 
       
   176 (* pretty printer *)
       
   177 
       
   178 fun pretty_tree (Num i) = Pretty.str (string_of_int i)
       
   179   | pretty_tree (Dec (i, j)) =
       
   180       Pretty.str (string_of_int i ^ "." ^ string_of_int j)
       
   181   | pretty_tree (Str s) =
       
   182       raw_explode s
       
   183       |> maps (fn "\"" => ["\\", "\""] | "\\" => ["\\", "\\"] | c => [c])
       
   184       |> implode
       
   185       |> enclose "\"" "\""
       
   186       |> Pretty.str
       
   187   | pretty_tree (Sym s) =
       
   188       if String.isPrefix "(" s (* for bit vector functions *) orelse
       
   189          forall is_sym (raw_explode s) then
       
   190         Pretty.str s
       
   191       else
       
   192         Pretty.str ("|" ^ s ^ "|")
       
   193   | pretty_tree (Key s) = Pretty.str (":" ^ s)
       
   194   | pretty_tree (S trees) =
       
   195       Pretty.enclose "(" ")" (Pretty.separate "" (map pretty_tree trees))
       
   196 
       
   197 val str_of = Pretty.str_of o pretty_tree
       
   198 
       
   199 end