adding theory to force code generation of ordering on characters to built-in ordering of characters in target languages
--- a/src/HOL/IsaMakefile Tue May 17 22:29:55 2011 +0200
+++ b/src/HOL/IsaMakefile Wed May 18 15:45:33 2011 +0200
@@ -437,7 +437,7 @@
Library/AssocList.thy Library/BigO.thy Library/Binomial.thy \
Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy \
Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy \
- Library/Code_Integer.thy Library/Code_Natural.thy \
+ Library/Code_Char_ord.thy Library/Code_Integer.thy Library/Code_Natural.thy \
Library/Code_Prolog.thy Tools/Predicate_Compile/code_prolog.ML \
Library/ContNotDenum.thy Library/Continuity.thy Library/Convex.thy \
Library/Countable.thy Library/Diagonalize.thy Library/Dlist.thy \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Code_Char_ord.thy Wed May 18 15:45:33 2011 +0200
@@ -0,0 +1,25 @@
+(* 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 "<=")
+ (Haskell infix 4 "<=")
+ (Scala infixl 4 "<=")
+ (Eval infixl 6 "<=")
+
+code_const "Orderings.less \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
+ (SML "!((_ : char) < _)")
+ (OCaml "<")
+ (Haskell infix 4 "<")
+ (Scala infixl 4 "<")
+ (Eval infixl 6 "<")
+
+end
\ No newline at end of file
--- a/src/HOL/Library/ROOT.ML Tue May 17 22:29:55 2011 +0200
+++ b/src/HOL/Library/ROOT.ML Wed May 18 15:45:33 2011 +0200
@@ -2,4 +2,4 @@
(* Classical Higher-order Logic -- batteries included *)
use_thys ["Library", "List_Prefix", "List_lexord", "Sublist_Order",
- "Code_Char_chr", "Code_Integer", "Efficient_Nat", "Executable_Set"(*, "Code_Prolog"*)];
+ "Code_Char_chr", "Code_Char_ord", "Code_Integer", "Efficient_Nat", "Executable_Set"(*, "Code_Prolog"*)];