src/HOL/Library/Char_ord.thy
author kleing
Fri, 27 May 2005 01:09:44 +0200
changeset 16095 f6af6b265d20
parent 15737 c7e522520910
child 17200 3a4d03d1a31b
permissions -rw-r--r--
put global isatest settings in one file, sourced by the other scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15737
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Library/Char_ord.thy
nipkow
parents:
diff changeset
     2
    ID:         $Id$
nipkow
parents:
diff changeset
     3
    Author:     Norbert Voelker
nipkow
parents:
diff changeset
     4
*)
nipkow
parents:
diff changeset
     5
nipkow
parents:
diff changeset
     6
header {* Instantiation of order classes type char *}
nipkow
parents:
diff changeset
     7
nipkow
parents:
diff changeset
     8
theory Char_ord
nipkow
parents:
diff changeset
     9
imports Product_ord
nipkow
parents:
diff changeset
    10
begin
nipkow
parents:
diff changeset
    11
nipkow
parents:
diff changeset
    12
text {* Conversions between nibbles and integers in [0..15]. *} 
nipkow
parents:
diff changeset
    13
nipkow
parents:
diff changeset
    14
consts 
nipkow
parents:
diff changeset
    15
  nibble_to_int:: "nibble \<Rightarrow> int"
nipkow
parents:
diff changeset
    16
  int_to_nibble:: "int \<Rightarrow> nibble"
nipkow
parents:
diff changeset
    17
nipkow
parents:
diff changeset
    18
primrec 
nipkow
parents:
diff changeset
    19
  "nibble_to_int Nibble0 = 0"  
nipkow
parents:
diff changeset
    20
  "nibble_to_int Nibble1 = 1" 
nipkow
parents:
diff changeset
    21
  "nibble_to_int Nibble2 = 2" 
nipkow
parents:
diff changeset
    22
  "nibble_to_int Nibble3 = 3" 
nipkow
parents:
diff changeset
    23
  "nibble_to_int Nibble4 = 4" 
nipkow
parents:
diff changeset
    24
  "nibble_to_int Nibble5 = 5" 
nipkow
parents:
diff changeset
    25
  "nibble_to_int Nibble6 = 6" 
nipkow
parents:
diff changeset
    26
  "nibble_to_int Nibble7 = 7" 
nipkow
parents:
diff changeset
    27
  "nibble_to_int Nibble8 = 8" 
nipkow
parents:
diff changeset
    28
  "nibble_to_int Nibble9 = 9" 
nipkow
parents:
diff changeset
    29
  "nibble_to_int NibbleA = 10" 
nipkow
parents:
diff changeset
    30
  "nibble_to_int NibbleB = 11" 
nipkow
parents:
diff changeset
    31
  "nibble_to_int NibbleC = 12" 
nipkow
parents:
diff changeset
    32
  "nibble_to_int NibbleD = 13" 
nipkow
parents:
diff changeset
    33
  "nibble_to_int NibbleE = 14" 
nipkow
parents:
diff changeset
    34
  "nibble_to_int NibbleF = 15"
nipkow
parents:
diff changeset
    35
nipkow
parents:
diff changeset
    36
defs 
nipkow
parents:
diff changeset
    37
  int_to_nibble_def:  
nipkow
parents:
diff changeset
    38
    "int_to_nibble x \<equiv> (let y = x mod 16 in 
nipkow
parents:
diff changeset
    39
       if y = 0 then Nibble0 else
nipkow
parents:
diff changeset
    40
       if y = 1 then Nibble1 else
nipkow
parents:
diff changeset
    41
       if y = 2 then Nibble2 else
nipkow
parents:
diff changeset
    42
       if y = 3 then Nibble3 else
nipkow
parents:
diff changeset
    43
       if y = 4 then Nibble4 else
nipkow
parents:
diff changeset
    44
       if y = 5 then Nibble5 else
nipkow
parents:
diff changeset
    45
       if y = 6 then Nibble6 else
nipkow
parents:
diff changeset
    46
       if y = 7 then Nibble7 else
nipkow
parents:
diff changeset
    47
       if y = 8 then Nibble8 else
nipkow
parents:
diff changeset
    48
       if y = 9 then Nibble9 else
