| 24999 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Author:     Florian Haftmann, TU Muenchen
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | header {* Monolithic strings (message strings) for code generation *}
 | 
|  |      6 | 
 | 
|  |      7 | theory Code_Message
 | 
|  |      8 | imports List
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | subsection {* Datatype of messages *}
 | 
|  |     12 | 
 | 
|  |     13 | datatype message_string = STR string
 | 
|  |     14 | 
 | 
|  |     15 | lemmas [code func del] = message_string.recs message_string.cases
 | 
|  |     16 | 
 | 
|  |     17 | lemma [code func]: "size (s\<Colon>message_string) = 0"
 | 
|  |     18 |   by (cases s) simp_all
 | 
|  |     19 | 
 | 
|  |     20 | subsection {* ML interface *}
 | 
|  |     21 | 
 | 
|  |     22 | ML {*
 | 
|  |     23 | structure Message_String =
 | 
|  |     24 | struct
 | 
|  |     25 | 
 | 
|  |     26 | fun mk s = @{term STR} $ HOLogic.mk_string s;
 | 
|  |     27 | 
 | 
|  |     28 | end;
 | 
|  |     29 | *}
 | 
|  |     30 | 
 | 
|  |     31 | 
 | 
|  |     32 | subsection {* Code serialization *}
 | 
|  |     33 | 
 | 
|  |     34 | code_type message_string
 | 
|  |     35 |   (SML "string")
 | 
|  |     36 |   (OCaml "string")
 | 
|  |     37 |   (Haskell "String")
 | 
|  |     38 | 
 | 
|  |     39 | setup {*
 | 
|  |     40 | let
 | 
|  |     41 |   val charr = @{const_name Char}
 | 
|  |     42 |   val nibbles = [@{const_name Nibble0}, @{const_name Nibble1},
 | 
|  |     43 |     @{const_name Nibble2}, @{const_name Nibble3},
 | 
|  |     44 |     @{const_name Nibble4}, @{const_name Nibble5},
 | 
|  |     45 |     @{const_name Nibble6}, @{const_name Nibble7},
 | 
|  |     46 |     @{const_name Nibble8}, @{const_name Nibble9},
 | 
|  |     47 |     @{const_name NibbleA}, @{const_name NibbleB},
 | 
|  |     48 |     @{const_name NibbleC}, @{const_name NibbleD},
 | 
|  |     49 |     @{const_name NibbleE}, @{const_name NibbleF}];
 | 
|  |     50 | in
 | 
|  |     51 |   fold (fn target => CodeTarget.add_pretty_message target
 | 
|  |     52 |     charr nibbles @{const_name Nil} @{const_name Cons} @{const_name STR})
 | 
|  |     53 |   ["SML", "OCaml", "Haskell"]
 | 
|  |     54 | end
 | 
|  |     55 | *}
 | 
|  |     56 | 
 | 
|  |     57 | code_reserved SML string
 | 
|  |     58 | code_reserved OCaml string
 | 
|  |     59 | 
 | 
|  |     60 | code_instance message_string :: eq
 | 
|  |     61 |   (Haskell -)
 | 
|  |     62 | 
 | 
|  |     63 | code_const "op = \<Colon> message_string \<Rightarrow> message_string \<Rightarrow> bool"
 | 
|  |     64 |   (SML "!((_ : string) = _)")
 | 
|  |     65 |   (OCaml "!((_ : string) = _)")
 | 
|  |     66 |   (Haskell infixl 4 "==")
 | 
|  |     67 | 
 | 
|  |     68 | end
 |