| author | wenzelm | 
| Thu, 15 Oct 2015 22:25:57 +0200 | |
| changeset 61456 | b521b8b400f7 | 
| parent 61348 | d7215449be83 | 
| child 61799 | 4cf66f21b764 | 
| permissions | -rw-r--r-- | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 1 | (* Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 2 | |
| 60758 | 3 | section \<open>Character and string types\<close> | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 4 | |
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 5 | theory String | 
| 55969 | 6 | imports Enum | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 7 | begin | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 8 | |
| 60758 | 9 | subsection \<open>Characters and strings\<close> | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 10 | |
| 61348 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 blanchet parents: 
61076diff
changeset | 11 | datatype (plugins del: transfer) nibble = | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 12 | Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7 | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 13 | | Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 14 | |
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 15 | lemma UNIV_nibble: | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 16 |   "UNIV = {Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7,
 | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 17 | Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF}" (is "_ = ?A") | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 18 | proof (rule UNIV_eq_I) | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 19 | fix x show "x \<in> ?A" by (cases x) simp_all | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 20 | qed | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 21 | |
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 22 | lemma size_nibble [code, simp]: | 
| 56846 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 23 | "size_nibble (x::nibble) = 0" | 
| 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 24 | "size (x::nibble) = 0" | 
| 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 25 | by (cases x, simp_all)+ | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 26 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 27 | instantiation nibble :: enum | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 28 | begin | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 29 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 30 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 31 | "Enum.enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 32 | Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 33 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 34 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 35 | "Enum.enum_all P \<longleftrightarrow> P Nibble0 \<and> P Nibble1 \<and> P Nibble2 \<and> P Nibble3 \<and> P Nibble4 \<and> P Nibble5 \<and> P Nibble6 \<and> P Nibble7 | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 36 | \<and> P Nibble8 \<and> P Nibble9 \<and> P NibbleA \<and> P NibbleB \<and> P NibbleC \<and> P NibbleD \<and> P NibbleE \<and> P NibbleF" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 37 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 38 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 39 | "Enum.enum_ex P \<longleftrightarrow> P Nibble0 \<or> P Nibble1 \<or> P Nibble2 \<or> P Nibble3 \<or> P Nibble4 \<or> P Nibble5 \<or> P Nibble6 \<or> P Nibble7 | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 40 | \<or> P Nibble8 \<or> P Nibble9 \<or> P NibbleA \<or> P NibbleB \<or> P NibbleC \<or> P NibbleD \<or> P NibbleE \<or> P NibbleF" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 41 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 42 | instance proof | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 43 | qed (simp_all only: enum_nibble_def enum_all_nibble_def enum_ex_nibble_def UNIV_nibble, simp_all) | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 44 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 45 | end | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 46 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 47 | lemma card_UNIV_nibble: | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 48 | "card (UNIV :: nibble set) = 16" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 49 | by (simp add: card_UNIV_length_enum enum_nibble_def) | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 50 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 51 | primrec nat_of_nibble :: "nibble \<Rightarrow> nat" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 52 | where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 53 | "nat_of_nibble Nibble0 = 0" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 54 | | "nat_of_nibble Nibble1 = 1" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 55 | | "nat_of_nibble Nibble2 = 2" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 56 | | "nat_of_nibble Nibble3 = 3" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 57 | | "nat_of_nibble Nibble4 = 4" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 58 | | "nat_of_nibble Nibble5 = 5" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 59 | | "nat_of_nibble Nibble6 = 6" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 60 | | "nat_of_nibble Nibble7 = 7" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 61 | | "nat_of_nibble Nibble8 = 8" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 62 | | "nat_of_nibble Nibble9 = 9" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 63 | | "nat_of_nibble NibbleA = 10" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 64 | | "nat_of_nibble NibbleB = 11" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 65 | | "nat_of_nibble NibbleC = 12" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 66 | | "nat_of_nibble NibbleD = 13" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 67 | | "nat_of_nibble NibbleE = 14" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 68 | | "nat_of_nibble NibbleF = 15" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 69 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 70 | definition nibble_of_nat :: "nat \<Rightarrow> nibble" where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 71 | "nibble_of_nat n = Enum.enum ! (n mod 16)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 72 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 73 | lemma nibble_of_nat_simps [simp]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 74 | "nibble_of_nat 0 = Nibble0" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 75 | "nibble_of_nat 1 = Nibble1" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 76 | "nibble_of_nat 2 = Nibble2" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 77 | "nibble_of_nat 3 = Nibble3" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 78 | "nibble_of_nat 4 = Nibble4" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 79 | "nibble_of_nat 5 = Nibble5" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 80 | "nibble_of_nat 6 = Nibble6" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 81 | "nibble_of_nat 7 = Nibble7" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 82 | "nibble_of_nat 8 = Nibble8" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 83 | "nibble_of_nat 9 = Nibble9" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 84 | "nibble_of_nat 10 = NibbleA" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 85 | "nibble_of_nat 11 = NibbleB" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 86 | "nibble_of_nat 12 = NibbleC" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 87 | "nibble_of_nat 13 = NibbleD" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 88 | "nibble_of_nat 14 = NibbleE" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 89 | "nibble_of_nat 15 = NibbleF" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 90 | unfolding nibble_of_nat_def by (simp_all add: enum_nibble_def) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 91 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 92 | lemma nibble_of_nat_of_nibble [simp]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 93 | "nibble_of_nat (nat_of_nibble x) = x" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 94 | by (cases x) (simp_all add: nibble_of_nat_def enum_nibble_def) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 95 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 96 | lemma nat_of_nibble_of_nat [simp]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 97 | "nat_of_nibble (nibble_of_nat n) = n mod 16" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 98 | by (cases "nibble_of_nat n") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 99 | (simp_all add: nibble_of_nat_def enum_nibble_def nth_equal_first_eq nth_non_equal_first_eq, arith) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 100 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 101 | lemma inj_nat_of_nibble: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 102 | "inj nat_of_nibble" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 103 | by (rule inj_on_inverseI) (rule nibble_of_nat_of_nibble) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 104 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 105 | lemma nat_of_nibble_eq_iff: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 106 | "nat_of_nibble x = nat_of_nibble y \<longleftrightarrow> x = y" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 107 | by (rule inj_eq) (rule inj_nat_of_nibble) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 108 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 109 | lemma nat_of_nibble_less_16: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 110 | "nat_of_nibble x < 16" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 111 | by (cases x) auto | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 112 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 113 | lemma nibble_of_nat_mod_16: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 114 | "nibble_of_nat (n mod 16) = nibble_of_nat n" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 115 | by (simp add: nibble_of_nat_def) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 116 | |
| 58310 | 117 | datatype char = Char nibble nibble | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 118 | -- "Note: canonical order of character encoding coincides with standard term ordering" | 
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 119 | |
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 120 | syntax | 
| 46483 | 121 |   "_Char" :: "str_position => char"    ("CHR _")
 | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 122 | |
