75647
|
1 |
(* Title: HOL/Library/Code_Abstract_Char.thy
|
|
2 |
Author: Florian Haftmann, TU Muenchen
|
|
3 |
Author: René Thiemann, UIBK
|
|
4 |
*)
|
|
5 |
|
|
6 |
theory Code_Abstract_Char
|
|
7 |
imports
|
|
8 |
Main
|
|
9 |
"HOL-Library.Char_ord"
|
|
10 |
begin
|
|
11 |
|
|
12 |
definition Chr :: \<open>integer \<Rightarrow> char\<close>
|
|
13 |
where [simp]: \<open>Chr = char_of\<close>
|
|
14 |
|
|
15 |
lemma char_of_integer_of_char [code abstype]:
|
|
16 |
\<open>Chr (integer_of_char c) = c\<close>
|
|
17 |
by (simp add: integer_of_char_def)
|
|
18 |
|
|
19 |
lemma char_of_integer_code [code]:
|
75666
|
20 |
\<open>integer_of_char (char_of_integer k) = (if 0 \<le> k \<and> k < 256 then k else k mod 256)\<close>
|
75937
|
21 |
by (simp add: integer_of_char_def char_of_integer_def integer_eq_iff integer_less_eq_iff integer_less_iff)
|
75647
|
22 |
|
75662
|
23 |
lemma of_char_code [code]:
|
|
24 |
\<open>of_char c = of_nat (nat_of_integer (integer_of_char c))\<close>
|
|
25 |
proof -
|
|
26 |
have \<open>int_of_integer (of_char c) = of_char c\<close>
|
|
27 |
by (cases c) simp
|
|
28 |
then show ?thesis
|
|
29 |
by (simp add: integer_of_char_def nat_of_integer_def of_nat_of_char)
|
|
30 |
qed
|
75647
|
31 |
|
75662
|
32 |
definition byte :: \<open>bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> integer\<close>
|
75647
|
33 |
where [simp]: \<open>byte b0 b1 b2 b3 b4 b5 b6 b7 = horner_sum of_bool 2 [b0, b1, b2, b3, b4, b5, b6, b7]\<close>
|
|
34 |
|
|
35 |
lemma byte_code [code]:
|
|
36 |
\<open>byte b0 b1 b2 b3 b4 b5 b6 b7 = (
|
|
37 |
let
|
|
38 |
s0 = if b0 then 1 else 0;
|
|
39 |
s1 = if b1 then s0 + 2 else s0;
|
|
40 |
s2 = if b2 then s1 + 4 else s1;
|
|
41 |
s3 = if b3 then s2 + 8 else s2;
|
|
42 |
s4 = if b4 then s3 + 16 else s3;
|
|
43 |
s5 = if b5 then s4 + 32 else s4;
|
|
44 |
s6 = if b6 then s5 + 64 else s5;
|
|
45 |
s7 = if b7 then s6 + 128 else s6
|
|
46 |
in s7)\<close>
|
|
47 |
by simp
|
|
48 |
|
|
49 |
lemma Char_code [code]:
|
|
50 |
\<open>integer_of_char (Char b0 b1 b2 b3 b4 b5 b6 b7) = byte b0 b1 b2 b3 b4 b5 b6 b7\<close>
|
|
51 |
by (simp add: integer_of_char_def)
|
75666
|
52 |
|
75647
|
53 |
lemma digit_0_code [code]:
|
|
54 |
\<open>digit0 c \<longleftrightarrow> bit (integer_of_char c) 0\<close>
|
|
55 |
by (cases c) (simp add: integer_of_char_def)
|
|
56 |
|
|
57 |
lemma digit_1_code [code]:
|
|
58 |
\<open>digit1 c \<longleftrightarrow> bit (integer_of_char c) 1\<close>
|
|
59 |
by (cases c) (simp add: integer_of_char_def)
|
|
60 |
|
|
61 |
lemma digit_2_code [code]:
|
|
62 |
\<open>digit2 c \<longleftrightarrow> bit (integer_of_char c) 2\<close>
|
|
63 |
by (cases c) (simp add: integer_of_char_def)
|
|
64 |
|
|
65 |
lemma digit_3_code [code]:
|
|
66 |
\<open>digit3 c \<longleftrightarrow> bit (integer_of_char c) 3\<close>
|
|
67 |
by (cases c) (simp add: integer_of_char_def)
|
|
68 |
|
|
69 |
lemma digit_4_code [code]:
|
|
70 |
\<open>digit4 c \<longleftrightarrow> bit (integer_of_char c) 4\<close>
|
|
71 |
by (cases c) (simp add: integer_of_char_def)
|
|
72 |
|
|
73 |
lemma digit_5_code [code]:
|
|
74 |
\<open>digit5 c \<longleftrightarrow> bit (integer_of_char c) 5\<close>
|
|
75 |
by (cases c) (simp add: integer_of_char_def)
|
|
76 |
|
|
77 |
lemma digit_6_code [code]:
|
|
78 |
\<open>digit6 c \<longleftrightarrow> bit (integer_of_char c) 6\<close>
|
|
79 |
by (cases c) (simp add: integer_of_char_def)
|
|
80 |
|
|
81 |
lemma digit_7_code [code]:
|
|
82 |
\<open>digit7 c \<longleftrightarrow> bit (integer_of_char c) 7\<close>
|
|
83 |
by (cases c) (simp add: integer_of_char_def)
|
|
84 |
|
|
85 |
lemma case_char_code [code]:
|
|
86 |
\<open>case_char f c = f (digit0 c) (digit1 c) (digit2 c) (digit3 c) (digit4 c) (digit5 c) (digit6 c) (digit7 c)\<close>
|
|
87 |
by (fact char.case_eq_if)
|
|
88 |
|
|
89 |
lemma rec_char_code [code]:
|
|
90 |
\<open>rec_char f c = f (digit0 c) (digit1 c) (digit2 c) (digit3 c) (digit4 c) (digit5 c) (digit6 c) (digit7 c)\<close>
|
|
91 |
by (cases c) simp
|
|
92 |
|
|
93 |
lemma char_of_code [code]:
|
|
94 |
\<open>integer_of_char (char_of a) =
|
|
95 |
byte (bit a 0) (bit a 1) (bit a 2) (bit a 3) (bit a 4) (bit a 5) (bit a 6) (bit a 7)\<close>
|
|
96 |
by (simp add: char_of_def integer_of_char_def)
|
|
97 |
|
|
98 |
lemma ascii_of_code [code]:
|
|
99 |
\<open>integer_of_char (String.ascii_of c) = (let k = integer_of_char c in if k < 128 then k else k - 128)\<close>
|
|
100 |
proof (cases \<open>of_char c < (128 :: integer)\<close>)
|
|
101 |
case True
|
|
102 |
moreover have \<open>(of_nat 0 :: integer) \<le> of_nat (of_char c)\<close>
|
|
103 |
by simp
|
|
104 |
then have \<open>(0 :: integer) \<le> of_char c\<close>
|
|
105 |
by (simp only: of_nat_0 of_nat_of_char)
|
|
106 |
ultimately show ?thesis
|
75937
|
107 |
by (simp add: Let_def integer_of_char_def take_bit_eq_mod integer_eq_iff integer_less_eq_iff integer_less_iff)
|
75647
|
108 |
next
|
|
109 |
case False
|
|
110 |
then have \<open>(128 :: integer) \<le> of_char c\<close>
|
|
111 |
by simp
|
|
112 |
moreover have \<open>of_nat (of_char c) < (of_nat 256 :: integer)\<close>
|
|
113 |
by (simp only: of_nat_less_iff) simp
|
|
114 |
then have \<open>of_char c < (256 :: integer)\<close>
|
|
115 |
by (simp add: of_nat_of_char)
|
|
116 |
moreover define k :: integer where \<open>k = of_char c - 128\<close>
|
|
117 |
then have \<open>of_char c = k + 128\<close>
|
|
118 |
by simp
|
|
119 |
ultimately show ?thesis
|
75937
|
120 |
by (simp add: Let_def integer_of_char_def take_bit_eq_mod integer_eq_iff integer_less_eq_iff integer_less_iff)
|
75647
|
121 |
qed
|
|
122 |
|
|
123 |
lemma equal_char_code [code]:
|
|
124 |
\<open>HOL.equal c d \<longleftrightarrow> integer_of_char c = integer_of_char d\<close>
|
|
125 |
by (simp add: integer_of_char_def equal)
|
|
126 |
|
|
127 |
lemma less_eq_char_code [code]:
|
|
128 |
\<open>c \<le> d \<longleftrightarrow> integer_of_char c \<le> integer_of_char d\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
|
|
129 |
proof -
|
|
130 |
have \<open>?P \<longleftrightarrow> of_nat (of_char c) \<le> (of_nat (of_char d) :: integer)\<close>
|
|
131 |
by (simp add: less_eq_char_def)
|
|
132 |
also have \<open>\<dots> \<longleftrightarrow> ?Q\<close>
|
|
133 |
by (simp add: of_nat_of_char integer_of_char_def)
|
|
134 |
finally show ?thesis .
|
|
135 |
qed
|
|
136 |
|
|
137 |
lemma less_char_code [code]:
|
|
138 |
\<open>c < d \<longleftrightarrow> integer_of_char c < integer_of_char d\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
|
|
139 |
proof -
|
|
140 |
have \<open>?P \<longleftrightarrow> of_nat (of_char c) < (of_nat (of_char d) :: integer)\<close>
|
|
141 |
by (simp add: less_char_def)
|
|
142 |
also have \<open>\<dots> \<longleftrightarrow> ?Q\<close>
|
|
143 |
by (simp add: of_nat_of_char integer_of_char_def)
|
|
144 |
finally show ?thesis .
|
|
145 |
qed
|
|
146 |
|
|
147 |
lemma absdef_simps:
|
|
148 |
\<open>horner_sum of_bool 2 [] = (0 :: integer)\<close>
|
|
149 |
\<open>horner_sum of_bool 2 (False # bs) = (0 :: integer) \<longleftrightarrow> horner_sum of_bool 2 bs = (0 :: integer)\<close>
|
|
150 |
\<open>horner_sum of_bool 2 (True # bs) = (1 :: integer) \<longleftrightarrow> horner_sum of_bool 2 bs = (0 :: integer)\<close>
|
|
151 |
\<open>horner_sum of_bool 2 (False # bs) = (numeral (Num.Bit0 n) :: integer) \<longleftrightarrow> horner_sum of_bool 2 bs = (numeral n :: integer)\<close>
|
|
152 |
\<open>horner_sum of_bool 2 (True # bs) = (numeral (Num.Bit1 n) :: integer) \<longleftrightarrow> horner_sum of_bool 2 bs = (numeral n :: integer)\<close>
|
|
153 |
by auto (auto simp only: numeral_Bit0 [of n] numeral_Bit1 [of n] mult_2 [symmetric] add.commute [of _ 1] add.left_cancel mult_cancel_left)
|
|
154 |
|
|
155 |
local_setup \<open>
|
|
156 |
let
|
|
157 |
val simps = @{thms absdef_simps integer_of_char_def of_char_Char numeral_One}
|
|
158 |
fun prove_eqn lthy n lhs def_eqn =
|
|
159 |
let
|
|
160 |
val eqn = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
|
|
161 |
(\<^term>\<open>integer_of_char\<close> $ lhs, HOLogic.mk_number \<^typ>\<open>integer\<close> n)
|
|
162 |
in
|
|
163 |
Goal.prove_future lthy [] [] eqn (fn {context = ctxt, ...} =>
|
|
164 |
unfold_tac ctxt (def_eqn :: simps))
|
|
165 |
end
|
|
166 |
fun define n =
|
|
167 |
let
|
|
168 |
val s = "Char_" ^ String_Syntax.hex n;
|
|
169 |
val b = Binding.name s;
|
|
170 |
val b_def = Thm.def_binding b;
|
|
171 |
val b_code = Binding.name (s ^ "_code");
|
|
172 |
in
|
|
173 |
Local_Theory.define ((b, Mixfix.NoSyn),
|
|
174 |
((Binding.empty, []), HOLogic.mk_char n))
|
|
175 |
#-> (fn (lhs, (_, raw_def_eqn)) =>
|
|
176 |
Local_Theory.note ((b_def, @{attributes [code_abbrev]}), [HOLogic.mk_obj_eq raw_def_eqn])
|
|
177 |
#-> (fn (_, [def_eqn]) => `(fn lthy => prove_eqn lthy n lhs def_eqn))
|
|
178 |
#-> (fn raw_code_eqn => Local_Theory.note ((b_code, []), [raw_code_eqn]))
|
|
179 |
#-> (fn (_, [code_eqn]) => Code.declare_abstract_eqn code_eqn))
|
|
180 |
end
|
|
181 |
in
|
|
182 |
fold define (0 upto 255)
|
|
183 |
end
|
|
184 |
\<close>
|
|
185 |
|
|
186 |
code_identifier
|
|
187 |
code_module Code_Abstract_Char \<rightharpoonup>
|
|
188 |
(SML) Str and (OCaml) Str and (Haskell) Str and (Scala) Str
|
|
189 |
|
|
190 |
end
|