| author | wenzelm | 
| Sat, 20 Feb 2021 23:01:35 +0100 | |
| changeset 73264 | 440546ea20e6 | 
| parent 72611 | c7bc3e70a8c7 | 
| child 74101 | d804e93ae9ff | 
| 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  | 
| 55969 | 6  | 
imports Enum  | 
| 
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  | 
||
| 71094 | 41  | 
context unique_euclidean_semiring_with_bit_shifts  | 
| 68028 | 42  | 
begin  | 
43  | 
||
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
44  | 
definition char_of :: \<open>'a \<Rightarrow> char\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
45  | 
where \<open>char_of n = Char (odd n) (bit n 1) (bit n 2) (bit n 3) (bit n 4) (bit n 5) (bit n 6) (bit n 7)\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
46  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
47  | 
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
 | 
48  | 
\<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
 | 
49  | 
using that by (simp add: char_of_def bit_take_bit_iff)  | 
| 68028 | 50  | 
|
51  | 
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
 | 
52  | 
\<open>char_of (of_char c) = c\<close>  | 
| 72024 | 53  | 
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
 | 
54  | 
|
| 68028 | 55  | 
lemma char_of_comp_of_char [simp]:  | 
56  | 
"char_of \<circ> of_char = id"  | 
|
57  | 
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
 | 
58  | 
|
| 68028 | 59  | 
lemma inj_of_char:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
60  | 
\<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
 | 
61  | 
proof (rule injI)  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
62  | 
fix c d  | 
| 68028 | 63  | 
assume "of_char c = of_char d"  | 
64  | 
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
 | 
65  | 
by simp  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
66  | 
then show "c = d"  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
67  | 
by simp  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
68  | 
qed  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
69  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
70  | 
lemma of_char_eqI:  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
71  | 
\<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
 | 
72  | 
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
 | 
73  | 
|
| 68028 | 74  | 
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
 | 
75  | 
\<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
 | 
76  | 
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
 | 
77  | 
|
| 68028 | 78  | 
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
 | 
79  | 
\<open>of_char (char_of a) = a mod 256\<close>  | 
| 68028 | 80  | 
proof -  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
81  | 
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
 | 
