src/HOL/Library/Code_Char.thy
changeset 61585 a9599d3d7610
parent 61076 bdc1e2f0a86a
child 62364 9209770bdcdf
equal deleted inserted replaced
61584:f06e5a5a4b46 61585:a9599d3d7610
   105     and (Scala) infixl 4 "<"
   105     and (Scala) infixl 4 "<"
   106     and (Eval) infixl 6 "<"
   106     and (Eval) infixl 6 "<"
   107 |  constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
   107 |  constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
   108     (SML) "!((_ : string) <= _)"
   108     (SML) "!((_ : string) <= _)"
   109     and (OCaml) "!((_ : string) <= _)"
   109     and (OCaml) "!((_ : string) <= _)"
   110     -- \<open>Order operations for @{typ String.literal} work in Haskell only 
   110     \<comment> \<open>Order operations for @{typ String.literal} work in Haskell only 
   111           if no type class instance needs to be generated, because String = [Char] in Haskell
   111           if no type class instance needs to be generated, because String = [Char] in Haskell
   112           and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
   112           and @{typ "char list"} need not have the same order as @{typ String.literal}.\<close>
   113     and (Haskell) infix 4 "<="
   113     and (Haskell) infix 4 "<="
   114     and (Scala) infixl 4 "<="
   114     and (Scala) infixl 4 "<="
   115     and (Eval) infixl 6 "<="
   115     and (Eval) infixl 6 "<="