author | Fabian Huch <huch@in.tum.de> |
Fri, 17 Jan 2025 12:17:37 +0100 | |
changeset 81822 | e7be7c4b871c |
parent 81761 | a1dc03194053 |
child 81986 | 3863850f4b0e |
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 |
|
60758 | 3 |
section \<open>Character and string types\<close> |
31051
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 |
74101 | 6 |
imports Enum Bit_Operations Code_Numeral |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
7 |
begin |
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
8 |
|
68028 | 9 |
subsection \<open>Strings as list of bytes\<close> |
10 |
||
11 |
text \<open> |
|
12 |
When modelling strings, we follow the approach given |
|
69593 | 13 |
in \<^url>\<open>https://utf8everywhere.org/\<close>: |
68028 | 14 |
|
15 |
\<^item> Strings are a list of bytes (8 bit). |
|
16 |
||
17 |
\<^item> Byte values from 0 to 127 are US-ASCII. |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
18 |
|
68028 | 19 |
\<^item> Byte values from 128 to 255 are uninterpreted blobs. |
20 |
\<close> |
|
21 |
||
22 |
subsubsection \<open>Bytes as datatype\<close> |
|
23 |
||
24 |
datatype char = |
|
25 |
Char (digit0: bool) (digit1: bool) (digit2: bool) (digit3: bool) |
|
26 |
(digit4: bool) (digit5: bool) (digit6: bool) (digit7: bool) |
|
27 |
||
28 |
context comm_semiring_1 |
|
29 |
begin |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
30 |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
31 |
definition of_char :: \<open>char \<Rightarrow> 'a\<close> |
72024 | 32 |
where \<open>of_char c = horner_sum of_bool 2 [digit0 c, digit1 c, digit2 c, digit3 c, digit4 c, digit5 c, digit6 c, digit7 c]\<close> |
68028 | 33 |
|
34 |
lemma of_char_Char [simp]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
35 |
\<open>of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) = |
72024 | 36 |
horner_sum of_bool 2 [b0, b1, b2, b3, b4, b5, b6, b7]\<close> |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
37 |
by (simp add: of_char_def) |
68028 | 38 |
|
39 |
end |
|
40 |
||
75622
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
41 |
lemma (in comm_semiring_1) of_nat_of_char: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
42 |
\<open>of_nat (of_char c) = of_char c\<close> |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
43 |
by (cases c) simp |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
44 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
45 |
lemma (in comm_ring_1) of_int_of_char: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
46 |
\<open>of_int (of_char c) = of_char c\<close> |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
47 |
by (cases c) simp |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
48 |
|
75662 | 49 |
lemma nat_of_char [simp]: |
50 |
\<open>nat (of_char c) = of_char c\<close> |
|
51 |
by (cases c) (simp only: of_char_Char nat_horner_sum) |
|
52 |
||
75622
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
53 |
|
78955 | 54 |
context linordered_euclidean_semiring_bit_operations |
68028 | 55 |
begin |
56 |
||
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
57 |
definition char_of :: \<open>'a \<Rightarrow> char\<close> |
75085 | 58 |
where \<open>char_of n = Char (bit n 0) (bit n 1) (bit n 2) (bit n 3) (bit n 4) (bit n 5) (bit n 6) (bit n 7)\<close> |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
59 |
|
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
60 |
lemma char_of_take_bit_eq: |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
61 |
\<open>char_of (take_bit n m) = char_of m\<close> if \<open>n \<ge> 8\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
62 |
using that by (simp add: char_of_def bit_take_bit_iff) |
68028 | 63 |
|
64 |
lemma char_of_char [simp]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
65 |
\<open>char_of (of_char c) = c\<close> |
72024 | 66 |
by (simp only: of_char_def char_of_def bit_horner_sum_bit_iff) simp |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
67 |
|
68028 | 68 |
lemma char_of_comp_of_char [simp]: |
69 |
"char_of \<circ> of_char = id" |
|
70 |
by (simp add: fun_eq_iff) |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
71 |
|
68028 | 72 |
lemma inj_of_char: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
73 |
\<open>inj of_char\<close> |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
74 |
proof (rule injI) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
75 |
fix c d |
68028 | 76 |
assume "of_char c = of_char d" |
77 |
then have "char_of (of_char c) = char_of (of_char d)" |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
78 |
by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
79 |
then show "c = d" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
80 |
by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
81 |
qed |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
82 |
|
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
83 |
lemma of_char_eqI: |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
84 |
\<open>c = d\<close> if \<open>of_char c = of_char d\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
85 |
using that inj_of_char by (simp add: inj_eq) |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
86 |
|
68028 | 87 |
lemma of_char_eq_iff [simp]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
88 |
\<open>of_char c = of_char d \<longleftrightarrow> c = d\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
89 |
by (auto intro: of_char_eqI) |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
90 |
|
68028 | 91 |
lemma of_char_of [simp]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
92 |
\<open>of_char (char_of a) = a mod 256\<close> |
68028 | 93 |
proof - |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
94 |
have \<open>[0..<8] = [0, Suc 0, 2, 3, 4, 5, 6, 7 :: nat]\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
95 |
by (simp add: upt_eq_Cons_conv) |
75085 | 96 |
then have \<open>[bit a 0, bit a 1, bit a 2, bit a 3, bit a 4, bit a 5, bit a 6, bit a 7] = map (bit a) [0..<8]\<close> |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
97 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
98 |
then have \<open>of_char (char_of a) = take_bit 8 a\<close> |
72024 | 99 |
by (simp only: char_of_def of_char_def char.sel horner_sum_bit_eq_take_bit) |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
100 |
then show ?thesis |
68028 | 101 |
by (simp add: take_bit_eq_mod) |
102 |
qed |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
103 |
|
68028 | 104 |
lemma char_of_mod_256 [simp]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
105 |
\<open>char_of (n mod 256) = char_of n\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
106 |
by (rule of_char_eqI) simp |
68028 | 107 |
|
108 |
lemma of_char_mod_256 [simp]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
109 |
\<open>of_char c mod 256 = of_char c\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
110 |
proof - |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
111 |
have \<open>of_char (char_of (of_char c)) mod 256 = of_char (char_of (of_char c))\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
112 |
by (simp only: of_char_of) simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
113 |
then show ?thesis |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
114 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
115 |
qed |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
116 |
|
68028 | 117 |
lemma char_of_quasi_inj [simp]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
118 |
\<open>char_of m = char_of n \<longleftrightarrow> m mod 256 = n mod 256\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
119 |
proof |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
120 |
assume ?Q |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
121 |
then show ?P |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
122 |
by (auto intro: of_char_eqI) |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
123 |
next |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
124 |
assume ?P |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
125 |
then have \<open>of_char (char_of m) = of_char (char_of n)\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
126 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
127 |
then show ?Q |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
128 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
129 |
qed |
68028 | 130 |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
131 |
lemma char_of_eq_iff: |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
132 |
\<open>char_of n = c \<longleftrightarrow> take_bit 8 n = of_char c\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
133 |
by (auto intro: of_char_eqI simp add: take_bit_eq_mod) |
68028 | 134 |
|
135 |
lemma char_of_nat [simp]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
136 |
\<open>char_of (of_nat n) = char_of n\<close> |
74309
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
haftmann
parents:
74108
diff
changeset
|
137 |
by (simp add: char_of_def String.char_of_def drop_bit_of_nat bit_simps possible_bit_def) |
68033
ad4b8b6892c3
uniform tagging for printable and non-printable literals
haftmann
parents:
68028
diff
changeset
|
138 |
|
68028 | 139 |
end |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
140 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
141 |
lemma inj_on_char_of_nat [simp]: |
68028 | 142 |
"inj_on char_of {0::nat..<256}" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
143 |
by (rule inj_onI) simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
144 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
145 |
lemma nat_of_char_less_256 [simp]: |
68028 | 146 |
"of_char c < (256 :: nat)" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
147 |
proof - |
68028 | 148 |
have "of_char c mod (256 :: nat) < 256" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
149 |
by arith |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
150 |
then show ?thesis by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
151 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
152 |
|
68028 | 153 |
lemma range_nat_of_char: |
154 |
"range of_char = {0::nat..<256}" |
|
155 |
proof (rule; rule) |
|
156 |
fix n :: nat |
|
157 |
assume "n \<in> range of_char" |
|
158 |
then show "n \<in> {0..<256}" |
|
159 |
by auto |
|
160 |
next |
|
161 |
fix n :: nat |
|
162 |
assume "n \<in> {0..<256}" |
|
163 |
then have "n = of_char (char_of n)" |
|
164 |
by simp |
|
165 |
then show "n \<in> range of_char" |
|
166 |
by (rule range_eqI) |
|
167 |
qed |
|
168 |
||
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
169 |
lemma UNIV_char_of_nat: |
68028 | 170 |
"UNIV = char_of ` {0::nat..<256}" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
171 |
proof - |
68028 | 172 |
have "range (of_char :: char \<Rightarrow> nat) = of_char ` char_of ` {0::nat..<256}" |
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
173 |
by (simp add: image_image range_nat_of_char) |
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
174 |
with inj_of_char [where ?'a = nat] show ?thesis |
68028 | 175 |
by (simp add: inj_image_eq_iff) |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
176 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
177 |
|
62597 | 178 |
lemma card_UNIV_char: |
179 |
"card (UNIV :: char set) = 256" |
|
180 |
by (auto simp add: UNIV_char_of_nat card_image) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
181 |
|
68028 | 182 |
context |
81113 | 183 |
includes lifting_syntax and integer.lifting and natural.lifting |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
184 |
begin |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
185 |
|
68028 | 186 |
lemma [transfer_rule]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
187 |
\<open>(pcr_integer ===> (=)) char_of char_of\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
188 |
by (unfold char_of_def) transfer_prover |
68028 | 189 |
|
190 |
lemma [transfer_rule]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
191 |
\<open>((=) ===> pcr_integer) of_char of_char\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
192 |
by (unfold of_char_def) transfer_prover |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
193 |
|
68028 | 194 |
lemma [transfer_rule]: |
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
195 |
\<open>(pcr_natural ===> (=)) char_of char_of\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
196 |
by (unfold char_of_def) transfer_prover |
68028 | 197 |
|
198 |
lemma [transfer_rule]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
199 |
\<open>((=) ===> pcr_natural) of_char of_char\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
200 |
by (unfold of_char_def) transfer_prover |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
201 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
202 |
end |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
203 |
|
68028 | 204 |
lifting_update integer.lifting |
205 |
lifting_forget integer.lifting |
|
64630
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
206 |
|
68028 | 207 |
lifting_update natural.lifting |
208 |
lifting_forget natural.lifting |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
209 |
|
75622
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
210 |
lemma size_char_eq_0 [simp, code]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
211 |
\<open>size c = 0\<close> for c :: char |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
212 |
by (cases c) simp |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
213 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
214 |
lemma size'_char_eq_0 [simp, code]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
215 |
\<open>size_char c = 0\<close> |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
216 |
by (cases c) simp |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
217 |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
218 |
syntax |
81124 | 219 |
"_Char" :: "str_position \<Rightarrow> char" (\<open>(\<open>open_block notation=\<open>literal char\<close>\<close>CHR _)\<close>) |
220 |
"_Char_ord" :: "num_const \<Rightarrow> char" (\<open>(\<open>open_block notation=\<open>literal char code\<close>\<close>CHR _)\<close>) |
|
80760 | 221 |
syntax_consts |
222 |
"_Char" "_Char_ord" \<rightleftharpoons> Char |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
223 |
|
42163
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
bulwahn
parents:
41750
diff
changeset
|
224 |
type_synonym string = "char list" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
225 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
226 |
syntax |
81124 | 227 |
"_String" :: "str_position \<Rightarrow> string" (\<open>(\<open>open_block notation=\<open>literal string\<close>\<close>_)\<close>) |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
228 |
|
69605 | 229 |
ML_file \<open>Tools/string_syntax.ML\<close> |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
230 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
231 |
instantiation char :: enum |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
232 |
begin |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
233 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
234 |
definition |
68028 | 235 |
"Enum.enum = [ |
236 |
CHR 0x00, CHR 0x01, CHR 0x02, CHR 0x03, |
|
62678 | 237 |
CHR 0x04, CHR 0x05, CHR 0x06, CHR 0x07, |
238 |
CHR 0x08, CHR 0x09, CHR ''\<newline>'', CHR 0x0B, |
|
239 |
CHR 0x0C, CHR 0x0D, CHR 0x0E, CHR 0x0F, |
|
240 |
CHR 0x10, CHR 0x11, CHR 0x12, CHR 0x13, |
|
241 |
CHR 0x14, CHR 0x15, CHR 0x16, CHR 0x17, |
|
242 |
CHR 0x18, CHR 0x19, CHR 0x1A, CHR 0x1B, |
|
243 |
CHR 0x1C, CHR 0x1D, CHR 0x1E, CHR 0x1F, |
|
244 |
CHR '' '', CHR ''!'', CHR 0x22, CHR ''#'', |
|
245 |
CHR ''$'', CHR ''%'', CHR ''&'', CHR 0x27, |
|
62597 | 246 |
CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', |
247 |
CHR '','', CHR ''-'', CHR ''.'', CHR ''/'', |
|
248 |
CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', |
|
249 |
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', |
|
250 |
CHR ''8'', CHR ''9'', CHR '':'', CHR '';'', |
|
251 |
CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', |
|
252 |
CHR ''@'', CHR ''A'', CHR ''B'', CHR ''C'', |
|
253 |
CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', |
|
254 |
CHR ''H'', CHR ''I'', CHR ''J'', CHR ''K'', |
|
255 |
CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', |
|
256 |
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', |
|
257 |
CHR ''T'', CHR ''U'', CHR ''V'', CHR ''W'', |
|
258 |
CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', |
|
62678 | 259 |
CHR 0x5C, CHR '']'', CHR ''^'', CHR ''_'', |
260 |
CHR 0x60, CHR ''a'', CHR ''b'', CHR ''c'', |
|
62597 | 261 |
CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', |
262 |
CHR ''h'', CHR ''i'', CHR ''j'', CHR ''k'', |
|
263 |
CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', |
|
264 |
CHR ''p'', CHR ''q'', CHR ''r'', CHR ''s'', |
|
265 |
CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', |
|
266 |
CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', |
|
62678 | 267 |
CHR ''|'', CHR ''}'', CHR ''~'', CHR 0x7F, |
268 |
CHR 0x80, CHR 0x81, CHR 0x82, CHR 0x83, |
|
269 |
CHR 0x84, CHR 0x85, CHR 0x86, CHR 0x87, |
|
270 |
CHR 0x88, CHR 0x89, CHR 0x8A, CHR 0x8B, |
|
271 |
CHR 0x8C, CHR 0x8D, CHR 0x8E, CHR 0x8F, |
|
272 |
CHR 0x90, CHR 0x91, CHR 0x92, CHR 0x93, |
|
273 |
CHR 0x94, CHR 0x95, CHR 0x96, CHR 0x97, |
|
274 |
CHR 0x98, CHR 0x99, CHR 0x9A, CHR 0x9B, |
|
275 |
CHR 0x9C, CHR 0x9D, CHR 0x9E, CHR 0x9F, |
|
276 |
CHR 0xA0, CHR 0xA1, CHR 0xA2, CHR 0xA3, |
|
277 |
CHR 0xA4, CHR 0xA5, CHR 0xA6, CHR 0xA7, |
|
278 |
CHR 0xA8, CHR 0xA9, CHR 0xAA, CHR 0xAB, |
|
279 |
CHR 0xAC, CHR 0xAD, CHR 0xAE, CHR 0xAF, |
|
280 |
CHR 0xB0, CHR 0xB1, CHR 0xB2, CHR 0xB3, |
|
281 |
CHR 0xB4, CHR 0xB5, CHR 0xB6, CHR 0xB7, |
|
282 |
CHR 0xB8, CHR 0xB9, CHR 0xBA, CHR 0xBB, |
|
283 |
CHR 0xBC, CHR 0xBD, CHR 0xBE, CHR 0xBF, |
|
284 |
CHR 0xC0, CHR 0xC1, CHR 0xC2, CHR 0xC3, |
|
285 |
CHR 0xC4, CHR 0xC5, CHR 0xC6, CHR 0xC7, |
|
286 |
CHR 0xC8, CHR 0xC9, CHR 0xCA, CHR 0xCB, |
|
287 |
CHR 0xCC, CHR 0xCD, CHR 0xCE, CHR 0xCF, |
|
288 |
CHR 0xD0, CHR 0xD1, CHR 0xD2, CHR 0xD3, |
|
289 |
CHR 0xD4, CHR 0xD5, CHR 0xD6, CHR 0xD7, |
|
290 |
CHR 0xD8, CHR 0xD9, CHR 0xDA, CHR 0xDB, |
|
291 |
CHR 0xDC, CHR 0xDD, CHR 0xDE, CHR 0xDF, |
|
292 |
CHR 0xE0, CHR 0xE1, CHR 0xE2, CHR 0xE3, |
|
293 |
CHR 0xE4, CHR 0xE5, CHR 0xE6, CHR 0xE7, |
|
294 |
CHR 0xE8, CHR 0xE9, CHR 0xEA, CHR 0xEB, |
|
295 |
CHR 0xEC, CHR 0xED, CHR 0xEE, CHR 0xEF, |
|
296 |
CHR 0xF0, CHR 0xF1, CHR 0xF2, CHR 0xF3, |
|
297 |
CHR 0xF4, CHR 0xF5, CHR 0xF6, CHR 0xF7, |
|
298 |
CHR 0xF8, CHR 0xF9, CHR 0xFA, CHR 0xFB, |
|
299 |
CHR 0xFC, CHR 0xFD, CHR 0xFE, CHR 0xFF]" |
|
31484 | 300 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
301 |
definition |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
302 |
"Enum.enum_all P \<longleftrightarrow> list_all P (Enum.enum :: char list)" |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
303 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
304 |
definition |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
305 |
"Enum.enum_ex P \<longleftrightarrow> list_ex P (Enum.enum :: char list)" |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
306 |
|
62597 | 307 |
lemma enum_char_unfold: |
68028 | 308 |
"Enum.enum = map char_of [0..<256]" |
62597 | 309 |
proof - |
68028 | 310 |
have "map (of_char :: char \<Rightarrow> nat) Enum.enum = [0..<256]" |
311 |
by (simp add: enum_char_def of_char_def upt_conv_Cons_Cons numeral_2_eq_2 [symmetric]) |
|
312 |
then have "map char_of (map (of_char :: char \<Rightarrow> nat) Enum.enum) = |
|
313 |
map char_of [0..<256]" |
|
314 |
by simp |
|
315 |
then show ?thesis |
|
316 |
by simp |
|
62597 | 317 |
qed |
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
318 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
319 |
instance proof |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
320 |
show UNIV: "UNIV = set (Enum.enum :: char list)" |
62597 | 321 |
by (simp add: enum_char_unfold UNIV_char_of_nat atLeast0LessThan) |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
322 |
show "distinct (Enum.enum :: char list)" |
62597 | 323 |
by (auto simp add: enum_char_unfold distinct_map intro: inj_onI) |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
324 |
show "\<And>P. Enum.enum_all P \<longleftrightarrow> Ball (UNIV :: char set) P" |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
325 |
by (simp add: UNIV enum_all_char_def list_all_iff) |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
326 |
show "\<And>P. Enum.enum_ex P \<longleftrightarrow> Bex (UNIV :: char set) P" |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
327 |
by (simp add: UNIV enum_ex_char_def list_ex_iff) |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
328 |
qed |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
329 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
330 |
end |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
331 |
|
68028 | 332 |
lemma linorder_char: |
333 |
"class.linorder (\<lambda>c d. of_char c \<le> (of_char d :: nat)) (\<lambda>c d. of_char c < (of_char d :: nat))" |
|
334 |
by standard auto |
|
335 |
||
336 |
text \<open>Optimized version for execution\<close> |
|
337 |
||
338 |
definition char_of_integer :: "integer \<Rightarrow> char" |
|
339 |
where [code_abbrev]: "char_of_integer = char_of" |
|
340 |
||
341 |
definition integer_of_char :: "char \<Rightarrow> integer" |
|
342 |
where [code_abbrev]: "integer_of_char = of_char" |
|
343 |
||
62597 | 344 |
lemma char_of_integer_code [code]: |
68028 | 345 |
"char_of_integer k = (let |
346 |
(q0, b0) = bit_cut_integer k; |
|
347 |
(q1, b1) = bit_cut_integer q0; |
|
348 |
(q2, b2) = bit_cut_integer q1; |
|
349 |
(q3, b3) = bit_cut_integer q2; |
|
350 |
(q4, b4) = bit_cut_integer q3; |
|
351 |
(q5, b5) = bit_cut_integer q4; |
|
352 |
(q6, b6) = bit_cut_integer q5; |
|
353 |
(_, b7) = bit_cut_integer q6 |
|
354 |
in Char b0 b1 b2 b3 b4 b5 b6 b7)" |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
355 |
by (simp add: bit_cut_integer_def char_of_integer_def char_of_def div_mult2_numeral_eq bit_iff_odd_drop_bit drop_bit_eq_div) |
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49834
diff
changeset
|
356 |
|
68028 | 357 |
lemma integer_of_char_code [code]: |
358 |
"integer_of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) = |
|
359 |
((((((of_bool b7 * 2 + of_bool b6) * 2 + |
|
360 |
of_bool b5) * 2 + of_bool b4) * 2 + |
|
361 |
of_bool b3) * 2 + of_bool b2) * 2 + |
|
362 |
of_bool b1) * 2 + of_bool b0" |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
363 |
by (simp add: integer_of_char_def of_char_def) |
66331 | 364 |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
365 |
|
68028 | 366 |
subsection \<open>Strings as dedicated type for target language code generation\<close> |
367 |
||
368 |
subsubsection \<open>Logical specification\<close> |
|
369 |
||
370 |
context |
|
371 |
begin |
|
372 |
||
373 |
qualified definition ascii_of :: "char \<Rightarrow> char" |
|
374 |
where "ascii_of c = Char (digit0 c) (digit1 c) (digit2 c) (digit3 c) (digit4 c) (digit5 c) (digit6 c) False" |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
375 |
|
68028 | 376 |
qualified lemma ascii_of_Char [simp]: |
377 |
"ascii_of (Char b0 b1 b2 b3 b4 b5 b6 b7) = Char b0 b1 b2 b3 b4 b5 b6 False" |
|
378 |
by (simp add: ascii_of_def) |
|
379 |
||
75622
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
380 |
qualified lemma digit0_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
381 |
"digit0 (String.ascii_of c) \<longleftrightarrow> digit0 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
382 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
383 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
384 |
qualified lemma digit1_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
385 |
"digit1 (String.ascii_of c) \<longleftrightarrow> digit1 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
386 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
387 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
388 |
qualified lemma digit2_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
389 |
"digit2 (String.ascii_of c) \<longleftrightarrow> digit2 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
390 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
391 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
392 |
qualified lemma digit3_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
393 |
"digit3 (String.ascii_of c) \<longleftrightarrow> digit3 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
394 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
395 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
396 |
qualified lemma digit4_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
397 |
"digit4 (String.ascii_of c) \<longleftrightarrow> digit4 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
398 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
399 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
400 |
qualified lemma digit5_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
401 |
"digit5 (String.ascii_of c) \<longleftrightarrow> digit5 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
402 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
403 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
404 |
qualified lemma digit6_ascii_of_iff [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
405 |
"digit6 (String.ascii_of c) \<longleftrightarrow> digit6 c" |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
406 |
by (simp add: String.ascii_of_def) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
407 |
|
68028 | 408 |
qualified lemma not_digit7_ascii_of [simp]: |
409 |
"\<not> digit7 (ascii_of c)" |
|
410 |
by (simp add: ascii_of_def) |
|
411 |
||
412 |
qualified lemma ascii_of_idem: |
|
413 |
"ascii_of c = c" if "\<not> digit7 c" |
|
414 |
using that by (cases c) simp |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
415 |
|
68028 | 416 |
qualified typedef literal = "{cs. \<forall>c\<in>set cs. \<not> digit7 c}" |
417 |
morphisms explode Abs_literal |
|
418 |
proof |
|
419 |
show "[] \<in> {cs. \<forall>c\<in>set cs. \<not> digit7 c}" |
|
420 |
by simp |
|
421 |
qed |
|
422 |
||
423 |
qualified setup_lifting type_definition_literal |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
424 |
|
68028 | 425 |
qualified lift_definition implode :: "string \<Rightarrow> literal" |
426 |
is "map ascii_of" |
|
427 |
by auto |
|
428 |
||
429 |
qualified lemma implode_explode_eq [simp]: |
|
430 |
"String.implode (String.explode s) = s" |
|
431 |
proof transfer |
|
432 |
fix cs |
|
433 |
show "map ascii_of cs = cs" if "\<forall>c\<in>set cs. \<not> digit7 c" |
|
434 |
using that |
|
435 |
by (induction cs) (simp_all add: ascii_of_idem) |
|
436 |
qed |
|
437 |
||
438 |
qualified lemma explode_implode_eq [simp]: |
|
439 |
"String.explode (String.implode cs) = map ascii_of cs" |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
440 |
by transfer rule |
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
441 |
|
68028 | 442 |
end |
443 |
||
78955 | 444 |
context linordered_euclidean_semiring_bit_operations |
75622
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
445 |
begin |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
446 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
447 |
context |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
448 |
begin |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
449 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
450 |
qualified lemma char_of_ascii_of [simp]: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
451 |
\<open>of_char (String.ascii_of c) = take_bit 7 (of_char c)\<close> |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
452 |
by (cases c) (simp only: String.ascii_of_Char of_char_Char take_bit_horner_sum_bit_eq, simp) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
453 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
454 |
qualified lemma ascii_of_char_of: |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
455 |
\<open>String.ascii_of (char_of a) = char_of (take_bit 7 a)\<close> |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
456 |
by (simp add: char_of_def bit_simps) |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
457 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
458 |
end |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
459 |
|
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
460 |
end |
53b61706749b
Centralized some char-related lemmas in distribution.
haftmann
parents:
75398
diff
changeset
|
461 |
|
68028 | 462 |
|
463 |
subsubsection \<open>Syntactic representation\<close> |
|
464 |
||
465 |
text \<open> |
|
466 |
Logical ground representations for literals are: |
|
467 |
||
69272 | 468 |
\<^enum> \<open>0\<close> for the empty literal; |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66190
diff
changeset
|
469 |
|
69272 | 470 |
\<^enum> \<open>Literal b0 \<dots> b6 s\<close> for a literal starting with one |
68028 | 471 |
character and continued by another literal. |
472 |
||
473 |
Syntactic representations for literals are: |
|
474 |
||
69272 | 475 |
\<^enum> Printable text as string prefixed with \<open>STR\<close>; |
68028 | 476 |
|
69272 | 477 |
\<^enum> A single ascii value as numerical hexadecimal value prefixed with \<open>STR\<close>. |
68028 | 478 |
\<close> |
479 |
||
480 |
instantiation String.literal :: zero |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
481 |
begin |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
482 |
|
68028 | 483 |
context |
484 |
begin |
|
485 |
||
486 |
qualified lift_definition zero_literal :: String.literal |
|
487 |
is Nil |
|
488 |
by simp |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
489 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
490 |
instance .. |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
491 |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
492 |
end |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
493 |
|
68028 | 494 |
end |
495 |
||
496 |
context |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
497 |
begin |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
498 |
|
68028 | 499 |
qualified abbreviation (output) empty_literal :: String.literal |
500 |
where "empty_literal \<equiv> 0" |
|
501 |
||
502 |
qualified lift_definition Literal :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal" |
|
503 |
is "\<lambda>b0 b1 b2 b3 b4 b5 b6 cs. Char b0 b1 b2 b3 b4 b5 b6 False # cs" |
|
504 |
by auto |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
505 |
|
68028 | 506 |
qualified lemma Literal_eq_iff [simp]: |
507 |
"Literal b0 b1 b2 b3 b4 b5 b6 s = Literal c0 c1 c2 c3 c4 c5 c6 t |
|
508 |
\<longleftrightarrow> (b0 \<longleftrightarrow> c0) \<and> (b1 \<longleftrightarrow> c1) \<and> (b2 \<longleftrightarrow> c2) \<and> (b3 \<longleftrightarrow> c3) |
|
509 |
\<and> (b4 \<longleftrightarrow> c4) \<and> (b5 \<longleftrightarrow> c5) \<and> (b6 \<longleftrightarrow> c6) \<and> s = t" |
|
510 |
by transfer simp |
|
511 |
||
512 |
qualified lemma empty_neq_Literal [simp]: |
|
513 |
"empty_literal \<noteq> Literal b0 b1 b2 b3 b4 b5 b6 s" |
|
514 |
by transfer simp |
|
515 |
||
516 |
qualified lemma Literal_neq_empty [simp]: |
|
517 |
"Literal b0 b1 b2 b3 b4 b5 b6 s \<noteq> empty_literal" |
|
518 |
by transfer simp |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
519 |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
520 |
end |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
521 |
|
68028 | 522 |
code_datatype "0 :: String.literal" String.Literal |
523 |
||
524 |
syntax |
|
81124 | 525 |
"_Literal" :: "str_position \<Rightarrow> String.literal" |
526 |
(\<open>(\<open>open_block notation=\<open>literal string\<close>\<close>STR _)\<close>) |
|
527 |
"_Ascii" :: "num_const \<Rightarrow> String.literal" |
|
528 |
(\<open>(\<open>open_block notation=\<open>literal char code\<close>\<close>STR _)\<close>) |
|
80760 | 529 |
syntax_consts |
530 |
"_Literal" "_Ascii" \<rightleftharpoons> String.Literal |
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
531 |
|
69605 | 532 |
ML_file \<open>Tools/literal.ML\<close> |
68028 | 533 |
|
52365
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
534 |
|
68028 | 535 |
subsubsection \<open>Operations\<close> |
536 |
||
537 |
instantiation String.literal :: plus |
|
67730 | 538 |
begin |
539 |
||
68028 | 540 |
context |
541 |
begin |
|
542 |
||
543 |
qualified lift_definition plus_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal" |
|
544 |
is "(@)" |
|
545 |
by auto |
|
67730 | 546 |
|
547 |
instance .. |
|
548 |
||
549 |
end |
|
550 |
||
68028 | 551 |
end |
67730 | 552 |
|
68028 | 553 |
instance String.literal :: monoid_add |
554 |
by (standard; transfer) simp_all |
|
555 |
||
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
556 |
lemma add_Literal_assoc: |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
557 |
\<open>String.Literal b0 b1 b2 b3 b4 b5 b6 t + s = String.Literal b0 b1 b2 b3 b4 b5 b6 (t + s)\<close> |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
558 |
by transfer simp |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
559 |
|
68028 | 560 |
instantiation String.literal :: size |
67729 | 561 |
begin |
562 |
||
68028 | 563 |
context |
564 |
includes literal.lifting |
|
565 |
begin |
|
566 |
||
567 |
lift_definition size_literal :: "String.literal \<Rightarrow> nat" |
|
568 |
is length . |
|
569 |
||
570 |
end |
|
67729 | 571 |
|
572 |
instance .. |
|
573 |
||
574 |
end |
|
575 |
||
68028 | 576 |
instantiation String.literal :: equal |
577 |
begin |
|
578 |
||
579 |
context |
|
580 |
begin |
|
581 |
||
582 |
qualified lift_definition equal_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" |
|
583 |
is HOL.equal . |
|
584 |
||
585 |
instance |
|
586 |
by (standard; transfer) (simp add: equal) |
|
587 |
||
588 |
end |
|
589 |
||
590 |
end |
|
591 |
||
592 |
instantiation String.literal :: linorder |
|
593 |
begin |
|
67729 | 594 |
|
68028 | 595 |
context |
596 |
begin |
|
597 |
||
598 |
qualified lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" |
|
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
599 |
is "ord.lexordp_eq (\<lambda>c d. of_char c < (of_char d :: nat))" |
68028 | 600 |
. |
601 |
||
602 |
qualified lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" |
|
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
603 |
is "ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat))" |
68028 | 604 |
. |
605 |
||
606 |
instance proof - |
|
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
607 |
from linorder_char interpret linorder "ord.lexordp_eq (\<lambda>c d. of_char c < (of_char d :: nat))" |
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
608 |
"ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat)) :: string \<Rightarrow> string \<Rightarrow> bool" |
68028 | 609 |
by (rule linorder.lexordp_linorder) |
610 |
show "PROP ?thesis" |
|
611 |
by (standard; transfer) (simp_all add: less_le_not_le linear) |
|
612 |
qed |
|
613 |
||
614 |
end |
|
615 |
||
616 |
end |
|
67730 | 617 |
|
68028 | 618 |
lemma infinite_literal: |
619 |
"infinite (UNIV :: String.literal set)" |
|
620 |
proof - |
|
621 |
define S where "S = range (\<lambda>n. replicate n CHR ''A'')" |
|
622 |
have "inj_on String.implode S" |
|
623 |
proof (rule inj_onI) |
|
624 |
fix cs ds |
|
625 |
assume "String.implode cs = String.implode ds" |
|
626 |
then have "String.explode (String.implode cs) = String.explode (String.implode ds)" |
|
627 |
by simp |
|
628 |
moreover assume "cs \<in> S" and "ds \<in> S" |
|
629 |
ultimately show "cs = ds" |
|
630 |
by (auto simp add: S_def) |
|
631 |
qed |
|
632 |
moreover have "infinite S" |
|
633 |
by (auto simp add: S_def dest: finite_range_imageI [of _ length]) |
|
634 |
ultimately have "infinite (String.implode ` S)" |
|
635 |
by (simp add: finite_image_iff) |
|
636 |
then show ?thesis |
|
637 |
by (auto intro: finite_subset) |
|
638 |
qed |
|
639 |
||
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
640 |
lemma add_literal_code [code]: |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
641 |
\<open>STR '''' + s = s\<close> |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
642 |
\<open>s + STR '''' = s\<close> |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
643 |
\<open>String.Literal b0 b1 b2 b3 b4 b5 b6 t + s = String.Literal b0 b1 b2 b3 b4 b5 b6 (t + s)\<close> |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
644 |
by (simp_all add: add_Literal_assoc) |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
645 |
|
68028 | 646 |
|
647 |
subsubsection \<open>Executable conversions\<close> |
|
648 |
||
649 |
context |
|
650 |
begin |
|
651 |
||
652 |
qualified lift_definition asciis_of_literal :: "String.literal \<Rightarrow> integer list" |
|
653 |
is "map of_char" |
|
654 |
. |
|
655 |
||
69879 | 656 |
qualified lemma asciis_of_zero [simp, code]: |
657 |
"asciis_of_literal 0 = []" |
|
658 |
by transfer simp |
|
659 |
||
660 |
qualified lemma asciis_of_Literal [simp, code]: |
|
661 |
"asciis_of_literal (String.Literal b0 b1 b2 b3 b4 b5 b6 s) = |
|
662 |
of_char (Char b0 b1 b2 b3 b4 b5 b6 False) # asciis_of_literal s " |
|
663 |
by transfer simp |
|
664 |
||
68028 | 665 |
qualified lift_definition literal_of_asciis :: "integer list \<Rightarrow> String.literal" |
666 |
is "map (String.ascii_of \<circ> char_of)" |
|
667 |
by auto |
|
55426
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
668 |
|
69879 | 669 |
qualified lemma literal_of_asciis_Nil [simp, code]: |
670 |
"literal_of_asciis [] = 0" |
|
671 |
by transfer simp |
|
672 |
||
673 |
qualified lemma literal_of_asciis_Cons [simp, code]: |
|
674 |
"literal_of_asciis (k # ks) = (case char_of k |
|
675 |
of Char b0 b1 b2 b3 b4 b5 b6 b7 \<Rightarrow> String.Literal b0 b1 b2 b3 b4 b5 b6 (literal_of_asciis ks))" |
|
676 |
by (simp add: char_of_def) (transfer, simp add: char_of_def) |
|
677 |
||
68028 | 678 |
qualified lemma literal_of_asciis_of_literal [simp]: |
679 |
"literal_of_asciis (asciis_of_literal s) = s" |
|
680 |
proof transfer |
|
681 |
fix cs |
|
682 |
assume "\<forall>c\<in>set cs. \<not> digit7 c" |
|
683 |
then show "map (String.ascii_of \<circ> char_of) (map of_char cs) = cs" |
|
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
684 |
by (induction cs) (simp_all add: String.ascii_of_idem) |
68028 | 685 |
qed |
686 |
||
687 |
qualified lemma explode_code [code]: |
|
688 |
"String.explode s = map char_of (asciis_of_literal s)" |
|
689 |
by transfer simp |
|
690 |
||
691 |
qualified lemma implode_code [code]: |
|
692 |
"String.implode cs = literal_of_asciis (map of_char cs)" |
|
693 |
by transfer simp |
|
64994
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
694 |
|
69879 | 695 |
qualified lemma equal_literal [code]: |
696 |
"HOL.equal (String.Literal b0 b1 b2 b3 b4 b5 b6 s) |
|
697 |
(String.Literal a0 a1 a2 a3 a4 a5 a6 r) |
|
698 |
\<longleftrightarrow> (b0 \<longleftrightarrow> a0) \<and> (b1 \<longleftrightarrow> a1) \<and> (b2 \<longleftrightarrow> a2) \<and> (b3 \<longleftrightarrow> a3) |
|
699 |
\<and> (b4 \<longleftrightarrow> a4) \<and> (b5 \<longleftrightarrow> a5) \<and> (b6 \<longleftrightarrow> a6) \<and> (s = r)" |
|
700 |
by (simp add: equal) |
|
68028 | 701 |
|
69879 | 702 |
end |
68028 | 703 |
|
704 |
||
705 |
subsubsection \<open>Technical code generation setup\<close> |
|
706 |
||
707 |
text \<open>Alternative constructor for generated computations\<close> |
|
708 |
||
709 |
context |
|
72170
7fa9605b226c
Another go with lex: now lexordp back in class ord
paulson <lp15@cam.ac.uk>
parents:
72164
diff
changeset
|
710 |
begin |
68028 | 711 |
|
712 |
qualified definition Literal' :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal" |
|
713 |
where [simp]: "Literal' = String.Literal" |
|
714 |
||
715 |
lemma [code]: |
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
716 |
\<open>Literal' b0 b1 b2 b3 b4 b5 b6 s = String.literal_of_asciis |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
717 |
[foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6] 0] + s\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
718 |
proof - |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
719 |
have \<open>foldr (\<lambda>b k. of_bool b + k * 2) [b0, b1, b2, b3, b4, b5, b6] 0 = of_char (Char b0 b1 b2 b3 b4 b5 b6 False)\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
720 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
721 |
moreover have \<open>Literal' b0 b1 b2 b3 b4 b5 b6 s = String.literal_of_asciis |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
722 |
[of_char (Char b0 b1 b2 b3 b4 b5 b6 False)] + s\<close> |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
723 |
by (unfold Literal'_def) (transfer, simp only: list.simps comp_apply char_of_char, simp) |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
724 |
ultimately show ?thesis |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
725 |
by simp |
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71094
diff
changeset
|
726 |
qed |
64994
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
727 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
728 |
lemma [code_computation_unfold]: |
68028 | 729 |
"String.Literal = Literal'" |
730 |
by simp |
|
64994
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
731 |
|
68028 | 732 |
end |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
733 |
|
81706 | 734 |
code_reserved |
735 |
(SML) string String Char Str_Literal |
|
736 |
and (OCaml) string String Char Str_Literal |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
737 |
and (Haskell) Str_Literal |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
738 |
and (Scala) String Str_Literal |
33237 | 739 |
|
75647
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75622
diff
changeset
|
740 |
code_identifier |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75622
diff
changeset
|
741 |
code_module String \<rightharpoonup> |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75622
diff
changeset
|
742 |
(SML) Str and (OCaml) Str and (Haskell) Str and (Scala) Str |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75622
diff
changeset
|
743 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
744 |
code_printing |
68028 | 745 |
type_constructor String.literal \<rightharpoonup> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
746 |
(SML) "string" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
747 |
and (OCaml) "string" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
748 |
and (Haskell) "String" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
749 |
and (Scala) "String" |
68028 | 750 |
| constant "STR ''''" \<rightharpoonup> |
751 |
(SML) "\"\"" |
|
752 |
and (OCaml) "\"\"" |
|
753 |
and (Haskell) "\"\"" |
|
754 |
and (Scala) "\"\"" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
755 |
|
60758 | 756 |
setup \<open> |
68028 | 757 |
fold Literal.add_code ["SML", "OCaml", "Haskell", "Scala"] |
60758 | 758 |
\<close> |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
759 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
760 |
code_printing |
75694 | 761 |
code_module "Str_Literal" \<rightharpoonup> |
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
762 |
(SML) \<open>structure Str_Literal : sig |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
763 |
type int = IntInf.int |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
764 |
val literal_of_asciis : int list -> string |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
765 |
val asciis_of_literal : string -> int list |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
766 |
end = struct |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
767 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
768 |
open IntInf; |
75694 | 769 |
|
770 |
fun map f [] = [] |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
771 |
| map f (x :: xs) = f x :: map f xs; (* deliberate clone not relying on List._ structure *) |
75694 | 772 |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
773 |
fun check_ascii k = |
75694 | 774 |
if 0 <= k andalso k < 128 |
775 |
then k |
|
776 |
else raise Fail "Non-ASCII character in literal"; |
|
777 |
||
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
778 |
val char_of_ascii = Char.chr o IntInf.toInt o (fn k => k mod 128); |
75694 | 779 |
|
780 |
val ascii_of_char = check_ascii o IntInf.fromInt o Char.ord; |
|
781 |
||
782 |
val literal_of_asciis = String.implode o map char_of_ascii; |
|
783 |
||
784 |
val asciis_of_literal = map ascii_of_char o String.explode; |
|
785 |
||
786 |
end;\<close> for constant String.literal_of_asciis String.asciis_of_literal |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
787 |
and (OCaml) \<open>module Str_Literal : sig |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
788 |
val literal_of_asciis : Z.t list -> string |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
789 |
val asciis_of_literal: string -> Z.t list |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
790 |
end = struct |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
791 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
792 |
(* deliberate clones not relying on List._ module *) |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
793 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
794 |
let rec length xs = match xs with |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
795 |
[] -> 0 |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
796 |
| x :: xs -> 1 + length xs;; |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
797 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
798 |
let rec nth xs n = match xs with |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
799 |
(x :: xs) -> if n <= 0 then x else nth xs (n - 1);; |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
800 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
801 |
let rec map_range f n = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
802 |
if n <= 0 |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
803 |
then [] |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
804 |
else |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
805 |
let m = n - 1 |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
806 |
in map_range f m @ [f m];; |
75694 | 807 |
|
808 |
let implode f xs = |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
809 |
String.init (length xs) (fun n -> f (nth xs n));; |
75694 | 810 |
|
811 |
let explode f s = |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
812 |
map_range (fun n -> f (String.get s n)) (String.length s);; |
75694 | 813 |
|
814 |
let z_128 = Z.of_int 128;; |
|
815 |
||
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
816 |
let check_ascii k = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
817 |
if 0 <= k && k < 128 |
75694 | 818 |
then k |
819 |
else failwith "Non-ASCII character in literal";; |
|
820 |
||
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
821 |
let char_of_ascii k = Char.chr (Z.to_int (Z.rem k z_128));; |
75694 | 822 |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
823 |
let ascii_of_char c = Z.of_int (check_ascii (Char.code c));; |
75694 | 824 |
|
825 |
let literal_of_asciis ks = implode char_of_ascii ks;; |
|
826 |
||
827 |
let asciis_of_literal s = explode ascii_of_char s;; |
|
828 |
||
829 |
end;;\<close> for constant String.literal_of_asciis String.asciis_of_literal |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
830 |
and (Haskell) \<open>module Str_Literal(literalOfAsciis, asciisOfLiteral) where |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
831 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
832 |
check_ascii :: Int -> Int |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
833 |
check_ascii k |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
834 |
| (0 <= k && k < 128) = k |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
835 |
| otherwise = error "Non-ASCII character in literal" |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
836 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
837 |
charOfAscii :: Integer -> Char |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
838 |
charOfAscii = toEnum . Prelude.fromInteger . (\k -> k `mod ` 128) |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
839 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
840 |
asciiOfChar :: Char -> Integer |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
841 |
asciiOfChar = toInteger . check_ascii . fromEnum |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
842 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
843 |
literalOfAsciis :: [Integer] -> [Char] |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
844 |
literalOfAsciis = map charOfAscii |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
845 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
846 |
asciisOfLiteral :: [Char] -> [Integer] |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
847 |
asciisOfLiteral = map asciiOfChar |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
848 |
\<close> for constant String.literal_of_asciis String.asciis_of_literal |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
849 |
and (Scala) \<open>object Str_Literal { |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
850 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
851 |
private def checkAscii(k : Int) : Int = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
852 |
0 <= k && k < 128 match { |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
853 |
case true => k |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
854 |
case false => sys.error("Non-ASCII character in literal") |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
855 |
} |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
856 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
857 |
private def charOfAscii(k : BigInt) : Char = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
858 |
(k % 128).charValue |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
859 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
860 |
private def asciiOfChar(c : Char) : BigInt = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
861 |
BigInt(checkAscii(c.toInt)) |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
862 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
863 |
def literalOfAsciis(ks : List[BigInt]) : String = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
864 |
ks.map(charOfAscii).mkString |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
865 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
866 |
def asciisOfLiteral(s : String) : List[BigInt] = |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
867 |
s.toList.map(asciiOfChar) |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
868 |
|
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
869 |
} |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
870 |
\<close> for constant String.literal_of_asciis String.asciis_of_literal |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
871 |
| constant \<open>(+) :: String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal\<close> \<rightharpoonup> |
68028 | 872 |
(SML) infixl 18 "^" |
873 |
and (OCaml) infixr 6 "^" |
|
874 |
and (Haskell) infixr 5 "++" |
|
875 |
and (Scala) infixl 7 "+" |
|
876 |
| constant String.literal_of_asciis \<rightharpoonup> |
|
75694 | 877 |
(SML) "Str'_Literal.literal'_of'_asciis" |
878 |
and (OCaml) "Str'_Literal.literal'_of'_asciis" |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
879 |
and (Haskell) "Str'_Literal.literalOfAsciis" |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
880 |
and (Scala) "Str'_Literal.literalOfAsciis" |
68028 | 881 |
| constant String.asciis_of_literal \<rightharpoonup> |
75694 | 882 |
(SML) "Str'_Literal.asciis'_of'_literal" |
883 |
and (OCaml) "Str'_Literal.asciis'_of'_literal" |
|
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
884 |
and (Haskell) "Str'_Literal.asciisOfLiteral" |
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
885 |
and (Scala) "Str'_Literal.asciisOfLiteral" |
68028 | 886 |
| class_instance String.literal :: equal \<rightharpoonup> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
887 |
(Haskell) - |
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
888 |
| constant \<open>HOL.equal :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool\<close> \<rightharpoonup> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
889 |
(SML) "!((_ : string) = _)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
890 |
and (OCaml) "!((_ : string) = _)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
891 |
and (Haskell) infix 4 "==" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
892 |
and (Scala) infixl 5 "==" |
81761
a1dc03194053
more correct code generation for string literals
haftmann
parents:
81706
diff
changeset
|
893 |
| constant \<open>(\<le>) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool\<close> \<rightharpoonup> |
68028 | 894 |
(SML) "!((_ : string) <= _)" |
895 |
and (OCaml) "!((_ : string) <= _)" |
|
69743 | 896 |
and (Haskell) infix 4 "<=" |
69593 | 897 |
\<comment> \<open>Order operations for \<^typ>\<open>String.literal\<close> work in Haskell only |
68028 | 898 |
if no type class instance needs to be generated, because String = [Char] in Haskell |
69593 | 899 |
and \<^typ>\<open>char list\<close> need not have the same order as \<^typ>\<open>String.literal\<close>.\<close> |
68028 | 900 |
and (Scala) infixl 4 "<=" |
901 |
and (Eval) infixl 6 "<=" |
|
902 |
| constant "(<) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup> |
|
903 |
(SML) "!((_ : string) < _)" |
|
904 |
and (OCaml) "!((_ : string) < _)" |
|
905 |
and (Haskell) infix 4 "<" |
|
906 |
and (Scala) infixl 4 "<" |
|
907 |
and (Eval) infixl 6 "<" |
|
908 |
||
909 |
||
910 |
subsubsection \<open>Code generation utility\<close> |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
911 |
|
60758 | 912 |
setup \<open>Sign.map_naming (Name_Space.mandatory_path "Code")\<close> |
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
913 |
|
68028 | 914 |
definition abort :: "String.literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a" |
915 |
where [simp]: "abort _ f = f ()" |
|
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
916 |
|
68028 | 917 |
declare [[code drop: Code.abort]] |
918 |
||
919 |
lemma abort_cong: |
|
920 |
"msg = msg' \<Longrightarrow> Code.abort msg f = Code.abort msg' f" |
|
921 |
by simp |
|
54317
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
922 |
|
60758 | 923 |
setup \<open>Sign.map_naming Name_Space.parent_path\<close> |
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
924 |
|
60758 | 925 |
setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm Code.abort_cong})\<close> |
54317
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
926 |
|
68028 | 927 |
code_printing |
928 |
constant Code.abort \<rightharpoonup> |
|
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
929 |
(SML) "!(raise/ Fail/ _)" |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
930 |
and (OCaml) "failwith" |
59483
ddb73392356e
explicit type annotation avoids problems with Haskell type inference
haftmann
parents:
58889
diff
changeset
|
931 |
and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)" |
68028 | 932 |
and (Scala) "!{/ sys.error((_));/ ((_)).apply(())/ }" |
933 |
||
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
934 |
|
68028 | 935 |
subsubsection \<open>Finally\<close> |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
936 |
|
68028 | 937 |
lifting_update literal.lifting |
938 |
lifting_forget literal.lifting |
|
57437 | 939 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
940 |
end |