refined HOL string theories and corresponding ML fragments
authorhaftmann
Wed, 06 May 2009 16:01:06 +0200
changeset 31048 ac146fc38b51
parent 31047 c13b0406c039
child 31049 396d4d6a1594
refined HOL string theories and corresponding ML fragments
src/HOL/Code_Eval.thy
src/HOL/Code_Message.thy
src/HOL/IsaMakefile
src/HOL/List.thy
src/HOL/Tools/hologic.ML
src/HOL/Tools/string_syntax.ML
src/HOL/Typerep.thy
--- a/src/HOL/Code_Eval.thy	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/Code_Eval.thy	Wed May 06 16:01:06 2009 +0200
@@ -33,7 +33,7 @@
 struct
 
 fun mk_term f g (Const (c, ty)) =
-      @{term Const} $ Message_String.mk c $ g ty
+      @{term Const} $ HOLogic.mk_message_string c $ g ty
   | mk_term f g (t1 $ t2) =
       @{term App} $ mk_term f g t1 $ mk_term f g t2
   | mk_term f g (Free v) = f v
@@ -154,7 +154,9 @@
   (Eval "Term.Const/ (_, _)" and "Term.$/ (_, _)")
 
 code_const "term_of \<Colon> message_string \<Rightarrow> term"
-  (Eval "Message'_String.mk")
+  (Eval "HOLogic.mk'_message'_string")
+
+code_reserved Eval HOLogic
 
 
 subsection {* Evaluation setup *}
--- a/src/HOL/Code_Message.thy	Wed May 06 16:01:05 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* Monolithic strings (message strings) for code generation *}
-
-theory Code_Message
-imports Plain "~~/src/HOL/List"
-begin
-
-subsection {* Datatype of messages *}
-
-datatype message_string = STR string
-
-lemmas [code del] = message_string.recs message_string.cases
-
-lemma [code]: "size (s\<Colon>message_string) = 0"
-  by (cases s) simp_all
-
-lemma [code]: "message_string_size (s\<Colon>message_string) = 0"
-  by (cases s) simp_all
-
-subsection {* ML interface *}
-
-ML {*
-structure Message_String =
-struct
-
-fun mk s = @{term STR} $ HOLogic.mk_string s;
-
-end;
-*}
-
-
-subsection {* Code serialization *}
-
-code_type message_string
-  (SML "string")
-  (OCaml "string")
-  (Haskell "String")
-
-setup {*
-  fold (fn target => add_literal_message @{const_name STR} target)
-    ["SML", "OCaml", "Haskell"]
-*}
-
-code_reserved SML string
-code_reserved OCaml string
-
-code_instance message_string :: eq
-  (Haskell -)
-
-code_const "eq_class.eq \<Colon> message_string \<Rightarrow> message_string \<Rightarrow> bool"
-  (SML "!((_ : string) = _)")
-  (OCaml "!((_ : string) = _)")
-  (Haskell infixl 4 "==")
-
-end
--- a/src/HOL/IsaMakefile	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/IsaMakefile	Wed May 06 16:01:06 2009 +0200
@@ -206,7 +206,6 @@
 MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
   ATP_Linkup.thy \
   Code_Eval.thy \
-  Code_Message.thy \
   Equiv_Relations.thy \
   Groebner_Basis.thy \
   Hilbert_Choice.thy \
@@ -220,6 +219,7 @@
   Presburger.thy \
   Recdef.thy \
   SetInterval.thy \
+  String.thy \
   $(SRC)/Provers/Arith/assoc_fold.ML \
   $(SRC)/Provers/Arith/cancel_numeral_factor.ML \
   $(SRC)/Provers/Arith/cancel_numerals.ML \
--- a/src/HOL/List.thy	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/List.thy	Wed May 06 16:01:06 2009 +0200
@@ -6,7 +6,6 @@
 
 theory List
 imports Plain Presburger Recdef ATP_Linkup
