src/HOL/Library/Code_Char.thy
author haftmann
Tue, 16 Sep 2008 16:13:09 +0200
changeset 28244 f433e544a855
parent 28228 7ebe8dc06cbb
child 28346 b8390cd56b8f
permissions -rw-r--r--
a sophisticated char/nibble conversion combinator
     1 (*  Title:      HOL/Library/Code_Char.thy
     2     ID:         $Id$
     3     Author:     Florian Haftmann
     4 *)
     5 
     6 header {* Code generation of pretty characters (and strings) *}
     7 
     8 theory Code_Char
     9 imports Plain "~~/src/HOL/List" "~~/src/HOL/Code_Eval"
    10 begin
    11 
    12 code_type char
    13   (SML "char")
    14   (OCaml "char")
    15   (Haskell "Char")
    16 
    17 setup {*
    18   fold (fn target => add_literal_char target) ["SML", "OCaml", "Haskell"] 
    19   #> add_literal_list_string "Haskell"
    20 *}
    21 
    22 code_instance char :: eq
    23   (Haskell -)
    24 
    25 code_reserved SML
    26   char
    27 
    28 code_reserved OCaml
    29   char
    30 
    31 code_const "op = \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
    32   (SML "!((_ : char) = _)")
    33   (OCaml "!((_ : char) = _)")
    34   (Haskell infixl 4 "==")
    35 
    36 code_const "Code_Eval.term_of \<Colon> char \<Rightarrow> term"
    37   (SML "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))")
    38 
    39 end