src/HOL/Library/Char_ord.thy
changeset 22805 1166a966e7b4
parent 22483 86064f2f2188
child 22845 5f9138bcb3d7
equal deleted inserted replaced
22804:d3c23b90c6c6 22805:1166a966e7b4
     1 (*  Title:      HOL/Library/Char_ord.thy
     1 (*  Title:      HOL/Library/Char_ord.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Norbert Voelker
     3     Author:     Norbert Voelker, Florian Haftmann
     4 *)
     4 *)
     5 
     5 
     6 header {* Order on characters *}
     6 header {* Order on characters *}
     7 
     7 
     8 theory Char_ord
     8 theory Char_ord
     9 imports Product_ord
     9 imports Product_ord Char_nat
    10 begin
    10 begin
    11 
    11 
    12 text {* Conversions between nibbles and integers in [0..15]. *}
    12 instance nibble :: linorder
       
    13   nibble_less_eq_def: "n \<le> m \<equiv> nat_of_nibble n \<le> nat_of_nibble m"
       
    14   nibble_less_def: "n < m \<equiv> nat_of_nibble n < nat_of_nibble m"
       
    15 proof
       
    16   fix n :: nibble show "n \<le> n" unfolding nibble_less_eq_def nibble_less_def by auto
       
    17 next
       
    18   fix n m q :: nibble
       
    19   assume "n \<le> m"
       
    20   and "m \<le> q"
       
    21   then show "n \<le> q" unfolding nibble_less_eq_def nibble_less_def by auto
       
    22 next
       
    23   fix n m :: nibble
       
    24   assume "n \<le> m"
       
    25   and "m \<le> n"
       
    26   then show "n = m" unfolding nibble_less_eq_def nibble_less_def by (auto simp add: nat_of_nibble_eq)
       
    27 next
       
    28   fix n m :: nibble
       
    29   show "n < m \<longleftrightarrow> n \<le> m \<and> n \<noteq> m"
       
    30   unfolding nibble_less_eq_def nibble_less_def less_le by (auto simp add: nat_of_nibble_eq)
       
    31 next
       
    32   fix n m :: nibble
       
    33   show "n \<le> m \<or> m \<le> n"
       
    34   unfolding nibble_less_eq_def by auto
       
    35 qed
    13 
    36 
    14 fun
    37 instance nibble :: distrib_lattice
    15   nibble_to_int:: "nibble \<Rightarrow> int" where
    38   "inf \<equiv> min"
    16   "nibble_to_int Nibble0 = 0"
    39   "sup \<equiv> max"
    17   | "nibble_to_int Nibble1 = 1"
    40   by default
    18   | "nibble_to_int Nibble2 = 2"
    41     (auto simp add: inf_nibble_def sup_nibble_def min_max.sup_inf_distrib1)
    19   | "nibble_to_int Nibble3 = 3"
       
    20   | "nibble_to_int Nibble4 = 4"
       
    21   | "nibble_to_int Nibble5 = 5"
       
    22   | "nibble_to_int Nibble6 = 6"
       
    23   | "nibble_to_int Nibble7 = 7"
       
    24   | "nibble_to_int Nibble8 = 8"
       
    25   | "nibble_to_int Nibble9 = 9"
       
    26   | "nibble_to_int NibbleA = 10"
       
    27   | "nibble_to_int NibbleB = 11"
       
    28   | "nibble_to_int NibbleC = 12"
       
    29   | "nibble_to_int NibbleD = 13"
       
    30   | "nibble_to_int NibbleE = 14"
       
    31   | "nibble_to_int NibbleF = 15"
       
    32 
       
    33 definition
       
    34   int_to_nibble :: "int \<Rightarrow> nibble" where
       
    35   "int_to_nibble x = (let y = x mod 16 in
       
    36     if y = 0 then Nibble0 else
       
    37     if y = 1 then Nibble1 else
       
    38     if y = 2 then Nibble2 else
       
    39     if y = 3 then Nibble3 else
       
    40     if y = 4 then Nibble4 else
       
    41     if y = 5 then Nibble5 else
       
    42     if y = 6 then Nibble6 else
       
    43     if y = 7 then Nibble7 else
       
    44     if y = 8 then Nibble8 else
       
    45     if y = 9 then Nibble9 else
       
    46     if y = 10 then NibbleA else
       
    47     if y = 11 then NibbleB else
       
    48     if y = 12 then NibbleC else
       
    49     if y = 13 then NibbleD else
       
    50     if y = 14 then NibbleE else
       
    51     NibbleF)"
       
    52 
       
    53 lemma int_to_nibble_nibble_to_int: "int_to_nibble (nibble_to_int x) = x"
       
    54   by (cases x) (auto simp: int_to_nibble_def Let_def)
       
    55 
       
    56 lemma inj_nibble_to_int: "inj nibble_to_int"
       
    57   by (rule inj_on_inverseI) (rule int_to_nibble_nibble_to_int)
       
    58 
       
    59 lemmas nibble_to_int_eq = inj_nibble_to_int [THEN inj_eq]
       
    60 
       
    61 lemma nibble_to_int_ge_0: "0 \<le> nibble_to_int x"
       
    62   by (cases x) auto
       
    63 
       
    64 lemma nibble_to_int_less_16: "nibble_to_int x < 16"
       
    65   by (cases x) auto
       
    66 
       
    67 text {* Conversion between chars and int pairs. *}
       
    68 
       
    69 fun
       
    70   char_to_int_pair :: "char \<Rightarrow> int \<times> int" where
       
    71   "char_to_int_pair (Char a b) = (nibble_to_int a, nibble_to_int b)"
       
    72 
       
    73 lemma inj_char_to_int_pair: "inj char_to_int_pair"
       
    74   apply (rule inj_onI)
       
    75   apply (case_tac x, case_tac y)
       
    76   apply (auto simp: nibble_to_int_eq)
       
    77   done
       
    78 
       
    79 lemmas char_to_int_pair_eq = inj_char_to_int_pair [THEN inj_eq]
       
    80 
       
    81 
       
    82 text {* Instantiation of order classes *}
       
    83 
       
    84 instance char :: ord
       
    85   char_le_def: "c \<le> d \<equiv> (char_to_int_pair c \<le> char_to_int_pair d)"
       
    86   char_less_def: "c < d \<equiv> (char_to_int_pair c < char_to_int_pair d)"  ..
       
    87 
       
    88 lemmas char_ord_defs = char_less_def char_le_def
       
    89 
       
    90 instance char :: order
       
    91   by default (auto simp: char_ord_defs char_to_int_pair_eq order_less_le)
       
    92 
    42 
    93 instance char :: linorder
    43 instance char :: linorder
    94   by default (auto simp: char_le_def)
    44   char_less_eq_def: "c1 \<le> c2 \<equiv> case c1 of Char n1 m1 \<Rightarrow> case c2 of Char n2 m2 \<Rightarrow>
       
    45     n1 < n2 \<or> n1 = n2 \<and> m1 \<le> m2"
       
    46   char_less_def:    "c1 < c2 \<equiv> case c1 of Char n1 m1 \<Rightarrow> case c2 of Char n2 m2 \<Rightarrow>
       
    47     n1 < n2 \<or> n1 = n2 \<and> m1 < m2"
       
    48   by default (auto simp: char_less_eq_def char_less_def split: char.splits)
       
    49 
       
    50 lemmas [code nofunc] = char_less_eq_def char_less_def
    95 
    51 
    96 instance char :: distrib_lattice
    52 instance char :: distrib_lattice
    97   "inf \<equiv> min"
    53   "inf \<equiv> min"
    98   "sup \<equiv> max"
    54   "sup \<equiv> max"
    99   by intro_classes
    55   by default
   100     (auto simp add: inf_char_def sup_char_def min_max.sup_inf_distrib1)
    56     (auto simp add: inf_char_def sup_char_def min_max.sup_inf_distrib1)
   101 
    57 
   102 
    58 lemma [simp, code func]:
   103 text {* code generator setup *}
    59   shows char_less_eq_simp: "Char n1 m1 \<le> Char n2 m2 \<longleftrightarrow> n1 < n2 \<or> n1 = n2 \<and> m1 \<le> m2"
   104 
    60   and char_less_simp:      "Char n1 m1 < Char n2 m2 \<longleftrightarrow> n1 < n2 \<or> n1 = n2 \<and> m1 < m2"
   105 code_const char_to_int_pair
    61   unfolding char_less_eq_def char_less_def by simp_all
   106   (SML "raise/ Fail/ \"char'_to'_int'_pair\"")
       
   107   (OCaml "failwith \"char'_to'_int'_pair\"")
       
   108   (Haskell "error/ \"char'_to'_int'_pair\"")
       
   109 
    62 
   110 end
    63 end