| author | nipkow | 
| Tue, 24 Mar 2009 14:09:24 +0100 | |
| changeset 30707 | b0391b9b7103 | 
| parent 28952 | 15a4b2cf8c34 | 
| child 31032 | 38901ed00ec3 | 
| permissions | -rw-r--r-- | 
| 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  | 
|
| 27487 | 8  | 
imports Plain "~~/src/HOL/List"  | 
| 24999 | 9  | 
begin  | 
10  | 
||
11  | 
subsection {* Datatype of messages *}
 | 
|
12  | 
||
13  | 
datatype message_string = STR string  | 
|
14  | 
||
| 28562 | 15  | 
lemmas [code del] = message_string.recs message_string.cases  | 
| 24999 | 16  | 
|
| 28562 | 17  | 
lemma [code]: "size (s\<Colon>message_string) = 0"  | 
| 24999 | 18  | 
by (cases s) simp_all  | 
19  | 
||
| 28562 | 20  | 
lemma [code]: "message_string_size (s\<Colon>message_string) = 0"  | 
| 25673 | 21  | 
by (cases s) simp_all  | 
22  | 
||
| 24999 | 23  | 
subsection {* ML interface *}
 | 
24  | 
||
25  | 
ML {*
 | 
|
26  | 
structure Message_String =  | 
|
27  | 
struct  | 
|
28  | 
||
29  | 
fun mk s = @{term STR} $ HOLogic.mk_string s;
 | 
|
30  | 
||
31  | 
end;  | 
|
32  | 
*}  | 
|
33  | 
||
34  | 
||
35  | 
subsection {* Code serialization *}
 | 
|
36  | 
||
37  | 
code_type message_string  | 
|
38  | 
(SML "string")  | 
|
39  | 
(OCaml "string")  | 
|
40  | 
(Haskell "String")  | 
|
41  | 
||
42  | 
setup {*
 | 
|
| 
28090
 
29af3c712d2b
distributed literal code generation out of central infrastructure
 
haftmann 
parents: 
28064 
diff
changeset
 | 
43  | 
  fold (fn target => add_literal_message @{const_name STR} target)
 | 
| 
 
29af3c712d2b
distributed literal code generation out of central infrastructure
 
haftmann 
parents: 
28064 
diff
changeset
 | 
44  | 
["SML", "OCaml", "Haskell"]  | 
| 24999 | 45  | 
*}  | 
46  | 
||
47  | 
code_reserved SML string  | 
|
48  | 
code_reserved OCaml string  | 
|
49  | 
||
50  | 
code_instance message_string :: eq  | 
|
51  | 
(Haskell -)  | 
|
52  | 
||
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28090 
diff
changeset
 | 
53  | 
code_const "eq_class.eq \<Colon> message_string \<Rightarrow> message_string \<Rightarrow> bool"  | 
| 24999 | 54  | 
(SML "!((_ : string) = _)")  | 
55  | 
(OCaml "!((_ : string) = _)")  | 
|
56  | 
(Haskell infixl 4 "==")  | 
|
57  | 
||
58  | 
end  |