# HG changeset patch # User haftmann # Date 1241629771 -7200 # Node ID 2cf6efca6c71096a67c5cdab7d31235f5ae91649 # Parent 841c9f67f9e7afe966e2ecd9b54d92da7d0ff59e proper structures for list and string code generation stuff diff -r 841c9f67f9e7 -r 2cf6efca6c71 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed May 06 19:09:14 2009 +0200 +++ b/src/HOL/IsaMakefile Wed May 06 19:09:31 2009 +0200 @@ -235,6 +235,7 @@ Tools/Groebner_Basis/normalizer.ML \ Tools/atp_manager.ML \ Tools/atp_wrapper.ML \ + Tools/list_code.ML \ Tools/meson.ML \ Tools/metis_tools.ML \ Tools/numeral.ML \ @@ -252,6 +253,7 @@ Tools/res_hol_clause.ML \ Tools/res_reconstruct.ML \ Tools/specification_package.ML \ + Tools/string_code.ML \ Tools/string_syntax.ML \ Tools/TFL/casesplit.ML \ Tools/TFL/dcterm.ML \ diff -r 841c9f67f9e7 -r 2cf6efca6c71 src/HOL/List.thy --- a/src/HOL/List.thy Wed May 06 19:09:14 2009 +0200 +++ b/src/HOL/List.thy Wed May 06 19:09:31 2009 +0200 @@ -6,6 +6,7 @@ theory List imports Plain Presburger Recdef ATP_Linkup +uses ("Tools/list_code.ML") begin datatype 'a list = @@ -3460,6 +3461,8 @@ subsubsection {* Setup *} +use "Tools/list_code.ML" + code_type list (SML "_ list") (OCaml "_ list") @@ -3470,11 +3473,6 @@ (OCaml "[]") (Haskell "[]") -code_const Cons - (SML infixr 7 "::") - (OCaml infixr 6 "::") - (Haskell infixr 5 ":") - code_instance list :: eq (Haskell -) @@ -3503,22 +3501,22 @@ and gen_list aG aT i = gen_list' aG aT i i; *} -consts_code Nil ("[]") consts_code Cons ("(_ ::/ _)") setup {* let - -fun list_codegen thy defs dep thyname b t gr = - let - val ts = HOLogic.dest_list t; - val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false - (fastype_of t) gr; - val (ps, gr'') = fold_map - (Codegen.invoke_codegen thy defs dep thyname false) ts gr' - in SOME (Pretty.list "[" "]" ps, gr'') end handle TERM _ => NONE; - -in Codegen.add_codegen "list_codegen" list_codegen end + fun list_codegen thy defs dep thyname b t gr = + let + val ts = HOLogic.dest_list t; + val (_, gr') = Codegen.invoke_tycodegen thy defs dep thyname false + (fastype_of t) gr; + val (ps, gr'') = fold_map + (Codegen.invoke_codegen thy defs dep thyname false) ts gr' + in SOME (Pretty.list "[" "]" ps, gr'') end handle TERM _ => NONE; +in + fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell"] + #> Codegen.add_codegen "list_codegen" list_codegen +end *} diff -r 841c9f67f9e7 -r 2cf6efca6c71 src/HOL/String.thy --- a/src/HOL/String.thy Wed May 06 19:09:14 2009 +0200 +++ b/src/HOL/String.thy Wed May 06 19:09:31 2009 +0200 @@ -4,7 +4,9 @@ theory String imports List -uses "Tools/string_syntax.ML" +uses + "Tools/string_syntax.ML" + ("Tools/string_code.ML") begin subsection {* Characters *} @@ -97,149 +99,7 @@ subsection {* Code generator *} -text {* This also covers pretty syntax for list literals. *} - -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"] -*} +use "Tools/string_code.ML" code_type message_string (SML "string") @@ -247,8 +107,7 @@ (Haskell "String") setup {* - fold (fn target => add_literal_message @{const_name STR} target) - ["SML", "OCaml", "Haskell"] + fold String_Code.add_literal_message ["SML", "OCaml", "Haskell"] *} code_instance message_string :: eq diff -r 841c9f67f9e7 -r 2cf6efca6c71 src/HOL/Tools/list_code.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/list_code.ML Wed May 06 19:09:31 2009 +0200 @@ -0,0 +1,52 @@ +(* Author: Florian Haftmann, TU Muenchen + +Code generation for list literals. +*) + +signature LIST_CODE = +sig + val implode_list: string -> string -> Code_Thingol.iterm -> Code_Thingol.iterm list option + val default_list: int * string + -> (Code_Printer.fixity -> Code_Thingol.iterm -> Pretty.T) + -> Code_Printer.fixity -> Code_Thingol.iterm -> Code_Thingol.iterm -> Pretty.T + val add_literal_list: string -> theory -> theory +end; + +structure List_Code : LIST_CODE = +struct + +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 + | _ => 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 add_literal_list target = + let + fun pretty literals [nil', cons'] pr thm vars fxy [(t1, _), (t2, _)] = + case Option.map (cons t1) (implode_list nil' cons' t2) + of SOME ts => + Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts) + | NONE => + default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2; + in Code_Target.add_syntax_const target + @{const_name Cons} (SOME (2, ([@{const_name Nil}, @{const_name Cons}], pretty))) + end + +end; diff -r 841c9f67f9e7 -r 2cf6efca6c71 src/HOL/Tools/string_code.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/string_code.ML Wed May 06 19:09:31 2009 +0200 @@ -0,0 +1,88 @@ +(* Author: Florian Haftmann, TU Muenchen + +Code generation for character and string literals. +*) + +signature STRING_CODE = +sig + val add_literal_list_string: string -> theory -> theory + val add_literal_char: string -> theory -> theory + val add_literal_message: string -> theory -> theory +end; + +structure String_Code : STRING_CODE = +struct + +open Basic_Code_Thingol; + +fun decode_char nibbles' tt = + 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 case tt + of (IConst (c1, _), IConst (c2, _)) => decode (idx c1) (idx c2) + | _ => NONE + end; + +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_filter implode_char ts; + in if length ts = length ts' + then (SOME o Code_Printer.str o mk_string o implode) ts' + else NONE + end; + +val cs_nibbles = [@{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}]; +val cs_summa = [@{const_name Nil}, @{const_name Cons}, @{const_name Char}] @ cs_nibbles; + +fun add_literal_list_string target = + let + fun pretty literals (nil' :: cons' :: char' :: nibbles') pr thm vars fxy [(t1, _), (t2, _)] = + case Option.map (cons t1) (List_Code.implode_list nil' cons' t2) + of SOME ts => (case implode_string char' nibbles' + (Code_Printer.literal_char literals) (Code_Printer.literal_string literals) ts + of SOME p => p + | NONE => + Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)) + | NONE => + List_Code.default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2; + in Code_Target.add_syntax_const target + @{const_name Cons} (SOME (2, (cs_summa, pretty))) + end; + +fun add_literal_char target = + let + fun pretty literals nibbles' _ thm _ _ [(t1, _), (t2, _)] = + case decode_char nibbles' (t1, t2) + of SOME c => (Code_Printer.str o Code_Printer.literal_char literals) c + | NONE => Code_Printer.nerror thm "Illegal character expression"; + in Code_Target.add_syntax_const target + @{const_name Char} (SOME (2, (cs_nibbles, pretty))) + end; + +fun add_literal_message target = + let + fun pretty literals (nil' :: cons' :: char' :: nibbles') _ thm _ _ [(t, _)] = + case List_Code.implode_list nil' cons' t + of SOME ts => (case implode_string char' nibbles' + (Code_Printer.literal_char literals) (Code_Printer.literal_string literals) ts + of SOME p => p + | NONE => Code_Printer.nerror thm "Illegal message expression") + | NONE => Code_Printer.nerror thm "Illegal message expression"; + in Code_Target.add_syntax_const target + @{const_name STR} (SOME (1, (cs_summa, pretty))) + end; + +end;