| 42163 
392fd6c4669c
renewing specifications in HOL: replacing types by type_synonym
 bulwahn parents: 
41750diff
changeset | 123 | type_synonym string = "char list" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 124 | |
| 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 125 | syntax | 
| 46483 | 126 |   "_String" :: "str_position => string"    ("_")
 | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 127 | |
| 48891 | 128 | ML_file "Tools/string_syntax.ML" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 129 | |
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 130 | lemma UNIV_char: | 
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60801diff
changeset | 131 | "UNIV = image (case_prod Char) (UNIV \<times> UNIV)" | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 132 | proof (rule UNIV_eq_I) | 
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60801diff
changeset | 133 | fix x show "x \<in> image (case_prod Char) (UNIV \<times> UNIV)" by (cases x) auto | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 134 | qed | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 135 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 136 | lemma size_char [code, simp]: | 
| 56846 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 137 | "size_char (c::char) = 0" | 
| 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 138 | "size (c::char) = 0" | 
| 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 blanchet parents: 
55969diff
changeset | 139 | by (cases c, simp)+ | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 140 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 141 | instantiation char :: enum | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 142 | begin | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 143 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 144 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 145 | "Enum.enum = [Char Nibble0 Nibble0, Char Nibble0 Nibble1, Char Nibble0 Nibble2, | 
| 31484 | 146 | Char Nibble0 Nibble3, Char Nibble0 Nibble4, Char Nibble0 Nibble5, | 
| 147 | Char Nibble0 Nibble6, Char Nibble0 Nibble7, Char Nibble0 Nibble8, | |
| 55015 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 wenzelm parents: 
54594diff
changeset | 148 | Char Nibble0 Nibble9, CHR ''\<newline>'', Char Nibble0 NibbleB, | 
| 31484 | 149 | Char Nibble0 NibbleC, Char Nibble0 NibbleD, Char Nibble0 NibbleE, | 
| 150 | Char Nibble0 NibbleF, Char Nibble1 Nibble0, Char Nibble1 Nibble1, | |
| 151 | Char Nibble1 Nibble2, Char Nibble1 Nibble3, Char Nibble1 Nibble4, | |
| 152 | Char Nibble1 Nibble5, Char Nibble1 Nibble6, Char Nibble1 Nibble7, | |
| 153 | Char Nibble1 Nibble8, Char Nibble1 Nibble9, Char Nibble1 NibbleA, | |
| 154 | Char Nibble1 NibbleB, Char Nibble1 NibbleC, Char Nibble1 NibbleD, | |
| 155 | Char Nibble1 NibbleE, Char Nibble1 NibbleF, CHR '' '', CHR ''!'', | |
| 156 | Char Nibble2 Nibble2, CHR ''#'', CHR ''$'', CHR ''%'', CHR ''&'', | |
| 157 |   Char Nibble2 Nibble7, CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', CHR '','',
 | |
