New
authornipkow
Fri, 15 Apr 2005 14:14:24 +0200
changeset 15737 c7e522520910
parent 15736 1bb0399a9517
child 15738 1c1d40ff875a
New
src/HOL/Library/Char_ord.thy
src/HOL/Library/List_lexord.thy
src/HOL/Library/Product_ord.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Char_ord.thy	Fri Apr 15 14:14:24 2005 +0200
@@ -0,0 +1,99 @@
+(*  Title:      HOL/Library/Char_ord.thy
+    ID:         $Id$
+    Author:     Norbert Voelker
+*)
+
+header {* Instantiation of order classes type char *}
+
+theory Char_ord
+imports Product_ord
+begin
+
+text {* Conversions between nibbles and integers in [0..15]. *} 
+
+consts 
+  nibble_to_int:: "nibble \<Rightarrow> int"
+  int_to_nibble:: "int \<Rightarrow> nibble"
+
+primrec 
+  "nibble_to_int Nibble0 = 0"  
+  "nibble_to_int Nibble1 = 1" 
+  "nibble_to_int Nibble2 = 2" 
+  "nibble_to_int Nibble3 = 3" 
+  "nibble_to_int Nibble4 = 4" 
+  "nibble_to_int Nibble5 = 5" 
+  "nibble_to_int Nibble6 = 6" 
+  "nibble_to_int Nibble7 = 7" 
+  "nibble_to_int Nibble8 = 8" 
+  "nibble_to_int Nibble9 = 9" 
+  "nibble_to_int NibbleA = 10" 
+  "nibble_to_int NibbleB = 11" 
+  "nibble_to_int NibbleC = 12" 
+  "nibble_to_int NibbleD = 13" 
+  "nibble_to_int NibbleE = 14" 
+  "nibble_to_int NibbleF = 15"
+
+defs 
+  int_to_nibble_def:  
+    "int_to_nibble x \<equiv> (let y = x mod 16 in 
+       if y = 0 then Nibble0 else
+       if y = 1 then Nibble1 else
+       if y = 2 then Nibble2 else
+       if y = 3 then Nibble3 else
+       if y = 4 then Nibble4 else
+       if y = 5 then Nibble5 else
+       if y = 6 then Nibble6 else
+       if y = 7 then Nibble7 else
+       if y = 8 then Nibble8 else
+       if y = 9 then Nibble9 else
+       if y = 10 then NibbleA else
+       if y = 11 then NibbleB else
+       if y = 12 then NibbleC else
+       if y = 13 then NibbleD else
+       if y = 14 then NibbleE else
+       NibbleF)"
+
+lemma int_to_nibble_nibble_to_int: "int_to_nibble(nibble_to_int x) = x"
+  by (case_tac x, auto simp: int_to_nibble_def Let_def)
+
+lemma inj_nibble_to_int: "inj nibble_to_int"
+  by (rule inj_on_inverseI, rule int_to_nibble_nibble_to_int)
+
+lemmas nibble_to_int_eq = inj_nibble_to_int [THEN inj_eq]
+
+lemma nibble_to_int_ge_0: "0 \<le> nibble_to_int x"
+  by (case_tac x, auto)
+
+lemma nibble_to_int_less_16: "nibble_to_int x < 16"
+  by (case_tac x, auto)
+
+text {* Conversion between chars and int pairs. *}
+
+consts 
+  char_to_int_pair :: "char \<Rightarrow> int \<times> int"
+primrec
+  "char_to_int_pair(Char a b) = (nibble_to_int a, nibble_to_int b)" 
+
+lemma inj_char_to_int_pair: "inj char_to_int_pair"
+  by (rule inj_onI, case_tac x, case_tac y, auto simp: nibble_to_int_eq)
+
+lemmas char_to_int_pair_eq = inj_char_to_int_pair [THEN inj_eq]
+
+text {* Instantiation of order classes *} 
+
+instance char :: ord ..
+
+defs (overloaded)
+  char_le_def:  "c \<le> d \<equiv> (char_to_int_pair c \<le> char_to_int_pair d)" 
+  char_less_def: "c < d \<equiv> (char_to_int_pair c < char_to_int_pair d)" 
+
+lemmas char_ord_defs = char_less_def char_le_def
+
+instance char::order
+  apply (intro_classes, unfold char_ord_defs)
+  by (auto simp: char_to_int_pair_eq order_less_le)
+
+instance char::linorder
+  by (intro_classes, unfold char_le_def, auto)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/List_lexord.thy	Fri Apr 15 14:14:24 2005 +0200
@@ -0,0 +1,53 @@
+(*  Title:      HOL/Library/List_lexord.thy
+    ID:         $Id$
+    Author:     Norbert Voelker
+*)
+
+header {* Instantiation of order classes for lexord on lists *}
+
+theory List_lexord
+imports Main
+begin
+
+instance list :: (ord) ord ..
+defs(overloaded)
+  list_le_def:  "(xs::('a::ord) list) \<le> ys \<equiv> (xs < ys \<or> xs = ys)" 
+  list_less_def: "(xs::('a::ord) list) < ys \<equiv> (xs,ys) \<in> lexord {(u,v). u < v}"
+
+lemmas list_ord_defs = list_less_def list_le_def
+
+instance list::(order)order
+  apply (intro_classes, unfold list_ord_defs)
+  apply (rule disjI2, safe)
+  apply (blast intro: lexord_trans transI order_less_trans)
+  apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
+  apply simp
+  apply (blast intro: lexord_trans transI order_less_trans)
+  apply (rule_tac r1 = "{(a::'a,b). a < b}" in lexord_irreflexive [THEN notE])
+  apply simp
+  by assumption
+
+instance list::(linorder)linorder
+  apply (intro_classes, unfold list_le_def list_less_def, safe)
+  apply (cut_tac x="x" and y="y" and  r = "{(a,b). a < b}"  in lexord_linear)
+  by (force, simp)
+
+lemma not_less_Nil[simp]: "~(x < [])";
+  by (unfold list_less_def, simp);
+
+lemma Nil_less_Cons[simp]: "[] < a # x";
+  by (unfold list_less_def, simp);
+
+lemma Cons_less_Cons[simp]: "(a # x < b # y) = (a < b | a = b & x < y)";
+  by (unfold list_less_def, simp);
+
+lemma le_Nil[simp]: "(x <= [])   = (x = [])";
+  by (unfold list_ord_defs, case_tac x, auto);
+
+lemma Nil_le_Cons[simp]: "([] <= x)";
+  by (unfold list_ord_defs, case_tac x, auto);
+
+lemma Cons_le_Cons[simp]: "(a # x <= b # y) = (a < b | a = b & x <= y)";
+  by (unfold list_ord_defs, auto);
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Product_ord.thy	Fri Apr 15 14:14:24 2005 +0200
@@ -0,0 +1,28 @@
+(*  Title:      HOL/Library/Product_ord.thy
+    ID:         $Id$
+    Author:     Norbert Voelker
+*)
+
+header {* Instantiation of order classes for product types *}
+
+theory Product_ord
+imports Main
+begin
+
+instance "*" :: (ord,ord) ord ..
+
+defs (overloaded)
+  prod_le_def: "(x \<le> y) \<equiv> (fst x < fst y) | (fst x = fst y & snd x \<le> snd y)" 
+  prod_less_def: "(x < y) \<equiv> (fst x < fst y) | (fst x = fst y & snd x < snd y)"
+
+
+lemmas prod_ord_defs = prod_less_def prod_le_def
+
+instance "*" :: (order,order) order 
+  apply (intro_classes, unfold prod_ord_defs)
+  by (auto intro: order_less_trans)
+
+instance "*":: (linorder,linorder)linorder
+  by (intro_classes, unfold prod_le_def, auto)
+
+end
\ No newline at end of file