nipkow
parents:
diff changeset
    49
       if y = 10 then NibbleA else
nipkow
parents:
diff changeset
    50
       if y = 11 then NibbleB else
nipkow
parents:
diff changeset
    51
       if y = 12 then NibbleC else
nipkow
parents:
diff changeset
    52
       if y = 13 then NibbleD else
nipkow
parents:
diff changeset
    53
       if y = 14 then NibbleE else
nipkow
parents:
diff changeset
    54
       NibbleF)"
nipkow
parents:
diff changeset
    55
nipkow
parents:
diff changeset
    56
lemma int_to_nibble_nibble_to_int: "int_to_nibble(nibble_to_int x) = x"
nipkow
parents:
diff changeset
    57
  by (case_tac x, auto simp: int_to_nibble_def Let_def)
nipkow
parents:
diff changeset
    58
nipkow
parents:
diff changeset
    59
lemma inj_nibble_to_int: "inj nibble_to_int"
nipkow
parents:
diff changeset
    60
  by (rule inj_on_inverseI, rule int_to_nibble_nibble_to_int)
nipkow
parents:
diff changeset
    61
nipkow
parents:
diff changeset
    62
lemmas nibble_to_int_eq = inj_nibble_to_int [THEN inj_eq]
nipkow
parents:
diff changeset
    63
nipkow
parents:
diff changeset
    64
lemma nibble_to_int_ge_0: "0 \<le> nibble_to_int x"
nipkow
parents:
diff changeset
    65
  by (case_tac x, auto)
nipkow
parents:
diff changeset
    66
nipkow
parents:
diff changeset
    67
lemma nibble_to_int_less_16: "nibble_to_int x < 16"
nipkow
parents:
diff changeset
    68
  by (case_tac x, auto)
nipkow
parents:
diff changeset
    69
nipkow
parents:
diff changeset
    70
text {* Conversion between chars and int pairs. *}
nipkow
parents:
diff changeset
    71
nipkow
parents:
diff changeset
    72
consts 
nipkow
parents:
diff changeset
    73
  char_to_int_pair :: "char \<Rightarrow> int \<times> int"
nipkow
parents:
diff changeset
    74
primrec
nipkow
parents:
diff changeset
    75
  "char_to_int_pair(Char a b) = (nibble_to_int a, nibble_to_int b)" 
nipkow
parents:
diff changeset
    76
nipkow
parents:
diff changeset
    77
lemma inj_char_to_int_pair: "inj char_to_int_pair"
nipkow
parents:
diff changeset
    78
  by (rule inj_onI, case_tac x, case_tac y, auto simp: nibble_to_int_eq)
nipkow
parents:
diff changeset
    79
nipkow
parents:
diff changeset
    80
lemmas char_to_int_pair_eq = inj_char_to_int_pair [THEN inj_eq]
nipkow
parents:
diff changeset
    81
nipkow
parents:
diff changeset
    82
text {* Instantiation of order classes *} 
nipkow
parents:
diff changeset
    83
nipkow
parents:
diff changeset
    84
instance char :: ord ..
nipkow
parents:
diff changeset
    85
nipkow
parents:
diff changeset
    86
defs (overloaded)
nipkow
parents:
diff changeset
    87
  char_le_def:  "c \<le> d \<equiv> (char_to_int_pair c \<le> char_to_int_pair d)" 
nipkow
parents:
diff changeset
    88
  char_less_def: "c < d \<equiv> (char_to_int_pair c < char_to_int_pair d)" 
nipkow
parents:
diff changeset
    89
nipkow
parents:
diff changeset
    90
lemmas char_ord_defs = char_less_def char_le_def
nipkow
parents:
diff changeset
    91
nipkow
parents:
diff changeset
    92
instance char::order
nipkow
parents:
diff changeset
    93
  apply (intro_classes, unfold char_ord_defs)
nipkow
parents:
diff changeset
    94
  by (auto simp: char_to_int_pair_eq order_less_le)
nipkow
parents:
diff changeset
    95
nipkow
parents:
diff changeset
    96
instance char::linorder
nipkow
parents:
diff changeset
    97
  by (intro_classes, unfold char_le_def, auto)
nipkow
parents:
diff changeset
    98
nipkow
parents:
diff changeset
    99
end