| 158 | CHR ''-'', CHR ''.'', CHR ''/'', CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', | |
| 159 | CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9'', CHR '':'', | |
| 160 | CHR '';'', CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', CHR ''@'', CHR ''A'', | |
| 161 | CHR ''B'', CHR ''C'', CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', CHR ''H'', | |
| 162 | CHR ''I'', CHR ''J'', CHR ''K'', CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', | |
| 163 | CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', CHR ''T'', CHR ''U'', CHR ''V'', | |
| 164 | CHR ''W'', CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', Char Nibble5 NibbleC, | |
| 165 | CHR '']'', CHR ''^'', CHR ''_'', Char Nibble6 Nibble0, CHR ''a'', CHR ''b'', | |
| 166 | CHR ''c'', CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', CHR ''h'', CHR ''i'', | |
| 167 | CHR ''j'', CHR ''k'', CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', CHR ''p'', | |
| 168 | CHR ''q'', CHR ''r'', CHR ''s'', CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', | |
| 169 |   CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', CHR ''|'', CHR ''}'', CHR ''~'',
 | |
| 170 | Char Nibble7 NibbleF, Char Nibble8 Nibble0, Char Nibble8 Nibble1, | |
| 171 | Char Nibble8 Nibble2, Char Nibble8 Nibble3, Char Nibble8 Nibble4, | |
| 172 | Char Nibble8 Nibble5, Char Nibble8 Nibble6, Char Nibble8 Nibble7, | |
| 173 | Char Nibble8 Nibble8, Char Nibble8 Nibble9, Char Nibble8 NibbleA, | |
| 174 | Char Nibble8 NibbleB, Char Nibble8 NibbleC, Char Nibble8 NibbleD, | |
| 175 | Char Nibble8 NibbleE, Char Nibble8 NibbleF, Char Nibble9 Nibble0, | |
| 176 | Char Nibble9 Nibble1, Char Nibble9 Nibble2, Char Nibble9 Nibble3, | |
| 177 | Char Nibble9 Nibble4, Char Nibble9 Nibble5, Char Nibble9 Nibble6, | |
| 178 | Char Nibble9 Nibble7, Char Nibble9 Nibble8, Char Nibble9 Nibble9, | |
| 179 | Char Nibble9 NibbleA, Char Nibble9 NibbleB, Char Nibble9 NibbleC, | |
| 180 | Char Nibble9 NibbleD, Char Nibble9 NibbleE, Char Nibble9 NibbleF, | |
| 181 | Char NibbleA Nibble0, Char NibbleA Nibble1, Char NibbleA Nibble2, | |
| 182 | Char NibbleA Nibble3, Char NibbleA Nibble4, Char NibbleA Nibble5, | |
| 183 | Char NibbleA Nibble6, Char NibbleA Nibble7, Char NibbleA Nibble8, | |
| 184 | Char NibbleA Nibble9, Char NibbleA NibbleA, Char NibbleA NibbleB, | |
| 185 | Char NibbleA NibbleC, Char NibbleA NibbleD, Char NibbleA NibbleE, | |
| 186 | Char NibbleA NibbleF, Char NibbleB Nibble0, Char NibbleB Nibble1, | |
| 187 | Char NibbleB Nibble2, Char NibbleB Nibble3, Char NibbleB Nibble4, | |
| 188 | Char NibbleB Nibble5, Char NibbleB Nibble6, Char NibbleB Nibble7, | |
| 189 | Char NibbleB Nibble8, Char NibbleB Nibble9, Char NibbleB NibbleA, | |
| 190 | Char NibbleB NibbleB, Char NibbleB NibbleC, Char NibbleB NibbleD, | |
| 191 | Char NibbleB NibbleE, Char NibbleB NibbleF, Char NibbleC Nibble0, | |
| 192 | Char NibbleC Nibble1, Char NibbleC Nibble2, Char NibbleC Nibble3, | |
| 193 | Char NibbleC Nibble4, Char NibbleC Nibble5, Char NibbleC Nibble6, | |
| 194 | Char NibbleC Nibble7, Char NibbleC Nibble8, Char NibbleC Nibble9, | |
| 195 | Char NibbleC NibbleA, Char NibbleC NibbleB, Char NibbleC NibbleC, | |
| 196 | Char NibbleC NibbleD, Char NibbleC NibbleE, Char NibbleC NibbleF, | |
| 197 | Char NibbleD Nibble0, Char NibbleD Nibble1, Char NibbleD Nibble2, | |
| 198 | Char NibbleD Nibble3, Char NibbleD Nibble4, Char NibbleD Nibble5, | |
| 199 | Char NibbleD Nibble6, Char NibbleD Nibble7, Char NibbleD Nibble8, | |
| 200 | Char NibbleD Nibble9, Char NibbleD NibbleA, Char NibbleD NibbleB, | |
| 201 | Char NibbleD NibbleC, Char NibbleD NibbleD, Char NibbleD NibbleE, | |
| 202 | Char NibbleD NibbleF, Char NibbleE Nibble0, Char NibbleE Nibble1, | |
| 203 | Char NibbleE Nibble2, Char NibbleE Nibble3, Char NibbleE Nibble4, | |
| 204 | Char NibbleE Nibble5, Char NibbleE Nibble6, Char NibbleE Nibble7, | |
| 205 | Char NibbleE Nibble8, Char NibbleE Nibble9, Char NibbleE NibbleA, | |
| 206 | Char NibbleE NibbleB, Char NibbleE NibbleC, Char NibbleE NibbleD, | |
| 207 | Char NibbleE NibbleE, Char NibbleE NibbleF, Char NibbleF Nibble0, | |
| 208 | Char NibbleF Nibble1, Char NibbleF Nibble2, Char NibbleF Nibble3, | |
| 209 | Char NibbleF Nibble4, Char NibbleF Nibble5, Char NibbleF Nibble6, | |
| 210 | Char NibbleF Nibble7, Char NibbleF Nibble8, Char NibbleF Nibble9, | |
| 211 | Char NibbleF NibbleA, Char NibbleF NibbleB, Char NibbleF NibbleC, | |
| 212 | Char NibbleF NibbleD, Char NibbleF NibbleE, Char NibbleF NibbleF]" | |
| 213 | ||
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 214 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 215 | "Enum.enum_all P \<longleftrightarrow> list_all P (Enum.enum :: char list)" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 216 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 217 | definition | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 218 | "Enum.enum_ex P \<longleftrightarrow> list_ex P (Enum.enum :: char list)" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 219 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 220 | lemma enum_char_product_enum_nibble: | 
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60801diff
changeset | 221 | "(Enum.enum :: char list) = map (case_prod Char) (List.product Enum.enum Enum.enum)" | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 222 | by (simp add: enum_char_def enum_nibble_def) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 223 | |
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 224 | instance proof | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 225 | show UNIV: "UNIV = set (Enum.enum :: char list)" | 
| 57247 | 226 | by (simp add: enum_char_product_enum_nibble UNIV_char enum_UNIV) | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 227 | show "distinct (Enum.enum :: char list)" | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 228 | by (auto intro: inj_onI simp add: enum_char_product_enum_nibble distinct_map distinct_product enum_distinct) | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 229 | show "\<And>P. Enum.enum_all P \<longleftrightarrow> Ball (UNIV :: char set) P" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 230 | by (simp add: UNIV enum_all_char_def list_all_iff) | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 231 | show "\<And>P. Enum.enum_ex P \<longleftrightarrow> Bex (UNIV :: char set) P" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 232 | by (simp add: UNIV enum_ex_char_def list_ex_iff) | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 233 | qed | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 234 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 235 | end | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 236 | |
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 237 | lemma card_UNIV_char: | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 238 | "card (UNIV :: char set) = 256" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 239 | by (simp add: card_UNIV_length_enum enum_char_def) | 
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 240 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 241 | definition nat_of_char :: "char \<Rightarrow> nat" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 242 | where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 243 | "nat_of_char c = (case c of Char x y \<Rightarrow> nat_of_nibble x * 16 + nat_of_nibble y)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 244 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 245 | lemma nat_of_char_Char: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 246 | "nat_of_char (Char x y) = nat_of_nibble x * 16 + nat_of_nibble y" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 247 | by (simp add: nat_of_char_def) | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 248 | |
| 60758 | 249 | setup \<open> | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 250 | let | 
| 59631 | 251 |   val nibbles = map_range (Thm.cterm_of @{context} o HOLogic.mk_nibble) 16;
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51160diff
changeset | 252 | val simpset = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51160diff
changeset | 253 |     put_simpset HOL_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51160diff
changeset | 254 |       addsimps @{thms nat_of_nibble.simps mult_0 mult_1 add_0 add_0_right arith_simps numeral_plus_one};
 | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 255 | fun mk_code_eqn x y = | 
