author | haftmann |
Tue, 28 Jan 2025 13:02:42 +0100 | |
changeset 81999 | 513f8fa74c82 |
parent 81761 | a1dc03194053 |
permissions | -rw-r--r-- |
81761 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
2 |
||
3 |
section \<open>Test of target-language string literal operations\<close> |
|
4 |
||
5 |
theory Generate_Target_String_Literals |
|
6 |
imports |
|
7 |
"HOL-Library.Code_Test" |
|
8 |
begin |
|
9 |
||
81999
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
10 |
context |
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
11 |
begin |
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
12 |
|
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
13 |
qualified definition computations where |
81761 | 14 |
\<open>computations = ( |
15 |
STR ''abc'' + STR 0x20 + STR ''def'', |
|
16 |
String.implode ''abc'', |
|
17 |
String.explode STR ''abc'', |
|
18 |
String.literal_of_asciis [10, 20, 30, 40, 1111, 2222, 3333], |
|
19 |
size (STR ''abc''), |
|
20 |
STR ''def'' \<le> STR ''abc'', |
|
21 |
STR ''abc'' < STR ''def'' |
|
22 |
)\<close> |
|
23 |
||
81999
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
24 |
qualified definition results where |
81761 | 25 |
\<open>results = ( |
26 |
STR ''abc def'', |
|
27 |
STR ''abc'', |
|
28 |
''abc'', |
|
29 |
STR ''\<newline>'' + STR 0x14 + STR 0x1E + STR ''(W.'' + STR 0x05, |
|
30 |
3, |
|
31 |
False, |
|
32 |
True |
|
33 |
)\<close> |
|
34 |
||
81999
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
35 |
qualified definition check where |
81761 | 36 |
\<open>check \<longleftrightarrow> computations = results\<close> |
37 |
||
38 |
lemma check |
|
39 |
by code_simp |
|
40 |
||
41 |
lemma check |
|
42 |
by normalization |
|
43 |
||
44 |
lemma check |
|
45 |
by eval |
|
46 |
||
47 |
test_code check in Scala |
|
48 |
||
49 |
end |
|
81999
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
50 |
|
513f8fa74c82
more explicit tests for non-PolyML SML platforms
haftmann
parents:
81761
diff
changeset
|
51 |
end |