--- a/src/HOL/Library/Code_Char.thy Wed Nov 27 10:43:51 2013 +0100
+++ b/src/HOL/Library/Code_Char.thy Wed Nov 27 10:54:44 2013 +0100
@@ -100,6 +100,9 @@
| constant "Orderings.less_eq :: String.literal \<Rightarrow> String.literal \<Rightarrow> bool" \<rightharpoonup>
(SML) "!((_ : string) <= _)"
and (OCaml) "!((_ : string) <= _)"
+ -- {* Order operations for @{typ String.literal} work in Haskell only
+ if no type class instance needs to be generated, because String = [Char] in Haskell
+ and @{typ "char list"} need not have the same order as @{typ String.literal}. *}
and (Haskell) infix 4 "<="
and (Scala) infixl 4 "<="
and (Eval) infixl 6 "<="
--- a/src/HOL/Library/List_lexord.thy Wed Nov 27 10:43:51 2013 +0100
+++ b/src/HOL/Library/List_lexord.thy Wed Nov 27 10:54:44 2013 +0100
@@ -118,6 +118,4 @@
"(x\<Colon>'a\<Colon>{equal, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
by simp_all
-code_printing class_instance String.literal :: ord \<rightharpoonup> (Haskell) -
-
end