| 60801 | 256 |     Thm.instantiate' [] [SOME x, SOME y] @{thm nat_of_char_Char}
 | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 257 | |> simplify simpset; | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 258 | val code_eqns = map_product mk_code_eqn nibbles nibbles; | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 259 | in | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 260 | Global_Theory.note_thmss "" | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 261 |     [((@{binding nat_of_char_simps}, []), [(code_eqns, [])])]
 | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 262 | #> snd | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 263 | end | 
| 60758 | 264 | \<close> | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 265 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 266 | declare nat_of_char_simps [code] | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 267 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 268 | lemma nibble_of_nat_of_char_div_16: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 269 | "nibble_of_nat (nat_of_char c div 16) = (case c of Char x y \<Rightarrow> x)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 270 | by (cases c) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57437diff
changeset | 271 | (simp add: nat_of_char_def add.commute nat_of_nibble_less_16) | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 272 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 273 | lemma nibble_of_nat_nat_of_char: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 274 | "nibble_of_nat (nat_of_char c) = (case c of Char x y \<Rightarrow> y)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 275 | proof (cases c) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 276 | case (Char x y) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 277 | then have *: "nibble_of_nat ((nat_of_nibble y + nat_of_nibble x * 16) mod 16) = y" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 278 | by (simp add: nibble_of_nat_mod_16) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 279 | then have "nibble_of_nat (nat_of_nibble y + nat_of_nibble x * 16) = y" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 280 | by (simp only: nibble_of_nat_mod_16) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57437diff
changeset | 281 | with Char show ?thesis by (simp add: nat_of_char_def add.commute) | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 282 | qed | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 283 | |
| 60758 | 284 | code_datatype Char -- \<open>drop case certificate for char\<close> | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 285 | |
| 55416 | 286 | lemma case_char_code [code]: | 
| 287 | "case_char f c = (let n = nat_of_char c in f (nibble_of_nat (n div 16)) (nibble_of_nat n))" | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 288 | by (cases c) | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55428diff
changeset | 289 | (simp only: Let_def nibble_of_nat_of_char_div_16 nibble_of_nat_nat_of_char char.case) | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 290 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 291 | lemma [code]: | 
| 55416 | 292 | "rec_char = case_char" | 
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 293 | by (simp add: fun_eq_iff split: char.split) | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49948diff
changeset | 294 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 295 | definition char_of_nat :: "nat \<Rightarrow> char" where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 296 | "char_of_nat n = Enum.enum ! (n mod 256)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 297 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 298 | lemma char_of_nat_Char_nibbles: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 299 | "char_of_nat n = Char (nibble_of_nat (n div 16)) (nibble_of_nat n)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 300 | proof - | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 301 | from mod_mult2_eq [of n 16 16] | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 302 | have "n mod 256 = 16 * (n div 16 mod 16) + n mod 16" by simp | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 303 | then show ?thesis | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 304 | by (simp add: char_of_nat_def enum_char_product_enum_nibble card_UNIV_nibble | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57437diff
changeset | 305 | card_UNIV_length_enum [symmetric] nibble_of_nat_def product_nth add.commute) | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 306 | qed | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 307 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 308 | lemma char_of_nat_of_char [simp]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 309 | "char_of_nat (nat_of_char c) = c" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 310 | by (cases c) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 311 | (simp add: char_of_nat_Char_nibbles nibble_of_nat_of_char_div_16 nibble_of_nat_nat_of_char) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 312 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 313 | lemma nat_of_char_of_nat [simp]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 314 | "nat_of_char (char_of_nat n) = n mod 256" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 315 | proof - | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 316 | have "n mod 256 = n mod (16 * 16)" by simp | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 317 | then have "n div 16 mod 16 * 16 + n mod 16 = n mod 256" by (simp only: mod_mult2_eq) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 318 | then show ?thesis | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 319 | by (cases "char_of_nat n") (auto simp add: nat_of_char_def char_of_nat_Char_nibbles) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 320 | qed | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 321 | |
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 322 | lemma char_of_nat_nibble_pair [simp]: | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 323 | "char_of_nat (nat_of_nibble a * 16 + nat_of_nibble b) = Char a b" | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 324 | by (simp add: nat_of_char_Char [symmetric]) | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 325 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 326 | lemma inj_nat_of_char: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 327 | "inj nat_of_char" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 328 | by (rule inj_on_inverseI) (rule char_of_nat_of_char) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 329 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 330 | lemma nat_of_char_eq_iff: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 331 | "nat_of_char c = nat_of_char d \<longleftrightarrow> c = d" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 332 | by (rule inj_eq) (rule inj_nat_of_char) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 333 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 334 | lemma nat_of_char_less_256: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 335 | "nat_of_char c < 256" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 336 | proof (cases c) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 337 | case (Char x y) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 338 | with nat_of_nibble_less_16 [of x] nat_of_nibble_less_16 [of y] | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 339 | show ?thesis by (simp add: nat_of_char_def) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 340 | qed | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 341 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 342 | lemma char_of_nat_mod_256: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 343 | "char_of_nat (n mod 256) = char_of_nat n" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 344 | proof - | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 345 | from nibble_of_nat_mod_16 [of "n mod 256"] | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 346 | have "nibble_of_nat (n mod 256) = nibble_of_nat (n mod 256 mod 16)" by simp | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 347 | with nibble_of_nat_mod_16 [of n] | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 348 | have *: "nibble_of_nat (n mod 256) = nibble_of_nat n" by (simp add: mod_mod_cancel) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 349 | have "n mod 256 = n mod (16 * 16)" by simp | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 350 | then have **: "n mod 256 = n div 16 mod 16 * 16 + n mod 16" by (simp only: mod_mult2_eq) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 351 | show ?thesis | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 352 | by (simp add: char_of_nat_Char_nibbles *) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 353 | (simp add: div_add1_eq nibble_of_nat_mod_16 [of "n div 16"] **) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49972diff
changeset | 354 | qed | 
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 355 | |
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 356 | |
| 60758 | 357 | subsection \<open>Strings as dedicated type\<close> | 
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 358 | |
| 49834 | 359 | typedef literal = "UNIV :: string set" | 
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 360 | morphisms explode STR .. | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 361 | |
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 362 | setup_lifting type_definition_literal | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 363 | |
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 364 | lemma STR_inject' [simp]: | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 365 | "STR s = STR t \<longleftrightarrow> s = t" | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 366 | by transfer rule | 
| 54594 
a2d1522cdd54
setup lifting/transfer for String.literal
 Andreas Lochbihler parents: 
