src/HOL/Library/ML_String.thy
author wenzelm
Tue, 18 Sep 2007 18:05:34 +0200
changeset 24633 0a3a02066244
parent 24219 e558fe311376
child 24717 56ba87ec8d31
permissions -rw-r--r--
moved Tools/integer.ML to Pure/General/integer.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     1
(*  ID:         $Id$
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     3
*)
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     4
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     5
header {* Monolithic strings for ML *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     6
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     7
theory ML_String
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     8
imports List
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
     9
begin
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    10
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    11
subsection {* Motivation *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    12
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    13
text {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    14
  Strings are represented in HOL as list of characters.
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    15
  For code generation to Haskell, this is no problem
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    16
  since in Haskell "abc" is equivalent to ['a', 'b', 'c'].
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    17
  On the other hand, in ML all strings have to
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    18
  be represented as list of characters which
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    19
  is awkward to read. This theory provides a distinguished
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    20
  datatype for strings which then by convention
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    21
  are serialized as monolithic ML strings.
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    22
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    23
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    24
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    25
subsection {* Datatype of monolithic strings *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    26
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    27
datatype ml_string = STR string
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    28
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    29
lemmas [code func del] = ml_string.recs ml_string.cases
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    30
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    31
lemma [code func]: "size (s\<Colon>ml_string) = 0"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    32
  by (cases s) simp_all
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    33
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    34
subsection {* ML interface *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    35
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    36
ML {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    37
structure ML_String =
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    38
struct
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    39
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    40
fun mk s = @{term STR} $ HOLogic.mk_string s;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    41
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    42
end;
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    43
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    44
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    45
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    46
subsection {* Code serialization *}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    47
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    48
code_type ml_string
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    49
  (SML "string")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    50
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    51
setup {*
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    52
let
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    53
  val charr = @{const_name Char}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    54
  val nibbles = [@{const_name Nibble0}, @{const_name Nibble1},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    55
    @{const_name Nibble2}, @{const_name Nibble3},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    56
    @{const_name Nibble4}, @{const_name Nibble5},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    57
    @{const_name Nibble6}, @{const_name Nibble7},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    58
    @{const_name Nibble8}, @{const_name Nibble9},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    59
    @{const_name NibbleA}, @{const_name NibbleB},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    60
    @{const_name NibbleC}, @{const_name NibbleD},
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    61
    @{const_name NibbleE}, @{const_name NibbleF}];
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    62
in
24219
e558fe311376 new structure for code generator modules
haftmann
parents: 23854
diff changeset
    63
  CodeTarget.add_pretty_ml_string "SML"
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    64
    charr nibbles @{const_name Nil} @{const_name Cons} @{const_name STR}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    65
end
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    66
*}
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    67
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    68
code_reserved SML string
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    69
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    70
code_const "op = \<Colon> ml_string \<Rightarrow> ml_string \<Rightarrow> bool"
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    71
  (SML "!((_ : string) = _)")
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    72
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents:
diff changeset
    73
end