src/HOL/Library/Code_Char_chr.thy
author haftmann
Wed, 08 Dec 2010 15:05:46 +0100
changeset 41082 9ff94e7cc3b3
parent 39302 d7728f65b353
child 48431 6efff142bb54
permissions -rw-r--r--
bot comes before top, inf before sup etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24999
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Code_Char_chr.thy
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann
haftmann
parents:
diff changeset
     3
*)
haftmann
parents:
diff changeset
     4
haftmann
parents:
diff changeset
     5
header {* Code generation of pretty characters with character codes *}
haftmann
parents:
diff changeset
     6
haftmann
parents:
diff changeset
     7
theory Code_Char_chr
30663
0b6aff7451b2 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 28562
diff changeset
     8
imports Char_nat Code_Char Code_Integer Main
24999
haftmann
parents:
diff changeset
     9
begin
haftmann
parents:
diff changeset
    10
haftmann
parents:
diff changeset
    11
definition
haftmann
parents:
diff changeset
    12
  "int_of_char = int o nat_of_char"
haftmann
parents:
diff changeset
    13
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27368
diff changeset
    14
lemma [code]:
24999
haftmann
parents:
diff changeset
    15
  "nat_of_char = nat o int_of_char"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    16
  unfolding int_of_char_def by (simp add: fun_eq_iff)
24999
haftmann
parents:
diff changeset
    17
haftmann
parents:
diff changeset
    18
definition
haftmann
parents:
diff changeset
    19
  "char_of_int = char_of_nat o nat"
haftmann
parents:
diff changeset
    20
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27368
diff changeset
    21
lemma [code]:
24999
haftmann
parents:
diff changeset
    22
  "char_of_nat = char_of_int o int"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    23
  unfolding char_of_int_def by (simp add: fun_eq_iff)
24999
haftmann
parents:
diff changeset
    24
haftmann
parents:
diff changeset
    25
code_const int_of_char and char_of_int
haftmann
parents:
diff changeset
    26
  (SML "!(IntInf.fromInt o Char.ord)" and "!(Char.chr o IntInf.toInt)")
haftmann
parents:
diff changeset
    27
  (OCaml "Big'_int.big'_int'_of'_int (Char.code _)" and "Char.chr (Big'_int.int'_of'_big'_int _)")
37222
4d984bc33c66 added Scala code setup
haftmann
parents: 32069
diff changeset
    28
  (Haskell "toInteger (fromEnum (_ :: Char))" and "!(let chr k | (0 <= k && k < 256) = toEnum k :: Char in chr . fromInteger)")
4d984bc33c66 added Scala code setup
haftmann
parents: 32069
diff changeset
    29
  (Scala "BigInt(_.toInt)" and "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else error(\"character value out of range\"))")
24999
haftmann
parents:
diff changeset
    30
37222
4d984bc33c66 added Scala code setup
haftmann
parents: 32069
diff changeset
    31
end