author | eberlm <eberlm@in.tum.de> |
Thu, 17 Aug 2017 18:19:16 +0200 | |
changeset 66448 | 97ad7a583457 |
parent 66331 | f773691617c0 |
child 67091 | 1393c2340eec |
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 |
|
60758 | 9 |
subsection \<open>Characters and strings\<close> |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
10 |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
11 |
subsubsection \<open>Characters as finite algebraic type\<close> |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
12 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
13 |
typedef char = "{n::nat. n < 256}" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
14 |
morphisms nat_of_char Abs_char |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
15 |
proof |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
16 |
show "Suc 0 \<in> {n. n < 256}" by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
17 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
18 |
|
66331 | 19 |
setup_lifting type_definition_char |
20 |
||
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
21 |
definition char_of_nat :: "nat \<Rightarrow> char" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
22 |
where |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
23 |
"char_of_nat n = Abs_char (n mod 256)" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
24 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
25 |
lemma char_cases [case_names char_of_nat, cases type: char]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
26 |
"(\<And>n. c = char_of_nat n \<Longrightarrow> n < 256 \<Longrightarrow> P) \<Longrightarrow> P" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
27 |
by (cases c) (simp add: char_of_nat_def) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
28 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
29 |
lemma char_of_nat_of_char [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
30 |
"char_of_nat (nat_of_char c) = c" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
31 |
by (cases c) (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
32 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
33 |
lemma inj_nat_of_char: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
34 |
"inj nat_of_char" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
35 |
proof (rule injI) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
36 |
fix c d |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
37 |
assume "nat_of_char c = nat_of_char d" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
38 |
then have "char_of_nat (nat_of_char c) = char_of_nat (nat_of_char d)" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
39 |
by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
40 |
then show "c = d" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
41 |
by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
42 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
43 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
44 |
lemma nat_of_char_eq_iff [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
45 |
"nat_of_char c = nat_of_char d \<longleftrightarrow> c = d" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
46 |
by (fact nat_of_char_inject) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
47 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
48 |
lemma nat_of_char_of_nat [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
49 |
"nat_of_char (char_of_nat n) = n mod 256" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
50 |
by (cases "char_of_nat n") (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
51 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
52 |
lemma char_of_nat_mod_256 [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
53 |
"char_of_nat (n mod 256) = char_of_nat n" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
54 |
by (cases "char_of_nat (n mod 256)") (simp add: char_of_nat_def) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
55 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
56 |
lemma char_of_nat_quasi_inj [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
57 |
"char_of_nat m = char_of_nat n \<longleftrightarrow> m mod 256 = n mod 256" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
58 |
by (simp add: char_of_nat_def Abs_char_inject) |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
59 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
60 |
lemma inj_on_char_of_nat [simp]: |
62597 | 61 |
"inj_on char_of_nat {..<256}" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
62 |
by (rule inj_onI) simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
63 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
64 |
lemma nat_of_char_mod_256 [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
65 |
"nat_of_char c mod 256 = nat_of_char c" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
66 |
by (cases c) simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
67 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
68 |
lemma nat_of_char_less_256 [simp]: |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
69 |
"nat_of_char c < 256" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
70 |
proof - |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
71 |
have "nat_of_char c mod 256 < 256" |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
72 |
by arith |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
73 |
then show ?thesis by simp |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
74 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
75 |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
76 |
lemma UNIV_char_of_nat: |
62597 | 77 |
"UNIV = char_of_nat ` {..<256}" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
78 |
proof - |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
79 |
{ fix c |
62597 | 80 |
have "c \<in> char_of_nat ` {..<256}" |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
81 |
by (cases c) auto |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
82 |
} then show ?thesis by auto |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
83 |
qed |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
84 |
|
62597 | 85 |
lemma card_UNIV_char: |
86 |
"card (UNIV :: char set) = 256" |
|
87 |
by (auto simp add: UNIV_char_of_nat card_image) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
88 |
|
62597 | 89 |
lemma range_nat_of_char: |
90 |
"range nat_of_char = {..<256}" |
|
91 |
by (auto simp add: UNIV_char_of_nat image_image image_def) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
92 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
93 |
|
62597 | 94 |
subsubsection \<open>Character literals as variant of numerals\<close> |
95 |
||
96 |
instantiation char :: zero |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
97 |
begin |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
98 |
|
62597 | 99 |
definition zero_char :: char |
100 |
where "0 = char_of_nat 0" |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
101 |
|
62597 | 102 |
instance .. |
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
103 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
104 |
end |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
105 |
|
62597 | 106 |
definition Char :: "num \<Rightarrow> char" |
107 |
where "Char k = char_of_nat (numeral k)" |
|
108 |
||
109 |
code_datatype "0 :: char" Char |
|
110 |
||
111 |
lemma nat_of_char_zero [simp]: |
|
112 |
"nat_of_char 0 = 0" |
|
113 |
by (simp add: zero_char_def) |
|
114 |
||
115 |
lemma nat_of_char_Char [simp]: |
|
116 |
"nat_of_char (Char k) = numeral k mod 256" |
|
117 |
by (simp add: Char_def) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
118 |
|
64630
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
119 |
lemma Char_eq_Char_iff: |
62597 | 120 |
"Char k = Char l \<longleftrightarrow> numeral k mod (256 :: nat) = numeral l mod 256" (is "?P \<longleftrightarrow> ?Q") |
121 |
proof - |
|
122 |
have "?P \<longleftrightarrow> nat_of_char (Char k) = nat_of_char (Char l)" |
|
123 |
by (rule sym, rule inj_eq) (fact inj_nat_of_char) |
|
124 |
also have "nat_of_char (Char k) = nat_of_char (Char l) \<longleftrightarrow> ?Q" |
|
125 |
by (simp only: Char_def nat_of_char_of_nat) |
|
126 |
finally show ?thesis . |
|
127 |
qed |
|
128 |
||
64630
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
129 |
lemma zero_eq_Char_iff: |
62597 | 130 |
"0 = Char k \<longleftrightarrow> numeral k mod (256 :: nat) = 0" |
131 |
by (auto simp add: zero_char_def Char_def) |
|
132 |
||
64630
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
133 |
lemma Char_eq_zero_iff: |
62597 | 134 |
"Char k = 0 \<longleftrightarrow> numeral k mod (256 :: nat) = 0" |
135 |
by (auto simp add: zero_char_def Char_def) |
|
136 |
||
64630
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
137 |
simproc_setup char_eq |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
138 |
("Char m = Char n" | "Char m = 0" | "0 = Char n") = \<open> |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
139 |
let |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
140 |
val ss = put_simpset HOL_ss @{context} |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
141 |
|> fold Simplifier.add_simp @{thms Char_eq_Char_iff zero_eq_Char_iff Char_eq_zero_iff cong_exp_iff_simps} |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
142 |
|> simpset_of |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
143 |
in |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
144 |
fn _ => fn ctxt => SOME o Simplifier.rewrite (put_simpset ss ctxt) |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
145 |
end |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
146 |
\<close> |
96015aecfeba
emphasize dedicated rewrite rules for congruences
haftmann
parents:
63950
diff
changeset
|
147 |
|
62597 | 148 |
definition integer_of_char :: "char \<Rightarrow> integer" |
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
149 |
where |
62597 | 150 |
"integer_of_char = integer_of_nat \<circ> nat_of_char" |
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
151 |
|
62597 | 152 |
definition char_of_integer :: "integer \<Rightarrow> char" |
153 |
where |
|
154 |
"char_of_integer = char_of_nat \<circ> nat_of_integer" |
|
155 |
||
156 |
lemma integer_of_char_zero [simp, code]: |
|
157 |
"integer_of_char 0 = 0" |
|
158 |
by (simp add: integer_of_char_def integer_of_nat_def) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
159 |
|
62597 | 160 |
lemma integer_of_char_Char [simp]: |
161 |
"integer_of_char (Char k) = numeral k mod 256" |
|
162 |
by (simp only: integer_of_char_def integer_of_nat_def comp_apply nat_of_char_Char map_fun_def |
|
63950
cdc1e59aa513
syntactic type class for operation mod named after mod;
haftmann
parents:
62678
diff
changeset
|
163 |
id_apply zmod_int modulo_integer.abs_eq [symmetric]) simp |
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
164 |
|
66190 | 165 |
lemma integer_of_char_Char_code [code]: |
166 |
"integer_of_char (Char k) = integer_of_num k mod 256" |
|
167 |
by simp |
|
168 |
||
62597 | 169 |
lemma nat_of_char_code [code]: |
170 |
"nat_of_char = nat_of_integer \<circ> integer_of_char" |
|
171 |
by (simp add: integer_of_char_def fun_eq_iff) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
172 |
|
62597 | 173 |
lemma char_of_nat_code [code]: |
174 |
"char_of_nat = char_of_integer \<circ> integer_of_nat" |
|
175 |
by (simp add: char_of_integer_def fun_eq_iff) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
176 |
|
62597 | 177 |
instantiation char :: equal |
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
178 |
begin |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
179 |
|
62597 | 180 |
definition equal_char |
181 |
where "equal_char (c :: char) d \<longleftrightarrow> c = d" |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
182 |
|
62597 | 183 |
instance |
184 |
by standard (simp add: equal_char_def) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
185 |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
186 |
end |
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
187 |
|
62597 | 188 |
lemma equal_char_simps [code]: |
189 |
"HOL.equal (0::char) 0 \<longleftrightarrow> True" |
|
190 |
"HOL.equal (Char k) (Char l) \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) (numeral l mod 256)" |
|
191 |
"HOL.equal 0 (Char k) \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) 0" |
|
192 |
"HOL.equal (Char k) 0 \<longleftrightarrow> HOL.equal (numeral k mod 256 :: nat) 0" |
|
193 |
by (simp_all only: equal refl Char_eq_Char_iff zero_eq_Char_iff Char_eq_zero_iff) |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
194 |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
195 |
syntax |
62597 | 196 |
"_Char" :: "str_position \<Rightarrow> char" ("CHR _") |
62678 | 197 |
"_Char_ord" :: "num_const \<Rightarrow> char" ("CHR _") |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
198 |
|
42163
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
bulwahn
parents:
41750
diff
changeset
|
199 |
type_synonym string = "char list" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
200 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
201 |
syntax |
46483 | 202 |
"_String" :: "str_position => string" ("_") |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
203 |
|
48891 | 204 |
ML_file "Tools/string_syntax.ML" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
205 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
206 |
instantiation char :: enum |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
207 |
begin |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
208 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
209 |
definition |
62678 | 210 |
"Enum.enum = [0, CHR 0x01, CHR 0x02, CHR 0x03, |
211 |
CHR 0x04, CHR 0x05, CHR 0x06, CHR 0x07, |
|
212 |
CHR 0x08, CHR 0x09, CHR ''\<newline>'', CHR 0x0B, |
|
213 |
CHR 0x0C, CHR 0x0D, CHR 0x0E, CHR 0x0F, |
|
214 |
CHR 0x10, CHR 0x11, CHR 0x12, CHR 0x13, |
|
215 |
CHR 0x14, CHR 0x15, CHR 0x16, CHR 0x17, |
|
216 |
CHR 0x18, CHR 0x19, CHR 0x1A, CHR 0x1B, |
|
217 |
CHR 0x1C, CHR 0x1D, CHR 0x1E, CHR 0x1F, |
|
218 |
CHR '' '', CHR ''!'', CHR 0x22, CHR ''#'', |
|
219 |
CHR ''$'', CHR ''%'', CHR ''&'', CHR 0x27, |
|
62597 | 220 |
CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', |
221 |
CHR '','', CHR ''-'', CHR ''.'', CHR ''/'', |
|
222 |
CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', |
|
223 |
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', |
|
224 |
CHR ''8'', CHR ''9'', CHR '':'', CHR '';'', |
|
225 |
CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', |
|
226 |
CHR ''@'', CHR ''A'', CHR ''B'', CHR ''C'', |
|
227 |
CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', |
|
228 |
CHR ''H'', CHR ''I'', CHR ''J'', CHR ''K'', |
|
229 |
CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', |
|
230 |
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', |
|
231 |
CHR ''T'', CHR ''U'', CHR ''V'', CHR ''W'', |
|
232 |
CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', |
|
62678 | 233 |
CHR 0x5C, CHR '']'', CHR ''^'', CHR ''_'', |
234 |
CHR 0x60, CHR ''a'', CHR ''b'', CHR ''c'', |
|
62597 | 235 |
CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', |
236 |
CHR ''h'', CHR ''i'', CHR ''j'', CHR ''k'', |
|
237 |
CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', |
|
238 |
CHR ''p'', CHR ''q'', CHR ''r'', CHR ''s'', |
|
239 |
CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', |
|
240 |
CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', |
|
62678 | 241 |
CHR ''|'', CHR ''}'', CHR ''~'', CHR 0x7F, |
242 |
CHR 0x80, CHR 0x81, CHR 0x82, CHR 0x83, |
|
243 |
CHR 0x84, CHR 0x85, CHR 0x86, CHR 0x87, |
|
244 |
CHR 0x88, CHR 0x89, CHR 0x8A, CHR 0x8B, |
|
245 |
CHR 0x8C, CHR 0x8D, CHR 0x8E, CHR 0x8F, |
|
246 |
CHR 0x90, CHR 0x91, CHR 0x92, CHR 0x93, |
|
247 |
CHR 0x94, CHR 0x95, CHR 0x96, CHR 0x97, |
|
248 |
CHR 0x98, CHR 0x99, CHR 0x9A, CHR 0x9B, |
|
249 |
CHR 0x9C, CHR 0x9D, CHR 0x9E, CHR 0x9F, |
|
250 |
CHR 0xA0, CHR 0xA1, CHR 0xA2, CHR 0xA3, |
|
251 |
CHR 0xA4, CHR 0xA5, CHR 0xA6, CHR 0xA7, |
|
252 |
CHR 0xA8, CHR 0xA9, CHR 0xAA, CHR 0xAB, |
|
253 |
CHR 0xAC, CHR 0xAD, CHR 0xAE, CHR 0xAF, |
|
254 |
CHR 0xB0, CHR 0xB1, CHR 0xB2, CHR 0xB3, |
|
255 |
CHR 0xB4, CHR 0xB5, CHR 0xB6, CHR 0xB7, |
|
256 |
CHR 0xB8, CHR 0xB9, CHR 0xBA, CHR 0xBB, |
|
257 |
CHR 0xBC, CHR 0xBD, CHR 0xBE, CHR 0xBF, |
|
258 |
CHR 0xC0, CHR 0xC1, CHR 0xC2, CHR 0xC3, |
|
259 |
CHR 0xC4, CHR 0xC5, CHR 0xC6, CHR 0xC7, |
|
260 |
CHR 0xC8, CHR 0xC9, CHR 0xCA, CHR 0xCB, |
|
261 |
CHR 0xCC, CHR 0xCD, CHR 0xCE, CHR 0xCF, |
|
262 |
CHR 0xD0, CHR 0xD1, CHR 0xD2, CHR 0xD3, |
|
263 |
CHR 0xD4, CHR 0xD5, CHR 0xD6, CHR 0xD7, |
|
264 |
CHR 0xD8, CHR 0xD9, CHR 0xDA, CHR 0xDB, |
|
265 |
CHR 0xDC, CHR 0xDD, CHR 0xDE, CHR 0xDF, |
|
266 |
CHR 0xE0, CHR 0xE1, CHR 0xE2, CHR 0xE3, |
|
267 |
CHR 0xE4, CHR 0xE5, CHR 0xE6, CHR 0xE7, |
|
268 |
CHR 0xE8, CHR 0xE9, CHR 0xEA, CHR 0xEB, |
|
269 |
CHR 0xEC, CHR 0xED, CHR 0xEE, CHR 0xEF, |
|
270 |
CHR 0xF0, CHR 0xF1, CHR 0xF2, CHR 0xF3, |
|
271 |
CHR 0xF4, CHR 0xF5, CHR 0xF6, CHR 0xF7, |
|
272 |
CHR 0xF8, CHR 0xF9, CHR 0xFA, CHR 0xFB, |
|
273 |
CHR 0xFC, CHR 0xFD, CHR 0xFE, CHR 0xFF]" |
|
31484 | 274 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
275 |
definition |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
276 |
"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
|
277 |
|
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_ex P \<longleftrightarrow> list_ex P (Enum.enum :: char list)" |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
280 |
|
62597 | 281 |
lemma enum_char_unfold: |
282 |
"Enum.enum = map char_of_nat [0..<256]" |
|
283 |
proof - |
|
284 |
have *: "Suc (Suc 0) = 2" by simp |
|
285 |
have "map nat_of_char Enum.enum = [0..<256]" |
|
286 |
by (simp add: enum_char_def upt_conv_Cons_Cons *) |
|
287 |
then have "map char_of_nat (map nat_of_char Enum.enum) = |
|
288 |
map char_of_nat [0..<256]" by simp |
|
289 |
then show ?thesis by (simp add: comp_def) |
|
290 |
qed |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
291 |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
292 |
instance proof |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
293 |
show UNIV: "UNIV = set (Enum.enum :: char list)" |
62597 | 294 |
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
|
295 |
show "distinct (Enum.enum :: char list)" |
62597 | 296 |
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
|
297 |
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
|
298 |
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
|
299 |
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
|
300 |
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
|
301 |
qed |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
302 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
303 |
end |
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
304 |
|
62597 | 305 |
lemma char_of_integer_code [code]: |
306 |
"char_of_integer n = Enum.enum ! (nat_of_integer n mod 256)" |
|
307 |
by (simp add: char_of_integer_def enum_char_unfold) |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49834
diff
changeset
|
308 |
|
66331 | 309 |
lifting_update char.lifting |
310 |
lifting_forget char.lifting |
|
311 |
||
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
312 |
|
60758 | 313 |
subsection \<open>Strings as dedicated type\<close> |
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
314 |
|
49834 | 315 |
typedef literal = "UNIV :: string set" |
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
316 |
morphisms explode STR .. |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
317 |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
318 |
setup_lifting type_definition_literal |
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
319 |
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
320 |
lemma STR_inject' [simp]: |
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
321 |
"STR s = STR t \<longleftrightarrow> s = t" |
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
322 |
by transfer rule |
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
323 |
|
57437 | 324 |
definition implode :: "string \<Rightarrow> String.literal" |
325 |
where |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
326 |
[code del]: "implode = STR" |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66190
diff
changeset
|
327 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
328 |
instantiation literal :: size |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
329 |
begin |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
330 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
331 |
definition size_literal :: "literal \<Rightarrow> nat" |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
332 |
where |
61076 | 333 |
[code]: "size_literal (s::literal) = 0" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
334 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
335 |
instance .. |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
336 |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
337 |
end |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
338 |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
339 |
instantiation literal :: equal |
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
340 |
begin |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
341 |
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
342 |
lift_definition equal_literal :: "literal \<Rightarrow> literal \<Rightarrow> bool" is "op =" . |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
343 |
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
344 |
instance by intro_classes (transfer, simp) |
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
345 |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
346 |
end |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
347 |
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
348 |
declare equal_literal.rep_eq[code] |
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
349 |
|
52365
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
350 |
lemma [code nbe]: |
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
351 |
fixes s :: "String.literal" |
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
352 |
shows "HOL.equal s s \<longleftrightarrow> True" |
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
353 |
by (simp add: equal) |
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
354 |
|
55426
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
355 |
lifting_update literal.lifting |
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
356 |
lifting_forget literal.lifting |
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
357 |
|
64994
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
358 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
359 |
subsection \<open>Dedicated conversion for generated computations\<close> |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
360 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
361 |
definition char_of_num :: "num \<Rightarrow> char" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
362 |
where "char_of_num = char_of_nat o nat_of_num" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
363 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
364 |
lemma [code_computation_unfold]: |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
365 |
"Char = char_of_num" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
366 |
by (simp add: fun_eq_iff char_of_num_def nat_of_num_numeral Char_def) |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
367 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64630
diff
changeset
|
368 |
|
60758 | 369 |
subsection \<open>Code generator\<close> |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
370 |
|
48891 | 371 |
ML_file "Tools/string_code.ML" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
372 |
|
33237 | 373 |
code_reserved SML string |
374 |
code_reserved OCaml string |
|
34886 | 375 |
code_reserved Scala string |
33237 | 376 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
377 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
378 |
type_constructor literal \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
379 |
(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
|
380 |
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
|
381 |
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
|
382 |
and (Scala) "String" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
383 |
|
60758 | 384 |
setup \<open> |
34886 | 385 |
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"] |
60758 | 386 |
\<close> |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
387 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
388 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
389 |
class_instance literal :: equal \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
390 |
(Haskell) - |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
391 |
| constant "HOL.equal :: literal \<Rightarrow> literal \<Rightarrow> bool" \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
392 |
(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
|
393 |
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
|
394 |
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
|
395 |
and (Scala) infixl 5 "==" |
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
396 |
|
60758 | 397 |
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
|
398 |
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
399 |
definition abort :: "literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a" |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
400 |
where [simp, code del]: "abort _ f = f ()" |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
401 |
|
54317
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
402 |
lemma abort_cong: "msg = msg' ==> Code.abort msg f = Code.abort msg' f" |
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
403 |
by simp |
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
404 |
|
60758 | 405 |
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
|
406 |
|
60758 | 407 |
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
|
408 |
|
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
409 |
code_printing constant Code.abort \<rightharpoonup> |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
410 |
(SML) "!(raise/ Fail/ _)" |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
411 |
and (OCaml) "failwith" |
59483
ddb73392356e
explicit type annotation avoids problems with Haskell type inference
haftmann
parents:
58889
diff
changeset
|
412 |
and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)" |
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
413 |
and (Scala) "!{/ sys.error((_));/ ((_)).apply(())/ }" |
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
414 |
|
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
35123
diff
changeset
|
415 |
hide_type (open) literal |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
416 |
|
57437 | 417 |
hide_const (open) implode explode |
418 |
||
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
419 |
end |