src/HOL/Library/Char_ord.thy
author wenzelm
Wed, 25 Jun 2025 16:35:25 +0200
changeset 82768 8f866fd6fae1
parent 75662 ed15f2cd4f7d
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15737
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Library/Char_ord.thy
22805
1166a966e7b4 moved stuff to Char_nat.thy
haftmann
parents: 22483
diff changeset
     2
    Author:     Norbert Voelker, Florian Haftmann
15737
nipkow
parents:
diff changeset
     3
*)
nipkow
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>Order on characters\<close>
15737
nipkow
parents:
diff changeset
     6
nipkow
parents:
diff changeset
     7
theory Char_ord
63462
c1fe30f2bc32 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
     8
  imports Main
15737
nipkow
parents:
diff changeset
     9
begin
nipkow
parents:
diff changeset
    10
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    11
instantiation char :: linorder
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    12
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    13
75601
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    14
definition less_eq_char :: \<open>char \<Rightarrow> char \<Rightarrow> bool\<close>
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    15
  where \<open>c1 \<le> c2 \<longleftrightarrow> of_char c1 \<le> (of_char c2 :: nat)\<close>
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 67399
diff changeset
    16
75601
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    17
definition less_char :: \<open>char \<Rightarrow> char \<Rightarrow> bool\<close>
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    18
  where \<open>c1 < c2 \<longleftrightarrow> of_char c1 < (of_char c2 :: nat)\<close>
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 67399
diff changeset
    19
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    20
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    21
instance
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61076
diff changeset
    22
  by standard (auto simp add: less_eq_char_def less_char_def)
22805
1166a966e7b4 moved stuff to Char_nat.thy
haftmann
parents: 22483
diff changeset
    23
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    24
end
15737
nipkow
parents:
diff changeset
    25
75601
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    26
lemma less_eq_char_simp [simp, code]:
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    27
  \<open>Char b0 b1 b2 b3 b4 b5 b6 b7 \<le> Char c0 c1 c2 c3 c4 c5 c6 c7
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    28
    \<longleftrightarrow> lexordp_eq [b7, b6, b5, b4, b3, b2, b1, b0] [c7, c6, c5, c4, c3, c2, c1, c0]\<close>
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    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
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    31
lemma less_char_simp [simp, code]:
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    32
  \<open>Char b0 b1 b2 b3 b4 b5 b6 b7 < Char c0 c1 c2 c3 c4 c5 c6 c7
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    33
    \<longleftrightarrow> ord_class.lexordp [b7, b6, b5, b4, b3, b2, b1, b0] [c7, c6, c5, c4, c3, c2, c1, c0]\<close>
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    34
  by (simp only: less_char_def of_char_def char.sel horner_sum_less_iff_lexordp list.size) simp
63462
c1fe30f2bc32 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
    35
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    36
instantiation char :: distrib_lattice
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    37
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    38
75601
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    39
definition \<open>(inf :: char \<Rightarrow> _) = min\<close>
5ec227251b07 Avoid calculations where not necessary.
haftmann
parents: 75600
diff changeset
    40
definition \<open>(sup :: char \<Rightarrow> _) = max\<close>
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    41
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    42
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    43
  by standard (auto simp add: inf_char_def sup_char_def max_min_distrib2)
22483
86064f2f2188 added instance for lattice
haftmann
parents: 21911
diff changeset
    44
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    45
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
    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