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