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