54317diff
changeset | 367 | |
| 57437 | 368 | definition implode :: "string \<Rightarrow> String.literal" | 
| 369 | where | |
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 370 | [code del]: "implode = STR" | 
| 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
59483diff
changeset | 371 | |
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 372 | instantiation literal :: size | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 373 | begin | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 374 | |
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 375 | definition size_literal :: "literal \<Rightarrow> nat" | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 376 | where | 
| 61076 | 377 | [code]: "size_literal (s::literal) = 0" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 378 | |
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 379 | instance .. | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 380 | |
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 381 | end | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 382 | |
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 383 | instantiation literal :: equal | 
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 384 | begin | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 385 | |
| 54594 
a2d1522cdd54
setup lifting/transfer for String.literal
 Andreas Lochbihler parents: 
54317diff
changeset | 386 | lift_definition equal_literal :: "literal \<Rightarrow> literal \<Rightarrow> bool" is "op =" . | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 387 | |
| 54594 
a2d1522cdd54
setup lifting/transfer for String.literal
 Andreas Lochbihler parents: 
54317diff
changeset | 388 | instance by intro_classes (transfer, simp) | 
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 389 | |
| 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 390 | end | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 391 | |
| 54594 
a2d1522cdd54
setup lifting/transfer for String.literal
 Andreas Lochbihler parents: 
