| 15737 |      1 | (*  Title:      HOL/Library/Char_ord.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Norbert Voelker
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
| 17200 |      6 | header {* Order on characters *}
 | 
| 15737 |      7 | 
 | 
|  |      8 | theory Char_ord
 | 
|  |      9 | imports Product_ord
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
| 17200 |     12 | text {* Conversions between nibbles and integers in [0..15]. *}
 | 
| 15737 |     13 | 
 | 
| 17200 |     14 | consts
 | 
| 15737 |     15 |   nibble_to_int:: "nibble \<Rightarrow> int"
 | 
|  |     16 |   int_to_nibble:: "int \<Rightarrow> nibble"
 | 
|  |     17 | 
 | 
| 17200 |     18 | primrec
 | 
|  |     19 |   "nibble_to_int Nibble0 = 0"
 | 
|  |     20 |   "nibble_to_int Nibble1 = 1"
 | 
|  |     21 |   "nibble_to_int Nibble2 = 2"
 | 
|  |     22 |   "nibble_to_int Nibble3 = 3"
 | 
|  |     23 |   "nibble_to_int Nibble4 = 4"
 | 
|  |     24 |   "nibble_to_int Nibble5 = 5"
 | 
|  |     25 |   "nibble_to_int Nibble6 = 6"
 | 
|  |     26 |   "nibble_to_int Nibble7 = 7"
 | 
|  |     27 |   "nibble_to_int Nibble8 = 8"
 | 
|  |     28 |   "nibble_to_int Nibble9 = 9"
 | 
|  |     29 |   "nibble_to_int NibbleA = 10"
 | 
|  |     30 |   "nibble_to_int NibbleB = 11"
 | 
|  |     31 |   "nibble_to_int NibbleC = 12"
 | 
|  |     32 |   "nibble_to_int NibbleD = 13"
 | 
|  |     33 |   "nibble_to_int NibbleE = 14"
 | 
| 15737 |     34 |   "nibble_to_int NibbleF = 15"
 | 
|  |     35 | 
 | 
| 17200 |     36 | defs
 | 
|  |     37 |   int_to_nibble_def:
 | 
|  |     38 |     "int_to_nibble x \<equiv> (let y = x mod 16 in
 | 
| 15737 |     39 |        if y = 0 then Nibble0 else
 | 
|  |     40 |        if y = 1 then Nibble1 else
 | 
|  |     41 |        if y = 2 then Nibble2 else
 | 
|  |     42 |        if y = 3 then Nibble3 else
 | 
|  |     43 |        if y = 4 then Nibble4 else
 | 
|  |     44 |        if y = 5 then Nibble5 else
 | 
|  |     45 |        if y = 6 then Nibble6 else
 | 
|  |     46 |        if y = 7 then Nibble7 else
 | 
|  |     47 |        if y = 8 then Nibble8 else
 | 
|  |     48 |        if y = 9 then Nibble9 else
 | 
|  |     49 |        if y = 10 then NibbleA else
 | 
|  |     50 |        if y = 11 then NibbleB else
 | 
|  |     51 |        if y = 12 then NibbleC else
 | 
|  |     52 |        if y = 13 then NibbleD else
 | 
|  |     53 |        if y = 14 then NibbleE else
 | 
|  |     54 |        NibbleF)"
 | 
|  |     55 | 
 | 
|  |     56 | lemma int_to_nibble_nibble_to_int: "int_to_nibble(nibble_to_int x) = x"
 | 
| 17200 |     57 |   by (cases x) (auto simp: int_to_nibble_def Let_def)
 | 
| 15737 |     58 | 
 | 
|  |     59 | lemma inj_nibble_to_int: "inj nibble_to_int"
 | 
| 17200 |     60 |   by (rule inj_on_inverseI) (rule int_to_nibble_nibble_to_int)
 | 
| 15737 |     61 | 
 | 
|  |     62 | lemmas nibble_to_int_eq = inj_nibble_to_int [THEN inj_eq]
 | 
|  |     63 | 
 | 
|  |     64 | lemma nibble_to_int_ge_0: "0 \<le> nibble_to_int x"
 | 
| 17200 |     65 |   by (cases x) auto
 | 
| 15737 |     66 | 
 | 
|  |     67 | lemma nibble_to_int_less_16: "nibble_to_int x < 16"
 | 
| 17200 |     68 |   by (cases x) auto
 | 
| 15737 |     69 | 
 | 
|  |     70 | text {* Conversion between chars and int pairs. *}
 | 
|  |     71 | 
 | 
| 17200 |     72 | consts
 | 
| 15737 |     73 |   char_to_int_pair :: "char \<Rightarrow> int \<times> int"
 | 
|  |     74 | primrec
 | 
| 17200 |     75 |   "char_to_int_pair (Char a b) = (nibble_to_int a, nibble_to_int b)"
 | 
| 15737 |     76 | 
 | 
|  |     77 | lemma inj_char_to_int_pair: "inj char_to_int_pair"
 | 
| 17200 |     78 |   apply (rule inj_onI)
 | 
|  |     79 |   apply (case_tac x, case_tac y)
 | 
|  |     80 |   apply (auto simp: nibble_to_int_eq)
 | 
|  |     81 |   done
 | 
| 15737 |     82 | 
 | 
|  |     83 | lemmas char_to_int_pair_eq = inj_char_to_int_pair [THEN inj_eq]
 | 
|  |     84 | 
 | 
| 17200 |     85 | text {* Instantiation of order classes *}
 | 
| 15737 |     86 | 
 | 
|  |     87 | instance char :: ord ..
 | 
|  |     88 | 
 | 
|  |     89 | defs (overloaded)
 | 
| 17200 |     90 |   char_le_def: "c \<le> d \<equiv> (char_to_int_pair c \<le> char_to_int_pair d)"
 | 
|  |     91 |   char_less_def: "c < d \<equiv> (char_to_int_pair c < char_to_int_pair d)"
 | 
| 15737 |     92 | 
 | 
|  |     93 | lemmas char_ord_defs = char_less_def char_le_def
 | 
|  |     94 | 
 | 
| 17200 |     95 | instance char :: order
 | 
|  |     96 |   apply intro_classes
 | 
|  |     97 |      apply (unfold char_ord_defs)
 | 
|  |     98 |      apply (auto simp: char_to_int_pair_eq order_less_le)
 | 
|  |     99 |   done
 | 
| 15737 |    100 | 
 | 
|  |    101 | instance char::linorder
 | 
| 17200 |    102 |   apply intro_classes
 | 
|  |    103 |   apply (unfold char_le_def)
 | 
|  |    104 |   apply auto
 | 
|  |    105 |   done
 | 
| 15737 |    106 | 
 | 
| 17200 |    107 | end
 |