src/HOL/Library/Char_nat.thy
author haftmann
Fri, 18 Jul 2008 18:25:53 +0200
changeset 27651 16a26996c30e
parent 27487 c8a6ce181805
child 28562 4e74209f113e
permissions -rw-r--r--
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Char_nat.thy
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     2
    ID:         $Id$
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     3
    Author:     Norbert Voelker, Florian Haftmann
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     4
*)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     5
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     6
header {* Mapping between characters and natural numbers *}
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     7
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
     8
theory Char_nat
27487
c8a6ce181805 absolute imports of HOL/*.thy theories
haftmann
parents: 27368
diff changeset
     9
imports Plain "~~/src/HOL/List"
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    10
begin
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    11
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    12
text {* Conversions between nibbles and natural numbers in [0..15]. *}
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    13
26152
cf2cccf17d6d some more primrec
haftmann
parents: 26086
diff changeset
    14
primrec
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    15
  nat_of_nibble :: "nibble \<Rightarrow> nat" where
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
    16
    "nat_of_nibble Nibble0 = 0"
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    17
  | "nat_of_nibble Nibble1 = 1"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    18
  | "nat_of_nibble Nibble2 = 2"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    19
  | "nat_of_nibble Nibble3 = 3"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    20
  | "nat_of_nibble Nibble4 = 4"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    21
  | "nat_of_nibble Nibble5 = 5"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    22
  | "nat_of_nibble Nibble6 = 6"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    23
  | "nat_of_nibble Nibble7 = 7"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    24
  | "nat_of_nibble Nibble8 = 8"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    25
  | "nat_of_nibble Nibble9 = 9"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    26
  | "nat_of_nibble NibbleA = 10"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    27
  | "nat_of_nibble NibbleB = 11"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    28
  | "nat_of_nibble NibbleC = 12"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    29
  | "nat_of_nibble NibbleD = 13"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    30
  | "nat_of_nibble NibbleE = 14"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    31
  | "nat_of_nibble NibbleF = 15"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    32
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    33
definition
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    34
  nibble_of_nat :: "nat \<Rightarrow> nibble" where
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    35
  "nibble_of_nat x = (let y = x mod 16 in
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    36
    if y = 0 then Nibble0 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    37
    if y = 1 then Nibble1 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    38
    if y = 2 then Nibble2 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    39
    if y = 3 then Nibble3 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    40
    if y = 4 then Nibble4 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    41
    if y = 5 then Nibble5 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    42
    if y = 6 then Nibble6 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    43
    if y = 7 then Nibble7 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    44
    if y = 8 then Nibble8 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    45
    if y = 9 then Nibble9 else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    46
    if y = 10 then NibbleA else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    47
    if y = 11 then NibbleB else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    48
    if y = 12 then NibbleC else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    49
    if y = 13 then NibbleD else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    50
    if y = 14 then NibbleE else
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    51
    NibbleF)"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    52
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    53
lemma nibble_of_nat_norm:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    54
  "nibble_of_nat (n mod 16) = nibble_of_nat n"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    55
  unfolding nibble_of_nat_def Let_def by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    56
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    57
lemmas [code func] = nibble_of_nat_norm [symmetric]
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    58
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    59
lemma nibble_of_nat_simps [simp]:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    60
  "nibble_of_nat  0 = Nibble0"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    61
  "nibble_of_nat  1 = Nibble1"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    62
  "nibble_of_nat  2 = Nibble2"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    63
  "nibble_of_nat  3 = Nibble3"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    64
  "nibble_of_nat  4 = Nibble4"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    65
  "nibble_of_nat  5 = Nibble5"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    66
  "nibble_of_nat  6 = Nibble6"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    67
  "nibble_of_nat  7 = Nibble7"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    68
  "nibble_of_nat  8 = Nibble8"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    69
  "nibble_of_nat  9 = Nibble9"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    70
  "nibble_of_nat 10 = NibbleA"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    71
  "nibble_of_nat 11 = NibbleB"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    72
  "nibble_of_nat 12 = NibbleC"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    73
  "nibble_of_nat 13 = NibbleD"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    74
  "nibble_of_nat 14 = NibbleE"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    75
  "nibble_of_nat 15 = NibbleF"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    76
  unfolding nibble_of_nat_def Let_def by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    77
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    78
lemmas nibble_of_nat_code [code func] = nibble_of_nat_simps
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 23394
diff changeset
    79
  [simplified nat_number Let_def not_neg_number_of_Pls neg_number_of_Bit0 neg_number_of_Bit1 if_False add_0 add_Suc]
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    80
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    81
lemma nibble_of_nat_of_nibble: "nibble_of_nat (nat_of_nibble n) = n"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    82
  by (cases n) (simp_all only: nat_of_nibble.simps nibble_of_nat_simps)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    83
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    84
lemma nat_of_nibble_of_nat: "nat_of_nibble (nibble_of_nat n) = n mod 16"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    85
proof -
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
    86
  have nibble_nat_enum:
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
    87
    "n mod 16 \<in> {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15}"
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    88
  proof -
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    89
    have set_unfold: "\<And>n. {0..Suc n} = insert (Suc n) {0..n}" by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    90
    have "(n\<Colon>nat) mod 16 \<in> {0..Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    91
      (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))))))))}" by simp
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    92
    from this [simplified set_unfold atLeastAtMost_singleton]
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    93
    show ?thesis by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    94
  qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    95
  then show ?thesis unfolding nibble_of_nat_def Let_def
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    96
  by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    97
qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    98
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
    99
lemma inj_nat_of_nibble: "inj nat_of_nibble"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   100
  by (rule inj_on_inverseI) (rule nibble_of_nat_of_nibble)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   101
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   102
lemma nat_of_nibble_eq: "nat_of_nibble n = nat_of_nibble m \<longleftrightarrow> n = m"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   103
  by (rule inj_eq) (rule inj_nat_of_nibble)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   104
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   105
lemma nat_of_nibble_less_16: "nat_of_nibble n < 16"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   106
  by (cases n) auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   107
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   108
lemma nat_of_nibble_div_16: "nat_of_nibble n div 16 = 0"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   109
  by (cases n) auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   110
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   111
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   112
text {* Conversion between chars and nats. *}
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   113
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   114
definition
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   115
  nibble_pair_of_nat :: "nat \<Rightarrow> nibble \<times> nibble" where
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   116
  "nibble_pair_of_nat n = (nibble_of_nat (n div 16), nibble_of_nat (n mod 16))"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   117
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   118
lemma nibble_of_pair [code func]:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   119
  "nibble_pair_of_nat n = (nibble_of_nat (n div 16), nibble_of_nat n)"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   120
  unfolding nibble_of_nat_norm [of n, symmetric] nibble_pair_of_nat_def ..
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   121
26152
cf2cccf17d6d some more primrec
haftmann
parents: 26086
diff changeset
   122
primrec
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   123
  nat_of_char :: "char \<Rightarrow> nat" where
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   124
  "nat_of_char (Char n m) = nat_of_nibble n * 16 + nat_of_nibble m"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   125
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   126
lemmas [simp del] = nat_of_char.simps
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   127
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   128
definition
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   129
  char_of_nat :: "nat \<Rightarrow> char" where
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   130
  char_of_nat_def: "char_of_nat n = split Char (nibble_pair_of_nat n)"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   131
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   132
lemma Char_char_of_nat:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   133
  "Char n m = char_of_nat (nat_of_nibble n * 16 + nat_of_nibble m)"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   134
  unfolding char_of_nat_def Let_def nibble_pair_of_nat_def
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   135
  by (auto simp add: div_add1_eq mod_add1_eq nat_of_nibble_div_16 nibble_of_nat_norm nibble_of_nat_of_nibble)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   136
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   137
lemma char_of_nat_of_char:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   138
  "char_of_nat (nat_of_char c) = c"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   139
  by (cases c) (simp add: nat_of_char.simps, simp add: Char_char_of_nat)
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   140
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   141
lemma nat_of_char_of_nat:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   142
  "nat_of_char (char_of_nat n) = n mod 256"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   143
proof -
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   144
  from mod_div_equality [of n, symmetric, of 16]
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   145
  have mod_mult_self3: "\<And>m k n \<Colon> nat. (k * n + m) mod n = m mod n"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   146
  proof -
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   147
    fix m k n :: nat
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   148
    show "(k * n + m) mod n = m mod n"
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   149
      by (simp only: mod_mult_self1 [symmetric, of m n k] add_commute)
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   150
  qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   151
  from mod_div_decomp [of n 256] obtain k l where n: "n = k * 256 + l"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   152
    and k: "k = n div 256" and l: "l = n mod 256" by blast
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   153
  have 16: "(0::nat) < 16" by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   154
  have 256: "(256 :: nat) = 16 * 16" by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   155
  have l_256: "l mod 256 = l" using l by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   156
  have l_div_256: "l div 16 * 16 mod 256 = l div 16 * 16"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   157
    using l by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   158
  have aux2: "(k * 256 mod 16 + l mod 16) div 16 = 0"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27487
diff changeset
   159
    unfolding 256 mult_assoc [symmetric] mod_mult_self2_is_0 by simp
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   160
  have aux3: "(k * 256 + l) div 16 = k * 16 + l div 16"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   161
    unfolding div_add1_eq [of "k * 256" l 16] aux2 256
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   162
      mult_assoc [symmetric] div_mult_self_is_m [OF 16] by simp
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   163
  have aux4: "(k * 256 + l) mod 16 = l mod 16"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   164
    unfolding 256 mult_assoc [symmetric] mod_mult_self3 ..
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   165
  show ?thesis
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   166
    by (simp add: nat_of_char.simps char_of_nat_def nibble_of_pair
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   167
      nat_of_nibble_of_nat mod_mult_distrib
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   168
      n aux3 mod_mult_self3 l_256 aux4 mod_add1_eq [of "256 * k"] l_div_256)
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   169
qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   170
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   171
lemma nibble_pair_of_nat_char:
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   172
  "nibble_pair_of_nat (nat_of_char (Char n m)) = (n, m)"
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   173
proof -
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   174
  have nat_of_nibble_256:
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   175
    "\<And>n m. (nat_of_nibble n * 16 + nat_of_nibble m) mod 256 =
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   176
      nat_of_nibble n * 16 + nat_of_nibble m"
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   177
  proof -
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   178
    fix n m
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   179
    have nat_of_nibble_less_eq_15: "\<And>n. nat_of_nibble n \<le> 15"
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   180
      using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: nat_number)
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   181
    have less_eq_240: "nat_of_nibble n * 16 \<le> 240"
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   182
      using nat_of_nibble_less_eq_15 by auto
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   183
    have "nat_of_nibble n * 16 + nat_of_nibble m \<le> 240 + 15"
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   184
      by (rule add_le_mono [of _ 240 _ 15]) (auto intro: nat_of_nibble_less_eq_15 less_eq_240)
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   185
    then have "nat_of_nibble n * 16 + nat_of_nibble m < 256" (is "?rhs < _") by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   186
    then show "?rhs mod 256 = ?rhs" by auto
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   187
  qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   188
  show ?thesis
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   189
    unfolding nibble_pair_of_nat_def Char_char_of_nat nat_of_char_of_nat nat_of_nibble_256
474ff28210c0 tuned proofs;
wenzelm
parents: 22799
diff changeset
   190
    by (simp add: add_commute nat_of_nibble_div_16 nibble_of_nat_norm nibble_of_nat_of_nibble)
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   191
qed
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   192
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   193
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   194
text {* Code generator setup *}
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   195
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   196
code_modulename SML
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   197
  Char_nat List
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   198
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   199
code_modulename OCaml
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   200
  Char_nat List
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   201
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   202
code_modulename Haskell
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   203
  Char_nat List
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   204
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents:
diff changeset
   205
end