-uses "Tools/string_syntax.ML"
 begin
 
 datatype 'a list =
@@ -3433,77 +3432,6 @@
 by (auto simp add: set_Cons_def intro: listrel.intros) 
 
 
-subsection{*Miscellany*}
-
-subsubsection {* Characters and strings *}
-
-datatype nibble =
-    Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7
-  | Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF
-
-lemma UNIV_nibble:
-  "UNIV = {Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7,
-    Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF}" (is "_ = ?A")
-proof (rule UNIV_eq_I)
-  fix x show "x \<in> ?A" by (cases x) simp_all
-qed
-
-instance nibble :: finite
-  by default (simp add: UNIV_nibble)
-
-datatype char = Char nibble nibble
-  -- "Note: canonical order of character encoding coincides with standard term ordering"
-
-lemma UNIV_char:
-  "UNIV = image (split Char) (UNIV \<times> UNIV)"
-proof (rule UNIV_eq_I)
-  fix x show "x \<in> image (split Char) (UNIV \<times> UNIV)" by (cases x) auto
-qed
-
-instance char :: finite
-  by default (simp add: UNIV_char)
-
-lemma size_char [code, simp]:
-  "size (c::char) = 0" by (cases c) simp
-
-lemma char_size [code, simp]:
-  "char_size (c::char) = 0" by (cases c) simp
-
-primrec nibble_pair_of_char :: "char \<Rightarrow> nibble \<times> nibble" where
-  "nibble_pair_of_char (Char n m) = (n, m)"
-
-declare nibble_pair_of_char.simps [code del]
-
-setup {*
-let
-  val nibbles = map (Thm.cterm_of @{theory} o HOLogic.mk_nibble) (0 upto 15);
-  val thms = map_product
-   (fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
-      nibbles nibbles;
-in
-  PureThy.note_thmss Thm.lemmaK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
-  #-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
-end
-*}
-
-lemma char_case_nibble_pair [code, code inline]:
-  "char_case f = split f o nibble_pair_of_char"
-  by (simp add: expand_fun_eq split: char.split)
-
-lemma char_rec_nibble_pair [code, code inline]:
-  "char_rec f = split f o nibble_pair_of_char"
-  unfolding char_case_nibble_pair [symmetric]
-  by (simp add: expand_fun_eq split: char.split)
-
-types string = "char list"
-
-syntax
-  "_Char" :: "xstr => char"    ("CHR _")
-  "_String" :: "xstr => string"    ("_")
-
-setup StringSyntax.setup
-
-
 subsection {* Size function *}
 
 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
@@ -3527,10 +3455,38 @@
   "(\<And>x. x \<in> set xs \<Longrightarrow> f x < g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
 by (induct xs) force+
 
+
 subsection {* Code generator *}
 
 subsubsection {* Setup *}
 
+code_type list
+  (SML "_ list")
+  (OCaml "_ list")
+  (Haskell "![_]")
+
+code_const Nil
+  (SML "[]")
+  (OCaml "[]")
+  (Haskell "[]")
+
+code_const Cons
+  (SML infixr 7 "::")
+  (OCaml infixr 6 "::")
+  (Haskell infixr 5 ":")
+
+code_instance list :: eq
+  (Haskell -)
+
+code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
+  (Haskell infixl 4 "==")
+
+code_reserved SML
+  list
+
+code_reserved OCaml
+  list
+
 types_code
   "list" ("_ list")
 attach (term_of) {*
@@ -3546,181 +3502,9 @@
    (1, fn () => ([], fn () => HOLogic.nil_const aT))] ()
 and gen_list aG aT i = gen_list' aG aT i i;
 *}
