| author | Andreas Lochbihler | 
| Fri, 27 Sep 2013 08:59:22 +0200 | |
| changeset 53943 | 2b761d9a74f5 | 
| parent 52435 | 6646bb548c6b | 
| child 55147 | bce3dbc11f95 | 
| permissions | -rw-r--r-- | 
| 37744 | 1  | 
(* Title: HOL/Tools/string_code.ML  | 
2  | 
Author: Florian Haftmann, TU Muenchen  | 
|
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
Code generation for character and string literals.  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
signature STRING_CODE =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
9  | 
val add_literal_list_string: string -> theory -> theory  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
10  | 
val add_literal_char: string -> theory -> theory  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31055 
diff
changeset
 | 
11  | 
val add_literal_string: string -> theory -> theory  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
12  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
14  | 
structure String_Code : STRING_CODE =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
15  | 
struct  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
open Basic_Code_Thingol;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
fun decode_char nibbles' tt =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
20  | 
let  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
21  | 
fun idx c = find_index (curry (op =) c) nibbles';  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
22  | 
fun decode ~1 _ = NONE  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
23  | 
| decode _ ~1 = NONE  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
24  | 
| decode n m = SOME (chr (n * 16 + m));  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
25  | 
in case tt  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
38923 
diff
changeset
 | 
26  | 
   of (IConst { name = c1, ... }, IConst { name = c2, ... }) => decode (idx c1) (idx c2)
 | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
27  | 
| _ => NONE  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
28  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
|
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
48072 
diff
changeset
 | 
30  | 
fun implode_string literals char' nibbles' ts =  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
31  | 
let  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
38923 
diff
changeset
 | 
32  | 
    fun implode_char (IConst { name = c, ... } `$ t1 `$ t2) =
 | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
33  | 
if c = char' then decode_char nibbles' (t1, t2) else NONE  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
34  | 
| implode_char _ = NONE;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
35  | 
val ts' = map_filter implode_char ts;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
36  | 
in if length ts = length ts'  | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
48072 
diff
changeset
 | 
37  | 
then (SOME o Code_Printer.str o Code_Printer.literal_string literals o implode) ts'  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
else NONE  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
39  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
41  | 
val cs_nibbles = [@{const_name Nibble0}, @{const_name Nibble1},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
42  | 
  @{const_name Nibble2}, @{const_name Nibble3},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
43  | 
  @{const_name Nibble4}, @{const_name Nibble5},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
  @{const_name Nibble6}, @{const_name Nibble7},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
45  | 
  @{const_name Nibble8}, @{const_name Nibble9},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
46  | 
  @{const_name NibbleA}, @{const_name NibbleB},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
47  | 
  @{const_name NibbleC}, @{const_name NibbleD},
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
48  | 
  @{const_name NibbleE}, @{const_name NibbleF}];
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
val cs_summa = [@{const_name Nil}, @{const_name Cons}, @{const_name Char}] @ cs_nibbles;
 | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
51  | 
fun add_literal_list_string target =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
52  | 
let  | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
48072 
diff
changeset
 | 
53  | 
fun pretty literals (nil' :: cons' :: char' :: nibbles') pr _ vars fxy [(t1, _), (t2, _)] =  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
54  | 
case Option.map (cons t1) (List_Code.implode_list nil' cons' t2)  | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
48072 
diff
changeset
 | 
55  | 
of SOME ts => (case implode_string literals char' nibbles' ts  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
56  | 
of SOME p => p  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
57  | 
| NONE =>  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
58  | 
Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts))  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
59  | 
| NONE =>  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
60  | 
List_Code.default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
61  | 
in  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
62  | 
    Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
63  | 
[(target, SOME (Code_Printer.complex_const_syntax (2, (cs_summa, pretty))))]))  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
64  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
66  | 
fun add_literal_char target =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
67  | 
let  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
68  | 
fun pretty literals nibbles' _ thm _ _ [(t1, _), (t2, _)] =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
69  | 
case decode_char nibbles' (t1, t2)  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
70  | 
of SOME c => (Code_Printer.str o Code_Printer.literal_char literals) c  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
31205 
diff
changeset
 | 
71  | 
| NONE => Code_Printer.eqn_error thm "Illegal character expression";  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
72  | 
in  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
73  | 
    Code_Target.set_printings (Code_Symbol.Constant (@{const_name Char},
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
74  | 
[(target, SOME (Code_Printer.complex_const_syntax (2, (cs_nibbles, pretty))))]))  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
75  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
76  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31055 
diff
changeset
 | 
77  | 
fun add_literal_string target =  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
78  | 
let  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
79  | 
fun pretty literals (nil' :: cons' :: char' :: nibbles') _ thm _ _ [(t, _)] =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
80  | 
case List_Code.implode_list nil' cons' t  | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
48072 
diff
changeset
 | 
81  | 
of SOME ts => (case implode_string literals char' nibbles' ts  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
82  | 
of SOME p => p  | 
| 
33989
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
31205 
diff
changeset
 | 
83  | 
| NONE => Code_Printer.eqn_error thm "Illegal message expression")  | 
| 
 
cb136b5f6050
more speaking function names for Code_Printer; added doublesemicolon
 
haftmann 
parents: 
31205 
diff
changeset
 | 
84  | 
| NONE => Code_Printer.eqn_error thm "Illegal message expression";  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
85  | 
in  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
86  | 
    Code_Target.set_printings (Code_Symbol.Constant (@{const_name STR},
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
87  | 
[(target, SOME (Code_Printer.complex_const_syntax (1, (cs_summa, pretty))))]))  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
88  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
90  | 
end;  |