| author | wenzelm | 
| Tue, 10 Jan 2017 16:03:50 +0100 | |
| changeset 64865 | 778c64c17363 | 
| parent 62597 | b3f2b8c906a6 | 
| child 67620 | 3817a93a3e5e | 
| 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  | 
|
| 62597 | 19  | 
fun decode_char_nonzero t =  | 
20  | 
case Numeral.dest_num t of  | 
|
21  | 
SOME n => if 0 < n andalso n < 256 then SOME n else NONE  | 
|
22  | 
| _ => NONE;  | 
|
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
23  | 
|
| 62597 | 24  | 
fun decode_char (IConst { sym = Code_Symbol.Constant @{const_name Groups.zero}, ... }) =
 | 
25  | 
SOME 0  | 
|
26  | 
  | decode_char (IConst { sym = Code_Symbol.Constant @{const_name Char}, ... } `$ t) =
 | 
|
27  | 
decode_char_nonzero t  | 
|
28  | 
| decode_char _ = NONE  | 
|
29  | 
||
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
30  | 
fun implode_string literals ts =  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
31  | 
let  | 
| 62597 | 32  | 
val is = map_filter decode_char ts;  | 
33  | 
in if length ts = length is  | 
|
34  | 
then (SOME o Code_Printer.str o Code_Printer.literal_string literals o implode o map chr) is  | 
|
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
35  | 
else NONE  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
36  | 
end;  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
fun add_literal_list_string target =  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
39  | 
let  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
40  | 
fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] =  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
41  | 
case Option.map (cons t1) (List_Code.implode_list t2)  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
42  | 
of SOME ts => (case implode_string literals ts  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
43  | 
of SOME p => p  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
| NONE =>  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
45  | 
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
 | 
46  | 
| NONE =>  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
47  | 
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
 | 
48  | 
in  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51160 
diff
changeset
 | 
49  | 
    Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
 | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
50  | 
[(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))]))  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
51  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
52  | 
|
| 55236 | 53  | 
fun add_literal_char target thy =  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
54  | 
let  | 
| 62597 | 55  | 
fun pr literals =  | 
56  | 
Code_Printer.str o Code_Printer.literal_char literals o chr;  | 
|
57  | 
fun pretty_zero literals _ _ _ _ [] =  | 
|
58  | 
pr literals 0  | 
|
59  | 
fun pretty_Char literals _ thm _ _ [(t, _)] =  | 
|
60  | 
case decode_char_nonzero t  | 
|
61  | 
of SOME i => pr literals i  | 
|
| 55236 | 62  | 
| NONE => Code_Printer.eqn_error thy 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
 | 
63  | 
in  | 
| 62597 | 64  | 
thy  | 
65  | 
    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name String.zero_char_inst.zero_char},
 | 
|
66  | 
[(target, SOME (Code_Printer.complex_const_syntax (0, pretty_zero)))]))  | 
|
67  | 
    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name Char},
 | 
|
68  | 
[(target, SOME (Code_Printer.complex_const_syntax (1, pretty_Char)))]))  | 
|
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
69  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
70  | 
|
| 55236 | 71  | 
fun add_literal_string target thy =  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
72  | 
let  | 
| 
55148
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
73  | 
fun pretty literals _ thm _ _ [(t, _)] =  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
74  | 
case List_Code.implode_list t  | 
| 
 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 
haftmann 
parents: 
55147 
diff
changeset
 | 
75  | 
of SOME ts => (case implode_string literals ts  | 
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
76  | 
of SOME p => p  | 
| 55236 | 77  | 
| NONE => Code_Printer.eqn_error thy thm "Illegal string literal expression")  | 
78  | 
| NONE => Code_Printer.eqn_error thy thm "Illegal string literal 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
 | 
79  | 
in  | 
| 62597 | 80  | 
thy  | 
81  | 
    |> Code_Target.set_printings (Code_Symbol.Constant (@{const_name STR},
 | 
|
82  | 
[(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))]))  | 
|
| 
31055
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
83  | 
end;  | 
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
2cf6efca6c71
proper structures for list and string code generation stuff
 
haftmann 
parents:  
diff
changeset
 | 
85  | 
end;  |