author | wenzelm |
Sun, 05 Sep 2010 21:41:24 +0200 | |
changeset 39134 | 917b4b6ba3d2 |
parent 38858 | 1920158cfa17 |
child 39198 | f967a16dfcdd |
permissions | -rw-r--r-- |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
1 |
(* Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
2 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
3 |
header {* Character and string types *} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
4 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
5 |
theory String |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
6 |
imports List |
31055
2cf6efca6c71
proper structures for list and string code generation stuff
haftmann
parents:
31051
diff
changeset
|
7 |
uses |
35115 | 8 |
("Tools/string_syntax.ML") |
31055
2cf6efca6c71
proper structures for list and string code generation stuff
haftmann
parents:
31051
diff
changeset
|
9 |
("Tools/string_code.ML") |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
10 |
begin |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
11 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
12 |
subsection {* Characters *} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
13 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
14 |
datatype nibble = |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
15 |
Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7 |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
16 |
| Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
17 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
18 |
lemma UNIV_nibble: |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
19 |
"UNIV = {Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
20 |
Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF}" (is "_ = ?A") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
21 |
proof (rule UNIV_eq_I) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
22 |
fix x show "x \<in> ?A" by (cases x) simp_all |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
23 |
qed |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
24 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
25 |
instance nibble :: finite |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
26 |
by default (simp add: UNIV_nibble) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
27 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
28 |
datatype char = Char nibble nibble |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
29 |
-- "Note: canonical order of character encoding coincides with standard term ordering" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
30 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
31 |
lemma UNIV_char: |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
32 |
"UNIV = image (split Char) (UNIV \<times> UNIV)" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
33 |
proof (rule UNIV_eq_I) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
34 |
fix x show "x \<in> image (split Char) (UNIV \<times> UNIV)" by (cases x) auto |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
35 |
qed |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
36 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
37 |
instance char :: finite |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
38 |
by default (simp add: UNIV_char) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
39 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
40 |
lemma size_char [code, simp]: |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
41 |
"size (c::char) = 0" by (cases c) simp |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
42 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
43 |
lemma char_size [code, simp]: |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
44 |
"char_size (c::char) = 0" by (cases c) simp |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
45 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
46 |
primrec nibble_pair_of_char :: "char \<Rightarrow> nibble \<times> nibble" where |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
47 |
"nibble_pair_of_char (Char n m) = (n, m)" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
48 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
49 |
setup {* |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
50 |
let |
33063 | 51 |
val nibbles = map_range (Thm.cterm_of @{theory} o HOLogic.mk_nibble) 16; |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
52 |
val thms = map_product |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
53 |
(fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps}) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
54 |
nibbles nibbles; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
55 |
in |
38808 | 56 |
PureThy.note_thmss Thm.definitionK [((@{binding nibble_pair_of_char_simps}, []), [(thms, [])])] |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
57 |
#-> (fn [(_, thms)] => fold_rev Code.add_eqn thms) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
58 |
end |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
59 |
*} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
60 |
|
32069
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
haftmann
parents:
31998
diff
changeset
|
61 |
lemma char_case_nibble_pair [code, code_unfold]: |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
62 |
"char_case f = split f o nibble_pair_of_char" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
63 |
by (simp add: expand_fun_eq split: char.split) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
64 |
|
32069
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
haftmann
parents:
31998
diff
changeset
|
65 |
lemma char_rec_nibble_pair [code, code_unfold]: |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
66 |
"char_rec f = split f o nibble_pair_of_char" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
67 |
unfolding char_case_nibble_pair [symmetric] |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
68 |
by (simp add: expand_fun_eq split: char.split) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
69 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
70 |
syntax |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
71 |
"_Char" :: "xstr => char" ("CHR _") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
72 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
73 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
74 |
subsection {* Strings *} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
75 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
76 |
types string = "char list" |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
77 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
78 |
syntax |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
79 |
"_String" :: "xstr => string" ("_") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
80 |
|
35115 | 81 |
use "Tools/string_syntax.ML" |
35123 | 82 |
setup String_Syntax.setup |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
83 |
|
31484 | 84 |
definition chars :: string where |
85 |
"chars = [Char Nibble0 Nibble0, Char Nibble0 Nibble1, Char Nibble0 Nibble2, |
|
86 |
Char Nibble0 Nibble3, Char Nibble0 Nibble4, Char Nibble0 Nibble5, |
|
87 |
Char Nibble0 Nibble6, Char Nibble0 Nibble7, Char Nibble0 Nibble8, |
|
88 |
Char Nibble0 Nibble9, Char Nibble0 NibbleA, Char Nibble0 NibbleB, |
|
89 |
Char Nibble0 NibbleC, Char Nibble0 NibbleD, Char Nibble0 NibbleE, |
|
90 |
Char Nibble0 NibbleF, Char Nibble1 Nibble0, Char Nibble1 Nibble1, |
|
91 |
Char Nibble1 Nibble2, Char Nibble1 Nibble3, Char Nibble1 Nibble4, |
|
92 |
Char Nibble1 Nibble5, Char Nibble1 Nibble6, Char Nibble1 Nibble7, |
|
93 |
Char Nibble1 Nibble8, Char Nibble1 Nibble9, Char Nibble1 NibbleA, |
|
94 |
Char Nibble1 NibbleB, Char Nibble1 NibbleC, Char Nibble1 NibbleD, |
|
95 |
Char Nibble1 NibbleE, Char Nibble1 NibbleF, CHR '' '', CHR ''!'', |
|
96 |
Char Nibble2 Nibble2, CHR ''#'', CHR ''$'', CHR ''%'', CHR ''&'', |
|
97 |
Char Nibble2 Nibble7, CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', CHR '','', |
|
98 |
CHR ''-'', CHR ''.'', CHR ''/'', CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', |
|
99 |
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9'', CHR '':'', |
|
100 |
CHR '';'', CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', CHR ''@'', CHR ''A'', |
|
101 |
CHR ''B'', CHR ''C'', CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', CHR ''H'', |
|
102 |
CHR ''I'', CHR ''J'', CHR ''K'', CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', |
|
103 |
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', CHR ''T'', CHR ''U'', CHR ''V'', |
|
104 |
CHR ''W'', CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', Char Nibble5 NibbleC, |
|
105 |
CHR '']'', CHR ''^'', CHR ''_'', Char Nibble6 Nibble0, CHR ''a'', CHR ''b'', |
|
106 |
CHR ''c'', CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', CHR ''h'', CHR ''i'', |
|
107 |
CHR ''j'', CHR ''k'', CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', CHR ''p'', |
|
108 |
CHR ''q'', CHR ''r'', CHR ''s'', CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', |
|
109 |
CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', CHR ''|'', CHR ''}'', CHR ''~'', |
|
110 |
Char Nibble7 NibbleF, Char Nibble8 Nibble0, Char Nibble8 Nibble1, |
|
111 |
Char Nibble8 Nibble2, Char Nibble8 Nibble3, Char Nibble8 Nibble4, |
|
112 |
Char Nibble8 Nibble5, Char Nibble8 Nibble6, Char Nibble8 Nibble7, |
|
113 |
Char Nibble8 Nibble8, Char Nibble8 Nibble9, Char Nibble8 NibbleA, |
|
114 |
Char Nibble8 NibbleB, Char Nibble8 NibbleC, Char Nibble8 NibbleD, |
|
115 |
Char Nibble8 NibbleE, Char Nibble8 NibbleF, Char Nibble9 Nibble0, |
|
116 |
Char Nibble9 Nibble1, Char Nibble9 Nibble2, Char Nibble9 Nibble3, |
|
117 |
Char Nibble9 Nibble4, Char Nibble9 Nibble5, Char Nibble9 Nibble6, |
|
118 |
Char Nibble9 Nibble7, Char Nibble9 Nibble8, Char Nibble9 Nibble9, |
|
119 |
Char Nibble9 NibbleA, Char Nibble9 NibbleB, Char Nibble9 NibbleC, |
|
120 |
Char Nibble9 NibbleD, Char Nibble9 NibbleE, Char Nibble9 NibbleF, |
|
121 |
Char NibbleA Nibble0, Char NibbleA Nibble1, Char NibbleA Nibble2, |
|
122 |
Char NibbleA Nibble3, Char NibbleA Nibble4, Char NibbleA Nibble5, |
|
123 |
Char NibbleA Nibble6, Char NibbleA Nibble7, Char NibbleA Nibble8, |
|
124 |
Char NibbleA Nibble9, Char NibbleA NibbleA, Char NibbleA NibbleB, |
|
125 |
Char NibbleA NibbleC, Char NibbleA NibbleD, Char NibbleA NibbleE, |
|
126 |
Char NibbleA NibbleF, Char NibbleB Nibble0, Char NibbleB Nibble1, |
|
127 |
Char NibbleB Nibble2, Char NibbleB Nibble3, Char NibbleB Nibble4, |
|
128 |
Char NibbleB Nibble5, Char NibbleB Nibble6, Char NibbleB Nibble7, |
|
129 |
Char NibbleB Nibble8, Char NibbleB Nibble9, Char NibbleB NibbleA, |
|
130 |
Char NibbleB NibbleB, Char NibbleB NibbleC, Char NibbleB NibbleD, |
|
131 |
Char NibbleB NibbleE, Char NibbleB NibbleF, Char NibbleC Nibble0, |
|
132 |
Char NibbleC Nibble1, Char NibbleC Nibble2, Char NibbleC Nibble3, |
|
133 |
Char NibbleC Nibble4, Char NibbleC Nibble5, Char NibbleC Nibble6, |
|
134 |
Char NibbleC Nibble7, Char NibbleC Nibble8, Char NibbleC Nibble9, |
|
135 |
Char NibbleC NibbleA, Char NibbleC NibbleB, Char NibbleC NibbleC, |
|
136 |
Char NibbleC NibbleD, Char NibbleC NibbleE, Char NibbleC NibbleF, |
|
137 |
Char NibbleD Nibble0, Char NibbleD Nibble1, Char NibbleD Nibble2, |
|
138 |
Char NibbleD Nibble3, Char NibbleD Nibble4, Char NibbleD Nibble5, |
|
139 |
Char NibbleD Nibble6, Char NibbleD Nibble7, Char NibbleD Nibble8, |
|
140 |
Char NibbleD Nibble9, Char NibbleD NibbleA, Char NibbleD NibbleB, |
|
141 |
Char NibbleD NibbleC, Char NibbleD NibbleD, Char NibbleD NibbleE, |
|
142 |
Char NibbleD NibbleF, Char NibbleE Nibble0, Char NibbleE Nibble1, |
|
143 |
Char NibbleE Nibble2, Char NibbleE Nibble3, Char NibbleE Nibble4, |
|
144 |
Char NibbleE Nibble5, Char NibbleE Nibble6, Char NibbleE Nibble7, |
|
145 |
Char NibbleE Nibble8, Char NibbleE Nibble9, Char NibbleE NibbleA, |
|
146 |
Char NibbleE NibbleB, Char NibbleE NibbleC, Char NibbleE NibbleD, |
|
147 |
Char NibbleE NibbleE, Char NibbleE NibbleF, Char NibbleF Nibble0, |
|
148 |
Char NibbleF Nibble1, Char NibbleF Nibble2, Char NibbleF Nibble3, |
|
149 |
Char NibbleF Nibble4, Char NibbleF Nibble5, Char NibbleF Nibble6, |
|
150 |
Char NibbleF Nibble7, Char NibbleF Nibble8, Char NibbleF Nibble9, |
|
151 |
Char NibbleF NibbleA, Char NibbleF NibbleB, Char NibbleF NibbleC, |
|
152 |
Char NibbleF NibbleD, Char NibbleF NibbleE, Char NibbleF NibbleF]" |
|
153 |
||
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
154 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
155 |
subsection {* Strings as dedicated datatype *} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
156 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
157 |
datatype literal = STR string |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
158 |
|
33237 | 159 |
declare literal.cases [code del] literal.recs [code del] |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
160 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
161 |
lemma [code]: "size (s\<Colon>literal) = 0" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
162 |
by (cases s) simp_all |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
163 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
164 |
lemma [code]: "literal_size (s\<Colon>literal) = 0" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
165 |
by (cases s) simp_all |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
166 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
167 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
168 |
subsection {* Code generator *} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
169 |
|
31055
2cf6efca6c71
proper structures for list and string code generation stuff
haftmann
parents:
31051
diff
changeset
|
170 |
use "Tools/string_code.ML" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
171 |
|
33237 | 172 |
code_reserved SML string |
173 |
code_reserved OCaml string |
|
34886 | 174 |
code_reserved Scala string |
33237 | 175 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
176 |
code_type literal |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
177 |
(SML "string") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
178 |
(OCaml "string") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
179 |
(Haskell "String") |
34886 | 180 |
(Scala "String") |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
181 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
182 |
setup {* |
34886 | 183 |
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"] |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
184 |
*} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
185 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37743
diff
changeset
|
186 |
code_instance literal :: equal |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
187 |
(Haskell -) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
188 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37743
diff
changeset
|
189 |
code_const "HOL.equal \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
190 |
(SML "!((_ : string) = _)") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
191 |
(OCaml "!((_ : string) = _)") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
192 |
(Haskell infixl 4 "==") |
34886 | 193 |
(Scala infixl 5 "==") |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
194 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
195 |
types_code |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
196 |
"char" ("string") |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
197 |
attach (term_of) {* |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
198 |
val term_of_char = HOLogic.mk_char o ord; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
199 |
*} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
200 |
attach (test) {* |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
201 |
fun gen_char i = |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
202 |
let val j = random_range (ord "a") (Int.min (ord "a" + i, ord "z")) |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
203 |
in (chr j, fn () => HOLogic.mk_char j) end; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
204 |
*} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
205 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
206 |
setup {* |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
207 |
let |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
208 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
209 |
fun char_codegen thy defs dep thyname b t gr = |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
210 |
let |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
211 |
val i = HOLogic.dest_char t; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
212 |
val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
213 |
(fastype_of t) gr; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
214 |
in SOME (Codegen.str (ML_Syntax.print_string (chr i)), gr') |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
215 |
end handle TERM _ => NONE; |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
216 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
217 |
in Codegen.add_codegen "char_codegen" char_codegen end |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
218 |
*} |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
219 |
|
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
35123
diff
changeset
|
220 |
hide_type (open) literal |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
221 |
|
37743 | 222 |
|
223 |
text {* Code generator setup *} |
|
224 |
||
225 |
code_modulename SML |
|
226 |
String String |
|
227 |
||
228 |
code_modulename OCaml |
|
229 |
String String |
|
230 |
||
231 |
code_modulename Haskell |
|
232 |
String String |
|
233 |
||
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
234 |
end |