--- a/src/HOL/List.thy Fri Feb 20 10:14:31 2009 +0100
+++ b/src/HOL/List.thy Fri Feb 20 10:14:32 2009 +0100
@@ -3564,52 +3564,51 @@
open Basic_Code_Thingol;
-fun implode_list (nil', cons') t =
- let
- fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
- if c = cons'
- then SOME (t1, t2)
- else NONE
- | dest_cons _ = NONE;
- val (ts, t') = Code_Thingol.unfoldr dest_cons t;
- in case t'
- of IConst (c, _) => if c = nil' then SOME ts else NONE
+fun implode_list naming t = case pairself
+ (Code_Thingol.lookup_const naming) (@{const_name Nil}, @{const_name Cons})
+ of (SOME nil', SOME cons') => let
+ fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
+ if c = cons'
+ then SOME (t1, t2)
+ else NONE
+ | dest_cons _ = NONE;
+ val (ts, t') = Code_Thingol.unfoldr dest_cons t;
+ in case t'
+ of IConst (c, _) => if c = nil' then SOME ts else NONE
+ | _ => NONE
+ end
| _ => NONE
- end;
-
-fun decode_char nibbles' (IConst (c1, _), IConst (c2, _)) =
- let
- fun idx c = find_index (curry (op =) c) nibbles';
- fun decode ~1 _ = NONE
- | decode _ ~1 = NONE
- | decode n m = SOME (chr (n * 16 + m));
- in decode (idx c1) (idx c2) end
- | decode_char _ _ = NONE;
-
-fun implode_string (char', nibbles') mk_char mk_string ts =
- let
- fun implode_char (IConst (c, _) `$ t1 `$ t2) =
- if c = char' then decode_char nibbles' (t1, t2) else NONE
- | implode_char _ = NONE;
- val ts' = map implode_char ts;
- in if forall is_some ts'
- then (SOME o Code_Printer.str o mk_string o implode o map_filter I) ts'
- else NONE
- end;
-
-fun list_names naming = pairself (the o Code_Thingol.lookup_const naming)
- (@{const_name Nil}, @{const_name Cons});
-fun char_name naming = (the o Code_Thingol.lookup_const naming)
- @{const_name Char}
-fun nibble_names naming = map (the o Code_Thingol.lookup_const naming)
- [@{const_name Nibble0}, @{const_name Nibble1},
+
+fun decode_char naming (IConst (c1, _), IConst (c2, _)) = (case map_filter
+ (Code_Thingol.lookup_const naming)[@{const_name Nibble0}, @{const_name Nibble1},
@{const_name Nibble2}, @{const_name Nibble3},
@{const_name Nibble4}, @{const_name Nibble5},
@{const_name Nibble6}, @{const_name Nibble7},
@{const_name Nibble8}, @{const_name Nibble9},
@{const_name NibbleA}, @{const_name NibbleB},
@{const_name NibbleC}, @{const_name NibbleD},
- @{const_name NibbleE}, @{const_name NibbleF}];
+ @{const_name NibbleE}, @{const_name NibbleF}]
+ of nibbles' as [_, _, _, _, _, _, _, _, _, _, _, _, _, _, _, _] => let
+ fun idx c = find_index (curry (op =) c) nibbles';
+ fun decode ~1 _ = NONE
+ | decode _ ~1 = NONE
+ | decode n m = SOME (chr (n * 16 + m));
+ in decode (idx c1) (idx c2) end
+ | _ => NONE)
+ | decode_char _ _ = NONE
+
+fun implode_string naming mk_char mk_string ts = case
+ Code_Thingol.lookup_const naming @{const_name Char}
+ of SOME char' => let
+ fun implode_char (IConst (c, _) `$ t1 `$ t2) =
+ if c = char' then decode_char naming (t1, t2) else NONE
+ | implode_char _ = NONE;
+ val ts' = map implode_char ts;
+ in if forall is_some ts'
+ then (SOME o Code_Printer.str o mk_string o implode o map_filter I) ts'
+ else NONE
+ end
+ | _ => NONE;
fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy [
@@ -3622,7 +3621,7 @@
let
val mk_list = Code_Printer.literal_list literals;
fun pretty pr naming thm vars fxy [(t1, _), (t2, _)] =
- case Option.map (cons t1) (implode_list (list_names naming) t2)
+ case Option.map (cons t1) (implode_list naming t2)
of SOME ts => mk_list (map (pr vars Code_Printer.NOBR) ts)
| NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
in (2, pretty) end;
@@ -3633,8 +3632,8 @@
val mk_char = Code_Printer.literal_char literals;
val mk_string = Code_Printer.literal_string literals;
fun pretty pr naming thm vars fxy [(t1, _), (t2, _)] =
- case Option.map (cons t1) (implode_list (list_names naming) t2)
- of SOME ts => (case implode_string (char_name naming, nibble_names naming) mk_char mk_string ts
+ case Option.map (cons t1) (implode_list naming t2)
+ of SOME ts => (case implode_string naming mk_char mk_string ts
of SOME p => p
| NONE => mk_list (map (pr vars Code_Printer.NOBR) ts))
| NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
@@ -3644,7 +3643,7 @@
let
val mk_char = Code_Printer.literal_char literals;
fun pretty _ naming thm _ _ [(t1, _), (t2, _)] =
- case decode_char (nibble_names naming) (t1, t2)
+ case decode_char naming (t1, t2)
of SOME c => (Code_Printer.str o mk_char) c
| NONE => Code_Printer.nerror thm "Illegal character expression";
in (2, pretty) end;
@@ -3654,8 +3653,8 @@
val mk_char = Code_Printer.literal_char literals;
val mk_string = Code_Printer.literal_string literals;
fun pretty _ naming thm _ _ [(t, _)] =
- case implode_list (list_names naming) t
- of SOME ts => (case implode_string (char_name naming, nibble_names naming) mk_char mk_string ts
+ case implode_list naming t
+ of SOME ts => (case implode_string naming mk_char mk_string ts
of SOME p => p
| NONE => Code_Printer.nerror thm "Illegal message expression")
| NONE => Code_Printer.nerror thm "Illegal message expression";