src/HOL/ex/String.thy
author berghofe
Tue, 25 Jun 1996 13:11:29 +0200
changeset 1824 44254696843a
parent 1384 007ad29ce6ca
child 3795 e687069e7257
permissions -rw-r--r--
Changed argument order of nat_rec.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/String.thy
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     3
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
Hex chars. Strings.
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
String = List +
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     9
datatype
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    10
  nibble = H00 | H01 | H02 | H03 | H04 | H05 | H06 | H07
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    11
         | H08 | H09 | H0A | H0B | H0C | H0D | H0E | H0F
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    12
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    13
datatype
977
5d57287e5e1e changed syntax of datatype declarations (curried types for constructor
clasohm
parents: 969
diff changeset
    14
  char = Char nibble nibble
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    15
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    16
types
1384
007ad29ce6ca removed quotes from types section
clasohm
parents: 1376
diff changeset
    17
  string = char list
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    18
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    19
syntax
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 977
diff changeset
    20
  "_Char"       :: xstr => char       ("CHR _")
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 977
diff changeset
    21
  "_String"     :: xstr => string     ("_")
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    22
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    23
end
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    24
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    25
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    26
ML
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    27
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    28
local
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    29
  open Syntax;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    30
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    31
  val ssquote = enclose "''" "''";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    32
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    33
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    34
  (* chars *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    35
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    36
  val zero = ord "0";
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    37
  val ten = ord "A" - 10;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    38
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    39
  fun mk_nib n =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    40
    const ("H0" ^ chr (n + (if n <= 9 then zero else ten)));
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    41
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    42
  fun dest_nib (Const (c, _)) =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    43
        (case explode c of
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    44
          ["H", "0", h] => ord h - (if h <= "9" then zero else ten)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    45
        | _ => raise Match)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    46
    | dest_nib _ = raise Match;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    47
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    48
  fun dest_nibs t1 t2 = chr (dest_nib t1 * 16 + dest_nib t2);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    49
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    50
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    51
  fun mk_char c =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    52
    const "Char" $ mk_nib (ord c div 16) $ mk_nib (ord c mod 16);
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    53
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    54
  fun dest_char (Const ("Char", _) $ t1 $ t2) = dest_nibs t1 t2
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    55
    | dest_char _ = raise Match;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    56
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    57
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    58
  fun char_tr (*"_Char"*) [Free (c, _)] =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    59
        if size c = 1 then mk_char c
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    60
        else error ("Bad character: " ^ quote c)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    61
    | char_tr (*"_Char"*) ts = raise_term "char_tr" ts;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    62
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    63
  fun char_tr' (*"Char"*) [t1, t2] =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    64
        const "_Char" $ free (ssquote (dest_nibs t1 t2))
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    65
    | char_tr' (*"Char"*) _ = raise Match;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    66
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    67
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    68
  (* strings *)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    69
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    70
  fun mk_string [] = const constrainC $ const "[]" $ const "string"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    71
    | mk_string (t :: ts) = const "op #" $ t $ mk_string ts;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    72
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    73
  fun dest_string (Const ("[]", _)) = []
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    74
    | dest_string (Const ("op #", _) $ c $ cs) = dest_char c :: dest_string cs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    75
    | dest_string _ = raise Match;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    76
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    77
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    78
  fun string_tr (*"_String"*) [Free (txt, _)] =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    79
        mk_string (map mk_char (explode txt))
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    80
    | string_tr (*"_String"*) ts = raise_term "string_tr" ts;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    81
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    82
  fun cons_tr' (*"op #"*) [c, cs] =
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    83
        const "_String" $ free (ssquote (implode (dest_char c :: dest_string cs)))
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    84
    | cons_tr' (*"op #"*) ts = raise Match;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    85
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    86
in
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    87
  val parse_translation = [("_Char", char_tr), ("_String", string_tr)];
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    88
  val print_translation = [("Char", char_tr'), ("op #", cons_tr')];
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    89
end;
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    90