| author | haftmann | 
| Fri, 23 Mar 2007 09:40:47 +0100 | |
| changeset 22505 | e2d378a97905 | 
| parent 22016 | e086b4e846b8 | 
| child 22520 | ebe95b0242b3 | 
| permissions | -rw-r--r-- | 
| 20400 | 1 | (* ID: $Id$ | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Monolithic strings for ML  *}
 | |
| 6 | ||
| 7 | theory MLString | |
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 8 | imports List | 
| 20400 | 9 | begin | 
| 10 | ||
| 11 | subsection {* Motivation *}
 | |
| 12 | ||
| 13 | text {*
 | |
| 14 | Strings are represented in HOL as list of characters. | |
| 15 | For code generation to Haskell, this is no problem | |
| 16 | since in Haskell "abc" is equivalent to ['a', 'b', 'c']. | |
| 17 | On the other hand, in ML all strings have to | |
| 18 | be represented as list of characters which | |
| 19 | is awkward to read. This theory provides a distinguished | |
| 20 | datatype for strings which then by convention | |
| 21 | are serialized as monolithic ML strings. Note | |
| 22 | that in Haskell these strings are identified | |
| 23 | with Haskell strings. | |
| 24 | *} | |
| 25 | ||
| 26 | ||
| 27 | subsection {* Datatype of monolithic strings *}
 | |
| 28 | ||
| 29 | datatype ml_string = STR string | |
| 30 | ||
| 31 | consts | |
| 32 | explode :: "ml_string \<Rightarrow> string" | |
| 33 | ||
| 34 | primrec | |
| 35 | "explode (STR s) = s" | |
| 36 | ||
| 37 | ||
| 38 | subsection {* ML interface *}
 | |
| 39 | ||
| 40 | ML {*
 | |
| 41 | structure MLString = | |
| 42 | struct | |
| 43 | ||
| 44 | local | |
| 45 | val thy = the_context (); | |
| 46 | val const_STR = Sign.intern_const thy "STR"; | |
| 47 | in | |
| 48 | val typ_ml_string = Type (Sign.intern_type thy "ml_string", []); | |
| 21455 | 49 | fun term_ml_string s = Const (const_STR, HOLogic.stringT --> typ_ml_string) | 
| 50 | $ HOLogic.mk_string s | |
| 20400 | 51 | end; | 
| 52 | ||
| 53 | end; | |
| 54 | *} | |
| 55 | ||
| 56 | ||
| 57 | subsection {* Code serialization *}
 | |
| 58 | ||
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 59 | code_type ml_string | 
| 21113 | 60 | (SML "string") | 
| 61 | (Haskell "String") | |
| 20400 | 62 | |
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 63 | code_const STR | 
| 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 64 | (Haskell "_") | 
| 20400 | 65 | |
| 66 | setup {*
 | |
| 21756 | 67 | CodegenSerializer.add_pretty_ml_string "SML" | 
| 68 | "List.list.Nil" "List.list.Cons" "MLString.ml_string.STR" | |
| 69 | ML_Syntax.print_char ML_Syntax.print_string "String.implode" | |
| 20400 | 70 | *} | 
| 71 | ||
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 72 | code_const explode | 
| 21113 | 73 | (SML "String.explode") | 
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 74 | (Haskell "_") | 
| 20400 | 75 | |
| 21079 | 76 | code_reserved SML string explode | 
| 77 | ||
| 22016 | 78 | text {* disable something ugly *}
 | 
| 79 | ||
| 80 | code_const "ml_string_rec" and "ml_string_case" and "size \<Colon> ml_string \<Rightarrow> nat" | |
| 81 | (SML "!((_); (_); raise Fail \"ml'_string'_rec\")" | |
| 82 | and "!((_); (_); raise Fail \"ml'_string'_case\")" | |
| 83 | and "!((_); raise Fail \"size'_ml'_string\")") | |
| 84 | (OCaml "!((_); (_); failwith \"ml'_string'_rec\")" | |
| 85 | and "!((_); (_); failwith \"ml'_string'_case\")" | |
| 86 | and "!((_); failwith \"size'_ml'_string\")") | |
| 87 | ||
| 20453 
855f07fabd76
final syntax for some Isar code generator keywords
 haftmann parents: 
20439diff
changeset | 88 | end |