| author | haftmann | 
| Sun, 08 Oct 2017 22:28:21 +0200 | |
| changeset 66805 | 274b4edca859 | 
| parent 65884 | d76937b773d9 | 
| 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> | 
| 65884 | 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 | 
| 62597 | 24 | constant integer_of_char \<rightharpoonup> | 
| 25 | (SML) "!(IntInf.fromInt o Char.ord)" | |
| 26 | and (OCaml) "Big'_int.big'_int'_of'_int (Char.code _)" | |
| 27 | and (Haskell) "Prelude.toInteger (Prelude.fromEnum (_ :: Prelude.Char))" | |
| 28 | and (Scala) "BigInt(_.toInt)" | |
| 29 | | constant char_of_integer \<rightharpoonup> | |
| 30 | (SML) "!(Char.chr o IntInf.toInt)" | |
| 31 | and (OCaml) "Char.chr (Big'_int.int'_of'_big'_int _)" | |
| 32 | and (Haskell) "!(let chr k | (0 <= k && k < 256) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)" | |
| 65884 | 33 | and (Scala) "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else sys.error(\"character value out of range\"))" | 
| 62597 | 34 | | class_instance char :: equal \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 35 | (Haskell) - | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 36 | | 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 | 37 | (SML) "!((_ : char) = _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 38 | 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 | 39 | 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 | 40 | and (Scala) infixl 5 "==" | 
| 61076 | 41 | | 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 | 42 | (Eval) "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))" | 
| 24999 | 43 | |
| 44 | code_reserved SML | |
| 45 | char | |
| 46 | ||
| 47 | code_reserved OCaml | |
| 48 | char | |
| 49 | ||
| 37222 | 50 | code_reserved Scala | 
| 51 | char | |
| 52 | ||
| 33234 | 53 | code_reserved SML String | 
| 54 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 55 | code_printing | 
| 57437 | 56 | 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 | 57 | (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 | 58 | 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 | 59 | and (Haskell) "_" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 60 | and (Scala) "!(\"\" ++/ _)" | 
| 57437 | 61 | | 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 | 62 | (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 | 63 | 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 | 64 | and (Haskell) "_" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 65 | and (Scala) "!(_.toList)" | 
| 33234 | 66 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 67 | code_printing | 
| 62597 | 68 | constant "Orderings.less_eq :: char \<Rightarrow> char \<Rightarrow> bool" \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 69 | (SML) "!((_ : char) <= _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 70 | 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 | 71 | 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 | 72 | 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 | 73 | 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 | 74 | | 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 | 75 | (SML) "!((_ : char) < _)" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51542diff
changeset | 76 | 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 | 77 | 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 | 78 | 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 | 79 | and (Eval) infixl 6 "<" | 
| 54596 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 80 | | 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 | 81 | (SML) "!((_ : string) <= _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 82 | and (OCaml) "!((_ : string) <= _)" | 
| 65884 | 83 |     \<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 | 84 | if no type class instance needs to be generated, because String = [Char] in Haskell | 
| 60500 | 85 |           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 | 86 | and (Haskell) infix 4 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 87 | and (Scala) infixl 4 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 88 | and (Eval) infixl 6 "<=" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 89 | | 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 | 90 | (SML) "!((_ : string) < _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 91 | and (OCaml) "!((_ : string) < _)" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 92 | and (Haskell) infix 4 "<" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 93 | and (Scala) infixl 4 "<" | 
| 
368c70ee1f46
implement comparisons on String.literal by target-language comparisons
 Andreas Lochbihler parents: 
52435diff
changeset | 94 | and (Eval) infixl 6 "<" | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
48431diff
changeset | 95 | |
| 24999 | 96 | end |