| author | eberlm |
| Fri, 26 Feb 2016 14:58:07 +0100 | |
| changeset 62425 | d0936b500bf5 |
| parent 62364 | 9209770bdcdf |
| child 62580 | 7011429f44f9 |
| permissions | -rw-r--r-- |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
1 |
(* Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
2 |
|
| 60758 | 3 |
section \<open>Character and string types\<close> |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
4 |
|
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
5 |
theory String |
| 55969 | 6 |
imports Enum |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
7 |
begin |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
8 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
9 |
lemma upt_conv_Cons_Cons: -- \<open>no precondition\<close> -- \<open>FIXME move\<close> |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
10 |
"m # n # ns = [m..<q] \<longleftrightarrow> n # ns = [Suc m..<q]" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
11 |
proof (cases "m < q") |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
12 |
case False then show ?thesis by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
13 |
next |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
14 |
case True then show ?thesis by (simp add: upt_conv_Cons) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
15 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
16 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
17 |
|
| 60758 | 18 |
subsection \<open>Characters and strings\<close> |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
19 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
20 |
subsubsection \<open>Characters as finite algebraic type\<close> |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
21 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
22 |
typedef char = "{n::nat. n < 256}"
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
23 |
morphisms nat_of_char Abs_char |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
24 |
proof |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
25 |
show "Suc 0 \<in> {n. n < 256}" by simp
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
26 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
27 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
28 |
definition char_of_nat :: "nat \<Rightarrow> char" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
29 |
where |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
30 |
"char_of_nat n = Abs_char (n mod 256)" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
31 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
32 |
lemma char_cases [case_names char_of_nat, cases type: char]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
33 |
"(\<And>n. c = char_of_nat n \<Longrightarrow> n < 256 \<Longrightarrow> P) \<Longrightarrow> P" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
34 |
by (cases c) (simp add: char_of_nat_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
35 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
36 |
lemma char_of_nat_of_char [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
37 |
"char_of_nat (nat_of_char c) = c" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
38 |
by (cases c) (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
39 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
40 |
lemma inj_nat_of_char: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
41 |
"inj nat_of_char" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
42 |
proof (rule injI) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
43 |
fix c d |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
44 |
assume "nat_of_char c = nat_of_char d" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
45 |
then have "char_of_nat (nat_of_char c) = char_of_nat (nat_of_char d)" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
46 |
by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
47 |
then show "c = d" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
48 |
by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
49 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
50 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
51 |
lemma nat_of_char_eq_iff [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
52 |
"nat_of_char c = nat_of_char d \<longleftrightarrow> c = d" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
53 |
by (fact nat_of_char_inject) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
54 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
55 |
lemma nat_of_char_of_nat [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
56 |
"nat_of_char (char_of_nat n) = n mod 256" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
57 |
by (cases "char_of_nat n") (simp add: char_of_nat_def Abs_char_inject Abs_char_inverse) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
58 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
59 |
lemma char_of_nat_mod_256 [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
60 |
"char_of_nat (n mod 256) = char_of_nat n" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
61 |
by (cases "char_of_nat (n mod 256)") (simp add: char_of_nat_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
62 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
63 |
lemma char_of_nat_quasi_inj [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
64 |
"char_of_nat m = char_of_nat n \<longleftrightarrow> m mod 256 = n mod 256" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
65 |
by (simp add: char_of_nat_def Abs_char_inject) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
66 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
67 |
lemma inj_on_char_of_nat [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
68 |
"inj_on char_of_nat {0..<256}"
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
69 |
by (rule inj_onI) simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
70 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
71 |
lemma nat_of_char_mod_256 [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
72 |
"nat_of_char c mod 256 = nat_of_char c" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
73 |
by (cases c) simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
74 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
75 |
lemma nat_of_char_less_256 [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
76 |
"nat_of_char c < 256" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
77 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
78 |
have "nat_of_char c mod 256 < 256" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
79 |
by arith |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
80 |
then show ?thesis by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
81 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
82 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
83 |
lemma UNIV_char_of_nat: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
84 |
"UNIV = char_of_nat ` {0..<256}"
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
85 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
86 |
{ fix c
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
87 |
have "c \<in> char_of_nat ` {0..<256}"
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
88 |
by (cases c) auto |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
89 |
} then show ?thesis by auto |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
90 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
91 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
92 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
93 |
subsubsection \<open>Traditional concrete representation of characters\<close> |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
94 |
|
|
61348
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
blanchet
parents:
61076
diff
changeset
|
95 |
datatype (plugins del: transfer) nibble = |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
96 |
Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7 |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
97 |
| Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
98 |
|
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
99 |
lemma UNIV_nibble: |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
100 |
"UNIV = {Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7,
|
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
101 |
Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF}" (is "_ = ?A") |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
102 |
proof (rule UNIV_eq_I) |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
103 |
fix x show "x \<in> ?A" by (cases x) simp_all |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
104 |
qed |
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
105 |
|
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
106 |
lemma size_nibble [code, simp]: |
|
56846
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents:
55969
diff
changeset
|
107 |
"size_nibble (x::nibble) = 0" |
|
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents:
55969
diff
changeset
|
108 |
"size (x::nibble) = 0" |
|
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
blanchet
parents:
55969
diff
changeset
|
109 |
by (cases x, simp_all)+ |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
110 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
111 |
instantiation nibble :: enum |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
112 |
begin |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
113 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
114 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
115 |
"Enum.enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
116 |
Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
117 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
118 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
119 |
"Enum.enum_all P \<longleftrightarrow> P Nibble0 \<and> P Nibble1 \<and> P Nibble2 \<and> P Nibble3 \<and> P Nibble4 \<and> P Nibble5 \<and> P Nibble6 \<and> P Nibble7 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
120 |
\<and> P Nibble8 \<and> P Nibble9 \<and> P NibbleA \<and> P NibbleB \<and> P NibbleC \<and> P NibbleD \<and> P NibbleE \<and> P NibbleF" |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
121 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
122 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
123 |
"Enum.enum_ex P \<longleftrightarrow> P Nibble0 \<or> P Nibble1 \<or> P Nibble2 \<or> P Nibble3 \<or> P Nibble4 \<or> P Nibble5 \<or> P Nibble6 \<or> P Nibble7 |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
124 |
\<or> P Nibble8 \<or> P Nibble9 \<or> P NibbleA \<or> P NibbleB \<or> P NibbleC \<or> P NibbleD \<or> P NibbleE \<or> P NibbleF" |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
125 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
126 |
instance proof |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
127 |
qed (simp_all only: enum_nibble_def enum_all_nibble_def enum_ex_nibble_def UNIV_nibble, simp_all) |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
128 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
129 |
end |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
130 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
131 |
lemma card_UNIV_nibble: |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
132 |
"card (UNIV :: nibble set) = 16" |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
133 |
by (simp add: card_UNIV_length_enum enum_nibble_def) |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
134 |
|
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
135 |
primrec nat_of_nibble :: "nibble \<Rightarrow> nat" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
136 |
where |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
137 |
"nat_of_nibble Nibble0 = 0" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
138 |
| "nat_of_nibble Nibble1 = 1" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
139 |
| "nat_of_nibble Nibble2 = 2" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
140 |
| "nat_of_nibble Nibble3 = 3" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
141 |
| "nat_of_nibble Nibble4 = 4" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
142 |
| "nat_of_nibble Nibble5 = 5" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
143 |
| "nat_of_nibble Nibble6 = 6" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
144 |
| "nat_of_nibble Nibble7 = 7" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
145 |
| "nat_of_nibble Nibble8 = 8" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
146 |
| "nat_of_nibble Nibble9 = 9" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
147 |
| "nat_of_nibble NibbleA = 10" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
148 |
| "nat_of_nibble NibbleB = 11" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
149 |
| "nat_of_nibble NibbleC = 12" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
150 |
| "nat_of_nibble NibbleD = 13" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
151 |
| "nat_of_nibble NibbleE = 14" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
152 |
| "nat_of_nibble NibbleF = 15" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
153 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
154 |
definition nibble_of_nat :: "nat \<Rightarrow> nibble" where |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
155 |
"nibble_of_nat n = Enum.enum ! (n mod 16)" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
156 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
157 |
lemma nibble_of_nat_simps [simp]: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
158 |
"nibble_of_nat 0 = Nibble0" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
159 |
"nibble_of_nat 1 = Nibble1" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
160 |
"nibble_of_nat 2 = Nibble2" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
161 |
"nibble_of_nat 3 = Nibble3" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
162 |
"nibble_of_nat 4 = Nibble4" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
163 |
"nibble_of_nat 5 = Nibble5" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
164 |
"nibble_of_nat 6 = Nibble6" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
165 |
"nibble_of_nat 7 = Nibble7" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
166 |
"nibble_of_nat 8 = Nibble8" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
167 |
"nibble_of_nat 9 = Nibble9" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
168 |
"nibble_of_nat 10 = NibbleA" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
169 |
"nibble_of_nat 11 = NibbleB" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
170 |
"nibble_of_nat 12 = NibbleC" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
171 |
"nibble_of_nat 13 = NibbleD" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
172 |
"nibble_of_nat 14 = NibbleE" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
173 |
"nibble_of_nat 15 = NibbleF" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
174 |
unfolding nibble_of_nat_def by (simp_all add: enum_nibble_def) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
175 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
176 |
lemma nibble_of_nat_of_nibble [simp]: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
177 |
"nibble_of_nat (nat_of_nibble x) = x" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
178 |
by (cases x) (simp_all add: nibble_of_nat_def enum_nibble_def) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
179 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
180 |
lemma nat_of_nibble_of_nat [simp]: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
181 |
"nat_of_nibble (nibble_of_nat n) = n mod 16" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
182 |
by (cases "nibble_of_nat n") |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
183 |
(simp_all add: nibble_of_nat_def enum_nibble_def nth_equal_first_eq nth_non_equal_first_eq, arith) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
184 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
185 |
lemma inj_nat_of_nibble: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
186 |
"inj nat_of_nibble" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
187 |
by (rule inj_on_inverseI) (rule nibble_of_nat_of_nibble) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
188 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
189 |
lemma nat_of_nibble_eq_iff: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
190 |
"nat_of_nibble x = nat_of_nibble y \<longleftrightarrow> x = y" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
191 |
by (rule inj_eq) (rule inj_nat_of_nibble) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
192 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
193 |
lemma nat_of_nibble_less_16: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
194 |
"nat_of_nibble x < 16" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
195 |
by (cases x) auto |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
196 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
197 |
lemma nibble_of_nat_mod_16: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
198 |
"nibble_of_nat (n mod 16) = nibble_of_nat n" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
199 |
by (simp add: nibble_of_nat_def) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
200 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
201 |
context |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
202 |
begin |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
203 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
204 |
local_setup \<open>Local_Theory.map_background_naming (Name_Space.add_path "char")\<close> |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
205 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
206 |
definition Char :: "nibble \<Rightarrow> nibble \<Rightarrow> char" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
207 |
where |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
208 |
"Char x y = char_of_nat (nat_of_nibble x * 16 + nat_of_nibble y)" |
| 61799 | 209 |
\<comment> "Note: canonical order of character encoding coincides with standard term ordering" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
210 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
211 |
end |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
212 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
213 |
lemma nat_of_char_Char: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
214 |
"nat_of_char (Char x y) = nat_of_nibble x * 16 + nat_of_nibble y" (is "_ = ?rhs") |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
215 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
216 |
have "?rhs < 256" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
217 |
using nat_of_nibble_less_16 [of x] nat_of_nibble_less_16 [of y] |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
218 |
by arith |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
219 |
then show ?thesis by (simp add: Char_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
220 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
221 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
222 |
lemma char_of_nat_Char_nibbles: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
223 |
"char_of_nat n = Char (nibble_of_nat (n div 16)) (nibble_of_nat n)" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
224 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
225 |
from mod_mult2_eq [of n 16 16] |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
226 |
have "n mod 256 = 16 * (n div 16 mod 16) + n mod 16" by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
227 |
then show ?thesis |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
228 |
by (simp add: Char_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
229 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
230 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
231 |
lemma char_of_nat_nibble_pair [simp]: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
232 |
"char_of_nat (nat_of_nibble a * 16 + nat_of_nibble b) = Char a b" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
233 |
by (simp add: nat_of_char_Char [symmetric]) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
234 |
|
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
235 |
free_constructors char for Char |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
236 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
237 |
fix P c |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
238 |
assume hyp: "\<And>x y. c = Char x y \<Longrightarrow> P" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
239 |
have "nat_of_char c \<le> 255" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
240 |
using nat_of_char_less_256 [of c] by arith |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
241 |
then have "nat_of_char c div 16 \<le> 255 div 16" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
242 |
by (auto intro: div_le_mono2) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
243 |
then have "nat_of_char c div 16 < 16" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
244 |
by auto |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
245 |
then have "nat_of_char c div 16 mod 16 = nat_of_char c div 16" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
246 |
by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
247 |
then have "c = Char (nibble_of_nat (nat_of_char c div 16)) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
248 |
(nibble_of_nat (nat_of_char c mod 16))" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
249 |
by (simp add: Char_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
250 |
then show P by (rule hyp) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
251 |
next |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
252 |
fix x y z w |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
253 |
have "Char x y = Char z w |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
254 |
\<longleftrightarrow> nat_of_char (Char x y) = nat_of_char (Char z w)" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
255 |
by auto |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
256 |
also have "\<dots> \<longleftrightarrow> nat_of_nibble x = nat_of_nibble z \<and> nat_of_nibble y = nat_of_nibble w" (is "?P \<longleftrightarrow> ?Q \<and> ?R") |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
257 |
proof |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
258 |
assume "?Q \<and> ?R" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
259 |
then show ?P by (simp add: nat_of_char_Char) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
260 |
next |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
261 |
assume ?P |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
262 |
then have ?Q |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
263 |
using nat_of_nibble_less_16 [of y] nat_of_nibble_less_16 [of w] |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
264 |
by (simp add: nat_of_char_Char) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
265 |
moreover with \<open>?P\<close> have ?R by (simp add: nat_of_char_Char) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
266 |
ultimately show "?Q \<and> ?R" .. |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
267 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
268 |
also have "\<dots> \<longleftrightarrow> x = z \<and> y = w" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
269 |
by (simp add: nat_of_nibble_eq_iff) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
270 |
finally show "Char x y = Char z w \<longleftrightarrow> x = z \<and> y = w" . |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
271 |
qed |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
272 |
|
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
273 |
syntax |
| 46483 | 274 |
"_Char" :: "str_position => char" ("CHR _")
|
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
275 |
|
|
42163
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
bulwahn
parents:
41750
diff
changeset
|
276 |
type_synonym string = "char list" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
277 |
|
|
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
278 |
syntax |
| 46483 | 279 |
"_String" :: "str_position => string" ("_")
|
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
280 |
|
| 48891 | 281 |
ML_file "Tools/string_syntax.ML" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
282 |
|
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
283 |
lemma UNIV_char: |
|
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60801
diff
changeset
|
284 |
"UNIV = image (case_prod Char) (UNIV \<times> UNIV)" |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
285 |
proof (rule UNIV_eq_I) |
|
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60801
diff
changeset
|
286 |
fix x show "x \<in> image (case_prod Char) (UNIV \<times> UNIV)" by (cases x) auto |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
287 |
qed |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
288 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
289 |
instantiation char :: enum |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
290 |
begin |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
291 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
292 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
293 |
"Enum.enum = [Char Nibble0 Nibble0, Char Nibble0 Nibble1, Char Nibble0 Nibble2, |
| 31484 | 294 |
Char Nibble0 Nibble3, Char Nibble0 Nibble4, Char Nibble0 Nibble5, |
295 |
Char Nibble0 Nibble6, Char Nibble0 Nibble7, Char Nibble0 Nibble8, |
|
|
55015
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
wenzelm
parents:
54594
diff
changeset
|
296 |
Char Nibble0 Nibble9, CHR ''\<newline>'', Char Nibble0 NibbleB, |
| 31484 | 297 |
Char Nibble0 NibbleC, Char Nibble0 NibbleD, Char Nibble0 NibbleE, |
298 |
Char Nibble0 NibbleF, Char Nibble1 Nibble0, Char Nibble1 Nibble1, |
|
299 |
Char Nibble1 Nibble2, Char Nibble1 Nibble3, Char Nibble1 Nibble4, |
|
300 |
Char Nibble1 Nibble5, Char Nibble1 Nibble6, Char Nibble1 Nibble7, |
|
301 |
Char Nibble1 Nibble8, Char Nibble1 Nibble9, Char Nibble1 NibbleA, |
|
302 |
Char Nibble1 NibbleB, Char Nibble1 NibbleC, Char Nibble1 NibbleD, |
|
303 |
Char Nibble1 NibbleE, Char Nibble1 NibbleF, CHR '' '', CHR ''!'', |
|
304 |
Char Nibble2 Nibble2, CHR ''#'', CHR ''$'', CHR ''%'', CHR ''&'', |
|
305 |
Char Nibble2 Nibble7, CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', CHR '','',
|
|
306 |
CHR ''-'', CHR ''.'', CHR ''/'', CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', |
|
307 |
CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9'', CHR '':'', |
|
308 |
CHR '';'', CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', CHR ''@'', CHR ''A'', |
|
309 |
CHR ''B'', CHR ''C'', CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', CHR ''H'', |
|
310 |
CHR ''I'', CHR ''J'', CHR ''K'', CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', |
|
311 |
CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', CHR ''T'', CHR ''U'', CHR ''V'', |
|
312 |
CHR ''W'', CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', Char Nibble5 NibbleC, |
|
313 |
CHR '']'', CHR ''^'', CHR ''_'', Char Nibble6 Nibble0, CHR ''a'', CHR ''b'', |
|
314 |
CHR ''c'', CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', CHR ''h'', CHR ''i'', |
|
315 |
CHR ''j'', CHR ''k'', CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', CHR ''p'', |
|
316 |
CHR ''q'', CHR ''r'', CHR ''s'', CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', |
|
317 |
CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', CHR ''|'', CHR ''}'', CHR ''~'',
|
|
318 |
Char Nibble7 NibbleF, Char Nibble8 Nibble0, Char Nibble8 Nibble1, |
|
319 |
Char Nibble8 Nibble2, Char Nibble8 Nibble3, Char Nibble8 Nibble4, |
|
320 |
Char Nibble8 Nibble5, Char Nibble8 Nibble6, Char Nibble8 Nibble7, |
|
321 |
Char Nibble8 Nibble8, Char Nibble8 Nibble9, Char Nibble8 NibbleA, |
|
322 |
Char Nibble8 NibbleB, Char Nibble8 NibbleC, Char Nibble8 NibbleD, |
|
323 |
Char Nibble8 NibbleE, Char Nibble8 NibbleF, Char Nibble9 Nibble0, |
|
324 |
Char Nibble9 Nibble1, Char Nibble9 Nibble2, Char Nibble9 Nibble3, |
|
325 |
Char Nibble9 Nibble4, Char Nibble9 Nibble5, Char Nibble9 Nibble6, |
|
326 |
Char Nibble9 Nibble7, Char Nibble9 Nibble8, Char Nibble9 Nibble9, |
|
327 |
Char Nibble9 NibbleA, Char Nibble9 NibbleB, Char Nibble9 NibbleC, |
|
328 |
Char Nibble9 NibbleD, Char Nibble9 NibbleE, Char Nibble9 NibbleF, |
|
329 |
Char NibbleA Nibble0, Char NibbleA Nibble1, Char NibbleA Nibble2, |
|
330 |
Char NibbleA Nibble3, Char NibbleA Nibble4, Char NibbleA Nibble5, |
|
331 |
Char NibbleA Nibble6, Char NibbleA Nibble7, Char NibbleA Nibble8, |
|
332 |
Char NibbleA Nibble9, Char NibbleA NibbleA, Char NibbleA NibbleB, |
|
333 |
Char NibbleA NibbleC, Char NibbleA NibbleD, Char NibbleA NibbleE, |
|
334 |
Char NibbleA NibbleF, Char NibbleB Nibble0, Char NibbleB Nibble1, |
|
335 |
Char NibbleB Nibble2, Char NibbleB Nibble3, Char NibbleB Nibble4, |
|
336 |
Char NibbleB Nibble5, Char NibbleB Nibble6, Char NibbleB Nibble7, |
|
337 |
Char NibbleB Nibble8, Char NibbleB Nibble9, Char NibbleB NibbleA, |
|
338 |
Char NibbleB NibbleB, Char NibbleB NibbleC, Char NibbleB NibbleD, |
|
339 |
Char NibbleB NibbleE, Char NibbleB NibbleF, Char NibbleC Nibble0, |
|
340 |
Char NibbleC Nibble1, Char NibbleC Nibble2, Char NibbleC Nibble3, |
|
341 |
Char NibbleC Nibble4, Char NibbleC Nibble5, Char NibbleC Nibble6, |
|
342 |
Char NibbleC Nibble7, Char NibbleC Nibble8, Char NibbleC Nibble9, |
|
343 |
Char NibbleC NibbleA, Char NibbleC NibbleB, Char NibbleC NibbleC, |
|
344 |
Char NibbleC NibbleD, Char NibbleC NibbleE, Char NibbleC NibbleF, |
|
345 |
Char NibbleD Nibble0, Char NibbleD Nibble1, Char NibbleD Nibble2, |
|
346 |
Char NibbleD Nibble3, Char NibbleD Nibble4, Char NibbleD Nibble5, |
|
347 |
Char NibbleD Nibble6, Char NibbleD Nibble7, Char NibbleD Nibble8, |
|
348 |
Char NibbleD Nibble9, Char NibbleD NibbleA, Char NibbleD NibbleB, |
|
349 |
Char NibbleD NibbleC, Char NibbleD NibbleD, Char NibbleD NibbleE, |
|
350 |
Char NibbleD NibbleF, Char NibbleE Nibble0, Char NibbleE Nibble1, |
|
351 |
Char NibbleE Nibble2, Char NibbleE Nibble3, Char NibbleE Nibble4, |
|
352 |
Char NibbleE Nibble5, Char NibbleE Nibble6, Char NibbleE Nibble7, |
|
353 |
Char NibbleE Nibble8, Char NibbleE Nibble9, Char NibbleE NibbleA, |
|
354 |
Char NibbleE NibbleB, Char NibbleE NibbleC, Char NibbleE NibbleD, |
|
355 |
Char NibbleE NibbleE, Char NibbleE NibbleF, Char NibbleF Nibble0, |
|
356 |
Char NibbleF Nibble1, Char NibbleF Nibble2, Char NibbleF Nibble3, |
|
357 |
Char NibbleF Nibble4, Char NibbleF Nibble5, Char NibbleF Nibble6, |
|
358 |
Char NibbleF Nibble7, Char NibbleF Nibble8, Char NibbleF Nibble9, |
|
359 |
Char NibbleF NibbleA, Char NibbleF NibbleB, Char NibbleF NibbleC, |
|
360 |
Char NibbleF NibbleD, Char NibbleF NibbleE, Char NibbleF NibbleF]" |
|
361 |
||
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
362 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
363 |
"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
|
364 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
365 |
definition |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
366 |
"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
|
367 |
|
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
368 |
lemma enum_char_product_enum_nibble: |
|
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60801
diff
changeset
|
369 |
"(Enum.enum :: char list) = map (case_prod Char) (List.product Enum.enum Enum.enum)" |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
370 |
by (simp add: enum_char_def enum_nibble_def) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
371 |
|
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
372 |
instance proof |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
373 |
show UNIV: "UNIV = set (Enum.enum :: char list)" |
| 57247 | 374 |
by (simp add: enum_char_product_enum_nibble UNIV_char enum_UNIV) |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
375 |
show "distinct (Enum.enum :: char list)" |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
376 |
by (auto intro: inj_onI simp add: enum_char_product_enum_nibble distinct_map distinct_product enum_distinct) |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
377 |
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
|
378 |
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
|
379 |
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
|
380 |
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
|
381 |
qed |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
382 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
383 |
end |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
384 |
|
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
385 |
lemma card_UNIV_char: |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
386 |
"card (UNIV :: char set) = 256" |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
387 |
by (simp add: card_UNIV_length_enum enum_char_def) |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49834
diff
changeset
|
388 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
389 |
lemma enum_char_unfold: |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
390 |
"Enum.enum = map char_of_nat [0..<256]" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
391 |
proof - |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
392 |
have *: "Suc (Suc 0) = 2" by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
393 |
have "map nat_of_char Enum.enum = [0..<256]" |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
394 |
by (simp add: enum_char_def nat_of_char_Char upt_conv_Cons_Cons *) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
395 |
then have "map char_of_nat (map nat_of_char Enum.enum) = |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
396 |
map char_of_nat [0..<256]" by simp |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
397 |
then show ?thesis by (simp add: comp_def) |
|
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
398 |
qed |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
399 |
|
| 60758 | 400 |
setup \<open> |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
401 |
let |
| 59631 | 402 |
val nibbles = map_range (Thm.cterm_of @{context} o HOLogic.mk_nibble) 16;
|
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51160
diff
changeset
|
403 |
val simpset = |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51160
diff
changeset
|
404 |
put_simpset HOL_ss @{context}
|
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51160
diff
changeset
|
405 |
addsimps @{thms nat_of_nibble.simps mult_0 mult_1 add_0 add_0_right arith_simps numeral_plus_one};
|
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
406 |
fun mk_code_eqn x y = |
| 60801 | 407 |
Thm.instantiate' [] [SOME x, SOME y] @{thm nat_of_char_Char}
|
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
408 |
|> simplify simpset; |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
409 |
val code_eqns = map_product mk_code_eqn nibbles nibbles; |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
410 |
in |
|
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
411 |
Global_Theory.note_thmss "" |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
412 |
[((@{binding nat_of_char_simps}, []), [(code_eqns, [])])]
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
413 |
#> snd |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
414 |
end |
| 60758 | 415 |
\<close> |
|
49972
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
haftmann
parents:
49948
diff
changeset
|
416 |
|
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
417 |
declare nat_of_char_simps [code] |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
418 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
419 |
lemma nibble_of_nat_of_char_div_16: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
420 |
"nibble_of_nat (nat_of_char c div 16) = (case c of Char x y \<Rightarrow> x)" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
421 |
by (cases c) |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
422 |
(simp add: nat_of_char_Char add.commute nat_of_nibble_less_16) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
423 |
|
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
424 |
lemma nibble_of_nat_nat_of_char: |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
425 |
"nibble_of_nat (nat_of_char c) = (case c of Char x y \<Rightarrow> y)" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
426 |
proof (cases c) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
427 |
case (Char x y) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
428 |
then have *: "nibble_of_nat ((nat_of_nibble y + nat_of_nibble x * 16) mod 16) = y" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
429 |
by (simp add: nibble_of_nat_mod_16) |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
430 |
then have "nibble_of_nat (nat_of_nibble y + nat_of_nibble x * 16) = y" |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
431 |
by (simp only: nibble_of_nat_mod_16) |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
432 |
with Char show ?thesis by (simp add: nat_of_char_Char add.commute) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
433 |
qed |
|
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
434 |
|
| 61799 | 435 |
code_datatype Char \<comment> \<open>drop case certificate for char\<close> |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
436 |
|
| 55416 | 437 |
lemma case_char_code [code]: |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
438 |
"char f c = (let n = nat_of_char c in f (nibble_of_nat (n div 16)) (nibble_of_nat n))" |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
439 |
by (cases c) |
|
55642
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents:
55428
diff
changeset
|
440 |
(simp only: Let_def nibble_of_nat_of_char_div_16 nibble_of_nat_nat_of_char char.case) |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
441 |
|
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
442 |
lemma char_of_nat_enum [code]: |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49972
diff
changeset
|
443 |
"char_of_nat n = Enum.enum ! (n mod 256)" |
|
62364
9209770bdcdf
more direct bootstrap of char type, still retaining the nibble representation for syntax
haftmann
parents:
61799
diff
changeset
|
444 |
by (simp add: enum_char_unfold) |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49834
diff
changeset
|
445 |
|
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
446 |
|
| 60758 | 447 |
subsection \<open>Strings as dedicated type\<close> |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
448 |
|
| 49834 | 449 |
typedef literal = "UNIV :: string set" |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
450 |
morphisms explode STR .. |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
451 |
|
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
452 |
setup_lifting type_definition_literal |
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
453 |
|
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
454 |
lemma STR_inject' [simp]: |
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
455 |
"STR s = STR t \<longleftrightarrow> s = t" |
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
456 |
by transfer rule |
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
457 |
|
| 57437 | 458 |
definition implode :: "string \<Rightarrow> String.literal" |
459 |
where |
|
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
460 |
[code del]: "implode = STR" |
|
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
59483
diff
changeset
|
461 |
|
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
462 |
instantiation literal :: size |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
463 |
begin |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
464 |
|
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
465 |
definition size_literal :: "literal \<Rightarrow> nat" |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
466 |
where |
| 61076 | 467 |
[code]: "size_literal (s::literal) = 0" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
468 |
|
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
469 |
instance .. |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
470 |
|
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
471 |
end |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
472 |
|
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
473 |
instantiation literal :: equal |
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
474 |
begin |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
475 |
|
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
476 |
lift_definition equal_literal :: "literal \<Rightarrow> literal \<Rightarrow> bool" is "op =" . |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
477 |
|
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
478 |
instance by intro_classes (transfer, simp) |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
479 |
|
|
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
480 |
end |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
481 |
|
|
54594
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
482 |
declare equal_literal.rep_eq[code] |
|
a2d1522cdd54
setup lifting/transfer for String.literal
Andreas Lochbihler
parents:
54317
diff
changeset
|
483 |
|
|
52365
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
484 |
lemma [code nbe]: |
|
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
485 |
fixes s :: "String.literal" |
|
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
486 |
shows "HOL.equal s s \<longleftrightarrow> True" |
|
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
487 |
by (simp add: equal) |
|
95186e6e4bf4
reflexive nbe equation for equality on String.literal
haftmann
parents:
51717
diff
changeset
|
488 |
|
|
55426
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
489 |
lifting_update literal.lifting |
|
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
490 |
lifting_forget literal.lifting |
|
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
Andreas Lochbihler
parents:
55015
diff
changeset
|
491 |
|
| 60758 | 492 |
subsection \<open>Code generator\<close> |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
493 |
|
| 48891 | 494 |
ML_file "Tools/string_code.ML" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
495 |
|
| 33237 | 496 |
code_reserved SML string |
497 |
code_reserved OCaml string |
|
| 34886 | 498 |
code_reserved Scala string |
| 33237 | 499 |
|
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
500 |
code_printing |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
501 |
type_constructor literal \<rightharpoonup> |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
502 |
(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
|
503 |
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
|
504 |
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
|
505 |
and (Scala) "String" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
506 |
|
| 60758 | 507 |
setup \<open> |
| 34886 | 508 |
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"] |
| 60758 | 509 |
\<close> |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
510 |
|
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
511 |
code_printing |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
512 |
class_instance literal :: equal \<rightharpoonup> |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
513 |
(Haskell) - |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
514 |
| constant "HOL.equal :: literal \<Rightarrow> literal \<Rightarrow> bool" \<rightharpoonup> |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52365
diff
changeset
|
515 |
(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
|
516 |
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
|
517 |
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
|
518 |
and (Scala) infixl 5 "==" |
|
31051
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
haftmann
parents:
diff
changeset
|
519 |
|
| 60758 | 520 |
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
|
521 |
|
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
522 |
definition abort :: "literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
523 |
where [simp, code del]: "abort _ f = f ()" |
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
524 |
|
|
54317
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
525 |
lemma abort_cong: "msg = msg' ==> Code.abort msg f = Code.abort msg' f" |
|
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
526 |
by simp |
|
da932f511746
add congruence rule to prevent code_simp from looping
Andreas Lochbihler
parents:
52910
diff
changeset
|
527 |
|
| 60758 | 528 |
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
|
529 |
|
| 60758 | 530 |
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
|
531 |
|
|
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
532 |
code_printing constant Code.abort \<rightharpoonup> |
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
533 |
(SML) "!(raise/ Fail/ _)" |
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
534 |
and (OCaml) "failwith" |
|
59483
ddb73392356e
explicit type annotation avoids problems with Haskell type inference
haftmann
parents:
58889
diff
changeset
|
535 |
and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)" |
|
52910
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
536 |
and (Scala) "!{/ sys.error((_));/ ((_)).apply(())/ }"
|
|
7bfe0df532a9
abort execution of generated code with explicit exception message
Andreas Lochbihler
parents:
52435
diff
changeset
|
537 |
|
|
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
35123
diff
changeset
|
538 |
hide_type (open) literal |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31176
diff
changeset
|
539 |
|
| 57437 | 540 |
hide_const (open) implode explode |
541 |
||
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
39198
diff
changeset
|
542 |
end |