haftmann@22799: (* Title: HOL/Library/Char_nat.thy haftmann@22799: ID: $Id$ haftmann@22799: Author: Norbert Voelker, Florian Haftmann haftmann@22799: *) haftmann@22799: haftmann@22799: header {* Mapping between characters and natural numbers *} haftmann@22799: haftmann@22799: theory Char_nat haftmann@22799: imports List haftmann@22799: begin haftmann@22799: haftmann@22799: text {* Conversions between nibbles and natural numbers in [0..15]. *} haftmann@22799: haftmann@22799: fun haftmann@22799: nat_of_nibble :: "nibble \ nat" where haftmann@22799: "nat_of_nibble Nibble0 = 0" haftmann@22799: | "nat_of_nibble Nibble1 = 1" haftmann@22799: | "nat_of_nibble Nibble2 = 2" haftmann@22799: | "nat_of_nibble Nibble3 = 3" haftmann@22799: | "nat_of_nibble Nibble4 = 4" haftmann@22799: | "nat_of_nibble Nibble5 = 5" haftmann@22799: | "nat_of_nibble Nibble6 = 6" haftmann@22799: | "nat_of_nibble Nibble7 = 7" haftmann@22799: | "nat_of_nibble Nibble8 = 8" haftmann@22799: | "nat_of_nibble Nibble9 = 9" haftmann@22799: | "nat_of_nibble NibbleA = 10" haftmann@22799: | "nat_of_nibble NibbleB = 11" haftmann@22799: | "nat_of_nibble NibbleC = 12" haftmann@22799: | "nat_of_nibble NibbleD = 13" haftmann@22799: | "nat_of_nibble NibbleE = 14" haftmann@22799: | "nat_of_nibble NibbleF = 15" haftmann@22799: haftmann@22799: definition haftmann@22799: nibble_of_nat :: "nat \ nibble" where haftmann@22799: "nibble_of_nat x = (let y = x mod 16 in haftmann@22799: if y = 0 then Nibble0 else haftmann@22799: if y = 1 then Nibble1 else haftmann@22799: if y = 2 then Nibble2 else haftmann@22799: if y = 3 then Nibble3 else haftmann@22799: if y = 4 then Nibble4 else haftmann@22799: if y = 5 then Nibble5 else haftmann@22799: if y = 6 then Nibble6 else haftmann@22799: if y = 7 then Nibble7 else haftmann@22799: if y = 8 then Nibble8 else haftmann@22799: if y = 9 then Nibble9 else haftmann@22799: if y = 10 then NibbleA else haftmann@22799: if y = 11 then NibbleB else haftmann@22799: if y = 12 then NibbleC else haftmann@22799: if y = 13 then NibbleD else haftmann@22799: if y = 14 then NibbleE else haftmann@22799: NibbleF)" haftmann@22799: haftmann@22799: lemma nibble_of_nat_norm: haftmann@22799: "nibble_of_nat (n mod 16) = nibble_of_nat n" haftmann@22799: unfolding nibble_of_nat_def Let_def by auto haftmann@22799: haftmann@22799: lemmas [code func] = nibble_of_nat_norm [symmetric] haftmann@22799: haftmann@22799: lemma nibble_of_nat_simps [simp]: haftmann@22799: "nibble_of_nat 0 = Nibble0" haftmann@22799: "nibble_of_nat 1 = Nibble1" haftmann@22799: "nibble_of_nat 2 = Nibble2" haftmann@22799: "nibble_of_nat 3 = Nibble3" haftmann@22799: "nibble_of_nat 4 = Nibble4" haftmann@22799: "nibble_of_nat 5 = Nibble5" haftmann@22799: "nibble_of_nat 6 = Nibble6" haftmann@22799: "nibble_of_nat 7 = Nibble7" haftmann@22799: "nibble_of_nat 8 = Nibble8" haftmann@22799: "nibble_of_nat 9 = Nibble9" haftmann@22799: "nibble_of_nat 10 = NibbleA" haftmann@22799: "nibble_of_nat 11 = NibbleB" haftmann@22799: "nibble_of_nat 12 = NibbleC" haftmann@22799: "nibble_of_nat 13 = NibbleD" haftmann@22799: "nibble_of_nat 14 = NibbleE" haftmann@22799: "nibble_of_nat 15 = NibbleF" haftmann@22799: unfolding nibble_of_nat_def Let_def by auto haftmann@22799: haftmann@22799: lemmas nibble_of_nat_code [code func] = nibble_of_nat_simps haftmann@22799: [simplified nat_number Let_def not_neg_number_of_Pls neg_number_of_BIT if_False add_0 add_Suc] haftmann@22799: haftmann@22799: lemma nibble_of_nat_of_nibble: "nibble_of_nat (nat_of_nibble n) = n" haftmann@22799: by (cases n) (simp_all only: nat_of_nibble.simps nibble_of_nat_simps) haftmann@22799: haftmann@22799: lemma nat_of_nibble_of_nat: "nat_of_nibble (nibble_of_nat n) = n mod 16" haftmann@22799: proof - haftmann@22799: have nibble_nat_enum: "n mod 16 \ {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15}" haftmann@22799: proof - haftmann@22799: have set_unfold: "\n. {0..Suc n} = insert (Suc n) {0..n}" by auto haftmann@22799: have "(n\nat) mod 16 \ {0..Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc haftmann@22799: (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))))))))}" by simp haftmann@22799: from this [simplified set_unfold atLeastAtMost_singleton] haftmann@22799: show ?thesis by auto haftmann@22799: qed haftmann@22799: then show ?thesis unfolding nibble_of_nat_def Let_def haftmann@22799: by auto haftmann@22799: qed haftmann@22799: haftmann@22799: lemma inj_nat_of_nibble: "inj nat_of_nibble" haftmann@22799: by (rule inj_on_inverseI) (rule nibble_of_nat_of_nibble) haftmann@22799: haftmann@22799: lemma nat_of_nibble_eq: "nat_of_nibble n = nat_of_nibble m \ n = m" haftmann@22799: by (rule inj_eq) (rule inj_nat_of_nibble) haftmann@22799: haftmann@22799: lemma nat_of_nibble_less_16: "nat_of_nibble n < 16" haftmann@22799: by (cases n) auto haftmann@22799: haftmann@22799: lemma nat_of_nibble_div_16: "nat_of_nibble n div 16 = 0" haftmann@22799: by (cases n) auto haftmann@22799: haftmann@22799: haftmann@22799: text {* Conversion between chars and nats. *} haftmann@22799: haftmann@22799: definition haftmann@22799: nibble_pair_of_nat :: "nat \ nibble \ nibble" haftmann@22799: where haftmann@22799: "nibble_pair_of_nat n = (nibble_of_nat (n div 16), nibble_of_nat (n mod 16))" haftmann@22799: haftmann@22799: lemma nibble_of_pair [code func]: haftmann@22799: "nibble_pair_of_nat n = (nibble_of_nat (n div 16), nibble_of_nat n)" haftmann@22799: unfolding nibble_of_nat_norm [of n, symmetric] nibble_pair_of_nat_def .. haftmann@22799: haftmann@22799: fun haftmann@22799: nat_of_char :: "char \ nat" where haftmann@22799: "nat_of_char (Char n m) = nat_of_nibble n * 16 + nat_of_nibble m" haftmann@22799: haftmann@22799: lemmas [simp del] = nat_of_char.simps haftmann@22799: haftmann@22799: definition haftmann@22799: char_of_nat :: "nat \ char" where haftmann@22799: char_of_nat_def: "char_of_nat n = split Char (nibble_pair_of_nat n)" haftmann@22799: haftmann@22799: lemma Char_char_of_nat: haftmann@22799: "Char n m = char_of_nat (nat_of_nibble n * 16 + nat_of_nibble m)" haftmann@22799: unfolding char_of_nat_def Let_def nibble_pair_of_nat_def haftmann@22799: by (auto simp add: div_add1_eq mod_add1_eq nat_of_nibble_div_16 nibble_of_nat_norm nibble_of_nat_of_nibble) haftmann@22799: haftmann@22799: lemma char_of_nat_of_char: haftmann@22799: "char_of_nat (nat_of_char c) = c" haftmann@22799: by (cases c) (simp add: nat_of_char.simps, simp add: Char_char_of_nat) haftmann@22799: haftmann@22799: lemma nat_of_char_of_nat: haftmann@22799: "nat_of_char (char_of_nat n) = n mod 256" haftmann@22799: proof - haftmann@22799: from mod_div_equality [of n, symmetric, of 16] haftmann@22799: have mod_mult_self3: "\m k n \ nat. (k * n + m) mod n = m mod n" haftmann@22799: proof - haftmann@22799: fix m k n :: nat haftmann@22799: show "(k * n + m) mod n = m mod n" haftmann@22799: by (simp only: mod_mult_self1 [symmetric, of m n k] add_commute) haftmann@22799: qed haftmann@22799: from mod_div_decomp [of n 256] obtain k l where n: "n = k * 256 + l" haftmann@22799: and k: "k = n div 256" and l: "l = n mod 256" by blast haftmann@22799: have 16: "(0::nat) < 16" by auto haftmann@22799: have 256: "(256 :: nat) = 16 * 16" by auto haftmann@22799: have l_256: "l mod 256 = l" using l by auto haftmann@22799: have l_div_256: "l div 16 * 16 mod 256 = l div 16 * 16" haftmann@22799: using l by auto haftmann@22799: have aux2: "(k * 256 mod 16 + l mod 16) div 16 = 0" haftmann@22799: unfolding 256 mult_assoc [symmetric] mod_mult_self_is_0 by simp haftmann@22799: have aux3: "(k * 256 + l) div 16 = k * 16 + l div 16" haftmann@22799: unfolding div_add1_eq [of "k * 256" l 16] aux2 256 haftmann@22799: mult_assoc [symmetric] div_mult_self_is_m [OF 16] by simp haftmann@22799: have aux4: "(k * 256 + l) mod 16 = l mod 16" haftmann@22799: unfolding 256 mult_assoc [symmetric] mod_mult_self3 .. haftmann@22799: show ?thesis haftmann@22799: by (simp add: nat_of_char.simps char_of_nat_def nibble_of_pair nat_of_nibble_of_nat mod_mult_distrib haftmann@22799: n aux3 mod_mult_self3 l_256 aux4 mod_add1_eq [of "256 * k"] l_div_256) haftmann@22799: qed haftmann@22799: haftmann@22799: lemma nibble_pair_of_nat_char: haftmann@22799: "nibble_pair_of_nat (nat_of_char (Char n m)) = (n, m)" haftmann@22799: proof - haftmann@22799: have nat_of_nibble_256: haftmann@22799: "\n m. (nat_of_nibble n * 16 + nat_of_nibble m) mod 256 = nat_of_nibble n * 16 + nat_of_nibble m" haftmann@22799: proof - haftmann@22799: fix n m haftmann@22799: have nat_of_nibble_less_eq_15: "\n. nat_of_nibble n \ 15" haftmann@22799: using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: nat_number) haftmann@22799: have less_eq_240: "nat_of_nibble n * 16 \ 240" using nat_of_nibble_less_eq_15 by auto haftmann@22799: have "nat_of_nibble n * 16 + nat_of_nibble m \ 240 + 15" haftmann@22799: by (rule add_le_mono [of _ 240 _ 15]) (auto intro: nat_of_nibble_less_eq_15 less_eq_240) haftmann@22799: then have "nat_of_nibble n * 16 + nat_of_nibble m < 256" (is "?rhs < _") by auto haftmann@22799: then show "?rhs mod 256 = ?rhs" by auto haftmann@22799: qed haftmann@22799: show ?thesis haftmann@22799: unfolding nibble_pair_of_nat_def Char_char_of_nat nat_of_char_of_nat nat_of_nibble_256 haftmann@22799: by (simp add: add_commute nat_of_nibble_div_16 nibble_of_nat_norm nibble_of_nat_of_nibble) haftmann@22799: qed haftmann@22799: haftmann@22799: haftmann@22799: text {* Code generator setup *} haftmann@22799: haftmann@22799: code_modulename SML haftmann@22799: Char_nat List haftmann@22799: haftmann@22799: code_modulename OCaml haftmann@22799: Char_nat List haftmann@22799: haftmann@22799: code_modulename Haskell haftmann@22799: Char_nat List haftmann@22799: haftmann@22799: end