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