author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 75662 | ed15f2cd4f7d |
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 |
|
63462 | 8 |
imports Main |
15737 | 9 |
begin |
10 |
||
25764 | 11 |
instantiation char :: linorder |
12 |
begin |
|
13 |
||
75601 | 14 |
definition less_eq_char :: \<open>char \<Rightarrow> char \<Rightarrow> bool\<close> |
15 |
where \<open>c1 \<le> c2 \<longleftrightarrow> of_char c1 \<le> (of_char c2 :: nat)\<close> |
|
68028 | 16 |
|
75601 | 17 |
definition less_char :: \<open>char \<Rightarrow> char \<Rightarrow> bool\<close> |
18 |
where \<open>c1 < c2 \<longleftrightarrow> of_char c1 < (of_char c2 :: nat)\<close> |
|
68028 | 19 |
|
25764 | 20 |
|
60679 | 21 |
instance |
62597 | 22 |
by standard (auto simp add: less_eq_char_def less_char_def) |
22805 | 23 |
|
25764 | 24 |
end |
15737 | 25 |
|
75601 | 26 |
lemma less_eq_char_simp [simp, code]: |
27 |
\<open>Char b0 b1 b2 b3 b4 b5 b6 b7 \<le> Char c0 c1 c2 c3 c4 c5 c6 c7 |
|
28 |
\<longleftrightarrow> lexordp_eq [b7, b6, b5, b4, b3, b2, b1, b0] [c7, c6, c5, c4, c3, c2, c1, c0]\<close> |
|
29 |
by (simp only: less_eq_char_def of_char_def char.sel horner_sum_less_eq_iff_lexordp_eq list.size) simp |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
48926
diff
changeset
|
30 |
|
75601 | 31 |
lemma less_char_simp [simp, code]: |
32 |
\<open>Char b0 b1 b2 b3 b4 b5 b6 b7 < Char c0 c1 c2 c3 c4 c5 c6 c7 |
|
33 |
\<longleftrightarrow> ord_class.lexordp [b7, b6, b5, b4, b3, b2, b1, b0] [c7, c6, c5, c4, c3, c2, c1, c0]\<close> |
|
34 |
by (simp only: less_char_def of_char_def char.sel horner_sum_less_iff_lexordp list.size) simp |
|
63462 | 35 |
|
25764 | 36 |
instantiation char :: distrib_lattice |
37 |
begin |
|
38 |
||
75601 | 39 |
definition \<open>(inf :: char \<Rightarrow> _) = min\<close> |
40 |
definition \<open>(sup :: char \<Rightarrow> _) = max\<close> |
|
25764 | 41 |
|
60679 | 42 |
instance |
43 |
by standard (auto simp add: inf_char_def sup_char_def max_min_distrib2) |
|
22483 | 44 |
|
25764 | 45 |
end |
46 |
||
75647
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75601
diff
changeset
|
47 |
code_identifier |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75601
diff
changeset
|
48 |
code_module Char_ord \<rightharpoonup> |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75601
diff
changeset
|
49 |
(SML) Str and (OCaml) Str and (Haskell) Str and (Scala) Str |
34cd1d210b92
officical abstract characters for code generation
haftmann
parents:
75601
diff
changeset
|
50 |
|
54595
a2eeeb335a48
instantiate linorder for String.literal by lexicographic order
Andreas Lochbihler
parents:
51160
diff
changeset
|
51 |
end |