| author | paulson | 
| Mon, 11 Jan 2016 22:14:15 +0000 | |
| changeset 62131 | 1baed43f453e | 
| parent 61585 | a9599d3d7610 | 
| child 62364 | 9209770bdcdf | 
| permissions | -rw-r--r-- | 
| 24999 | 1 | (* Title: HOL/Library/Code_Char.thy | 
| 2 | Author: Florian Haftmann | |
| 3 | *) | |
| 4 | ||
| 60500 | 5 | section \<open>Code generation of pretty characters (and strings)\<close> | 
| 24999 | 6 | |
| 7 | theory Code_Char | |
| 51542 | 8 | imports Main Char_ord | 
| 24999 | 9 | begin | 
| 10 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 11 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 12 | type_constructor char \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 13 | (SML) "char" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 14 | and (OCaml) "char" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 15 | and (Haskell) "Prelude.Char" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 16 | and (Scala) "Char" | 
| 24999 | 17 | |
| 60500 | 18 | setup \<open> | 
| 37222 | 19 | fold String_Code.add_literal_char ["SML", "OCaml", "Haskell", "Scala"] | 
| 31053 | 20 | #> String_Code.add_literal_list_string "Haskell" | 
| 60500 | 21 | \<close> | 
| 24999 | 22 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 23 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 24 | class_instance char :: equal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 25 | (Haskell) - | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 26 | | constant "HOL.equal :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 27 | (SML) "!((_ : char) = _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 28 | and (OCaml) "!((_ : char) = _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 29 | and (Haskell) infix 4 "==" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 30 | and (Scala) infixl 5 "==" | 
| 61076 | 31 | | constant "Code_Evaluation.term_of :: char \<Rightarrow> term" \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 32 | (Eval) "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))" | 
| 24999 | 33 | |
| 34 | code_reserved SML | |
| 35 | char | |
| 36 | ||
| 37 | code_reserved OCaml | |
| 38 | char | |
| 39 | ||
| 37222 | 40 | code_reserved Scala | 
| 41 | char | |
| 42 | ||
| 33234 | 43 | code_reserved SML String | 
| 44 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 45 | code_printing | 
| 57437 | 46 | constant String.implode \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 47 | (SML) "String.implode" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 48 | and (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)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 49 | and (Haskell) "_" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 50 | and (Scala) "!(\"\" ++/ _)" | 
| 57437 | 51 | | constant String.explode \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 52 | (SML) "String.explode" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 53 | and (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) [])" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 54 | and (Haskell) "_" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 55 | and (Scala) "!(_.toList)" | 
| 33234 | 56 | |
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 57 | |
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 58 | definition integer_of_char :: "char \<Rightarrow> integer" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 59 | where | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 60 | "integer_of_char = integer_of_nat o nat_of_char" | 
| 
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 char_of_integer :: "integer \<Rightarrow> char" | 
| 
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 | "char_of_integer = char_of_nat \<circ> nat_of_integer" | 
| 
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 | lemma [code]: | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 67 | "nat_of_char = nat_of_integer o integer_of_char" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 68 | by (simp add: integer_of_char_def fun_eq_iff) | 
| 
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 | "char_of_nat = char_of_integer o integer_of_nat" | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 72 | by (simp add: char_of_integer_def fun_eq_iff) | 
| 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 73 | |
| 55427 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 74 | lemmas integer_of_char_code [code] = | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 75 | nat_of_char_simps[ | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 76 | THEN arg_cong[where f="integer_of_nat"], | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 77 | unfolded integer_of_nat_numeral integer_of_nat_1 integer_of_nat_0, | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 78 | folded fun_cong[OF integer_of_char_def, unfolded o_apply]] | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 79 | |
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 80 | lemma char_of_integer_code [code]: | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 81 | "char_of_integer n = enum_class.enum ! (nat_of_integer n mod 256)" | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 82 | by(simp add: char_of_integer_def char_of_nat_def) | 
| 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 Andreas Lochbihler parents: 
54599diff
changeset | 83 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 84 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 85 | constant integer_of_char \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 86 | (SML) "!(IntInf.fromInt o Char.ord)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 87 | and (OCaml) "Big'_int.big'_int'_of'_int (Char.code _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 88 | and (Haskell) "Prelude.toInteger (Prelude.fromEnum (_ :: Prelude.Char))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 89 | and (Scala) "BigInt(_.toInt)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 90 | | constant char_of_integer \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 91 | (SML) "!(Char.chr o IntInf.toInt)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 92 | and (OCaml) "Char.chr (Big'_int.int'_of'_big'_int _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 93 | and (Haskell) "!(let chr k | (0 <= k && k < 256) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 94 | and (Scala) "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else error(\"character value out of range\"))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 95 | | constant "Orderings.less_eq :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 96 | (SML) "!((_ : char) <= _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 97 | and (OCaml) "!((_ : char) <= _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 98 | and (Haskell) infix 4 "<=" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 99 | and (Scala) infixl 4 "<=" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 100 | and (Eval) infixl 6 "<=" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 101 | | constant "Orderings.less :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 102 | (SML) "!((_ : char) < _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 103 | and (OCaml) "!((_ : char) < _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 104 | and (Haskell) infix 4 "<" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 105 | and (Scala) infixl 4 "<" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 106 | and (Eval) infixl 6 "<" | 
| 54596 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 107 | | constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup> | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 108 | (SML) "!((_ : string) <= _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 109 | and (OCaml) "!((_ : string) <= _)" | 
| 61585 | 110 |     \<comment> \<open>Order operations for @{typ String.literal} work in Haskell only 
 | 
| 54599 
17d76426c7da
revert 4af7c82463d3 and document type class problem in Haskell
 Andreas Lochbihler parents: 
54596diff
changeset | 111 | if no type class instance needs to be generated, because String = [Char] in Haskell | 
| 60500 | 112 |           and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
 | 
| 54596 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 113 | and (Haskell) infix 4 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 114 | and (Scala) infixl 4 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 115 | and (Eval) infixl 6 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 116 | | constant "Orderings.less :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup> | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 117 | (SML) "!((_ : string) < _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 118 | and (OCaml) "!((_ : string) < _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 119 | and (Haskell) infix 4 "<" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 120 | and (Scala) infixl 4 "<" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 121 | and (Eval) infixl 6 "<" | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 122 | |
| 24999 | 123 | end | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 124 |