-  "char" ("string")
-attach (term_of) {*
-val term_of_char = HOLogic.mk_char o ord;
-*}
-attach (test) {*
-fun gen_char i =
-  let val j = random_range (ord "a") (Int.min (ord "a" + i, ord "z"))
-  in (chr j, fn () => HOLogic.mk_char j) end;
-*}
-
-consts_code "Cons" ("(_ ::/ _)")
-
-code_type list
-  (SML "_ list")
-  (OCaml "_ list")
-  (Haskell "![_]")
-
-code_reserved SML
-  list
-
-code_reserved OCaml
-  list
-
-code_const Nil
-  (SML "[]")
-  (OCaml "[]")
-  (Haskell "[]")
-
-ML {*
-local
-
-open Basic_Code_Thingol;
-
-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
-
-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}]
-   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 [
-    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 naming thm vars fxy [(t1, _), (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;
-
-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 naming thm vars fxy [(t1, _), (t2, _)] =
-      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;
-  in (2, pretty) end;
-
-fun pretty_char literals =
-  let
-    val mk_char = Code_Printer.literal_char literals;
-    fun pretty _ naming thm _ _ [(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;
-
-fun pretty_message literals =
-  let
-    val mk_char = Code_Printer.literal_char literals;
-    val mk_string = Code_Printer.literal_string literals;
-    fun pretty _ naming thm _ _ [(t, _)] =
-      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";
-  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 => add_literal_list target) ["SML", "OCaml", "Haskell"]
-*}
-
-code_instance list :: eq
-  (Haskell -)
-
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
-  (Haskell infixl 4 "==")
+
+consts_code Nil ("[]")
+consts_code Cons ("(_ ::/ _)")
 
 setup {*
 let
@@ -3734,18 +3518,7 @@
       (Codegen.invoke_codegen thy defs dep thyname false) ts gr'
   in SOME (Pretty.list "[" "]" ps, gr'') end handle TERM _ => NONE;
 
-fun char_codegen thy defs dep thyname b t gr =
-  let
-    val i = HOLogic.dest_char t;
-    val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false
-      (fastype_of t) gr;
-  in SOME (Codegen.str (ML_Syntax.print_string (chr i)), gr')
-  end handle TERM _ => NONE;
-
-in
-  Codegen.add_codegen "list_codegen" list_codegen
-  #> Codegen.add_codegen "char_codegen" char_codegen
-end;
+in Codegen.add_codegen "list_codegen" list_codegen end
 *}
 
 
--- a/src/HOL/Tools/hologic.ML	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/Tools/hologic.ML	Wed May 06 16:01:06 2009 +0200
@@ -116,6 +116,9 @@
   val stringT: typ
   val mk_string: string -> term
   val dest_string: term -> string
+  val message_stringT: typ
+  val mk_message_string: string -> term
+  val dest_message_string: term -> string
 end;
 
 structure HOLogic: HOLOGIC =
@@ -510,44 +513,6 @@
 val realT = Type ("RealDef.real", []);
 
 
-(* nibble *)
-
-val nibbleT = Type ("List.nibble", []);
-
-fun mk_nibble n =
-  let val s =
-    if 0 <= n andalso n <= 9 then chr (n + ord "0")
-    else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10)
-    else raise TERM ("mk_nibble", [])
-  in Const ("List.nibble.Nibble" ^ s, nibbleT) end;
-
-fun dest_nibble t =
-  let fun err () = raise TERM ("dest_nibble", [t]) in
-    (case try (unprefix "List.nibble.Nibble" o fst o Term.dest_Const) t of
-      NONE => err ()
-    | SOME c =>
-        if size c <> 1 then err ()
-        else if "0" <= c andalso c <= "9" then ord c - ord "0"
-        else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10
-        else err ())
-  end;
-
-
-(* char *)
-
-val charT = Type ("List.char", []);
-
-fun mk_char n =
-  if 0 <= n andalso n <= 255 then
-    Const ("List.char.Char", nibbleT --> nibbleT --> charT) $
-      mk_nibble (n div 16) $ mk_nibble (n mod 16)
-  else raise TERM ("mk_char", []);
-
-fun dest_char (Const ("List.char.Char", _) $ t $ u) =
-      dest_nibble t * 16 + dest_nibble u
-  | dest_char t = raise TERM ("dest_char", [t]);
-
-
 (* list *)
 
 fun listT T = Type ("List.list", [T]);
@@ -570,11 +535,60 @@
   | dest_list t = raise TERM ("dest_list", [t]);
 
 
+(* nibble *)
+
+val nibbleT = Type ("String.nibble", []);
+
+fun mk_nibble n =
+  let val s =
+    if 0 <= n andalso n <= 9 then chr (n + ord "0")
+    else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10)
+    else raise TERM ("mk_nibble", [])
+  in Const ("String.nibble.Nibble" ^ s, nibbleT) end;
+
+fun dest_nibble t =
+  let fun err () = raise TERM ("dest_nibble", [t]) in
+    (case try (unprefix "String.nibble.Nibble" o fst o Term.dest_Const) t of
+      NONE => err ()
+    | SOME c =>
+        if size c <> 1 then err ()
+        else if "0" <= c andalso c <= "9" then ord c - ord "0"
+        else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10
+        else err ())
+  end;
+
+
+(* char *)
+
+val charT = Type ("String.char", []);
+
+fun mk_char n =
+  if 0 <= n andalso n <= 255 then
+    Const ("String.char.Char", nibbleT --> nibbleT --> charT) $
+      mk_nibble (n div 16) $ mk_nibble (n mod 16)
+  else raise TERM ("mk_char", []);
+
+fun dest_char (Const ("String.char.Char", _) $ t $ u) =
+      dest_nibble t * 16 + dest_nibble u
+  | dest_char t = raise TERM ("dest_char", [t]);
+
+
 (* string *)
 
-val stringT = Type ("List.string", []);
+val stringT = Type ("String.string", []);
 
 val mk_string = mk_list charT o map (mk_char o ord) o explode;
 val dest_string = implode o map (chr o dest_char) o dest_list;
 
+
+(* message_string *)
+
+val message_stringT = Type ("String.message_string", []);
+
+fun mk_message_string s = Const ("String.message_string.STR", stringT --> message_stringT)
+      $ mk_string s;
+fun dest_message_string (Const ("String.message_string.STR", _) $ t) =
+      dest_string t
+  | dest_message_string t = raise TERM ("dest_message_string", [t]);
+
 end;
--- a/src/HOL/Tools/string_syntax.ML	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/Tools/string_syntax.ML	Wed May 06 16:01:06 2009 +0200
@@ -15,12 +15,14 @@
 
 (* nibble *)
 
+val nib_prefix = "String.nibble.";
+
 val mk_nib =
-  Syntax.Constant o unprefix "List.nibble." o
+  Syntax.Constant o unprefix nib_prefix o
   fst o Term.dest_Const o HOLogic.mk_nibble;
 
 fun dest_nib (Syntax.Constant c) =
-  HOLogic.dest_nibble (Const ("List.nibble." ^ c, dummyT))
+  HOLogic.dest_nibble (Const (nib_prefix ^ c, dummyT))
     handle TERM _ => raise Match;
 
 
--- a/src/HOL/Typerep.thy	Wed May 06 16:01:05 2009 +0200
+++ b/src/HOL/Typerep.thy	Wed May 06 16:01:06 2009 +0200
@@ -1,11 +1,9 @@
-(*  Title:      HOL/Typerep.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
+(* Author: Florian Haftmann, TU Muenchen *)
 
 header {* Reflecting Pure types into HOL *}
 
 theory Typerep
-imports Plain List Code_Message
+imports Plain String
 begin
 
 datatype typerep = Typerep message_string "typerep list"
@@ -42,7 +40,7 @@
 struct
 
 fun mk f (Type (tyco, tys)) =
-      @{term Typerep} $ Message_String.mk tyco
+      @{term Typerep} $ HOLogic.mk_message_string tyco
         $ HOLogic.mk_list @{typ typerep} (map (mk f) tys)
   | mk f (TFree v) =
       f v;