82  | 
by (simp add: upt_eq_Cons_conv)  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
83  | 
then have \<open>[odd a, 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>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
84  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
85  | 
then have \<open>of_char (char_of a) = take_bit 8 a\<close>  | 
| 72024 | 86  | 
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
 | 
87  | 
then show ?thesis  | 
| 68028 | 88  | 
by (simp add: take_bit_eq_mod)  | 
89  | 
qed  | 
|
| 
62364
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
90  | 
|
| 68028 | 91  | 
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
 | 
92  | 
\<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
 | 
93  | 
by (rule of_char_eqI) simp  | 
| 68028 | 94  | 
|
95  | 
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
 | 
96  | 
\<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
 | 
97  | 
proof -  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
98  | 
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
 | 
99  | 
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
 | 
100  | 
then show ?thesis  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
101  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
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_quasi_inj [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 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
 | 
106  | 
proof  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
107  | 
assume ?Q  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
108  | 
then show ?P  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
109  | 
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
 | 
110  | 
next  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
111  | 
assume ?P  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
112  | 
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
 | 
113  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
114  | 
then show ?Q  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
115  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
116  | 
qed  | 
| 68028 | 117  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
118  | 
lemma char_of_eq_iff:  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
119  | 
\<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
 | 
120  | 
by (auto intro: of_char_eqI simp add: take_bit_eq_mod)  | 
| 68028 | 121  | 
|
122  | 
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
 | 
123  | 
\<open>char_of (of_nat n) = char_of n\<close>  | 
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72170 
diff
changeset
 | 
124  | 
by (simp add: char_of_def String.char_of_def drop_bit_of_nat bit_simps)  | 
| 
68033
 
ad4b8b6892c3
uniform tagging for printable and non-printable literals
 
haftmann 
parents: 
68028 
diff
changeset
 | 
125  | 
|
| 68028 | 126  | 
end  | 
| 
62364
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
127  | 
|
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
128  | 
lemma inj_on_char_of_nat [simp]:  | 
| 68028 | 129  | 
  "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
 | 
130  | 
by (rule inj_onI) simp  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
131  | 
|
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
132  | 
lemma nat_of_char_less_256 [simp]:  | 
| 68028 | 133  | 
"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
 | 
134  | 
proof -  | 
| 68028 | 135  | 
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
 | 
136  | 
by arith  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
137  | 
then show ?thesis by simp  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
138  | 
qed  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
139  | 
|
| 68028 | 140  | 
lemma range_nat_of_char:  | 
141  | 
  "range of_char = {0::nat..<256}"
 | 
|
142  | 
proof (rule; rule)  | 
|
143  | 
fix n :: nat  | 
|
144  | 
assume "n \<in> range of_char"  | 
|
145  | 
  then show "n \<in> {0..<256}"
 | 
|
146  | 
by auto  | 
|
147  | 
next  | 
|
148  | 
fix n :: nat  | 
|
149  | 
  assume "n \<in> {0..<256}"
 | 
|
150  | 
then have "n = of_char (char_of n)"  | 
|
151  | 
by simp  | 
|
152  | 
then show "n \<in> range of_char"  | 
|
153  | 
by (rule range_eqI)  | 
|
154  | 
qed  | 
|
155  | 
||
| 
62364
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
156  | 
lemma UNIV_char_of_nat:  | 
| 68028 | 157  | 
  "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
 | 
158  | 
proof -  | 
| 68028 | 159  | 
  have "range (of_char :: char \<Rightarrow> nat) = of_char ` char_of ` {0::nat..<256}"
 | 
160  | 
by (auto simp add: range_nat_of_char intro!: image_eqI)  | 
|
| 
72170
 
7fa9605b226c
Another go with lex: now lexordp back in class ord
 
paulson <lp15@cam.ac.uk> 
parents: 
72164 
diff
changeset
 | 
161  | 
with inj_of_char [where ?'a = nat] show ?thesis  | 
| 68028 | 162  | 
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
 | 
163  | 
qed  | 
| 
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
164  | 
|
| 62597 | 165  | 
lemma card_UNIV_char:  | 
166  | 
"card (UNIV :: char set) = 256"  | 
|
167  | 
by (auto simp add: UNIV_char_of_nat card_image)  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
168  | 
|
| 68028 | 169  | 
context  | 
170  | 
includes lifting_syntax integer.lifting natural.lifting  | 
|
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
171  | 
begin  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
172  | 
|
| 68028 | 173  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
174  | 
\<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
 | 
175  | 
by (unfold char_of_def) transfer_prover  | 
| 68028 | 176  | 
|
177  | 
lemma [transfer_rule]:  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
178  | 
\<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
 | 
179  | 
by (unfold of_char_def) transfer_prover  | 
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
180  | 
|
| 68028 | 181  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
182  | 
\<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
 | 
183  | 
by (unfold char_of_def) transfer_prover  | 
| 68028 | 184  | 
|
185  | 
lemma [transfer_rule]:  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
186  | 
\<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
 | 
187  | 
by (unfold of_char_def) transfer_prover  | 
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
188  | 
|
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
189  | 
end  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
190  | 
|
| 68028 | 191  | 
lifting_update integer.lifting  | 
192  | 
lifting_forget integer.lifting  | 
|
| 
64630
 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 
haftmann 
parents: 
63950 
diff
changeset
 | 
193  | 
|
| 68028 | 194  | 
lifting_update natural.lifting  | 
195  | 
lifting_forget natural.lifting  | 
|
| 
62364
 
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
 
haftmann 
parents: 
61799 
diff
changeset
 | 
196  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
197  | 
syntax  | 
| 62597 | 198  | 
  "_Char" :: "str_position \<Rightarrow> char"    ("CHR _")
 | 
| 62678 | 199  | 
  "_Char_ord" :: "num_const \<Rightarrow> char"   ("CHR _")
 | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
200  | 
|
| 
42163
 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 
bulwahn 
parents: 
41750 
diff
changeset
 | 
201  | 
type_synonym string = "char list"  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
203  | 
syntax  | 
| 68028 | 204  | 
  "_String" :: "str_position \<Rightarrow> string"    ("_")
 | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
205  | 
|
| 69605 | 206  | 
ML_file \<open>Tools/string_syntax.ML\<close>  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
207  | 
|
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
208  | 
instantiation char :: enum  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
209  | 
begin  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
210  | 
|
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
211  | 
definition  | 
| 68028 | 212  | 
"Enum.enum = [  | 
213  | 
CHR 0x00, CHR 0x01, CHR 0x02, CHR 0x03,  | 
|
| 62678 | 214  | 
CHR 0x04, CHR 0x05, CHR 0x06, CHR 0x07,  | 
215  | 
CHR 0x08, CHR 0x09, CHR ''\<newline>'', CHR 0x0B,  | 
|
216  | 
CHR 0x0C, CHR 0x0D, CHR 0x0E, CHR 0x0F,  | 
|
217  | 
CHR 0x10, CHR 0x11, CHR 0x12, CHR 0x13,  | 
|
218  | 
CHR 0x14, CHR 0x15, CHR 0x16, CHR 0x17,  | 
|
219  | 
CHR 0x18, CHR 0x19, CHR 0x1A, CHR 0x1B,  | 
|
220  | 
CHR 0x1C, CHR 0x1D, CHR 0x1E, CHR 0x1F,  | 
|
221  | 
CHR '' '', CHR ''!'', CHR 0x22, CHR ''#'',  | 
|
222  | 
CHR ''$'', CHR ''%'', CHR ''&'', CHR 0x27,  | 
|
| 62597 | 223  | 
    CHR ''('', CHR '')'', CHR ''*'', CHR ''+'',
 | 
224  | 
CHR '','', CHR ''-'', CHR ''.'', CHR ''/'',  | 
|
225  | 
CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'',  | 
|
226  | 
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'',  | 
|
227  | 
CHR ''8'', CHR ''9'', CHR '':'', CHR '';'',  | 
|
228  | 
CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'',  | 
|
229  | 
CHR ''@'', CHR ''A'', CHR ''B'', CHR ''C'',  | 
|
230  | 
CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'',  | 
|
231  | 
CHR ''H'', CHR ''I'', CHR ''J'', CHR ''K'',  | 
|
232  | 
CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'',  | 
|
233  | 
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'',  | 
|
234  | 
CHR ''T'', CHR ''U'', CHR ''V'', CHR ''W'',  | 
|
235  | 
CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['',  | 
|
| 62678 | 236  | 
CHR 0x5C, CHR '']'', CHR ''^'', CHR ''_'',  | 
237  | 
CHR 0x60, CHR ''a'', CHR ''b'', CHR ''c'',  | 
|
| 62597 | 238  | 
CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'',  | 
239  | 
CHR ''h'', CHR ''i'', CHR ''j'', CHR ''k'',  | 
|
240  | 
CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'',  | 
|
241  | 
CHR ''p'', CHR ''q'', CHR ''r'', CHR ''s'',  | 
|
242  | 
CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'',  | 
|
243  | 
    CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'',
 | 
|
| 62678 | 244  | 
CHR ''|'', CHR ''}'', CHR ''~'', CHR 0x7F,  | 
245  | 
CHR 0x80, CHR 0x81, CHR 0x82, CHR 0x83,  | 
|
246  | 
CHR 0x84, CHR 0x85, CHR 0x86, CHR 0x87,  | 
|
247  | 
CHR 0x88, CHR 0x89, CHR 0x8A, CHR 0x8B,  | 
|
248  | 
CHR 0x8C, CHR 0x8D, CHR 0x8E, CHR 0x8F,  | 
|
249  | 
CHR 0x90, CHR 0x91, CHR 0x92, CHR 0x93,  | 
|
250  | 
CHR 0x94, CHR 0x95, CHR 0x96, CHR 0x97,  | 
|
251  | 
CHR 0x98, CHR 0x99, CHR 0x9A, CHR 0x9B,  | 
|
252  | 
CHR 0x9C, CHR 0x9D, CHR 0x9E, CHR 0x9F,  | 
|
253  | 
CHR 0xA0, CHR 0xA1, CHR 0xA2, CHR 0xA3,  | 
|
254  | 
CHR 0xA4, CHR 0xA5, CHR 0xA6, CHR 0xA7,  | 
|
255  | 
CHR 0xA8, CHR 0xA9, CHR 0xAA, CHR 0xAB,  | 
|
256  | 
CHR 0xAC, CHR 0xAD, CHR 0xAE, CHR 0xAF,  | 
|
257  | 
CHR 0xB0, CHR 0xB1, CHR 0xB2, CHR 0xB3,  | 
|
258  | 
CHR 0xB4, CHR 0xB5, CHR 0xB6, CHR 0xB7,  | 
|
259  | 
CHR 0xB8, CHR 0xB9, CHR 0xBA, CHR 0xBB,  | 
|
260  | 
CHR 0xBC, CHR 0xBD, CHR 0xBE, CHR 0xBF,  | 
|
261  | 
CHR 0xC0, CHR 0xC1, CHR 0xC2, CHR 0xC3,  | 
|
262  | 
CHR 0xC4, CHR 0xC5, CHR 0xC6, CHR 0xC7,  | 
|
263  | 
CHR 0xC8, CHR 0xC9, CHR 0xCA, CHR 0xCB,  | 
|
264  | 
CHR 0xCC, CHR 0xCD, CHR 0xCE, CHR 0xCF,  | 
|
265  | 
CHR 0xD0, CHR 0xD1, CHR 0xD2, CHR 0xD3,  | 
|
266  | 
CHR 0xD4, CHR 0xD5, CHR 0xD6, CHR 0xD7,  | 
|
267  | 
CHR 0xD8, CHR 0xD9, CHR 0xDA, CHR 0xDB,  | 
|
268  | 
CHR 0xDC, CHR 0xDD, CHR 0xDE, CHR 0xDF,  | 
|
269  | 
CHR 0xE0, CHR 0xE1, CHR 0xE2, CHR 0xE3,  | 
|
270  | 
CHR 0xE4, CHR 0xE5, CHR 0xE6, CHR 0xE7,  | 
|
271  | 
CHR 0xE8, CHR 0xE9, CHR 0xEA, CHR 0xEB,  | 
|
272  | 
CHR 0xEC, CHR 0xED, CHR 0xEE, CHR 0xEF,  | 
|
273  | 
CHR 0xF0, CHR 0xF1, CHR 0xF2, CHR 0xF3,  | 
|
274  | 
CHR 0xF4, CHR 0xF5, CHR 0xF6, CHR 0xF7,  | 
|
275  | 
CHR 0xF8, CHR 0xF9, CHR 0xFA, CHR 0xFB,  | 
|
276  | 
CHR 0xFC, CHR 0xFD, CHR 0xFE, CHR 0xFF]"  | 
|
| 31484 | 277  | 
|
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
278  | 
definition  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
279  | 
"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
 | 
280  | 
|
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
281  | 
definition  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
282  | 
"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
 | 
283  | 
|
| 62597 | 284  | 
lemma enum_char_unfold:  | 
| 68028 | 285  | 
"Enum.enum = map char_of [0..<256]"  | 
| 62597 | 286  | 
proof -  | 
| 68028 | 287  | 
have "map (of_char :: char \<Rightarrow> nat) Enum.enum = [0..<256]"  | 
288  | 
by (simp add: enum_char_def of_char_def upt_conv_Cons_Cons numeral_2_eq_2 [symmetric])  | 
|
289  | 
then have "map char_of (map (of_char :: char \<Rightarrow> nat) Enum.enum) =  | 
|
290  | 
map char_of [0..<256]"  | 
|
291  | 
by simp  | 
|
292  | 
then show ?thesis  | 
|
293  | 
by simp  | 
|
| 62597 | 294  | 
qed  | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
49972 
diff
changeset
 | 
295  | 
|
| 
49972
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
296  | 
instance proof  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
297  | 
show UNIV: "UNIV = set (Enum.enum :: char list)"  | 
| 62597 | 298  | 
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
 | 
299  | 
show "distinct (Enum.enum :: char list)"  | 
| 62597 | 300  | 
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
 | 
301  | 
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
 | 
302  | 
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
 | 
303  | 
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
 | 
304  | 
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
 | 
305  | 
qed  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
306  | 
|
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
307  | 
end  | 
| 
 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 
haftmann 
parents: 
49948 
diff
changeset
 | 
308  | 
|
| 68028 | 309  | 
lemma linorder_char:  | 
310  | 
"class.linorder (\<lambda>c d. of_char c \<le> (of_char d :: nat)) (\<lambda>c d. of_char c < (of_char d :: nat))"  | 
|
311  | 
by standard auto  | 
|
312  | 
||
313  | 
text \<open>Optimized version for execution\<close>  | 
|
314  | 
||
315  | 
definition char_of_integer :: "integer \<Rightarrow> char"  | 
|
316  | 
where [code_abbrev]: "char_of_integer = char_of"  | 
|
317  | 
||
318  | 
definition integer_of_char :: "char \<Rightarrow> integer"  | 
|
319  | 
where [code_abbrev]: "integer_of_char = of_char"  | 
|
320  | 
||
| 62597 | 321  | 
lemma char_of_integer_code [code]:  | 
| 68028 | 322  | 
"char_of_integer k = (let  | 
323  | 
(q0, b0) = bit_cut_integer k;  | 
|
324  | 
(q1, b1) = bit_cut_integer q0;  | 
|
325  | 
(q2, b2) = bit_cut_integer q1;  | 
|
326  | 
(q3, b3) = bit_cut_integer q2;  | 
|
327  | 
(q4, b4) = bit_cut_integer q3;  | 
|
328  | 
(q5, b5) = bit_cut_integer q4;  | 
|
329  | 
(q6, b6) = bit_cut_integer q5;  | 
|
330  | 
(_, b7) = bit_cut_integer q6  | 
|
331  | 
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
 | 
332  | 
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
 | 
333  | 
|
| 68028 | 334  | 
lemma integer_of_char_code [code]:  | 
335  | 
"integer_of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) =  | 
|
336  | 
((((((of_bool b7 * 2 + of_bool b6) * 2 +  | 
|
337  | 
of_bool b5) * 2 + of_bool b4) * 2 +  | 
|
338  | 
of_bool b3) * 2 + of_bool b2) * 2 +  | 
|
339  | 
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
 | 
340  | 
by (simp add: integer_of_char_def of_char_def)  | 
| 66331 | 341  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
342  | 
|
| 68028 | 343  | 
subsection \<open>Strings as dedicated type for target language code generation\<close>  | 
344  | 
||
345  | 
subsubsection \<open>Logical specification\<close>  | 
|
346  | 
||
347  | 
context  | 
|
348  | 
begin  | 
|
349  | 
||
350  | 
qualified definition ascii_of :: "char \<Rightarrow> char"  | 
|
351  | 
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
 | 
352  | 
|
| 68028 | 353  | 
qualified lemma ascii_of_Char [simp]:  | 
354  | 
"ascii_of (Char b0 b1 b2 b3 b4 b5 b6 b7) = Char b0 b1 b2 b3 b4 b5 b6 False"  | 
|
355  | 
by (simp add: ascii_of_def)  | 
|
356  | 
||
357  | 
qualified lemma not_digit7_ascii_of [simp]:  | 
|
358  | 
"\<not> digit7 (ascii_of c)"  | 
|
359  | 
by (simp add: ascii_of_def)  | 
|
360  | 
||
361  | 
qualified lemma ascii_of_idem:  | 
|
362  | 
"ascii_of c = c" if "\<not> digit7 c"  | 
|
363  | 
using that by (cases c) simp  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
364  | 
|
| 68028 | 365  | 
qualified lemma char_of_ascii_of [simp]:  | 
366  | 
"of_char (ascii_of c) = take_bit 7 (of_char c :: nat)"  | 
|
| 72024 | 367  | 
by (cases c) (simp only: ascii_of_Char of_char_Char take_bit_horner_sum_bit_eq, simp)  | 
| 68028 | 368  | 
|
369  | 
qualified typedef literal = "{cs. \<forall>c\<in>set cs. \<not> digit7 c}"
 | 
|
370  | 
morphisms explode Abs_literal  | 
|
371  | 
proof  | 
|
372  | 
  show "[] \<in> {cs. \<forall>c\<in>set cs. \<not> digit7 c}"
 | 
|
373  | 
by simp  | 
|
374  | 
qed  | 
|
375  | 
||
376  | 
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
 | 
377  | 
|
| 68028 | 378  | 
qualified lift_definition implode :: "string \<Rightarrow> literal"  | 
379  | 
is "map ascii_of"  | 
|
380  | 
by auto  | 
|
381  | 
||
382  | 
qualified lemma implode_explode_eq [simp]:  | 
|
383  | 
"String.implode (String.explode s) = s"  | 
|
384  | 
proof transfer  | 
|
385  | 
fix cs  | 
|
386  | 
show "map ascii_of cs = cs" if "\<forall>c\<in>set cs. \<not> digit7 c"  | 
|
387  | 
using that  | 
|
388  | 
by (induction cs) (simp_all add: ascii_of_idem)  | 
|
389  | 
qed  | 
|
390  | 
||
391  | 
qualified lemma explode_implode_eq [simp]:  | 
|
392  | 
"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
 | 
393  | 
by transfer rule  | 
| 
54594
 
a2d1522cdd54
setup lifting/transfer for String.literal
 
Andreas Lochbihler 
parents: 
54317 
diff
changeset
 | 
394  | 
|
| 68028 | 395  | 
end  | 
396  | 
||
397  | 
||
398  | 
subsubsection \<open>Syntactic representation\<close>  | 
|
399  | 
||
400  | 
text \<open>  | 
|
401  | 
Logical ground representations for literals are:  | 
|
402  | 
||
| 69272 | 403  | 
\<^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
 | 
404  | 
|
| 69272 | 405  | 
\<^enum> \<open>Literal b0 \<dots> b6 s\<close> for a literal starting with one  | 
| 68028 | 406  | 
character and continued by another literal.  | 
407  | 
||
408  | 
Syntactic representations for literals are:  | 
|
409  | 
||
| 69272 | 410  | 
\<^enum> Printable text as string prefixed with \<open>STR\<close>;  | 
| 68028 | 411  | 
|
| 69272 | 412  | 
\<^enum> A single ascii value as numerical hexadecimal value prefixed with \<open>STR\<close>.  | 
| 68028 | 413  | 
\<close>  | 
414  | 
||
415  | 
instantiation String.literal :: zero  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
416  | 
begin  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
417  | 
|
| 68028 | 418  | 
context  | 
419  | 
begin  | 
|
420  | 
||
421  | 
qualified lift_definition zero_literal :: String.literal  | 
|
422  | 
is Nil  | 
|
423  | 
by simp  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
424  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
425  | 
instance ..  | 
| 
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
426  | 
|
| 
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
427  | 
end  | 
| 
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
428  | 
|
| 68028 | 429  | 
end  | 
430  | 
||
431  | 
context  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
432  | 
begin  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
433  | 
|
| 68028 | 434  | 
qualified abbreviation (output) empty_literal :: String.literal  | 
435  | 
where "empty_literal \<equiv> 0"  | 
|
436  | 
||
437  | 
qualified lift_definition Literal :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal"  | 
|
438  | 
is "\<lambda>b0 b1 b2 b3 b4 b5 b6 cs. Char b0 b1 b2 b3 b4 b5 b6 False # cs"  | 
|
439  | 
by auto  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
440  | 
|
| 68028 | 441  | 
qualified lemma Literal_eq_iff [simp]:  | 
442  | 
"Literal b0 b1 b2 b3 b4 b5 b6 s = Literal c0 c1 c2 c3 c4 c5 c6 t  | 
|
443  | 
\<longleftrightarrow> (b0 \<longleftrightarrow> c0) \<and> (b1 \<longleftrightarrow> c1) \<and> (b2 \<longleftrightarrow> c2) \<and> (b3 \<longleftrightarrow> c3)  | 
|
444  | 
\<and> (b4 \<longleftrightarrow> c4) \<and> (b5 \<longleftrightarrow> c5) \<and> (b6 \<longleftrightarrow> c6) \<and> s = t"  | 
|
445  | 
by transfer simp  | 
|
446  | 
||
447  | 
qualified lemma empty_neq_Literal [simp]:  | 
|
448  | 
"empty_literal \<noteq> Literal b0 b1 b2 b3 b4 b5 b6 s"  | 
|
449  | 
by transfer simp  | 
|
450  | 
||
451  | 
qualified lemma Literal_neq_empty [simp]:  | 
|
452  | 
"Literal b0 b1 b2 b3 b4 b5 b6 s \<noteq> empty_literal"  | 
|
453  | 
by transfer simp  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
454  | 
|
| 
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
455  | 
end  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
456  | 
|
| 68028 | 457  | 
code_datatype "0 :: String.literal" String.Literal  | 
458  | 
||
459  | 
syntax  | 
|
460  | 
  "_Literal" :: "str_position \<Rightarrow> String.literal"   ("STR _")
 | 
|
| 
68033
 
ad4b8b6892c3
uniform tagging for printable and non-printable literals
 
haftmann 
parents: 
68028 
diff
changeset
 | 
461  | 
  "_Ascii" :: "num_const \<Rightarrow> String.literal"        ("STR _")
 | 
| 
54594
 
a2d1522cdd54
setup lifting/transfer for String.literal
 
Andreas Lochbihler 
parents: 
54317 
diff
changeset
 | 
462  | 
|
| 69605 | 463  | 
ML_file \<open>Tools/literal.ML\<close>  | 
| 68028 | 464  | 
|
| 
52365
 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 
haftmann 
parents: 
51717 
diff
changeset
 | 
465  | 
|
| 68028 | 466  | 
subsubsection \<open>Operations\<close>  | 
467  | 
||
468  | 
instantiation String.literal :: plus  | 
|
| 67730 | 469  | 
begin  | 
470  | 
||
| 68028 | 471  | 
context  | 
472  | 
begin  | 
|
473  | 
||
474  | 
qualified lift_definition plus_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal"  | 
|
475  | 
is "(@)"  | 
|
476  | 
by auto  | 
|
| 67730 | 477  | 
|
478  | 
instance ..  | 
|
479  | 
||
480  | 
end  | 
|
481  | 
||
| 68028 | 482  | 
end  | 
| 67730 | 483  | 
|
| 68028 | 484  | 
instance String.literal :: monoid_add  | 
485  | 
by (standard; transfer) simp_all  | 
|
486  | 
||
487  | 
instantiation String.literal :: size  | 
|
| 67729 | 488  | 
begin  | 
489  | 
||
| 68028 | 490  | 
context  | 
491  | 
includes literal.lifting  | 
|
492  | 
begin  | 
|
493  | 
||
494  | 
lift_definition size_literal :: "String.literal \<Rightarrow> nat"  | 
|
495  | 
is length .  | 
|
496  | 
||
497  | 
end  | 
|
| 67729 | 498  | 
|
499  | 
instance ..  | 
|
500  | 
||
501  | 
end  | 
|
502  | 
||
| 68028 | 503  | 
instantiation String.literal :: equal  | 
504  | 
begin  | 
|
505  | 
||
506  | 
context  | 
|
507  | 
begin  | 
|
508  | 
||
509  | 
qualified lift_definition equal_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool"  | 
|
510  | 
is HOL.equal .  | 
|
511  | 
||
512  | 
instance  | 
|
513  | 
by (standard; transfer) (simp add: equal)  | 
|
514  | 
||
515  | 
end  | 
|
516  | 
||
517  | 
end  | 
|
518  | 
||
519  | 
instantiation String.literal :: linorder  | 
|
520  | 
begin  | 
|
| 67729 | 521  | 
|
| 68028 | 522  | 
context  | 
523  | 
begin  | 
|
524  | 
||
525  | 
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
 | 
526  | 
is "ord.lexordp_eq (\<lambda>c d. of_char c < (of_char d :: nat))"  | 
| 68028 | 527  | 
.  | 
528  | 
||
529  | 
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
 | 
530  | 
is "ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat))"  | 
| 68028 | 531  | 
.  | 
532  | 
||
533  | 
instance proof -  | 
|
| 
72170
 
7fa9605b226c
Another go with lex: now lexordp back in class ord
 
paulson <lp15@cam.ac.uk> 
parents: 
72164 
diff
changeset
 | 
534  | 
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
 | 
535  | 
"ord.lexordp (\<lambda>c d. of_char c < (of_char d :: nat)) :: string \<Rightarrow> string \<Rightarrow> bool"  | 
| 68028 | 536  | 
by (rule linorder.lexordp_linorder)  | 
537  | 
show "PROP ?thesis"  | 
|
538  | 
by (standard; transfer) (simp_all add: less_le_not_le linear)  | 
|
539  | 
qed  | 
|
540  | 
||
541  | 
end  | 
|
542  | 
||
543  | 
end  | 
|
| 67730 | 544  | 
|
| 68028 | 545  | 
lemma infinite_literal:  | 
546  | 
"infinite (UNIV :: String.literal set)"  | 
|
547  | 
proof -  | 
|
548  | 
define S where "S = range (\<lambda>n. replicate n CHR ''A'')"  | 
|
549  | 
have "inj_on String.implode S"  | 
|
550  | 
proof (rule inj_onI)  | 
|
551  | 
fix cs ds  | 
|
552  | 
assume "String.implode cs = String.implode ds"  | 
|
553  | 
then have "String.explode (String.implode cs) = String.explode (String.implode ds)"  | 
|
554  | 
by simp  | 
|
555  | 
moreover assume "cs \<in> S" and "ds \<in> S"  | 
|
556  | 
ultimately show "cs = ds"  | 
|
557  | 
by (auto simp add: S_def)  | 
|
558  | 
qed  | 
|
559  | 
moreover have "infinite S"  | 
|
560  | 
by (auto simp add: S_def dest: finite_range_imageI [of _ length])  | 
|
561  | 
ultimately have "infinite (String.implode ` S)"  | 
|
562  | 
by (simp add: finite_image_iff)  | 
|
563  | 
then show ?thesis  | 
|
564  | 
by (auto intro: finite_subset)  | 
|
565  | 
qed  | 
|
566  | 
||
567  | 
||
568  | 
subsubsection \<open>Executable conversions\<close>  | 
|
569  | 
||
570  | 
context  | 
|
571  | 
begin  | 
|
572  | 
||
573  | 
qualified lift_definition asciis_of_literal :: "String.literal \<Rightarrow> integer list"  | 
|
574  | 
is "map of_char"  | 
|
575  | 
.  | 
|
576  | 
||
| 69879 | 577  | 
qualified lemma asciis_of_zero [simp, code]:  | 
578  | 
"asciis_of_literal 0 = []"  | 
|
579  | 
by transfer simp  | 
|
580  | 
||
581  | 
qualified lemma asciis_of_Literal [simp, code]:  | 
|
582  | 
"asciis_of_literal (String.Literal b0 b1 b2 b3 b4 b5 b6 s) =  | 
|
583  | 
of_char (Char b0 b1 b2 b3 b4 b5 b6 False) # asciis_of_literal s "  | 
|
584  | 
by transfer simp  | 
|
585  | 
||
| 68028 | 586  | 
qualified lift_definition literal_of_asciis :: "integer list \<Rightarrow> String.literal"  | 
587  | 
is "map (String.ascii_of \<circ> char_of)"  | 
|
588  | 
by auto  | 
|
| 
55426
 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 
Andreas Lochbihler 
parents: 
55015 
diff
changeset
 | 
589  | 
|
| 69879 | 590  | 
qualified lemma literal_of_asciis_Nil [simp, code]:  | 
591  | 
"literal_of_asciis [] = 0"  | 
|
592  | 
by transfer simp  | 
|
593  | 
||
594  | 
qualified lemma literal_of_asciis_Cons [simp, code]:  | 
|
595  | 
"literal_of_asciis (k # ks) = (case char_of k  | 
|
596  | 
of Char b0 b1 b2 b3 b4 b5 b6 b7 \<Rightarrow> String.Literal b0 b1 b2 b3 b4 b5 b6 (literal_of_asciis ks))"  | 
|
597  | 
by (simp add: char_of_def) (transfer, simp add: char_of_def)  | 
|
598  | 
||
| 68028 | 599  | 
qualified lemma literal_of_asciis_of_literal [simp]:  | 
600  | 
"literal_of_asciis (asciis_of_literal s) = s"  | 
|
601  | 
proof transfer  | 
|
602  | 
fix cs  | 
|
603  | 
assume "\<forall>c\<in>set cs. \<not> digit7 c"  | 
|
604  | 
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
 | 
605  | 
by (induction cs) (simp_all add: String.ascii_of_idem)  | 
| 68028 | 606  | 
qed  | 
607  | 
||
608  | 
qualified lemma explode_code [code]:  | 
|
609  | 
"String.explode s = map char_of (asciis_of_literal s)"  | 
|
610  | 
by transfer simp  | 
|
611  | 
||
612  | 
qualified lemma implode_code [code]:  | 
|
613  | 
"String.implode cs = literal_of_asciis (map of_char cs)"  | 
|
614  | 
by transfer simp  | 
|
| 
64994
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64630 
diff
changeset
 | 
615  | 
|
| 69879 | 616  | 
qualified lemma equal_literal [code]:  | 
617  | 
"HOL.equal (String.Literal b0 b1 b2 b3 b4 b5 b6 s)  | 
|
618  | 
(String.Literal a0 a1 a2 a3 a4 a5 a6 r)  | 
|
619  | 
\<longleftrightarrow> (b0 \<longleftrightarrow> a0) \<and> (b1 \<longleftrightarrow> a1) \<and> (b2 \<longleftrightarrow> a2) \<and> (b3 \<longleftrightarrow> a3)  | 
|
620  | 
\<and> (b4 \<longleftrightarrow> a4) \<and> (b5 \<longleftrightarrow> a5) \<and> (b6 \<longleftrightarrow> a6) \<and> (s = r)"  | 
|
621  | 
by (simp add: equal)  | 
|
| 68028 | 622  | 
|
| 69879 | 623  | 
end  | 
| 68028 | 624  | 
|
625  | 
||
626  | 
subsubsection \<open>Technical code generation setup\<close>  | 
|
627  | 
||
628  | 
text \<open>Alternative constructor for generated computations\<close>  | 
|
629  | 
||
630  | 
context  | 
|
| 
72170
 
7fa9605b226c
Another go with lex: now lexordp back in class ord
 
paulson <lp15@cam.ac.uk> 
parents: 
72164 
diff
changeset
 | 
631  | 
begin  | 
| 68028 | 632  | 
|
633  | 
qualified definition Literal' :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> String.literal \<Rightarrow> String.literal"  | 
|
634  | 
where [simp]: "Literal' = String.Literal"  | 
|
635  | 
||
636  | 
lemma [code]:  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
637  | 
\<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
 | 
638  | 
[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
 | 
639  | 
proof -  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
640  | 
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
 | 
641  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
642  | 
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
 | 
643  | 
[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
 | 
644  | 
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
 | 
645  | 
ultimately show ?thesis  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
646  | 
by simp  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71094 
diff
changeset
 | 
647  | 
qed  | 
| 
64994
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64630 
diff
changeset
 | 
648  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64630 
diff
changeset
 | 
649  | 
lemma [code_computation_unfold]:  | 
| 68028 | 650  | 
"String.Literal = Literal'"  | 
651  | 
by simp  | 
|
| 
64994
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64630 
diff
changeset
 | 
652  | 
|
| 68028 | 653  | 
end  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
654  | 
|
| 69879 | 655  | 
code_reserved SML string String Char List  | 
| 68028 | 656  | 
code_reserved OCaml string String Char List  | 
657  | 
code_reserved Haskell Prelude  | 
|
| 34886 | 658  | 
code_reserved Scala string  | 
| 33237 | 659  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52365 
diff
changeset
 | 
660  | 
code_printing  | 
| 68028 | 661  | 
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
 | 
662  | 
(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
 | 
663  | 
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
 | 
664  | 
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
 | 
665  | 
and (Scala) "String"  | 
| 68028 | 666  | 
| constant "STR ''''" \<rightharpoonup>  | 
667  | 
(SML) "\"\""  | 
|
668  | 
and (OCaml) "\"\""  | 
|
669  | 
and (Haskell) "\"\""  | 
|
670  | 
and (Scala) "\"\""  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
671  | 
|
| 60758 | 672  | 
setup \<open>  | 
| 68028 | 673  | 
fold Literal.add_code ["SML", "OCaml", "Haskell", "Scala"]  | 
| 60758 | 674  | 
\<close>  | 
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
675  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52365 
diff
changeset
 | 
676  | 
code_printing  | 
| 68028 | 677  | 
constant "(+) :: String.literal \<Rightarrow> String.literal \<Rightarrow> String.literal" \<rightharpoonup>  | 
678  | 
(SML) infixl 18 "^"  | 
|
679  | 
and (OCaml) infixr 6 "^"  | 
|
680  | 
and (Haskell) infixr 5 "++"  | 
|
681  | 
and (Scala) infixl 7 "+"  | 
|
682  | 
| constant String.literal_of_asciis \<rightharpoonup>  | 
|
| 69879 | 683  | 
(SML) "!(String.implode/ o List.map (fn k => if 0 <= k andalso k < 128 then (Char.chr o IntInf.toInt) k else raise Fail \"Non-ASCII character in literal\"))"  | 
| 69743 | 684  | 
and (OCaml) "!(let xs = _  | 
685  | 
and chr k =  | 
|
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69879 
diff
changeset
 | 
686  | 
let l = Z.to'_int k  | 
| 69743 | 687  | 
in if 0 <= l && l < 128  | 
688  | 
then Char.chr l  | 
|
689  | 
else failwith \"Non-ASCII character in literal\"  | 
|
690  | 
in String.init (List.length xs) (List.nth (List.map chr xs)))"  | 
|
| 68028 | 691  | 
and (Haskell) "map/ (let chr k | (0 <= k && k < 128) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)"  | 
692  | 
and (Scala) "\"\"/ ++/ _.map((k: BigInt) => if (BigInt(0) <= k && k < BigInt(128)) k.charValue else sys.error(\"Non-ASCII character in literal\"))"  | 
|
693  | 
| constant String.asciis_of_literal \<rightharpoonup>  | 
|
| 69879 | 694  | 
(SML) "!(List.map (fn c => let val k = Char.ord c in if k < 128 then IntInf.fromInt k else raise Fail \"Non-ASCII character in literal\" end) /o String.explode)"  | 
| 69743 | 695  | 
and (OCaml) "!(let s = _ in let rec exp i l = if i < 0 then l else exp (i - 1) (let k = Char.code (String.get s i) in  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69879 
diff
changeset
 | 
696  | 
if k < 128 then Z.of'_int k :: l else failwith \"Non-ASCII character in literal\") in exp (String.length s - 1) [])"  | 
| 68028 | 697  | 
and (Haskell) "map/ (let ord k | (k < 128) = Prelude.toInteger k in ord . (Prelude.fromEnum :: Prelude.Char -> Prelude.Int))"  | 
698  | 
    and (Scala) "!(_.toList.map(c => { val k: Int = c.toInt; if (k < 128) BigInt(k) else sys.error(\"Non-ASCII character in literal\") }))"
 | 
|
699  | 
| 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
 | 
700  | 
(Haskell) -  | 
| 68028 | 701  | 
| constant "HOL.equal :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<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
 | 
702  | 
(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
 | 
703  | 
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
 | 
704  | 
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
 | 
705  | 
and (Scala) infixl 5 "=="  | 
| 68028 | 706  | 
| constant "(\<le>) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>  | 
707  | 
(SML) "!((_ : string) <= _)"  | 
|
708  | 
and (OCaml) "!((_ : string) <= _)"  | 
|
| 69743 | 709  | 
and (Haskell) infix 4 "<="  | 
| 69593 | 710  | 
\<comment> \<open>Order operations for \<^typ>\<open>String.literal\<close> work in Haskell only  | 
| 68028 | 711  | 
if no type class instance needs to be generated, because String = [Char] in Haskell  | 
| 69593 | 712  | 
and \<^typ>\<open>char list\<close> need not have the same order as \<^typ>\<open>String.literal\<close>.\<close>  | 
| 68028 | 713  | 
and (Scala) infixl 4 "<="  | 
714  | 
and (Eval) infixl 6 "<="  | 
|
715  | 
| constant "(<) :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>  | 
|
716  | 
(SML) "!((_ : string) < _)"  | 
|
717  | 
and (OCaml) "!((_ : string) < _)"  | 
|
718  | 
and (Haskell) infix 4 "<"  | 
|
719  | 
and (Scala) infixl 4 "<"  | 
|
720  | 
and (Eval) infixl 6 "<"  | 
|
721  | 
||
722  | 
||
723  | 
subsubsection \<open>Code generation utility\<close>  | 
|
| 
31051
 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 
haftmann 
parents:  
diff
changeset
 | 
724  | 
|
| 60758 | 725  | 
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
 | 
726  | 
|
| 68028 | 727  | 
definition abort :: "String.literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a"  | 
728  | 
where [simp]: "abort _ f = f ()"  | 
|
| 
52910
 
7bfe0df532a9
abort execution of generated code with explicit exception message
 
Andreas Lochbihler 
parents: 
52435 
diff
changeset
 | 
729  | 
|
| 68028 | 730  | 
declare [[code drop: Code.abort]]  | 
731  | 
||
732  | 
lemma abort_cong:  | 
|
733  | 
"msg = msg' \<Longrightarrow> Code.abort msg f = Code.abort msg' f"  | 
|
734  | 
by simp  | 
|
| 
54317
 
da932f511746
add congruence rule to prevent code_simp from looping
 
Andreas Lochbihler 
parents: 
52910 
diff
changeset
 | 
735  | 
|
| 60758 | 736  | 
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
 | 
737  | 
|
| 60758 | 738  | 
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
 | 
739  | 
|
| 68028 | 740  | 
code_printing  | 
741  | 
constant Code.abort \<rightharpoonup>  | 
|
| 
52910
 
7bfe0df532a9
abort execution of generated code with explicit exception message
 
Andreas Lochbihler 
parents: 
52435 
diff
changeset
 | 
742  | 
(SML) "!(raise/ Fail/ _)"  | 
| 
 
7bfe0df532a9
abort execution of generated code with explicit exception message
 
Andreas Lochbihler 
parents: 
52435 
diff
changeset
 | 
743  | 
and (OCaml) "failwith"  | 
| 
59483
 
ddb73392356e
explicit type annotation avoids problems with Haskell type inference
 
haftmann 
parents: 
58889 
diff
changeset
 | 
744  | 
and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)"  | 
| 68028 | 745  | 
    and (Scala) "!{/ sys.error((_));/ ((_)).apply(())/ }"
 | 
746  | 
||
| 
52910
 
7bfe0df532a9
abort execution of generated code with explicit exception message
 
Andreas Lochbihler 
parents: 
52435 
diff
changeset
 | 
747  | 
|
| 68028 | 748  | 
subsubsection \<open>Finally\<close>  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31176 
diff
changeset
 | 
749  | 
|
| 68028 | 750  | 
lifting_update literal.lifting  | 
751  | 
lifting_forget literal.lifting  | 
|
| 57437 | 752  | 
|
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
753  | 
end  |