| author | wenzelm | 
| Sat, 09 Jan 2010 16:31:19 +0100 | |
| changeset 34296 | 5f454603228b | 
| parent 33237 | 565ad811db21 | 
| child 34886 | 873c31d9f10d | 
| 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  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents: 
31051 
diff
changeset
 | 
8  | 
"Tools/string_syntax.ML"  | 
| 
 
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  | 
| 31176 | 56  | 
PureThy.note_thmss Thm.definitionK [((Binding.name "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  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
81  | 
setup StringSyntax.setup  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
82  | 
|
| 31484 | 83  | 
definition chars :: string where  | 
84  | 
"chars = [Char Nibble0 Nibble0, Char Nibble0 Nibble1, Char Nibble0 Nibble2,  | 
|
85  | 
Char Nibble0 Nibble3, Char Nibble0 Nibble4, Char Nibble0 Nibble5,  | 
|
86  | 
Char Nibble0 Nibble6, Char Nibble0 Nibble7, Char Nibble0 Nibble8,  | 
|
87  | 
Char Nibble0 Nibble9, Char Nibble0 NibbleA, Char Nibble0 NibbleB,  | 
|
88  | 
Char Nibble0 NibbleC, Char Nibble0 NibbleD, Char Nibble0 NibbleE,  | 
|
89  | 
Char Nibble0 NibbleF, Char Nibble1 Nibble0, Char Nibble1 Nibble1,  | 
|
90  | 
Char Nibble1 Nibble2, Char Nibble1 Nibble3, Char Nibble1 Nibble4,  | 
|
91  | 
Char Nibble1 Nibble5, Char Nibble1 Nibble6, Char Nibble1 Nibble7,  | 
|
92  | 
Char Nibble1 Nibble8, Char Nibble1 Nibble9, Char Nibble1 NibbleA,  | 
|
93  | 
Char Nibble1 NibbleB, Char Nibble1 NibbleC, Char Nibble1 NibbleD,  | 
|
94  | 
Char Nibble1 NibbleE, Char Nibble1 NibbleF, CHR '' '', CHR ''!'',  | 
|
95  | 
Char Nibble2 Nibble2, CHR ''#'', CHR ''$'', CHR ''%'', CHR ''&'',  | 
|
96  | 
  Char Nibble2 Nibble7, CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', CHR '','',
 | 
|
97  | 
CHR ''-'', CHR ''.'', CHR ''/'', CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'',  | 
|
98  | 
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9'', CHR '':'',  | 
|
99  | 
CHR '';'', CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', CHR ''@'', CHR ''A'',  | 
|
100  | 
CHR ''B'', CHR ''C'', CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', CHR ''H'',  | 
|
101  | 
CHR ''I'', CHR ''J'', CHR ''K'', CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'',  | 
|
102  | 
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', CHR ''T'', CHR ''U'', CHR ''V'',  | 
|
103  | 
CHR ''W'', CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', Char Nibble5 NibbleC,  | 
|
104  | 
CHR '']'', CHR ''^'', CHR ''_'', Char Nibble6 Nibble0, CHR ''a'', CHR ''b'',  | 
|
105  | 
CHR ''c'', CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', CHR ''h'', CHR ''i'',  | 
|
106  | 
CHR ''j'', CHR ''k'', CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', CHR ''p'',  | 
|
107  | 
CHR ''q'', CHR ''r'', CHR ''s'', CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'',  | 
|
108  | 
  CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', CHR ''|'', CHR ''}'', CHR ''~'',
 | 
|
109  | 
Char Nibble7 NibbleF, Char Nibble8 Nibble0, Char Nibble8 Nibble1,  | 
|
110  | 
Char Nibble8 Nibble2, Char Nibble8 Nibble3, Char Nibble8 Nibble4,  | 
|
111  | 
Char Nibble8 Nibble5, Char Nibble8 Nibble6, Char Nibble8 Nibble7,  | 
|
112  | 
Char Nibble8 Nibble8, Char Nibble8 Nibble9, Char Nibble8 NibbleA,  | 
|
113  | 
Char Nibble8 NibbleB, Char Nibble8 NibbleC, Char Nibble8 NibbleD,  | 
|
114  | 
Char Nibble8 NibbleE, Char Nibble8 NibbleF, Char Nibble9 Nibble0,  | 
|
115  | 
Char Nibble9 Nibble1, Char Nibble9 Nibble2, Char Nibble9 Nibble3,  | 
|
116  | 
Char Nibble9 Nibble4, Char Nibble9 Nibble5, Char Nibble9 Nibble6,  | 
|
117  | 
Char Nibble9 Nibble7, Char Nibble9 Nibble8, Char Nibble9 Nibble9,  | 
|
118  | 
Char Nibble9 NibbleA, Char Nibble9 NibbleB, Char Nibble9 NibbleC,  | 
|
119  | 
Char Nibble9 NibbleD, Char Nibble9 NibbleE, Char Nibble9 NibbleF,  | 
|
120  | 
Char NibbleA Nibble0, Char NibbleA Nibble1, Char NibbleA Nibble2,  | 
|
121  | 
Char NibbleA Nibble3, Char NibbleA Nibble4, Char NibbleA Nibble5,  | 
|
122  | 
Char NibbleA Nibble6, Char NibbleA Nibble7, Char NibbleA Nibble8,  | 
|
123  | 
Char NibbleA Nibble9, Char NibbleA NibbleA, Char NibbleA NibbleB,  | 
|
124  | 
Char NibbleA NibbleC, Char NibbleA NibbleD, Char NibbleA NibbleE,  | 
|
125  | 
Char NibbleA NibbleF, Char NibbleB Nibble0, Char NibbleB Nibble1,  | 
|
126  | 
Char NibbleB Nibble2, Char NibbleB Nibble3, Char NibbleB Nibble4,  | 
|
127  | 
Char NibbleB Nibble5, Char NibbleB Nibble6, Char NibbleB Nibble7,  | 
|
128  | 
Char NibbleB Nibble8, Char NibbleB Nibble9, Char NibbleB NibbleA,  | 
|
129  | 
Char NibbleB NibbleB, Char NibbleB NibbleC, Char NibbleB NibbleD,  | 
|
130  | 
Char NibbleB NibbleE, Char NibbleB NibbleF, Char NibbleC Nibble0,  | 
|
131  | 
Char NibbleC Nibble1, Char NibbleC Nibble2, Char NibbleC Nibble3,  | 
|
132  | 
Char NibbleC Nibble4, Char NibbleC Nibble5, Char NibbleC Nibble6,  | 
|
133  | 
Char NibbleC Nibble7, Char NibbleC Nibble8, Char NibbleC Nibble9,  | 
|
134  | 
Char NibbleC NibbleA, Char NibbleC NibbleB, Char NibbleC NibbleC,  | 
|
135  | 
Char NibbleC NibbleD, Char NibbleC NibbleE, Char NibbleC NibbleF,  | 
|
136  | 
Char NibbleD Nibble0, Char NibbleD Nibble1, Char NibbleD Nibble2,  | 
|
137  | 
Char NibbleD Nibble3, Char NibbleD Nibble4, Char NibbleD Nibble5,  | 
|
138  | 
Char NibbleD Nibble6, Char NibbleD Nibble7, Char NibbleD Nibble8,  | 
|
139  | 
Char NibbleD Nibble9, Char NibbleD NibbleA, Char NibbleD NibbleB,  | 
|
140  | 
Char NibbleD NibbleC, Char NibbleD NibbleD, Char NibbleD NibbleE,  | 
|
141  | 
Char NibbleD NibbleF, Char NibbleE Nibble0, Char NibbleE Nibble1,  | 
|
142  | 
Char NibbleE Nibble2, Char NibbleE Nibble3, Char NibbleE Nibble4,  | 
|
143  | 
Char NibbleE Nibble5, Char NibbleE Nibble6, Char NibbleE Nibble7,  | 
|
144  | 
Char NibbleE Nibble8, Char NibbleE Nibble9, Char NibbleE NibbleA,  | 
|
145  | 
Char NibbleE NibbleB, Char NibbleE NibbleC, Char NibbleE NibbleD,  | 
|
146  | 
Char NibbleE NibbleE, Char NibbleE NibbleF, Char NibbleF Nibble0,  | 
|
147  | 
Char NibbleF Nibble1, Char NibbleF Nibble2, Char NibbleF Nibble3,  | 
|
148  | 
Char NibbleF Nibble4, Char NibbleF Nibble5, Char NibbleF Nibble6,  | 
|
149  | 
Char NibbleF Nibble7, Char NibbleF Nibble8, Char NibbleF Nibble9,  | 
|
150  | 
Char NibbleF NibbleA, Char NibbleF NibbleB, Char NibbleF NibbleC,  | 
|
151  | 
Char NibbleF NibbleD, Char NibbleF NibbleE, Char NibbleF NibbleF]"  | 
|
152  | 
||
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
154  | 
subsection {* Strings as dedicated datatype *}
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
155  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
156  | 
datatype literal = STR string  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
157  | 
|
| 33237 | 158  | 
declare literal.cases [code del] literal.recs [code del]  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
159  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
160  | 
lemma [code]: "size (s\<Colon>literal) = 0"  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
161  | 
by (cases s) simp_all  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
162  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
163  | 
lemma [code]: "literal_size (s\<Colon>literal) = 0"  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
164  | 
by (cases s) simp_all  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
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  | 
subsection {* Code generator *}
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
168  | 
|
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents: 
31051 
diff
changeset
 | 
169  | 
use "Tools/string_code.ML"  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
170  | 
|
| 33237 | 171  | 
code_reserved SML string  | 
172  | 
code_reserved OCaml string  | 
|
173  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
174  | 
code_type literal  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
175  | 
(SML "string")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
176  | 
(OCaml "string")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
177  | 
(Haskell "String")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
178  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
179  | 
setup {*
 | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
180  | 
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell"]  | 
| 
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  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
183  | 
code_instance literal :: eq  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
184  | 
(Haskell -)  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
185  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
186  | 
code_const "eq_class.eq \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool"  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
187  | 
(SML "!((_ : string) = _)")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
188  | 
(OCaml "!((_ : string) = _)")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
189  | 
(Haskell infixl 4 "==")  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
192  | 
types_code  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
193  | 
  "char" ("string")
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
194  | 
attach (term_of) {*
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
195  | 
val term_of_char = HOLogic.mk_char o ord;  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
196  | 
*}  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
197  | 
attach (test) {*
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
198  | 
fun gen_char i =  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
199  | 
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
 | 
200  | 
in (chr j, fn () => HOLogic.mk_char j) end;  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
201  | 
*}  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
203  | 
setup {*
 | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
204  | 
let  | 
| 
 
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  | 
fun char_codegen thy defs dep thyname b t gr =  | 
| 
 
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  | 
val i = HOLogic.dest_char t;  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
209  | 
val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
210  | 
(fastype_of t) gr;  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
211  | 
in SOME (Codegen.str (ML_Syntax.print_string (chr i)), gr')  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
212  | 
end handle TERM _ => NONE;  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
214  | 
in Codegen.add_codegen "char_codegen" char_codegen end  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
215  | 
*}  | 
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
216  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
217  | 
hide (open) type literal  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
218  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
219  | 
end  |