| author | hoelzl | 
| Fri, 22 Mar 2013 10:41:43 +0100 | |
| changeset 51481 | ef949192e5d6 | 
| parent 51160 | 599ff65b85e2 | 
| child 51542 | 738598beeb26 | 
| permissions | -rw-r--r-- | 
| 24999 | 1 | (* Title: HOL/Library/Code_Char.thy | 
| 2 | Author: Florian Haftmann | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Code generation of pretty characters (and strings) *}
 | |
| 6 | ||
| 7 | theory Code_Char | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 8 | imports Main "~~/src/HOL/Library/Char_ord" | 
| 24999 | 9 | begin | 
| 10 | ||
| 11 | code_type char | |
| 12 | (SML "char") | |
| 13 | (OCaml "char") | |
| 48431 
6efff142bb54
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
 haftmann parents: 
42598diff
changeset | 14 | (Haskell "Prelude.Char") | 
| 37222 | 15 | (Scala "Char") | 
| 24999 | 16 | |
| 17 | setup {*
 | |
| 37222 | 18 | fold String_Code.add_literal_char ["SML", "OCaml", "Haskell", "Scala"] | 
| 31053 | 19 | #> String_Code.add_literal_list_string "Haskell" | 
| 24999 | 20 | *} | 
| 21 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37459diff
changeset | 22 | code_instance char :: equal | 
| 24999 | 23 | (Haskell -) | 
| 24 | ||
| 25 | code_reserved SML | |
| 26 | char | |
| 27 | ||
| 28 | code_reserved OCaml | |
| 29 | char | |
| 30 | ||
| 37222 | 31 | code_reserved Scala | 
| 32 | char | |
| 33 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37459diff
changeset | 34 | code_const "HOL.equal \<Colon> char \<Rightarrow> char \<Rightarrow> bool" | 
| 24999 | 35 | (SML "!((_ : char) = _)") | 
| 36 | (OCaml "!((_ : char) = _)") | |
| 39272 | 37 | (Haskell infix 4 "==") | 
| 37222 | 38 | (Scala infixl 5 "==") | 
| 24999 | 39 | |
| 32657 | 40 | code_const "Code_Evaluation.term_of \<Colon> char \<Rightarrow> term" | 
| 31046 | 41 | (Eval "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))") | 
| 28228 | 42 | |
| 33234 | 43 | |
| 44 | definition implode :: "string \<Rightarrow> String.literal" where | |
| 45 | "implode = STR" | |
| 46 | ||
| 47 | code_reserved SML String | |
| 48 | ||
| 49 | code_const implode | |
| 50 | (SML "String.implode") | |
| 42598 
85ca44488a29
adding code generation setup for String.implode and String.explode in OCaml (contributed by Andreas Lochbihler)
 bulwahn parents: 
42316diff
changeset | 51 | (OCaml "!(let l = _ in let res = String.create (List.length l) in let rec imp i = function | [] -> res | c :: l -> String.set res i c; imp (i + 1) l in imp 0 l)") | 
| 33234 | 52 | (Haskell "_") | 
| 37459 | 53 | (Scala "!(\"\" ++/ _)") | 
| 33234 | 54 | |
| 55 | code_const explode | |
| 56 | (SML "String.explode") | |
| 42598 
85ca44488a29
adding code generation setup for String.implode and String.explode in OCaml (contributed by Andreas Lochbihler)
 bulwahn parents: 
42316diff
changeset | 57 | (OCaml "!(let s = _ in let rec exp i l = if i < 0 then l else exp (i - 1) (String.get s i :: l) in exp (String.length s - 1) [])") | 
| 33234 | 58 | (Haskell "_") | 
| 37459 | 59 | (Scala "!(_.toList)") | 
| 33234 | 60 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 61 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 62 | definition integer_of_char :: "char \<Rightarrow> integer" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 63 | where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 64 | "integer_of_char = integer_of_nat o nat_of_char" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 65 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 66 | definition char_of_integer :: "integer \<Rightarrow> char" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 67 | where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 68 | "char_of_integer = char_of_nat \<circ> nat_of_integer" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 69 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 70 | lemma [code]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 71 | "nat_of_char = nat_of_integer o integer_of_char" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 72 | by (simp add: integer_of_char_def fun_eq_iff) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 73 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 74 | lemma [code]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 75 | "char_of_nat = char_of_integer o integer_of_nat" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 76 | by (simp add: char_of_integer_def fun_eq_iff) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 77 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 78 | code_const integer_of_char and char_of_integer | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 79 | (SML "!(IntInf.fromInt o Char.ord)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 80 | and "!(Char.chr o IntInf.toInt)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 81 | (OCaml "Big'_int.big'_int'_of'_int (Char.code _)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 82 | and "Char.chr (Big'_int.int'_of'_big'_int _)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 83 | (Haskell "Prelude.toInteger (Prelude.fromEnum (_ :: Prelude.Char))" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 84 | and "!(let chr k | (0 <= k && k < 256) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 85 | (Scala "BigInt(_.toInt)" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 86 | and "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else error(\"character value out of range\"))") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 87 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 88 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 89 | code_const "Orderings.less_eq \<Colon> char \<Rightarrow> char \<Rightarrow> bool" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 90 | (SML "!((_ : char) <= _)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 91 | (OCaml "!((_ : char) <= _)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 92 | (Haskell infix 4 "<=") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 93 | (Scala infixl 4 "<=") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 94 | (Eval infixl 6 "<=") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 95 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 96 | code_const "Orderings.less \<Colon> char \<Rightarrow> char \<Rightarrow> bool" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 97 | (SML "!((_ : char) < _)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 98 | (OCaml "!((_ : char) < _)") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 99 | (Haskell infix 4 "<") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 100 | (Scala infixl 4 "<") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 101 | (Eval infixl 6 "<") | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 102 | |
| 24999 | 103 | end | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 104 |