54317diff
changeset | 392 | declare equal_literal.rep_eq[code] | 
| 
a2d1522cdd54
setup lifting/transfer for String.literal
 Andreas Lochbihler parents: 
54317diff
changeset | 393 | |
| 52365 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 haftmann parents: 
51717diff
changeset | 394 | lemma [code nbe]: | 
| 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 haftmann parents: 
51717diff
changeset | 395 | fixes s :: "String.literal" | 
| 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 haftmann parents: 
51717diff
changeset | 396 | shows "HOL.equal s s \<longleftrightarrow> True" | 
| 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 haftmann parents: 
51717diff
changeset | 397 | by (simp add: equal) | 
| 
95186e6e4bf4
reflexive nbe equation for equality on String.literal
 haftmann parents: 
51717diff
changeset | 398 | |
| 55426 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
55015diff
changeset | 399 | lifting_update literal.lifting | 
| 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
55015diff
changeset | 400 | lifting_forget literal.lifting | 
| 
90f2ceed2828
make lifting setup for String.literal local to prevent transfer from replacing STR ''...'' literals
 Andreas Lochbihler parents: 
55015diff
changeset | 401 | |
| 60758 | 402 | subsection \<open>Code generator\<close> | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 403 | |
| 48891 | 404 | ML_file "Tools/string_code.ML" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 405 | |
| 33237 | 406 | code_reserved SML string | 
| 407 | code_reserved OCaml string | |
| 34886 | 408 | code_reserved Scala string | 
| 33237 | 409 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 410 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 411 | type_constructor literal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 412 | (SML) "string" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 413 | and (OCaml) "string" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 414 | and (Haskell) "String" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 415 | and (Scala) "String" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 416 | |
| 60758 | 417 | setup \<open> | 
| 34886 | 418 | fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"] | 
| 60758 | 419 | \<close> | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 420 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 421 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 422 | class_instance literal :: equal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 423 | (Haskell) - | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 424 | | constant "HOL.equal :: literal \<Rightarrow> literal \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 425 | (SML) "!((_ : string) = _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 426 | and (OCaml) "!((_ : string) = _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 427 | and (Haskell) infix 4 "==" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52365diff
changeset | 428 | and (Scala) infixl 5 "==" | 
| 31051 
4d9b52e0a48c
refined HOL string theories and corresponding ML fragments
 haftmann parents: diff
changeset | 429 | |
| 60758 | 430 | setup \<open>Sign.map_naming (Name_Space.mandatory_path "Code")\<close> | 
| 52910 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 431 | |
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 432 | definition abort :: "literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a" | 
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 433 | where [simp, code del]: "abort _ f = f ()" | 
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 434 | |
| 54317 
da932f511746
add congruence rule to prevent code_simp from looping
 Andreas Lochbihler parents: 
52910diff
changeset | 435 | lemma abort_cong: "msg = msg' ==> Code.abort msg f = Code.abort msg' f" | 
| 
da932f511746
add congruence rule to prevent code_simp from looping
 Andreas Lochbihler parents: 
52910diff
changeset | 436 | by simp | 
| 
da932f511746
add congruence rule to prevent code_simp from looping
 Andreas Lochbihler parents: 
52910diff
changeset | 437 | |
| 60758 | 438 | setup \<open>Sign.map_naming Name_Space.parent_path\<close> | 
| 52910 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 439 | |
| 60758 | 440 | setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm Code.abort_cong})\<close>
 | 
| 54317 
da932f511746
add congruence rule to prevent code_simp from looping
 Andreas Lochbihler parents: 
52910diff
changeset | 441 | |
| 52910 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 442 | code_printing constant Code.abort \<rightharpoonup> | 
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 443 | (SML) "!(raise/ Fail/ _)" | 
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 444 | and (OCaml) "failwith" | 
| 59483 
ddb73392356e
explicit type annotation avoids problems with Haskell type inference
 haftmann parents: 
58889diff
changeset | 445 | and (Haskell) "!(error/ ::/ forall a./ String -> (() -> a) -> a)" | 
| 52910 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 446 |     and (Scala) "!{/ sys.error((_));/  ((_)).apply(())/ }"
 | 
| 
7bfe0df532a9
abort execution of generated code with explicit exception message
 Andreas Lochbihler parents: 
52435diff
changeset | 447 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35123diff
changeset | 448 | hide_type (open) literal | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31176diff
changeset | 449 | |
| 57437 | 450 | hide_const (open) implode explode | 
| 451 | ||
| 39250 
548a3e5521ab
changing String.literal to a type instead of a datatype
 bulwahn parents: 
39198diff
changeset | 452 | end |