src/HOL/List.thy
changeset 28090 29af3c712d2b
parent 28072 a45e8c872dc1
child 28230 87feb146d3d1
--- a/src/HOL/List.thy	Tue Sep 02 20:07:51 2008 +0200
+++ b/src/HOL/List.thy	Tue Sep 02 20:38:17 2008 +0200
@@ -3449,10 +3449,145 @@
   (OCaml "[]")
   (Haskell "[]")
 
+ML {*
+local
+
+open Basic_Code_Thingol;
+val nil' = Code_Name.const @{theory} @{const_name Nil};
+val cons' = Code_Name.const @{theory} @{const_name Cons};
+val char' = Code_Name.const @{theory} @{const_name Char}
+val nibbles' = map (Code_Name.const @{theory})
+   [@{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}];
+
+fun implode_list 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
+    | _ => NONE
+  end;
+
+fun decode_char (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 mk_char mk_string ts =
+  let
+    fun implode_char (IConst (c, _) `$ t1 `$ t2) =
+          if c = char' then decode_char (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 default_list (target_fxy, target_cons) pr fxy t1 t2 =
+  Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy [
+    pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
+    Code_Printer.str target_cons,
+    pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
+  ];
+
+fun pretty_list literals =
+  let
+    val mk_list = Code_Printer.literal_list literals;
+    fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] =
+      case Option.map (cons t1) (implode_list 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;
+
+fun pretty_list_string literals =
+  let
+    val mk_list = Code_Printer.literal_list literals;
+    val mk_char = Code_Printer.literal_char literals;
+    val mk_string = Code_Printer.literal_string literals;
+    fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] =
+      case Option.map (cons t1) (implode_list t2)
+       of SOME ts => (case implode_string 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;
+  in (2, pretty) end;
+
+fun pretty_char literals =
+  let
+    val mk_char = Code_Printer.literal_char literals;
+    fun pretty _ thm _ _ _ [(t1, _), (t2, _)] =
+      case decode_char (t1, t2)
+       of SOME c => (Code_Printer.str o mk_char) c
+        | NONE => Code_Printer.nerror thm "Illegal character expression";
+  in (2, pretty) end;
+
+fun pretty_message literals =
+  let
+    val mk_char = Code_Printer.literal_char literals;
+    val mk_string = Code_Printer.literal_string literals;
+    fun pretty _ thm _ _ _ [(t, _)] =
+      case implode_list t
+       of SOME ts => (case implode_string 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";
+  in (1, pretty) end;
+
+in
+
+fun add_literal_list target thy =
+  let
+    val pr = pretty_list (Code_Target.the_literals thy target);
+  in
+    thy
+    |> Code_Target.add_syntax_const target @{const_name Cons} (SOME pr)
+  end;
+
+fun add_literal_list_string target thy =
+  let
+    val pr = pretty_list_string (Code_Target.the_literals thy target);
+  in
+    thy
+    |> Code_Target.add_syntax_const target @{const_name Cons} (SOME pr)
+  end;
+
+fun add_literal_char target thy =
+  let
+    val pr = pretty_char (Code_Target.the_literals thy target);
+  in
+    thy
+    |> Code_Target.add_syntax_const target @{const_name Char} (SOME pr)
+  end;
+
+fun add_literal_message str target thy =
+  let
+    val pr = pretty_message (Code_Target.the_literals thy target);
+  in
+    thy
+    |> Code_Target.add_syntax_const target str (SOME pr)
+  end;
+
+end;
+*}
+
 setup {*
-  fold (fn target => Code_Target.add_literal_list target
-    @{const_name Nil} @{const_name Cons}
-  ) ["SML", "OCaml", "Haskell"]
+  fold (fn target => add_literal_list target) ["SML", "OCaml", "Haskell"]
 *}
 
 code_instance list :: eq