src/HOL/Library/Code_Char_ord.thy
author bulwahn
Thu, 19 May 2011 18:09:20 +0200
changeset 42857 1e1b74448f6b
parent 42841 9079f49053e5
permissions -rw-r--r--
a deeper understanding of the code generation adaptation compared to 9079f49053e5

(*  Title:      HOL/Library/Code_Char_ord.thy
    Author:     Lukas Bulwahn, Florian Haftmann, Rene Thiemann
*)

header {* Code generation of orderings for pretty characters *}

theory Code_Char_ord
imports Code_Char Char_ord
begin
  
code_const "Orderings.less_eq \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
  (SML "!((_ : char) <= _)")
  (OCaml "!((_ : char) <= _)")
  (Haskell infix 4 "<=")
  (Scala infixl 4 "<=")
  (Eval infixl 6 "<=")

code_const "Orderings.less \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
  (SML "!((_ : char) < _)")
  (OCaml "!((_ : char) < _)")
  (Haskell infix 4 "<")
  (Scala infixl 4 "<")
  (Eval infixl 6 "<")

end