| author | wenzelm | 
| Thu, 07 Apr 2016 22:09:23 +0200 | |
| changeset 62912 | 745d31e63c21 | 
| parent 62597 | b3f2b8c906a6 | 
| child 63462 | c1fe30f2bc32 | 
| permissions | -rw-r--r-- | 
| 15737 | 1 | (* Title: HOL/Library/Char_ord.thy | 
| 22805 | 2 | Author: Norbert Voelker, Florian Haftmann | 
| 15737 | 3 | *) | 
| 4 | ||
| 60500 | 5 | section \<open>Order on characters\<close> | 
| 15737 | 6 | |
| 7 | theory Char_ord | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48926diff
changeset | 8 | imports Main | 
| 15737 | 9 | begin | 
| 10 | ||
| 25764 | 11 | instantiation char :: linorder | 
| 12 | begin | |
| 13 | ||
| 60679 | 14 | definition "c1 \<le> c2 \<longleftrightarrow> nat_of_char c1 \<le> nat_of_char c2" | 
| 15 | definition "c1 < c2 \<longleftrightarrow> nat_of_char c1 < nat_of_char c2" | |
| 25764 | 16 | |
| 60679 | 17 | instance | 
| 62597 | 18 | by standard (auto simp add: less_eq_char_def less_char_def) | 
| 22805 | 19 | |
| 25764 | 20 | end | 
| 15737 | 21 | |
| 62597 | 22 | lemma less_eq_char_simps: | 
| 23 | "(0 :: char) \<le> c" | |
| 24 | "Char k \<le> 0 \<longleftrightarrow> numeral k mod 256 = (0 :: nat)" | |
| 25 | "Char k \<le> Char l \<longleftrightarrow> numeral k mod 256 \<le> (numeral l mod 256 :: nat)" | |
| 26 | by (simp_all add: Char_def less_eq_char_def) | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48926diff
changeset | 27 | |
| 62597 | 28 | lemma less_char_simps: | 
| 29 | "\<not> c < (0 :: char)" | |
| 30 | "0 < Char k \<longleftrightarrow> (0 :: nat) < numeral k mod 256" | |
| 31 | "Char k < Char l \<longleftrightarrow> numeral k mod 256 < (numeral l mod 256 :: nat)" | |
| 32 | by (simp_all add: Char_def less_char_def) | |
| 33 | ||
| 25764 | 34 | instantiation char :: distrib_lattice | 
| 35 | begin | |
| 36 | ||
| 61076 | 37 | definition "(inf :: char \<Rightarrow> _) = min" | 
| 38 | definition "(sup :: char \<Rightarrow> _) = max" | |
| 25764 | 39 | |
| 60679 | 40 | instance | 
| 41 | by standard (auto simp add: inf_char_def sup_char_def max_min_distrib2) | |
| 22483 | 42 | |
| 25764 | 43 | end | 
| 44 | ||
| 54595 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 45 | instantiation String.literal :: linorder | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 46 | begin | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 47 | |
| 55426 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
54863diff
changeset | 48 | context includes literal.lifting begin | 
| 54595 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 49 | lift_definition less_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" is "ord.lexordp op <" . | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 50 | lift_definition less_eq_literal :: "String.literal \<Rightarrow> String.literal \<Rightarrow> bool" is "ord.lexordp_eq op <" . | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 51 | |
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 52 | instance | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 53 | proof - | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 54 | interpret linorder "ord.lexordp_eq op <" "ord.lexordp op < :: string \<Rightarrow> string \<Rightarrow> bool" | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 55 | by(rule linorder.lexordp_linorder[where less_eq="op \<le>"])(unfold_locales) | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 56 | show "PROP ?thesis" | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 57 | by(intro_classes)(transfer, simp add: less_le_not_le linear)+ | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 58 | qed | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 59 | |
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 60 | end | 
| 55426 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
54863diff
changeset | 61 | end | 
| 54595 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 62 | |
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 63 | lemma less_literal_code [code]: | 
| 57437 | 64 | "op < = (\<lambda>xs ys. ord.lexordp op < (String.explode xs) (String.explode ys))" | 
| 54595 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 65 | by(simp add: less_literal.rep_eq fun_eq_iff) | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 66 | |
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 67 | lemma less_eq_literal_code [code]: | 
| 57437 | 68 | "op \<le> = (\<lambda>xs ys. ord.lexordp_eq op < (String.explode xs) (String.explode ys))" | 
| 54595 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 69 | by(simp add: less_eq_literal.rep_eq fun_eq_iff) | 
| 
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
 Andreas Lochbihler parents: 
51160diff
changeset | 70 | |
| 55426 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
54863diff
changeset | 71 | lifting_update literal.lifting | 
| 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
54863diff
changeset | 72 | lifting_forget literal.lifting | 
| 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
54863diff
changeset | 73 | |
| 17200 | 74 | end |