--- a/src/HOL/HOL.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/HOL.thy Fri Aug 10 17:04:34 2007 +0200
@@ -183,7 +183,7 @@
True_or_False: "(P=True) | (P=False)"
defs
- Let_def [code func]: "Let s f == f(s)"
+ Let_def: "Let s f == f(s)"
if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)"
finalconsts
@@ -1701,6 +1701,8 @@
subsubsection {* Generic code generator setup *}
+setup "CodeName.setup #> CodeTarget.setup"
+
text {* operational equality for code generation *}
class eq (attach "op =") = type
@@ -1737,8 +1739,6 @@
lemmas [code] = imp_conv_disj
-lemmas [code func] = if_True if_False
-
instance bool :: eq ..
lemma [code func]:
@@ -1796,15 +1796,18 @@
text {* Let and If *}
+lemmas [code func] = Let_def if_True if_False
+
setup {*
- CodegenPackage.add_appconst (@{const_name Let}, CodegenPackage.appgen_let)
- #> CodegenPackage.add_appconst (@{const_name If}, CodegenPackage.appgen_if)
+ CodePackage.add_appconst (@{const_name Let}, CodePackage.appgen_let)
+ #> CodePackage.add_appconst (@{const_name If}, CodePackage.appgen_if)
*}
+
subsubsection {* Evaluation oracle *}
oracle eval_oracle ("term") = {* fn thy => fn t =>
- if CodegenPackage.satisfies thy (HOLogic.dest_Trueprop t) []
+ if CodePackage.satisfies thy (HOLogic.dest_Trueprop t) []
then t
else HOLogic.Trueprop $ HOLogic.true_const (*dummy*)
*}
--- a/src/HOL/Lambda/WeakNorm.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Lambda/WeakNorm.thy Fri Aug 10 17:04:34 2007 +0200
@@ -574,7 +574,7 @@
*}
setup {*
- CodegenSerializer.add_undefined "SML" "arbitrary" "(raise Fail \"arbitrary\")"
+ CodeTarget.add_undefined "SML" "arbitrary" "(raise Fail \"arbitrary\")"
*}
definition
--- a/src/HOL/Library/Eval.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Library/Eval.thy Fri Aug 10 17:04:34 2007 +0200
@@ -50,7 +50,7 @@
(Type (tyco,
map TFree (Name.names Name.context "'a" asorts))))]) arities, thy);
fun hook specs =
- DatatypeCodegen.prove_codetypes_arities (ClassPackage.intro_classes_tac [])
+ DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
(map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
[TypOf.class_typ_of] mk ((K o K) I)
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
@@ -112,7 +112,7 @@
fun hook specs =
if (fst o hd) specs = (fst o dest_Type) @{typ typ} then I
else
- DatatypeCodegen.prove_codetypes_arities (ClassPackage.intro_classes_tac [])
+ DatatypeCodegen.prove_codetypes_arities (Class.intro_classes_tac [])
(map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
[TermOf.class_term_of] ((K o K o pair) []) mk
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
@@ -165,7 +165,7 @@
val eval_ref = ref (NONE : term option);
fun eval_term thy t =
- CodegenPackage.eval_term
+ CodePackage.eval_term
thy (("Eval.eval_ref", eval_ref), TermOf.mk_term_of t);
fun print eval s = Toplevel.keep (fn state =>
--- a/src/HOL/Library/ML_Int.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Library/ML_Int.thy Fri Aug 10 17:04:34 2007 +0200
@@ -169,7 +169,7 @@
(SML "int")
setup {*
- CodegenSerializer.add_pretty_numeral "SML" false
+ CodeTarget.add_pretty_numeral "SML" false
(@{const_name number_of}, @{typ "int \<Rightarrow> ml_int"})
@{const_name Numeral.B0} @{const_name Numeral.B1}
@{const_name Numeral.Pls} @{const_name Numeral.Min}
--- a/src/HOL/Library/ML_String.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Library/ML_String.thy Fri Aug 10 17:04:34 2007 +0200
@@ -60,7 +60,7 @@
@{const_name NibbleC}, @{const_name NibbleD},
@{const_name NibbleE}, @{const_name NibbleF}];
in
- CodegenSerializer.add_pretty_ml_string "SML"
+ CodeTarget.add_pretty_ml_string "SML"
charr nibbles @{const_name Nil} @{const_name Cons} @{const_name STR}
end
*}
--- a/src/HOL/Library/Pretty_Char.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Library/Pretty_Char.thy Fri Aug 10 17:04:34 2007 +0200
@@ -26,9 +26,9 @@
@{const_name NibbleC}, @{const_name NibbleD},
@{const_name NibbleE}, @{const_name NibbleF}];
in
- fold (fn target => CodegenSerializer.add_pretty_char target charr nibbles)
+ fold (fn target => CodeTarget.add_pretty_char target charr nibbles)
["SML", "OCaml", "Haskell"]
- #> CodegenSerializer.add_pretty_list_string "Haskell"
+ #> CodeTarget.add_pretty_list_string "Haskell"
@{const_name Nil} @{const_name Cons} charr nibbles
end
*}
--- a/src/HOL/Library/Pretty_Int.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Library/Pretty_Int.thy Fri Aug 10 17:04:34 2007 +0200
@@ -24,7 +24,7 @@
(Haskell -)
setup {*
- fold (fn target => CodegenSerializer.add_pretty_numeral target true
+ fold (fn target => CodeTarget.add_pretty_numeral target true
(@{const_name number_of}, @{typ "int \<Rightarrow> int"})
@{const_name Numeral.B0} @{const_name Numeral.B1}
@{const_name Numeral.Pls} @{const_name Numeral.Min}
--- a/src/HOL/List.thy Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/List.thy Fri Aug 10 17:04:34 2007 +0200
@@ -2815,7 +2815,7 @@
(Haskell "[]")
setup {*
- fold (fn target => CodegenSerializer.add_pretty_list target
+ fold (fn target => CodeTarget.add_pretty_list target
@{const_name Nil} @{const_name Cons}
) ["SML", "OCaml", "Haskell"]
*}
--- a/src/HOL/Tools/datatype_codegen.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/datatype_codegen.ML Fri Aug 10 17:04:34 2007 +0200
@@ -538,7 +538,7 @@
(* registering code types in code generator *)
fun codetype_hook dtspecs =
- fold (fn (dtco, (_, spec)) => CodegenData.add_datatype (dtco, spec)) dtspecs;
+ fold (fn (dtco, (_, spec)) => Code.add_datatype (dtco, spec)) dtspecs;
(* instrumentalizing the sort algebra *)
@@ -553,7 +553,7 @@
val algebra' = algebra
|> fold (fn (tyco, _) =>
Sorts.add_arities pp (tyco, map (fn class => (class, map snd vs)) sort)) css;
- fun typ_sort_inst ty = CodegenConsts.typ_sort_inst algebra' (Logic.varifyT ty, sort);
+ fun typ_sort_inst ty = CodeUnit.typ_sort_inst algebra' (Logic.varifyT ty, sort);
val venv = Vartab.empty
|> fold (fn (v, sort) => Vartab.update_new ((v, 0), sort)) vs
|> fold (fn (_, cs) => fold (fn (_, tys) => fold typ_sort_inst tys) cs) css;
@@ -585,7 +585,7 @@
|> not (null arities) ? (
f arities css
#-> (fn defs =>
- ClassPackage.prove_instance_arity tac arities defs
+ Class.prove_instance_arity tac arities defs
#> after_qed arities css))
end;
@@ -600,10 +600,10 @@
val const = ("op =", SOME dtco);
val get_thms = (fn () => get_eq (Theory.deref thy_ref) dtco |> rev);
in
- CodegenData.add_funcl (const, CodegenData.lazy_thms get_thms) thy
+ Code.add_funcl (const, Susp.delay get_thms) thy
end;
in
- prove_codetypes_arities (ClassPackage.intro_classes_tac [])
+ prove_codetypes_arities (Class.intro_classes_tac [])
(map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
[HOLogic.class_eq] ((K o K o pair) []) ((K o K) (fold add_eq_thms specs))
end;
@@ -616,14 +616,14 @@
let
val {case_name, index, descr, ...} = DatatypePackage.the_datatype thy dtco;
in
- CodegenPackage.add_appconst (case_name, CodegenPackage.appgen_case dest_case_expr) thy
+ CodePackage.add_appconst (case_name, CodePackage.appgen_case dest_case_expr) thy
end;
fun add_datatype_case_defs dtco thy =
let
val {case_rewrites, ...} = DatatypePackage.the_datatype thy dtco
in
- fold_rev (CodegenData.add_func true) case_rewrites thy
+ fold_rev (Code.add_func true) case_rewrites thy
end;
val setup =
--- a/src/HOL/Tools/inductive_codegen.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/inductive_codegen.ML Fri Aug 10 17:04:34 2007 +0200
@@ -697,7 +697,7 @@
val setup =
add_codegen "inductive" inductive_codegen #>
- add_attribute "ind" (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
+ Code.add_attribute ("ind", Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
Scan.option (Args.$$$ "params" |-- Args.colon |-- Args.nat) >> uncurry add);
end;
--- a/src/HOL/Tools/inductive_set_package.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/inductive_set_package.ML Fri Aug 10 17:04:34 2007 +0200
@@ -531,8 +531,8 @@
"convert rule to set notation"),
("to_pred", Attrib.syntax (Attrib.thms >> to_pred_att),
"convert rule to predicate notation")] #>
- Codegen.add_attribute "ind_set"
- (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> code_ind_att) #>
+ Code.add_attribute ("ind_set",
+ Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> code_ind_att) #>
Codegen.add_preprocessor codegen_preproc #>
Attrib.add_attributes [("mono_set", Attrib.add_del_args mono_add_att mono_del_att,
"declaration of monotonicity rule for set operators")] #>
--- a/src/HOL/Tools/recfun_codegen.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/recfun_codegen.ML Fri Aug 10 17:04:34 2007 +0200
@@ -44,7 +44,7 @@
handle TERM _ => tap (fn _ => warn thm);
in
Thm.declaration_attribute (fn thm => Context.mapping
- (add thm #> CodegenData.add_func false thm) I)
+ (add thm #> Code.add_func false thm) I)
end;
fun del_thm thm thy =
@@ -58,7 +58,7 @@
end handle TERM _ => (warn thm; thy);
val del = Thm.declaration_attribute
- (fn thm => Context.mapping (del_thm thm #> CodegenData.del_func thm) I)
+ (fn thm => Context.mapping (del_thm thm #> Code.del_func thm) I)
fun del_redundant thy eqs [] = eqs
| del_redundant thy eqs (eq :: eqs') =
@@ -170,9 +170,8 @@
val setup =
- add_codegen "recfun" recfun_codegen #>
- add_attribute ""
- (Args.del |-- Scan.succeed del
+ add_codegen "recfun" recfun_codegen
+ #> Code.add_attribute ("", Args.del |-- Scan.succeed del
|| Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add);
end;
--- a/src/HOL/Tools/record_package.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/record_package.ML Fri Aug 10 17:04:34 2007 +0200
@@ -1512,8 +1512,8 @@
||>> PureThy.add_defs_i false (map Thm.no_attributes (ext_spec::dest_specs))
||>> PureThy.add_defs_i false (map Thm.no_attributes upd_specs)
|-> (fn args as ((_, dest_defs), upd_defs) =>
- fold (CodegenData.add_func false) dest_defs
- #> fold (CodegenData.add_func false) upd_defs
+ fold (Code.add_func false) dest_defs
+ #> fold (Code.add_func false) upd_defs
#> pair args);
val ((([abs_inject, abs_inverse, abs_induct], ext_def :: dest_defs), upd_defs), defs_thy) =
timeit_msg "record extension type/selector/update defs:" mk_defs;
@@ -1916,9 +1916,9 @@
||>> ((PureThy.add_defs_i false o map Thm.no_attributes)
[make_spec, fields_spec, extend_spec, truncate_spec])
|-> (fn defs as ((sel_defs, upd_defs), derived_defs) =>
- fold (CodegenData.add_func false) sel_defs
- #> fold (CodegenData.add_func false) upd_defs
- #> fold (CodegenData.add_func false) derived_defs
+ fold (Code.add_func false) sel_defs
+ #> fold (Code.add_func false) upd_defs
+ #> fold (Code.add_func false) derived_defs
#> pair defs)
val (((sel_defs, upd_defs), derived_defs), defs_thy) =
timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:"
--- a/src/HOL/Tools/typecopy_package.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/HOL/Tools/typecopy_package.ML Fri Aug 10 17:04:34 2007 +0200
@@ -141,8 +141,7 @@
(* hook for projection function code *)
-fun add_project (_ , {proj_def, ...} : info) =
- CodegenData.add_func true proj_def;
+fun add_project (_ , {proj_def, ...} : info) = Code.add_func true proj_def;
val setup = add_hook add_project;
--- a/src/Pure/IsaMakefile Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/IsaMakefile Fri Aug 10 17:04:34 2007 +0200
@@ -31,7 +31,8 @@
General/scan.ML General/secure.ML General/seq.ML General/source.ML \
General/stack.ML General/susp.ML General/symbol.ML General/table.ML \
General/url.ML Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML Isar/attrib.ML \
- Isar/auto_bind.ML Isar/calculation.ML Isar/constdefs.ML Isar/context_rules.ML \
+ Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML \
+ Isar/code_unit.ML Isar/code.ML Isar/constdefs.ML Isar/context_rules.ML \
Isar/element.ML Isar/find_theorems.ML Isar/induct_attrib.ML Isar/isar_cmd.ML \
Isar/isar_syn.ML Isar/local_defs.ML Isar/local_syntax.ML Isar/local_theory.ML \
Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/object_logic.ML \
@@ -40,6 +41,12 @@
Isar/proof_display.ML Isar/proof_history.ML Isar/rule_cases.ML \
Isar/rule_insts.ML Isar/session.ML Isar/skip_proof.ML Isar/spec_parse.ML \
Isar/specification.ML Isar/theory_target.ML Isar/toplevel.ML \
+ $(SRC)/Pure/codegen.ML \
+ $(SRC)/Tools/code/code_funcgr.ML \
+ $(SRC)/Tools/code/code_thingol.ML \
+ $(SRC)/Tools/code/code_target.ML \
+ $(SRC)/Tools/code/code_name.ML \
+ $(SRC)/Tools/code/code_package.ML \
ML-Systems/alice.ML ML-Systems/exn.ML ML-Systems/multithreading_dummy.ML \
ML-Systems/multithreading_polyml.ML ML-Systems/overloading_smlnj.ML \
ML-Systems/polyml-4.1.3.ML \
@@ -59,12 +66,10 @@
Syntax/syntax.ML Syntax/type_ext.ML Thy/html.ML Thy/latex.ML \
Thy/ml_context.ML Thy/present.ML Thy/term_style.ML Thy/thm_database.ML \
Thy/thm_deps.ML Thy/thy_edit.ML Thy/thy_header.ML Thy/thy_info.ML \
- Thy/thy_load.ML Thy/thy_output.ML Tools/ROOT.ML Tools/class_package.ML \
- Tools/codegen_consts.ML Tools/codegen_data.ML Tools/codegen_func.ML \
- Tools/codegen_funcgr.ML Tools/codegen_names.ML Tools/codegen_package.ML \
- Tools/codegen_serializer.ML Tools/codegen_thingol.ML Tools/invoke.ML \
+ Thy/thy_load.ML Thy/thy_output.ML Tools/ROOT.ML \
+ Tools/invoke.ML \
Tools/named_thms.ML \
- Tools/xml.ML Tools/xml_syntax.ML assumption.ML axclass.ML codegen.ML \
+ Tools/xml.ML Tools/xml_syntax.ML assumption.ML axclass.ML \
compress.ML config.ML conjunction.ML consts.ML context.ML context_position.ML \
conv.ML defs.ML display.ML drule.ML envir.ML fact_index.ML goal.ML library.ML \
logic.ML meta_simplifier.ML more_thm.ML morphism.ML name.ML net.ML \
--- a/src/Pure/Isar/ROOT.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Isar/ROOT.ML Fri Aug 10 17:04:34 2007 +0200
@@ -43,10 +43,9 @@
use "net_rules.ML";
use "induct_attrib.ML";
-(*code generator base*)
-use "../Tools/codegen_consts.ML";
-use "../Tools/codegen_func.ML";
-use "../Tools/codegen_data.ML";
+(*executable theory content*)
+use "code_unit.ML";
+use "code.ML";
(*derived theory and proof elements*)
use "local_theory.ML";
@@ -54,7 +53,7 @@
use "obtain.ML";
use "locale.ML";
use "spec_parse.ML";
-use "../Tools/class_package.ML";
+use "class.ML";
use "theory_target.ML";
use "specification.ML";
use "constdefs.ML";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/code.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,925 @@
+(* Title: Pure/Isar/code.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Abstract executable content of theory. Management of data dependent on
+executable content.
+*)
+
+signature CODE =
+sig
+ val add_func: bool -> thm -> theory -> theory
+ val del_func: thm -> theory -> theory
+ val add_funcl: CodeUnit.const * thm list Susp.T -> theory -> theory
+ val add_func_attr: bool -> Attrib.src
+ val add_inline: thm -> theory -> theory
+ val del_inline: thm -> theory -> theory
+ val add_inline_proc: string * (theory -> cterm list -> thm list) -> theory -> theory
+ val del_inline_proc: string -> theory -> theory
+ val add_preproc: string * (theory -> thm list -> thm list) -> theory -> theory
+ val del_preproc: string -> theory -> theory
+ val add_post: thm -> theory -> theory
+ val del_post: thm -> theory -> theory
+ val add_datatype: string * ((string * sort) list * (string * typ list) list)
+ -> theory -> theory
+ val add_datatype_consts: CodeUnit.const list -> theory -> theory
+ val add_datatype_consts_cmd: string list -> theory -> theory
+
+ val coregular_algebra: theory -> Sorts.algebra
+ val operational_algebra: theory -> (sort -> sort) * Sorts.algebra
+ val these_funcs: theory -> CodeUnit.const -> thm list
+ val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
+ val get_datatype_of_constr: theory -> CodeUnit.const -> string option
+ val default_typ: theory -> CodeUnit.const -> typ
+
+ val preprocess_conv: cterm -> thm
+ val postprocess_conv: cterm -> thm
+
+ val add_attribute: string * (Args.T list -> attribute * Args.T list) -> theory -> theory
+
+ val print_codesetup: theory -> unit
+end;
+
+signature CODE_DATA_ARGS =
+sig
+ type T
+ val empty: T
+ val merge: Pretty.pp -> T * T -> T
+ val purge: theory option -> CodeUnit.const list option -> T -> T
+end;
+
+signature CODE_DATA =
+sig
+ type T
+ val get: theory -> T
+ val change: theory -> (T -> T) -> T
+ val change_yield: theory -> (T -> 'a * T) -> 'a * T
+end;
+
+signature PRIVATE_CODE =
+sig
+ include CODE
+ val declare_data: Object.T -> (Pretty.pp -> Object.T * Object.T -> Object.T)
+ -> (theory option -> CodeUnit.const list option -> Object.T -> Object.T) -> serial
+ val get_data: serial * ('a -> Object.T) * (Object.T -> 'a)
+ -> theory -> 'a
+ val change_data: serial * ('a -> Object.T) * (Object.T -> 'a)
+ -> theory -> ('a -> 'a) -> 'a
+ val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a)
+ -> theory -> ('a -> 'b * 'a) -> 'b * 'a
+end;
+
+structure Code : PRIVATE_CODE =
+struct
+
+(** preliminaries **)
+
+structure Consttab = CodeUnit.Consttab;
+
+
+(* certificate theorems *)
+
+fun string_of_lthms r = case Susp.peek r
+ of SOME thms => (map string_of_thm o rev) thms
+ | NONE => ["[...]"];
+
+fun pretty_lthms ctxt r = case Susp.peek r
+ of SOME thms => map (ProofContext.pretty_thm ctxt) thms
+ | NONE => [Pretty.str "[...]"];
+
+fun certificate thy f r =
+ case Susp.peek r
+ of SOME thms => (Susp.value o f thy) thms
+ | NONE => let
+ val thy_ref = Theory.check_thy thy;
+ in Susp.delay (fn () => (f (Theory.deref thy_ref) o Susp.force) r) end;
+
+fun merge' _ ([], []) = (false, [])
+ | merge' _ ([], ys) = (true, ys)
+ | merge' eq (xs, ys) = fold_rev
+ (fn y => fn (t, xs) => (t orelse not (member eq xs y), insert eq y xs)) ys (false, xs);
+
+fun merge_alist eq_key eq (xys as (xs, ys)) =
+ if eq_list (eq_pair eq_key eq) (xs, ys)
+ then (false, xs)
+ else (true, AList.merge eq_key eq xys);
+
+val merge_thms = merge' Thm.eq_thm_prop;
+
+fun merge_lthms (r1, r2) =
+ if Susp.same (r1, r2)
+ then (false, r1)
+ else case Susp.peek r1
+ of SOME [] => (true, r2)
+ | _ => case Susp.peek r2
+ of SOME [] => (true, r1)
+ | _ => (apsnd (Susp.delay o K)) (merge_thms (Susp.force r1, Susp.force r2));
+
+
+(* pairs of (selected, deleted) defining equations *)
+
+type sdthms = thm list Susp.T * thm list;
+
+fun add_drop_redundant thm (sels, dels) =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val args_of = snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+ val args = args_of thm;
+ fun matches [] _ = true
+ | matches (Var _ :: xs) [] = matches xs []
+ | matches (_ :: _) [] = false
+ | matches (x :: xs) (y :: ys) = Pattern.matches thy (x, y) andalso matches xs ys;
+ fun drop thm' = not (matches args (args_of thm'))
+ orelse (warning ("code generator: dropping redundant defining equation\n" ^ string_of_thm thm'); false);
+ val (keeps, drops) = List.partition drop sels;
+ in (thm :: keeps, dels |> remove Thm.eq_thm_prop thm |> fold (insert Thm.eq_thm_prop) drops) end;
+
+fun add_thm thm (sels, dels) =
+ apfst Susp.value (add_drop_redundant thm (Susp.force sels, dels));
+
+fun add_lthms lthms (sels, []) =
+ (Susp.delay (fn () => fold add_drop_redundant
+ (Susp.force lthms) (Susp.force sels, []) |> fst), [])
+ (*FIXME*)
+ | add_lthms lthms (sels, dels) =
+ fold add_thm (Susp.force lthms) (sels, dels);
+
+fun del_thm thm (sels, dels) =
+ (Susp.value (remove Thm.eq_thm_prop thm (Susp.force sels)), thm :: dels);
+
+fun pretty_sdthms ctxt (sels, _) = pretty_lthms ctxt sels;
+
+fun merge_sdthms ((sels1, dels1), (sels2, dels2)) =
+ let
+ val (dels_t, dels) = merge_thms (dels1, dels2);
+ in if dels_t
+ then let
+ val (_, sels) = merge_thms
+ (subtract Thm.eq_thm_prop dels2 (Susp.force sels1), Susp.force sels2);
+ val (_, dels) = merge_thms
+ (subtract Thm.eq_thm_prop (Susp.force sels2) dels1, dels2);
+ in (true, ((Susp.delay o K) sels, dels)) end
+ else let
+ val (sels_t, sels) = merge_lthms (sels1, sels2);
+ in (sels_t, (sels, dels)) end
+ end;
+
+
+(* code attributes *)
+
+structure CodeAttr = TheoryDataFun (
+ type T = (string * (Args.T list -> attribute * Args.T list)) list;
+ val empty = [];
+ val copy = I;
+ val extend = I;
+ fun merge _ = AList.merge (op =) (K true);
+);
+
+fun add_attribute (attr as (name, _)) =
+ let
+ fun add_parser ("", parser) attrs = attrs @ [("", parser)]
+ | add_parser (name, parser) attrs = (name, Args.$$$ name |-- parser) :: attrs;
+ fun error "" = error ("Code attribute already declared")
+ | error name = error ("Code attribute " ^ name ^ " already declared")
+ in CodeAttr.map (fn attrs => if AList.defined (op =) attrs name
+ then error name else add_parser attr attrs)
+ end;
+
+val _ =
+ let
+ val code_attr = Attrib.syntax (Scan.peek (fn context =>
+ List.foldr op || Scan.fail (map snd (CodeAttr.get (Context.theory_of context)))));
+ in
+ Context.add_setup (Attrib.add_attributes
+ [("code", code_attr, "declare theorems for code generation")])
+ end;
+
+
+
+(** exeuctable content **)
+
+datatype thmproc = Preproc of {
+ inlines: thm list,
+ inline_procs: (string * (serial * (theory -> cterm list -> thm list))) list,
+ preprocs: (string * (serial * (theory -> thm list -> thm list))) list,
+ posts: thm list
+};
+
+fun mk_thmproc (((inlines, inline_procs), preprocs), posts) =
+ Preproc { inlines = inlines, inline_procs = inline_procs, preprocs = preprocs,
+ posts = posts };
+fun map_thmproc f (Preproc { inlines, inline_procs, preprocs, posts }) =
+ mk_thmproc (f (((inlines, inline_procs), preprocs), posts));
+fun merge_thmproc (Preproc { inlines = inlines1, inline_procs = inline_procs1,
+ preprocs = preprocs1, posts = posts1 },
+ Preproc { inlines = inlines2, inline_procs = inline_procs2,
+ preprocs = preprocs2, posts= posts2 }) =
+ let
+ val (touched1, inlines) = merge_thms (inlines1, inlines2);
+ val (touched2, inline_procs) = merge_alist (op =) (eq_fst (op =)) (inline_procs1, inline_procs2);
+ val (touched3, preprocs) = merge_alist (op =) (eq_fst (op =)) (preprocs1, preprocs2);
+ val (_, posts) = merge_thms (posts1, posts2);
+ in (touched1 orelse touched2 orelse touched3,
+ mk_thmproc (((inlines, inline_procs), preprocs), posts)) end;
+
+fun join_func_thms (tabs as (tab1, tab2)) =
+ let
+ val cs1 = Consttab.keys tab1;
+ val cs2 = Consttab.keys tab2;
+ val cs' = filter (member CodeUnit.eq_const cs2) cs1;
+ val cs'' = subtract (op =) cs' cs1 @ subtract (op =) cs' cs2;
+ val cs''' = ref [] : CodeUnit.const list ref;
+ fun merge c x = let val (touched, thms') = merge_sdthms x in
+ (if touched then cs''' := cons c (!cs''') else (); thms') end;
+ in (cs'' @ !cs''', Consttab.join merge tabs) end;
+fun merge_funcs (thms1, thms2) =
+ let
+ val (consts, thms) = join_func_thms (thms1, thms2);
+ in (SOME consts, thms) end;
+
+val eq_string = op = : string * string -> bool;
+val eq_co = op = : (string * typ list) * (string * typ list) -> bool;
+fun eq_dtyp ((vs1, cs1), (vs2, cs2)) =
+ gen_eq_set (eq_pair eq_string (gen_eq_set eq_string)) (vs1, vs2)
+ andalso gen_eq_set eq_co (cs1, cs2);
+fun merge_dtyps (tabs as (tab1, tab2)) =
+ let
+ val tycos1 = Symtab.keys tab1;
+ val tycos2 = Symtab.keys tab2;
+ val tycos' = filter (member eq_string tycos2) tycos1;
+ val new_types = not (gen_eq_set (op =) (tycos1, tycos2));
+ val diff_types = not (gen_eq_set (eq_pair (op =) eq_dtyp)
+ (AList.make (the o Symtab.lookup tab1) tycos',
+ AList.make (the o Symtab.lookup tab2) tycos'));
+ fun join _ (cos as (_, cos2)) = if eq_dtyp cos
+ then raise Symtab.SAME else cos2;
+ in ((new_types, diff_types), Symtab.join join tabs) end;
+
+datatype spec = Spec of {
+ funcs: sdthms Consttab.table,
+ dtyps: ((string * sort) list * (string * typ list) list) Symtab.table
+};
+
+fun mk_spec (funcs, dtyps) =
+ Spec { funcs = funcs, dtyps = dtyps };
+fun map_spec f (Spec { funcs = funcs, dtyps = dtyps }) =
+ mk_spec (f (funcs, dtyps));
+fun merge_spec (Spec { funcs = funcs1, dtyps = dtyps1 },
+ Spec { funcs = funcs2, dtyps = dtyps2 }) =
+ let
+ val (touched_cs, funcs) = merge_funcs (funcs1, funcs2);
+ val ((new_types, diff_types), dtyps) = merge_dtyps (dtyps1, dtyps2);
+ val touched = if new_types orelse diff_types then NONE else touched_cs;
+ in (touched, mk_spec (funcs, dtyps)) end;
+
+datatype exec = Exec of {
+ thmproc: thmproc,
+ spec: spec
+};
+
+fun mk_exec (thmproc, spec) =
+ Exec { thmproc = thmproc, spec = spec };
+fun map_exec f (Exec { thmproc = thmproc, spec = spec }) =
+ mk_exec (f (thmproc, spec));
+fun merge_exec (Exec { thmproc = thmproc1, spec = spec1 },
+ Exec { thmproc = thmproc2, spec = spec2 }) =
+ let
+ val (touched', thmproc) = merge_thmproc (thmproc1, thmproc2);
+ val (touched_cs, spec) = merge_spec (spec1, spec2);
+ val touched = if touched' then NONE else touched_cs;
+ in (touched, mk_exec (thmproc, spec)) end;
+val empty_exec = mk_exec (mk_thmproc ((([], []), []), []),
+ mk_spec (Consttab.empty, Symtab.empty));
+
+fun the_thmproc (Exec { thmproc = Preproc x, ...}) = x;
+fun the_spec (Exec { spec = Spec x, ...}) = x;
+val the_funcs = #funcs o the_spec;
+val the_dtyps = #dtyps o the_spec;
+val map_thmproc = map_exec o apfst o map_thmproc;
+val map_funcs = map_exec o apsnd o map_spec o apfst;
+val map_dtyps = map_exec o apsnd o map_spec o apsnd;
+
+
+(* data slots dependent on executable content *)
+
+(*private copy avoids potential conflict of table exceptions*)
+structure Datatab = TableFun(type key = int val ord = int_ord);
+
+local
+
+type kind = {
+ empty: Object.T,
+ merge: Pretty.pp -> Object.T * Object.T -> Object.T,
+ purge: theory option -> CodeUnit.const list option -> Object.T -> Object.T
+};
+
+val kinds = ref (Datatab.empty: kind Datatab.table);
+val kind_keys = ref ([]: serial list);
+
+fun invoke f k = case Datatab.lookup (! kinds) k
+ of SOME kind => f kind
+ | NONE => sys_error "Invalid code data identifier";
+
+in
+
+fun declare_data empty merge purge =
+ let
+ val k = serial ();
+ val kind = {empty = empty, merge = merge, purge = purge};
+ val _ = change kinds (Datatab.update (k, kind));
+ val _ = change kind_keys (cons k);
+ in k end;
+
+fun invoke_empty k = invoke (fn kind => #empty kind) k;
+
+fun invoke_merge_all pp = Datatab.join
+ (invoke (fn kind => #merge kind pp));
+
+fun invoke_purge_all thy_opt cs =
+ fold (fn k => Datatab.map_entry k
+ (invoke (fn kind => #purge kind thy_opt cs) k)) (! kind_keys);
+
+end; (*local*)
+
+
+(* theory store *)
+
+local
+
+type data = Object.T Datatab.table;
+
+structure CodeData = TheoryDataFun
+(
+ type T = exec * data ref;
+ val empty = (empty_exec, ref Datatab.empty : data ref);
+ fun copy (exec, data) = (exec, ref (! data));
+ val extend = copy;
+ fun merge pp ((exec1, data1), (exec2, data2)) =
+ let
+ val (touched, exec) = merge_exec (exec1, exec2);
+ val data1' = invoke_purge_all NONE touched (! data1);
+ val data2' = invoke_purge_all NONE touched (! data2);
+ val data = invoke_merge_all pp (data1', data2');
+ in (exec, ref data) end;
+);
+
+val _ = Context.add_setup CodeData.init;
+
+fun ch r f = let val x = f (! r) in (r := x; x) end;
+fun thy_data f thy = f ((snd o CodeData.get) thy);
+
+fun get_ensure_init kind data_ref =
+ case Datatab.lookup (! data_ref) kind
+ of SOME x => x
+ | NONE => let val y = invoke_empty kind
+ in (change data_ref (Datatab.update (kind, y)); y) end;
+
+in
+
+(* access to executable content *)
+
+val get_exec = fst o CodeData.get;
+
+fun map_exec_purge touched f thy =
+ CodeData.map (fn (exec, data) =>
+ (f exec, ref (invoke_purge_all (SOME thy) touched (! data)))) thy;
+
+
+(* access to data dependent on abstract executable content *)
+
+fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest);
+
+fun change_data (kind, mk, dest) =
+ let
+ fun chnge data_ref f =
+ let
+ val data = get_ensure_init kind data_ref;
+ val data' = f (dest data);
+ in (change data_ref (Datatab.update (kind, mk data')); data') end;
+ in thy_data chnge end;
+
+fun change_yield_data (kind, mk, dest) =
+ let
+ fun chnge data_ref f =
+ let
+ val data = get_ensure_init kind data_ref;
+ val (x, data') = f (dest data);
+ in (x, (change data_ref (Datatab.update (kind, mk data')); data')) end;
+ in thy_data chnge end;
+
+end; (*local*)
+
+
+(* print executable content *)
+
+fun print_codesetup thy =
+ let
+ val ctxt = ProofContext.init thy;
+ val exec = get_exec thy;
+ fun pretty_func (s, lthms) =
+ (Pretty.block o Pretty.fbreaks) (
+ Pretty.str s :: pretty_sdthms ctxt lthms
+ );
+ fun pretty_dtyp (s, []) =
+ Pretty.str s
+ | pretty_dtyp (s, cos) =
+ (Pretty.block o Pretty.breaks) (
+ Pretty.str s
+ :: Pretty.str "="
+ :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str c
+ | (c, tys) =>
+ (Pretty.block o Pretty.breaks)
+ (Pretty.str c :: Pretty.str "of" :: map (Pretty.quote o Sign.pretty_typ thy) tys)) cos)
+ );
+ val inlines = (#inlines o the_thmproc) exec;
+ val inline_procs = (map fst o #inline_procs o the_thmproc) exec;
+ val preprocs = (map fst o #preprocs o the_thmproc) exec;
+ val funs = the_funcs exec
+ |> Consttab.dest
+ |> (map o apfst) (CodeUnit.string_of_const thy)
+ |> sort (string_ord o pairself fst);
+ val dtyps = the_dtyps exec
+ |> Symtab.dest
+ |> map (fn (dtco, (vs, cos)) => (Sign.string_of_typ thy (Type (dtco, map TFree vs)), cos))
+ |> sort (string_ord o pairself fst)
+ in
+ (Pretty.writeln o Pretty.chunks) [
+ Pretty.block (
+ Pretty.str "defining equations:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map pretty_func) funs
+ ),
+ Pretty.block (
+ Pretty.str "inlining theorems:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map (ProofContext.pretty_thm ctxt)) inlines
+ ),
+ Pretty.block (
+ Pretty.str "inlining procedures:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map Pretty.str) inline_procs
+ ),
+ Pretty.block (
+ Pretty.str "preprocessors:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map Pretty.str) preprocs
+ ),
+ Pretty.block (
+ Pretty.str "datatypes:"
+ :: Pretty.fbrk
+ :: (Pretty.fbreaks o map pretty_dtyp) dtyps
+ )
+ ]
+ end;
+
+
+
+(** theorem transformation and certification **)
+
+fun common_typ_funcs [] = []
+ | common_typ_funcs [thm] = [thm]
+ | common_typ_funcs (thms as thm :: _) =
+ let
+ val thy = Thm.theory_of_thm thm;
+ fun incr_thm thm max =
+ let
+ val thm' = incr_indexes max thm;
+ val max' = Thm.maxidx_of thm' + 1;
+ in (thm', max') end;
+ val (thms', maxidx) = fold_map incr_thm thms 0;
+ val ty1 :: tys = map (snd o CodeUnit.head_func) thms';
+ fun unify ty env = Sign.typ_unify thy (ty1, ty) env
+ handle Type.TUNIFY =>
+ error ("Type unificaton failed, while unifying defining equations\n"
+ ^ (cat_lines o map Display.string_of_thm) thms
+ ^ "\nwith types\n"
+ ^ (cat_lines o map (CodeUnit.string_of_typ thy)) (ty1 :: tys));
+ val (env, _) = fold unify tys (Vartab.empty, maxidx)
+ val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
+ cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
+ in map (Thm.instantiate (instT, [])) thms' end;
+
+fun certify_const thy const thms =
+ let
+ fun cert thm = if CodeUnit.eq_const (const, fst (CodeUnit.head_func thm))
+ then thm else error ("Wrong head of defining equation,\nexpected constant "
+ ^ CodeUnit.string_of_const thy const ^ "\n" ^ string_of_thm thm)
+ in map cert thms end;
+
+
+
+(** operational sort algebra and class discipline **)
+
+local
+
+fun aggr_neutr f y [] = y
+ | aggr_neutr f y (x::xs) = aggr_neutr f (f y x) xs;
+
+fun aggregate f [] = NONE
+ | aggregate f (x::xs) = SOME (aggr_neutr f x xs);
+
+fun inter_sorts thy =
+ let
+ val algebra = Sign.classes_of thy;
+ val inters = curry (Sorts.inter_sort algebra);
+ in aggregate (map2 inters) end;
+
+fun specific_constraints thy (class, tyco) =
+ let
+ val vs = Name.invents Name.context "" (Sign.arity_number thy tyco);
+ val clsops = (these o Option.map snd o try (AxClass.params_of_class thy)) class;
+ val funcs = clsops
+ |> map (fn (clsop, _) => (clsop, SOME tyco))
+ |> map (Consttab.lookup ((the_funcs o get_exec) thy))
+ |> (map o Option.map) (Susp.force o fst)
+ |> maps these
+ |> map (Thm.transfer thy);
+ val sorts = map (map (snd o dest_TVar) o snd o dest_Type o the_single
+ o Sign.const_typargs thy o (fn ((c, _), ty) => (c, ty)) o CodeUnit.head_func) funcs;
+ in sorts end;
+
+fun weakest_constraints thy (class, tyco) =
+ let
+ val all_superclasses = class :: Graph.all_succs ((#classes o Sorts.rep_algebra o Sign.classes_of) thy) [class];
+ in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) all_superclasses)
+ of SOME sorts => sorts
+ | NONE => Sign.arity_sorts thy tyco [class]
+ end;
+
+fun strongest_constraints thy (class, tyco) =
+ let
+ val algebra = Sign.classes_of thy;
+ val all_subclasses = class :: Graph.all_preds ((#classes o Sorts.rep_algebra) algebra) [class];
+ val inst_subclasses = filter (can (Sorts.mg_domain algebra tyco) o single) all_subclasses;
+ in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) inst_subclasses)
+ of SOME sorts => sorts
+ | NONE => replicate
+ (Sign.arity_number thy tyco) (Sign.certify_sort thy (Sign.all_classes thy))
+ end;
+
+fun gen_classop_typ constr thy class (c, tyco) =
+ let
+ val (var, cs) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
+ val ty = (the o AList.lookup (op =) cs) c;
+ val sort_args = Name.names (Name.declare var Name.context) "'a"
+ (constr thy (class, tyco));
+ val ty_inst = Type (tyco, map TFree sort_args);
+ in Logic.varifyT (map_type_tfree (K ty_inst) ty) end;
+
+fun retrieve_algebra thy operational =
+ Sorts.subalgebra (Sign.pp thy) operational
+ (weakest_constraints thy)
+ (Sign.classes_of thy);
+
+in
+
+fun coregular_algebra thy = retrieve_algebra thy (K true) |> snd;
+fun operational_algebra thy =
+ let
+ fun add_iff_operational class =
+ can (AxClass.get_definition thy) class ? cons class;
+ val operational_classes = fold add_iff_operational (Sign.all_classes thy) []
+ in retrieve_algebra thy (member (op =) operational_classes) end;
+
+val classop_weakest_typ = gen_classop_typ weakest_constraints;
+val classop_strongest_typ = gen_classop_typ strongest_constraints;
+
+fun assert_func_typ thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ fun check_typ_classop class (const as (c, SOME tyco), thm) =
+ let
+ val (_, ty) = CodeUnit.head_func thm;
+ val ty_decl = classop_weakest_typ thy class (c, tyco);
+ val ty_strongest = classop_strongest_typ thy class (c, tyco);
+ fun constrain thm =
+ let
+ val max = Thm.maxidx_of thm + 1;
+ val ty_decl' = Logic.incr_tvar max ty_decl;
+ val (_, ty') = CodeUnit.head_func thm;
+ val (env, _) = Sign.typ_unify thy (ty_decl', ty') (Vartab.empty, max);
+ val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
+ cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
+ in Thm.instantiate (instT, []) thm end;
+ in if Sign.typ_instance thy (ty_strongest, ty)
+ then if Sign.typ_instance thy (ty, ty_decl)
+ then thm
+ else (warning ("Constraining type\n" ^ CodeUnit.string_of_typ thy ty
+ ^ "\nof defining equation\n"
+ ^ string_of_thm thm
+ ^ "\nto permitted most general type\n"
+ ^ CodeUnit.string_of_typ thy ty_decl);
+ constrain thm)
+ else CodeUnit.bad_thm ("Type\n" ^ CodeUnit.string_of_typ thy ty
+ ^ "\nof defining equation\n"
+ ^ string_of_thm thm
+ ^ "\nis incompatible with permitted least general type\n"
+ ^ CodeUnit.string_of_typ thy ty_strongest)
+ end
+ | check_typ_classop class ((c, NONE), thm) =
+ CodeUnit.bad_thm ("Illegal type for class operation " ^ quote c
+ ^ "\nin defining equation\n"
+ ^ string_of_thm thm);
+ fun check_typ_fun (const as (c, _), thm) =
+ let
+ val (_, ty) = CodeUnit.head_func thm;
+ val ty_decl = Sign.the_const_type thy c;
+ in if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
+ then thm
+ else CodeUnit.bad_thm ("Type\n" ^ CodeUnit.string_of_typ thy ty
+ ^ "\nof defining equation\n"
+ ^ string_of_thm thm
+ ^ "\nis incompatible with declared function type\n"
+ ^ CodeUnit.string_of_typ thy ty_decl)
+ end;
+ fun check_typ (const as (c, _), thm) =
+ case AxClass.class_of_param thy c
+ of SOME class => check_typ_classop class (const, thm)
+ | NONE => check_typ_fun (const, thm);
+ in check_typ (fst (CodeUnit.head_func thm), thm) end;
+
+val mk_func = CodeUnit.error_thm
+ (assert_func_typ o CodeUnit.mk_func);
+val mk_func_liberal = CodeUnit.warning_thm
+ (assert_func_typ o CodeUnit.mk_func);
+
+end;
+
+
+
+(** interfaces and attributes **)
+
+fun add_func true thm thy =
+ let
+ val func = mk_func thm;
+ val (const, _) = CodeUnit.head_func func;
+ in map_exec_purge (SOME [const]) (map_funcs
+ (Consttab.map_default
+ (const, (Susp.value [], [])) (add_thm func))) thy
+ end
+ | add_func false thm thy =
+ case mk_func_liberal thm
+ of SOME func => let
+ val (const, _) = CodeUnit.head_func func
+ in map_exec_purge (SOME [const]) (map_funcs
+ (Consttab.map_default
+ (const, (Susp.value [], [])) (add_thm func))) thy
+ end
+ | NONE => thy;
+
+fun delete_force msg key xs =
+ if AList.defined (op =) xs key then AList.delete (op =) key xs
+ else error ("No such " ^ msg ^ ": " ^ quote key);
+
+fun del_func thm thy =
+ let
+ val func = mk_func thm;
+ val (const, _) = CodeUnit.head_func func;
+ in map_exec_purge (SOME [const]) (map_funcs
+ (Consttab.map_entry
+ const (del_thm func))) thy
+ end;
+
+fun add_funcl (const, lthms) thy =
+ let
+ val lthms' = certificate thy (fn thy => certify_const thy const) lthms;
+ (*FIXME must check compatibility with sort algebra;
+ alas, naive checking results in non-termination!*)
+ in
+ map_exec_purge (SOME [const]) (map_funcs (Consttab.map_default (const, (Susp.value [], []))
+ (add_lthms lthms'))) thy
+ end;
+
+fun add_func_attr strict = Attrib.internal (fn _ => Thm.declaration_attribute
+ (fn thm => Context.mapping (add_func strict thm) I));
+
+local
+
+fun del_datatype tyco thy =
+ case Symtab.lookup ((the_dtyps o get_exec) thy) tyco
+ of SOME (vs, cos) => let
+ val consts = CodeUnit.consts_of_cos thy tyco vs cos;
+ in map_exec_purge (if null consts then NONE else SOME consts)
+ (map_dtyps (Symtab.delete tyco)) thy end
+ | NONE => thy;
+
+in
+
+fun add_datatype (tyco, (vs_cos as (vs, cos))) thy =
+ let
+ val consts = CodeUnit.consts_of_cos thy tyco vs cos;
+ in
+ thy
+ |> del_datatype tyco
+ |> map_exec_purge (SOME consts) (map_dtyps (Symtab.update_new (tyco, vs_cos)))
+ end;
+
+fun add_datatype_consts consts thy =
+ add_datatype (CodeUnit.cos_of_consts thy consts) thy;
+
+fun add_datatype_consts_cmd raw_cs thy =
+ add_datatype_consts (map (CodeUnit.read_const thy) raw_cs) thy
+
+end; (*local*)
+
+fun add_inline thm thy =
+ (map_exec_purge NONE o map_thmproc o apfst o apfst o apfst)
+ (insert Thm.eq_thm_prop (CodeUnit.error_thm CodeUnit.mk_rew thm)) thy;
+ (*fully applied in order to get right context for mk_rew!*)
+
+fun del_inline thm thy =
+ (map_exec_purge NONE o map_thmproc o apfst o apfst o apfst)
+ (remove Thm.eq_thm_prop (CodeUnit.error_thm CodeUnit.mk_rew thm)) thy;
+ (*fully applied in order to get right context for mk_rew!*)
+
+fun add_inline_proc (name, f) =
+ (map_exec_purge NONE o map_thmproc o apfst o apfst o apsnd)
+ (AList.update (op =) (name, (serial (), f)));
+
+fun del_inline_proc name =
+ (map_exec_purge NONE o map_thmproc o apfst o apfst o apsnd)
+ (delete_force "inline procedure" name);
+
+fun add_preproc (name, f) =
+ (map_exec_purge NONE o map_thmproc o apfst o apsnd)
+ (AList.update (op =) (name, (serial (), f)));
+
+fun del_preproc name =
+ (map_exec_purge NONE o map_thmproc o apfst o apsnd)
+ (delete_force "preprocessor" name);
+
+fun add_post thm thy =
+ (map_exec_purge NONE o map_thmproc o apsnd)
+ (insert Thm.eq_thm_prop (CodeUnit.error_thm CodeUnit.mk_rew thm)) thy;
+ (*fully applied in order to get right context for mk_rew!*)
+
+fun del_post thm thy =
+ (map_exec_purge NONE o map_thmproc o apsnd)
+ (remove Thm.eq_thm_prop (CodeUnit.error_thm CodeUnit.mk_rew thm)) thy;
+ (*fully applied in order to get right context for mk_rew!*)
+
+val _ = Context.add_setup
+ (let
+ fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+ fun add_simple_attribute (name, f) =
+ add_attribute (name, Scan.succeed (mk_attribute f));
+ fun add_del_attribute (name, (add, del)) =
+ add_attribute (name, Args.del |-- Scan.succeed (mk_attribute del)
+ || Scan.succeed (mk_attribute add))
+ in
+ add_del_attribute ("func", (add_func true, del_func))
+ #> add_del_attribute ("inline", (add_inline, del_inline))
+ #> add_del_attribute ("post", (add_post, del_post))
+ end);
+
+
+(** post- and preprocessing **)
+
+local
+
+fun gen_apply_inline_proc prep post thy f x =
+ let
+ val cts = prep x;
+ val rews = map CodeUnit.assert_rew (f thy cts);
+ in post rews x end;
+
+val apply_inline_proc = gen_apply_inline_proc (maps
+ ((fn [args, rhs] => rhs :: (snd o Drule.strip_comb) args) o snd o Drule.strip_comb o Thm.cprop_of))
+ (fn rews => map (CodeUnit.rewrite_func rews));
+val apply_inline_proc_cterm = gen_apply_inline_proc single
+ (MetaSimplifier.rewrite false);
+
+fun apply_preproc thy f [] = []
+ | apply_preproc thy f (thms as (thm :: _)) =
+ let
+ val (const, _) = CodeUnit.head_func thm;
+ val thms' = f thy thms;
+ in certify_const thy const thms' end;
+
+fun rhs_conv conv thm =
+ let
+ val thm' = (conv o Thm.rhs_of) thm;
+ in Thm.transitive thm thm' end
+
+in
+
+fun preprocess thy thms =
+ thms
+ |> fold (fn (_, (_, f)) => apply_preproc thy f) ((#preprocs o the_thmproc o get_exec) thy)
+ |> map (CodeUnit.rewrite_func ((#inlines o the_thmproc o get_exec) thy))
+ |> fold (fn (_, (_, f)) => apply_inline_proc thy f) ((#inline_procs o the_thmproc o get_exec) thy)
+(*FIXME - must check: rewrite rule, defining equation, proper constant |> map (snd o check_func false thy) *)
+ |> common_typ_funcs;
+
+fun preprocess_conv ct =
+ let
+ val thy = Thm.theory_of_cterm ct;
+ in
+ ct
+ |> MetaSimplifier.rewrite false ((#inlines o the_thmproc o get_exec) thy)
+ |> fold (fn (_, (_, f)) => rhs_conv (apply_inline_proc_cterm thy f))
+ ((#inline_procs o the_thmproc o get_exec) thy)
+ end;
+
+fun postprocess_conv ct =
+ let
+ val thy = Thm.theory_of_cterm ct;
+ in
+ ct
+ |> MetaSimplifier.rewrite false ((#posts o the_thmproc o get_exec) thy)
+ end;
+
+end; (*local*)
+
+fun get_datatype thy tyco =
+ case Symtab.lookup ((the_dtyps o get_exec) thy) tyco
+ of SOME spec => spec
+ | NONE => Sign.arity_number thy tyco
+ |> Name.invents Name.context "'a"
+ |> map (rpair [])
+ |> rpair [];
+
+fun get_datatype_of_constr thy const =
+ case CodeUnit.co_of_const' thy const
+ of SOME (tyco, (_, co)) => if member eq_co
+ (Symtab.lookup (((the_dtyps o get_exec) thy)) tyco
+ |> Option.map snd
+ |> the_default []) co then SOME tyco else NONE
+ | NONE => NONE;
+
+fun get_constr_typ thy const =
+ case get_datatype_of_constr thy const
+ of SOME tyco => let
+ val (vs, cos) = get_datatype thy tyco;
+ val (_, (_, (co, tys))) = CodeUnit.co_of_const thy const
+ in (tys ---> Type (tyco, map TFree vs))
+ |> map_atyps (fn TFree (v, _) => TFree (v, AList.lookup (op =) vs v |> the))
+ |> Logic.varifyT
+ |> SOME end
+ | NONE => NONE;
+
+fun default_typ_proto thy (const as (c, SOME tyco)) = classop_weakest_typ thy
+ ((the o AxClass.class_of_param thy) c) (c, tyco) |> SOME
+ | default_typ_proto thy (const as (c, NONE)) = case AxClass.class_of_param thy c
+ of SOME class => SOME (Term.map_type_tvar
+ (K (TVar (("'a", 0), [class]))) (Sign.the_const_type thy c))
+ | NONE => get_constr_typ thy const;
+
+local
+
+fun get_funcs thy const =
+ Consttab.lookup ((the_funcs o get_exec) thy) const
+ |> Option.map (Susp.force o fst)
+ |> these
+ |> map (Thm.transfer thy);
+
+in
+
+fun these_funcs thy const =
+ let
+ fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals
+ o ObjectLogic.drop_judgment thy o Thm.plain_prop_of);
+ in
+ get_funcs thy const
+ |> preprocess thy
+ |> drop_refl thy
+ end;
+
+fun default_typ thy (const as (c, _)) = case default_typ_proto thy const
+ of SOME ty => ty
+ | NONE => (case get_funcs thy const
+ of thm :: _ => snd (CodeUnit.head_func thm)
+ | [] => Sign.the_const_type thy c);
+
+end; (*local*)
+
+end; (*struct*)
+
+
+(** type-safe interfaces for data depedent on executable content **)
+
+functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA =
+struct
+
+type T = Data.T;
+exception Data of T;
+fun dest (Data x) = x
+
+val kind = Code.declare_data (Data Data.empty)
+ (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)))
+ (fn thy_opt => fn cs => fn Data x => Data (Data.purge thy_opt cs x));
+
+val data_op = (kind, Data, dest);
+
+val get = Code.get_data data_op;
+val change = Code.change_data data_op;
+fun change_yield thy = Code.change_yield_data data_op thy;
+
+end;
+
+structure Code : CODE =
+struct
+
+open Code;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/code_unit.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,461 @@
+(* Title: Pure/Isar/code_unit.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Basic units of code generation: Identifying (possibly overloaded) constants
+by name plus optional type constructor. Convenient data structures for constants.
+Defining equations ("func"s). Auxiliary.
+*)
+
+signature CODE_UNIT =
+sig
+ type const = string * string option (*constant name, possibly instance*)
+ val const_ord: const * const -> order
+ val eq_const: const * const -> bool
+ structure Consttab: TABLE
+ val const_of_cexpr: theory -> string * typ -> const
+ val string_of_typ: theory -> typ -> string
+ val string_of_const: theory -> const -> string
+ val read_bare_const: theory -> string -> string * typ
+ val read_const: theory -> string -> const
+ val read_const_exprs: theory -> (const list -> const list)
+ -> string list -> bool * const list
+
+ val co_of_const: theory -> const
+ -> string * ((string * sort) list * (string * typ list))
+ val co_of_const': theory -> const
+ -> (string * ((string * sort) list * (string * typ list))) option
+ val cos_of_consts: theory -> const list
+ -> string * ((string * sort) list * (string * typ list) list)
+ val const_of_co: theory -> string -> (string * sort) list
+ -> string * typ list -> const
+ val consts_of_cos: theory -> string -> (string * sort) list
+ -> (string * typ list) list -> const list
+ val no_args: theory -> const -> int
+
+ val typargs: theory -> string * typ -> typ list
+ val typ_sort_inst: Sorts.algebra -> typ * sort
+ -> sort Vartab.table -> sort Vartab.table
+
+ val assert_rew: thm -> thm
+ val mk_rew: thm -> thm
+ val mk_func: thm -> thm
+ val head_func: thm -> const * typ
+ val bad_thm: string -> 'a
+ val error_thm: (thm -> thm) -> thm -> thm
+ val warning_thm: (thm -> thm) -> thm -> thm option
+
+ val inst_thm: sort Vartab.table -> thm -> thm
+ val expand_eta: int -> thm -> thm
+ val rewrite_func: thm list -> thm -> thm
+ val norm_args: thm list -> thm list
+ val norm_varnames: (string -> string) -> (string -> string) -> thm list -> thm list
+end;
+
+structure CodeUnit: CODE_UNIT =
+struct
+
+
+(* auxiliary *)
+
+exception BAD_THM of string;
+fun bad_thm msg = raise BAD_THM msg;
+fun error_thm f thm = f thm handle BAD_THM msg => error msg;
+fun warning_thm f thm = SOME (f thm) handle BAD_THM msg
+ => (warning ("code generator: " ^ msg); NONE);
+
+
+(* basic data structures *)
+
+type const = string * string option;
+val const_ord = prod_ord fast_string_ord (option_ord string_ord);
+val eq_const = is_equal o const_ord;
+
+structure Consttab =
+ TableFun(
+ type key = const;
+ val ord = const_ord;
+ );
+
+fun string_of_typ thy = setmp show_sorts true (Sign.string_of_typ thy);
+
+
+(* conversion between constant expressions and constants *)
+
+fun const_of_cexpr thy (c_ty as (c, _)) =
+ case AxClass.class_of_param thy c
+ of SOME class => (case Sign.const_typargs thy c_ty
+ of [Type (tyco, _)] => if can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
+ then (c, SOME tyco)
+ else (c, NONE)
+ | [_] => (c, NONE))
+ | NONE => (c, NONE);
+
+fun string_of_const thy (c, NONE) = Sign.extern_const thy c
+ | string_of_const thy (c, SOME tyco) = Sign.extern_const thy c
+ ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco);
+
+
+(* reading constants as terms and wildcards pattern *)
+
+fun read_bare_const thy raw_t =
+ let
+ val t = Sign.read_term thy raw_t;
+ in case try dest_Const t
+ of SOME c_ty => c_ty
+ | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t)
+ end;
+
+fun read_const thy = const_of_cexpr thy o read_bare_const thy;
+
+local
+
+fun consts_of thy some_thyname =
+ let
+ val this_thy = Option.map theory some_thyname |> the_default thy;
+ val cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I)
+ ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) this_thy) [];
+ fun classop c = case AxClass.class_of_param thy c
+ of NONE => [(c, NONE)]
+ | SOME class => Symtab.fold
+ (fn (tyco, classes) => if AList.defined (op =) classes class
+ then cons (c, SOME tyco) else I)
+ ((#arities o Sorts.rep_algebra o Sign.classes_of) this_thy)
+ [(c, NONE)];
+ val consts = maps classop cs;
+ fun test_instance thy (class, tyco) =
+ can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
+ fun belongs_here thyname (c, NONE) =
+ not (exists (fn thy' => Sign.declared_const thy' c) (Theory.parents_of this_thy))
+ | belongs_here thyname (c, SOME tyco) =
+ let
+ val SOME class = AxClass.class_of_param thy c
+ in not (exists (fn thy' => test_instance thy' (class, tyco))
+ (Theory.parents_of this_thy))
+ end;
+ in case some_thyname
+ of NONE => consts
+ | SOME thyname => filter (belongs_here thyname) consts
+ end;
+
+fun read_const_expr thy "*" = ([], consts_of thy NONE)
+ | read_const_expr thy s = if String.isSuffix ".*" s
+ then ([], consts_of thy (SOME (unsuffix ".*" s)))
+ else ([read_const thy s], []);
+
+in
+
+fun read_const_exprs thy select exprs =
+ case (pairself flat o split_list o map (read_const_expr thy)) exprs
+ of (consts, []) => (false, consts)
+ | (consts, consts') => (true, consts @ select consts');
+
+end; (*local*)
+
+(* conversion between constants, constant expressions and datatype constructors *)
+
+fun const_of_co thy tyco vs (co, tys) =
+ const_of_cexpr thy (co, tys ---> Type (tyco, map TFree vs));
+
+fun consts_of_cos thy tyco vs cos =
+ let
+ val dty = Type (tyco, map TFree vs);
+ fun mk_co (co, tys) = const_of_cexpr thy (co, tys ---> dty);
+ in map mk_co cos end;
+
+local
+
+exception BAD of string;
+
+fun mg_typ_of_const thy (c, NONE) = Sign.the_const_type thy c
+ | mg_typ_of_const thy (c, SOME tyco) =
+ let
+ val SOME class = AxClass.class_of_param thy c;
+ val ty = Sign.the_const_type thy c;
+ (*an approximation*)
+ val sorts = Sorts.mg_domain (Sign.classes_of thy) tyco [class]
+ handle CLASS_ERROR => raise BAD ("No such instance: " ^ tyco ^ " :: " ^ class
+ ^ ",\nrequired for overloaded constant " ^ c);
+ val vs = Name.invents Name.context "'a" (length sorts);
+ in map_atyps (K (Type (tyco, map (fn v => TVar ((v, 0), [])) vs))) ty end;
+
+fun gen_co_of_const thy const =
+ let
+ val (c, _) = const;
+ val ty = (Logic.unvarifyT o mg_typ_of_const thy) const;
+ fun err () = raise BAD
+ ("Illegal type for datatype constructor: " ^ string_of_typ thy ty);
+ val (tys, ty') = strip_type ty;
+ val (tyco, vs) = ((apsnd o map) dest_TFree o dest_Type) ty'
+ handle TYPE _ => err ();
+ val sorts = if has_duplicates (eq_fst op =) vs then err ()
+ else map snd vs;
+ val vs_names = Name.invent_list [] "'a" (length vs);
+ val vs_map = map fst vs ~~ vs_names;
+ val vs' = vs_names ~~ sorts;
+ val tys' = (map o map_type_tfree) (fn (v, sort) =>
+ (TFree ((the o AList.lookup (op =) vs_map) v, sort))) tys
+ handle Option => err ();
+ in (tyco, (vs', (c, tys'))) end;
+
+in
+
+fun co_of_const thy const = gen_co_of_const thy const handle BAD msg => error msg;
+fun co_of_const' thy const = SOME (gen_co_of_const thy const) handle BAD msg => NONE;
+
+fun no_args thy = length o fst o strip_type o mg_typ_of_const thy;
+
+end;
+
+fun cos_of_consts thy consts =
+ let
+ val raw_cos = map (co_of_const thy) consts;
+ val (tyco, (vs_names, sorts_cos)) = if (length o distinct (eq_fst op =)) raw_cos = 1
+ then ((fst o hd) raw_cos, ((map fst o fst o snd o hd) raw_cos,
+ map ((apfst o map) snd o snd) raw_cos))
+ else error ("Term constructors not referring to the same type: "
+ ^ commas (map (string_of_const thy) consts));
+ val sorts = foldr1 ((uncurry o map2 o curry o Sorts.inter_sort) (Sign.classes_of thy))
+ (map fst sorts_cos);
+ val cos = map snd sorts_cos;
+ val vs = vs_names ~~ sorts;
+ in (tyco, (vs, cos)) end;
+
+
+(* dictionary values *)
+
+fun typargs thy (c_ty as (c, ty)) =
+ let
+ val opt_class = AxClass.class_of_param thy c;
+ val tys = Sign.const_typargs thy (c, ty);
+ in case (opt_class, tys)
+ of (SOME class, ty as [Type (tyco, tys')]) =>
+ if can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
+ then tys' else ty
+ | _ => tys
+ end;
+
+fun typ_sort_inst algebra =
+ let
+ val inters = Sorts.inter_sort algebra;
+ fun match _ [] = I
+ | match (TVar (v, S)) S' = Vartab.map_default (v, []) (fn S'' => inters (S, inters (S', S'')))
+ | match (Type (a, Ts)) S =
+ fold2 match Ts (Sorts.mg_domain algebra a S)
+ in uncurry match end;
+
+
+(* making rewrite theorems *)
+
+fun assert_rew thm =
+ let
+ val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
+ handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
+ | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
+ fun vars_of t = fold_aterms
+ (fn Var (v, _) => insert (op =) v
+ | Free _ => bad_thm ("Illegal free variable in rewrite theorem\n"
+ ^ Display.string_of_thm thm)
+ | _ => I) t [];
+ fun tvars_of t = fold_term_types
+ (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v
+ | TFree _ => bad_thm
+ ("Illegal free type variable in rewrite theorem\n" ^ Display.string_of_thm thm))) t [];
+ val lhs_vs = vars_of lhs;
+ val rhs_vs = vars_of rhs;
+ val lhs_tvs = tvars_of lhs;
+ val rhs_tvs = tvars_of lhs;
+ val _ = if null (subtract (op =) lhs_vs rhs_vs)
+ then ()
+ else bad_thm ("Free variables on right hand side of rewrite theorem\n"
+ ^ Display.string_of_thm thm);
+ val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
+ then ()
+ else bad_thm ("Free type variables on right hand side of rewrite theorem\n"
+ ^ Display.string_of_thm thm)
+ in thm end;
+
+fun mk_rew thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val ctxt = ProofContext.init thy;
+ in
+ thm
+ |> LocalDefs.meta_rewrite_rule ctxt
+ |> assert_rew
+ end;
+
+
+(* making defining equations *)
+
+fun assert_func thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val (head, args) = (strip_comb o fst o Logic.dest_equals
+ o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm;
+ val _ = case head of Const _ => () | _ =>
+ bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
+ val _ =
+ if has_duplicates (op =)
+ ((fold o fold_aterms) (fn Var (v, _) => cons v
+ | _ => I
+ ) args [])
+ then bad_thm ("Duplicated variables on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ else ()
+ fun check _ (Abs _) = bad_thm
+ ("Abstraction on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ | check 0 (Var _) = ()
+ | check _ (Var _) = bad_thm
+ ("Variable with application on left hand side of defining equation\n"
+ ^ Display.string_of_thm thm)
+ | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
+ | check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty
+ then bad_thm
+ ("Partially applied constant on left hand side of equation\n"
+ ^ Display.string_of_thm thm)
+ else ();
+ val _ = map (check 0) args;
+ in thm end;
+
+val mk_func = assert_func o mk_rew;
+
+fun head_func thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val (Const (c_ty as (_, ty))) = (fst o strip_comb o fst o Logic.dest_equals
+ o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm;
+ val const = const_of_cexpr thy c_ty;
+ in (const, ty) end;
+
+
+(* utilities *)
+
+fun inst_thm tvars' thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val tvars = (Term.add_tvars o Thm.prop_of) thm [];
+ fun mk_inst (tvar as (v, _)) = case Vartab.lookup tvars' v
+ of SOME sort => SOME (pairself (Thm.ctyp_of thy o TVar) (tvar, (v, sort)))
+ | NONE => NONE;
+ val insts = map_filter mk_inst tvars;
+ in Thm.instantiate (insts, []) thm end;
+
+fun expand_eta k thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
+ val (head, args) = strip_comb lhs;
+ val l = if k = ~1
+ then (length o fst o strip_abs) rhs
+ else Int.max (0, k - length args);
+ val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
+ fun get_name _ 0 used = ([], used)
+ | get_name (Abs (v, ty, t)) k used =
+ used
+ |> Name.variants [v]
+ ||>> get_name t (k - 1)
+ |>> (fn ([v'], vs') => (v', ty) :: vs')
+ | get_name t k used =
+ let
+ val (tys, _) = (strip_type o fastype_of) t
+ in case tys
+ of [] => raise TERM ("expand_eta", [t])
+ | ty :: _ =>
+ used
+ |> Name.variants [""]
+ |-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
+ #>> (fn vs' => (v, ty) :: vs'))
+ end;
+ val (vs, _) = get_name rhs l used;
+ val vs_refl = map (fn (v, ty) => Thm.reflexive (Thm.cterm_of thy (Var ((v, 0), ty)))) vs;
+ in
+ thm
+ |> fold (fn refl => fn thm => Thm.combination thm refl) vs_refl
+ |> Conv.fconv_rule Drule.beta_eta_conversion
+ end;
+
+fun rewrite_func rewrites thm =
+ let
+ val rewrite = MetaSimplifier.rewrite false rewrites;
+ val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o Thm.cprop_of) thm;
+ val Const ("==", _) = Thm.term_of ct_eq;
+ val (ct_f, ct_args) = Drule.strip_comb ct_lhs;
+ val rhs' = rewrite ct_rhs;
+ val args' = map rewrite ct_args;
+ val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1)
+ args' (Thm.reflexive ct_f));
+ in Thm.transitive (Thm.transitive lhs' thm) rhs' end;
+
+fun norm_args thms =
+ let
+ val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
+ val k = fold (curry Int.max o num_args_of o Thm.plain_prop_of) thms 0;
+ in
+ thms
+ |> map (expand_eta k)
+ |> map (Conv.fconv_rule Drule.beta_eta_conversion)
+ end;
+
+fun canonical_tvars purify_tvar thm =
+ let
+ val ctyp = Thm.ctyp_of (Thm.theory_of_thm thm);
+ fun tvars_subst_for thm = (fold_types o fold_atyps)
+ (fn TVar (v_i as (v, _), sort) => let
+ val v' = purify_tvar v
+ in if v = v' then I
+ else insert (op =) (v_i, (v', sort)) end
+ | _ => I) (prop_of thm) [];
+ fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
+ let
+ val ty = TVar (v_i, sort)
+ in
+ (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
+ end;
+ val maxidx = Thm.maxidx_of thm + 1;
+ val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
+ in Thm.instantiate (inst, []) thm end;
+
+fun canonical_vars purify_var thm =
+ let
+ val cterm = Thm.cterm_of (Thm.theory_of_thm thm);
+ fun vars_subst_for thm = fold_aterms
+ (fn Var (v_i as (v, _), ty) => let
+ val v' = purify_var v
+ in if v = v' then I
+ else insert (op =) (v_i, (v', ty)) end
+ | _ => I) (prop_of thm) [];
+ fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
+ let
+ val t = Var (v_i, ty)
+ in
+ (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
+ end;
+ val maxidx = Thm.maxidx_of thm + 1;
+ val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
+ in Thm.instantiate ([], inst) thm end;
+
+fun canonical_absvars purify_var thm =
+ let
+ val t = Thm.plain_prop_of thm;
+ val t' = Term.map_abs_vars purify_var t;
+ in Thm.rename_boundvars t t' thm end;
+
+fun norm_varnames purify_tvar purify_var thms =
+ let
+ fun burrow_thms f [] = []
+ | burrow_thms f thms =
+ thms
+ |> Conjunction.intr_balanced
+ |> f
+ |> Conjunction.elim_balanced (length thms)
+ in
+ thms
+ |> burrow_thms (canonical_tvars purify_tvar)
+ |> map (canonical_vars purify_var)
+ |> map (canonical_absvars purify_var)
+ |> map Drule.zero_var_indexes
+ end;
+
+end;
--- a/src/Pure/Isar/constdefs.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Isar/constdefs.ML Fri Aug 10 17:04:34 2007 +0200
@@ -51,7 +51,7 @@
thy
|> Sign.add_consts_i [(c, T, mx)]
|> PureThy.add_defs_i false [((name, def), atts)]
- |-> (fn [thm] => CodegenData.add_func false thm);
+ |-> (fn [thm] => Code.add_func false thm);
in ((c, T), thy') end;
fun gen_constdefs prep_vars prep_prop prep_att (raw_structs, specs) thy =
--- a/src/Pure/Isar/isar_syn.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Isar/isar_syn.ML Fri Aug 10 17:04:34 2007 +0200
@@ -89,7 +89,7 @@
(P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
P.!!! (P.list1 P.xname)) []
-- Scan.repeat (SpecParse.thm_name ":" -- (P.prop >> single))
- >> (fn (x, y) => Toplevel.theory (snd o ClassPackage.axclass_cmd x y)));
+ >> (fn (x, y) => Toplevel.theory (snd o Class.axclass_cmd x y)));
(* types *)
@@ -422,16 +422,16 @@
-- P.opt_begin
>> (fn (((bname, add_consts), (supclasses, elems)), begin) =>
Toplevel.begin_local_theory begin
- (ClassPackage.class_cmd bname supclasses elems add_consts #-> TheoryTarget.begin)));
+ (Class.class_cmd bname supclasses elems add_consts #-> TheoryTarget.begin)));
val instanceP =
OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal ((
P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname)
- >> ClassPackage.instance_class_cmd
+ >> Class.instance_class_cmd
|| P.$$$ "advanced" |-- P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname)
- >> ClassPackage.instance_sort_cmd
+ >> Class.instance_sort_cmd
|| P.and_list1 P.arity -- Scan.repeat (SpecParse.opt_thm_name ":" -- P.prop)
- >> (fn (arities, defs) => ClassPackage.instance_arity_cmd arities defs)
+ >> (fn (arities, defs) => Class.instance_arity_cmd arities defs)
) >> (Toplevel.print oo Toplevel.theory_to_proof));
end;
@@ -441,7 +441,7 @@
val code_datatypeP =
OuterSyntax.command "code_datatype" "define set of code datatype constructors" K.thy_decl
- (Scan.repeat1 P.term >> (Toplevel.theory o CodegenData.add_datatype_consts_cmd));
+ (Scan.repeat1 P.term >> (Toplevel.theory o Code.add_datatype_consts_cmd));
@@ -763,7 +763,7 @@
val print_classesP =
OuterSyntax.improper_command "print_classes" "print classes of this theory" K.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory
- o Toplevel.keep (ClassPackage.print_classes o Toplevel.theory_of)));
+ o Toplevel.keep (Class.print_classes o Toplevel.theory_of)));
val print_localeP =
OuterSyntax.improper_command "print_locale" "print locale expression in this theory" K.diag
@@ -884,7 +884,7 @@
OuterSyntax.improper_command "print_codesetup" "print code generator setup of this theory" K.diag
(Scan.succeed
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
- (CodegenData.print_codesetup o Toplevel.theory_of)));
+ (Code.print_codesetup o Toplevel.theory_of)));
(** system commands (for interactive mode only) **)
--- a/src/Pure/Isar/specification.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Isar/specification.ML Fri Aug 10 17:04:34 2007 +0200
@@ -130,7 +130,7 @@
|> LocalTheory.def Thm.definitionK ((x, mx), ((name ^ "_raw", []), rhs));
val ((b, [th']), lthy3) = lthy2
|> LocalTheory.note Thm.definitionK
- ((name, CodegenData.add_func_attr false :: atts), [prove lthy2 th]);
+ ((name, Code.add_func_attr false :: atts), [prove lthy2 th]);
val lhs' = Morphism.term (LocalTheory.target_morphism lthy3) lhs;
val _ = print_consts lthy3 (member (op =) (Term.add_frees lhs' [])) [(x, T)];
--- a/src/Pure/Proof/extraction.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Proof/extraction.ML Fri Aug 10 17:04:34 2007 +0200
@@ -741,7 +741,7 @@
(ProofChecker.thm_of_proof thy'
(fst (Proofterm.freeze_thaw_prf prf)))))), [])
|> snd
- |> fold (CodegenData.add_func false) def_thms
+ |> fold (Code.add_func false) def_thms
end
| SOME _ => thy);
--- a/src/Pure/Tools/ROOT.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Tools/ROOT.ML Fri Aug 10 17:04:34 2007 +0200
@@ -13,12 +13,10 @@
(*derived theory and proof elements*)
use "invoke.ML";
-(*code generator, 1st generation*)
+(*code generator*)
use "../codegen.ML";
-
-(*code generator, 2nd generation*)
-use "codegen_names.ML";
-use "codegen_funcgr.ML";
-use "codegen_thingol.ML";
-use "codegen_serializer.ML";
-use "codegen_package.ML";
+use "../../Tools/code/code_name.ML";
+use "../../Tools/code/code_funcgr.ML";
+use "../../Tools/code/code_thingol.ML";
+use "../../Tools/code/code_target.ML";
+use "../../Tools/code/code_package.ML";
--- a/src/Pure/Tools/codegen_consts.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,223 +0,0 @@
-(* Title: Pure/Tools/codegen_consts.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Identifying constants by name plus normalized type instantantiation schemes.
-Convenient data structures for constants. Auxiliary.
-*)
-
-signature CODEGEN_CONSTS =
-sig
- type const = string * string option (*constant name, possibly instance*)
- val const_ord: const * const -> order
- val eq_const: const * const -> bool
- structure Consttab: TABLE
- val const_of_cexpr: theory -> string * typ -> const
- val string_of_typ: theory -> typ -> string
- val string_of_const: theory -> const -> string
- val read_bare_const: theory -> string -> string * typ
- val read_const: theory -> string -> const
- val read_const_exprs: theory -> (const list -> const list)
- -> string list -> bool * const list
-
- val co_of_const: theory -> const
- -> string * ((string * sort) list * (string * typ list))
- val co_of_const': theory -> const
- -> (string * ((string * sort) list * (string * typ list))) option
- val cos_of_consts: theory -> const list
- -> string * ((string * sort) list * (string * typ list) list)
- val const_of_co: theory -> string -> (string * sort) list
- -> string * typ list -> const
- val consts_of_cos: theory -> string -> (string * sort) list
- -> (string * typ list) list -> const list
- val no_args: theory -> const -> int
-
- val typargs: theory -> string * typ -> typ list
- val typ_sort_inst: Sorts.algebra -> typ * sort
- -> sort Vartab.table -> sort Vartab.table
-end;
-
-structure CodegenConsts: CODEGEN_CONSTS =
-struct
-
-
-(* basic data structures *)
-
-type const = string * string option;
-val const_ord = prod_ord fast_string_ord (option_ord string_ord);
-val eq_const = is_equal o const_ord;
-
-structure Consttab =
- TableFun(
- type key = const;
- val ord = const_ord;
- );
-
-fun string_of_typ thy = setmp show_sorts true (Sign.string_of_typ thy);
-
-
-(* conversion between constant expressions and constants *)
-
-fun const_of_cexpr thy (c_ty as (c, _)) =
- case AxClass.class_of_param thy c
- of SOME class => (case Sign.const_typargs thy c_ty
- of [Type (tyco, _)] => if can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
- then (c, SOME tyco)
- else (c, NONE)
- | [_] => (c, NONE))
- | NONE => (c, NONE);
-
-fun string_of_const thy (c, NONE) = Sign.extern_const thy c
- | string_of_const thy (c, SOME tyco) = Sign.extern_const thy c
- ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco);
-
-
-(* reading constants as terms and wildcards pattern *)
-
-fun read_bare_const thy raw_t =
- let
- val t = Sign.read_term thy raw_t;
- in case try dest_Const t
- of SOME c_ty => c_ty
- | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t)
- end;
-
-fun read_const thy = const_of_cexpr thy o read_bare_const thy;
-
-local
-
-fun consts_of thy some_thyname =
- let
- val this_thy = Option.map theory some_thyname |> the_default thy;
- val cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I)
- ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) this_thy) [];
- fun classop c = case AxClass.class_of_param thy c
- of NONE => [(c, NONE)]
- | SOME class => Symtab.fold
- (fn (tyco, classes) => if AList.defined (op =) classes class
- then cons (c, SOME tyco) else I)
- ((#arities o Sorts.rep_algebra o Sign.classes_of) this_thy)
- [(c, NONE)];
- val consts = maps classop cs;
- fun test_instance thy (class, tyco) =
- can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
- fun belongs_here thyname (c, NONE) =
- not (exists (fn thy' => Sign.declared_const thy' c) (Theory.parents_of this_thy))
- | belongs_here thyname (c, SOME tyco) =
- let
- val SOME class = AxClass.class_of_param thy c
- in not (exists (fn thy' => test_instance thy' (class, tyco))
- (Theory.parents_of this_thy))
- end;
- in case some_thyname
- of NONE => consts
- | SOME thyname => filter (belongs_here thyname) consts
- end;
-
-fun read_const_expr thy "*" = ([], consts_of thy NONE)
- | read_const_expr thy s = if String.isSuffix ".*" s
- then ([], consts_of thy (SOME (unsuffix ".*" s)))
- else ([read_const thy s], []);
-
-in
-
-fun read_const_exprs thy select exprs =
- case (pairself flat o split_list o map (read_const_expr thy)) exprs
- of (consts, []) => (false, consts)
- | (consts, consts') => (true, consts @ select consts');
-
-end; (*local*)
-
-(* conversion between constants, constant expressions and datatype constructors *)
-
-fun const_of_co thy tyco vs (co, tys) =
- const_of_cexpr thy (co, tys ---> Type (tyco, map TFree vs));
-
-fun consts_of_cos thy tyco vs cos =
- let
- val dty = Type (tyco, map TFree vs);
- fun mk_co (co, tys) = const_of_cexpr thy (co, tys ---> dty);
- in map mk_co cos end;
-
-local
-
-exception BAD of string;
-
-fun mg_typ_of_const thy (c, NONE) = Sign.the_const_type thy c
- | mg_typ_of_const thy (c, SOME tyco) =
- let
- val SOME class = AxClass.class_of_param thy c;
- val ty = Sign.the_const_type thy c;
- (*an approximation*)
- val sorts = Sorts.mg_domain (Sign.classes_of thy) tyco [class]
- handle CLASS_ERROR => raise BAD ("No such instance: " ^ tyco ^ " :: " ^ class
- ^ ",\nrequired for overloaded constant " ^ c);
- val vs = Name.invents Name.context "'a" (length sorts);
- in map_atyps (K (Type (tyco, map (fn v => TVar ((v, 0), [])) vs))) ty end;
-
-fun gen_co_of_const thy const =
- let
- val (c, _) = const;
- val ty = (Logic.unvarifyT o mg_typ_of_const thy) const;
- fun err () = raise BAD
- ("Illegal type for datatype constructor: " ^ string_of_typ thy ty);
- val (tys, ty') = strip_type ty;
- val (tyco, vs) = ((apsnd o map) dest_TFree o dest_Type) ty'
- handle TYPE _ => err ();
- val sorts = if has_duplicates (eq_fst op =) vs then err ()
- else map snd vs;
- val vs_names = Name.invent_list [] "'a" (length vs);
- val vs_map = map fst vs ~~ vs_names;
- val vs' = vs_names ~~ sorts;
- val tys' = (map o map_type_tfree) (fn (v, sort) =>
- (TFree ((the o AList.lookup (op =) vs_map) v, sort))) tys
- handle Option => err ();
- in (tyco, (vs', (c, tys'))) end;
-
-in
-
-fun co_of_const thy const = gen_co_of_const thy const handle BAD msg => error msg;
-fun co_of_const' thy const = SOME (gen_co_of_const thy const) handle BAD msg => NONE;
-
-fun no_args thy = length o fst o strip_type o mg_typ_of_const thy;
-
-end;
-
-fun cos_of_consts thy consts =
- let
- val raw_cos = map (co_of_const thy) consts;
- val (tyco, (vs_names, sorts_cos)) = if (length o distinct (eq_fst op =)) raw_cos = 1
- then ((fst o hd) raw_cos, ((map fst o fst o snd o hd) raw_cos,
- map ((apfst o map) snd o snd) raw_cos))
- else error ("Term constructors not referring to the same type: "
- ^ commas (map (string_of_const thy) consts));
- val sorts = foldr1 ((uncurry o map2 o curry o Sorts.inter_sort) (Sign.classes_of thy))
- (map fst sorts_cos);
- val cos = map snd sorts_cos;
- val vs = vs_names ~~ sorts;
- in (tyco, (vs, cos)) end;
-
-
-(* dictionary values *)
-
-fun typargs thy (c_ty as (c, ty)) =
- let
- val opt_class = AxClass.class_of_param thy c;
- val tys = Sign.const_typargs thy (c, ty);
- in case (opt_class, tys)
- of (SOME class, ty as [Type (tyco, tys')]) =>
- if can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
- then tys' else ty
- | _ => tys
- end;
-
-fun typ_sort_inst algebra =
- let
- val inters = Sorts.inter_sort algebra;
- fun match _ [] = I
- | match (TVar (v, S)) S' = Vartab.map_default (v, []) (fn S'' => inters (S, inters (S', S'')))
- | match (Type (a, Ts)) S =
- fold2 match Ts (Sorts.mg_domain algebra a S)
- in uncurry match end;
-
-end;
--- a/src/Pure/Tools/codegen_data.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,892 +0,0 @@
-(* Title: Pure/Tools/codegen_data.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Abstract executable content of theory. Management of data dependent on
-executable content.
-*)
-
-signature CODEGEN_DATA =
-sig
- val lazy_thms: (unit -> thm list) -> thm list Susp.T
- val eval_always: bool ref
-
- val add_func: bool -> thm -> theory -> theory
- val del_func: thm -> theory -> theory
- val add_funcl: CodegenConsts.const * thm list Susp.T -> theory -> theory
- val add_func_attr: bool -> Attrib.src
- val add_inline: thm -> theory -> theory
- val del_inline: thm -> theory -> theory
- val add_inline_proc: string * (theory -> cterm list -> thm list) -> theory -> theory
- val del_inline_proc: string -> theory -> theory
- val add_preproc: string * (theory -> thm list -> thm list) -> theory -> theory
- val del_preproc: string -> theory -> theory
- val add_post: thm -> theory -> theory
- val del_post: thm -> theory -> theory
- val add_datatype: string * ((string * sort) list * (string * typ list) list)
- -> theory -> theory
- val add_datatype_consts: CodegenConsts.const list -> theory -> theory
- val add_datatype_consts_cmd: string list -> theory -> theory
-
- val coregular_algebra: theory -> Sorts.algebra
- val operational_algebra: theory -> (sort -> sort) * Sorts.algebra
- val these_funcs: theory -> CodegenConsts.const -> thm list
- val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
- val get_datatype_of_constr: theory -> CodegenConsts.const -> string option
- val default_typ: theory -> CodegenConsts.const -> typ
-
- val preprocess_conv: cterm -> thm
- val postprocess_conv: cterm -> thm
-
- val print_codesetup: theory -> unit
-
- val trace: bool ref
-end;
-
-signature CODE_DATA_ARGS =
-sig
- type T
- val empty: T
- val merge: Pretty.pp -> T * T -> T
- val purge: theory option -> CodegenConsts.const list option -> T -> T
-end;
-
-signature CODE_DATA =
-sig
- type T
- val get: theory -> T
- val change: theory -> (T -> T) -> T
- val change_yield: theory -> (T -> 'a * T) -> 'a * T
-end;
-
-signature PRIVATE_CODEGEN_DATA =
-sig
- include CODEGEN_DATA
- val declare_data: Object.T -> (Pretty.pp -> Object.T * Object.T -> Object.T)
- -> (theory option -> CodegenConsts.const list option -> Object.T -> Object.T) -> serial
- val get_data: serial * ('a -> Object.T) * (Object.T -> 'a)
- -> theory -> 'a
- val change_data: serial * ('a -> Object.T) * (Object.T -> 'a)
- -> theory -> ('a -> 'a) -> 'a
- val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a)
- -> theory -> ('a -> 'b * 'a) -> 'b * 'a
-end;
-
-structure CodegenData : PRIVATE_CODEGEN_DATA =
-struct
-
-(* auxiliary, diagnostics *)
-
-structure Consttab = CodegenConsts.Consttab;
-
-val trace = ref false;
-fun tracing f x = (if !trace then Output.tracing (f x) else (); x);
-
-
-(* lazy theorems, certificate theorems *)
-
-val eval_always = ref false;
-
-fun lazy_thms f = if !eval_always
- then Susp.value (f ())
- else Susp.delay f;
-
-fun string_of_lthms r = case Susp.peek r
- of SOME thms => (map string_of_thm o rev) thms
- | NONE => ["[...]"];
-
-fun pretty_lthms ctxt r = case Susp.peek r
- of SOME thms => map (ProofContext.pretty_thm ctxt) thms
- | NONE => [Pretty.str "[...]"];
-
-fun certificate thy f r =
- case Susp.peek r
- of SOME thms => (Susp.value o f thy) thms
- | NONE => let
- val thy_ref = Theory.check_thy thy;
- in lazy_thms (fn () => (f (Theory.deref thy_ref) o Susp.force) r) end;
-
-fun merge' _ ([], []) = (false, [])
- | merge' _ ([], ys) = (true, ys)
- | merge' eq (xs, ys) = fold_rev
- (fn y => fn (t, xs) => (t orelse not (member eq xs y), insert eq y xs)) ys (false, xs);
-
-fun merge_alist eq_key eq (xys as (xs, ys)) =
- if eq_list (eq_pair eq_key eq) (xs, ys)
- then (false, xs)
- else (true, AList.merge eq_key eq xys);
-
-val merge_thms = merge' Thm.eq_thm_prop;
-
-fun merge_lthms (r1, r2) =
- if Susp.same (r1, r2)
- then (false, r1)
- else case Susp.peek r1
- of SOME [] => (true, r2)
- | _ => case Susp.peek r2
- of SOME [] => (true, r1)
- | _ => (apsnd (lazy_thms o K)) (merge_thms (Susp.force r1, Susp.force r2));
-
-
-(* pairs of (selected, deleted) defining equations *)
-
-type sdthms = thm list Susp.T * thm list;
-
-fun add_drop_redundant thm (sels, dels) =
- let
- val thy = Thm.theory_of_thm thm;
- val args_of = snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
- val args = args_of thm;
- fun matches [] _ = true
- | matches (Var _ :: xs) [] = matches xs []
- | matches (_ :: _) [] = false
- | matches (x :: xs) (y :: ys) = Pattern.matches thy (x, y) andalso matches xs ys;
- fun drop thm' = not (matches args (args_of thm'))
- orelse (warning ("code generator: dropping redundant defining equation\n" ^ string_of_thm thm'); false);
- val (keeps, drops) = List.partition drop sels;
- in (thm :: keeps, dels |> remove Thm.eq_thm_prop thm |> fold (insert Thm.eq_thm_prop) drops) end;
-
-fun add_thm thm (sels, dels) =
- apfst Susp.value (add_drop_redundant thm (Susp.force sels, dels));
-
-fun add_lthms lthms (sels, []) =
- (lazy_thms (fn () => fold add_drop_redundant
- (Susp.force lthms) (Susp.force sels, []) |> fst), [])
- (*FIXME*)
- | add_lthms lthms (sels, dels) =
- fold add_thm (Susp.force lthms) (sels, dels);
-
-fun del_thm thm (sels, dels) =
- (Susp.value (remove Thm.eq_thm_prop thm (Susp.force sels)), thm :: dels);
-
-fun pretty_sdthms ctxt (sels, _) = pretty_lthms ctxt sels;
-
-fun merge_sdthms ((sels1, dels1), (sels2, dels2)) =
- let
- val (dels_t, dels) = merge_thms (dels1, dels2);
- in if dels_t
- then let
- val (_, sels) = merge_thms
- (subtract Thm.eq_thm_prop dels2 (Susp.force sels1), Susp.force sels2);
- val (_, dels) = merge_thms
- (subtract Thm.eq_thm_prop (Susp.force sels2) dels1, dels2);
- in (true, ((lazy_thms o K) sels, dels)) end
- else let
- val (sels_t, sels) = merge_lthms (sels1, sels2);
- in (sels_t, (sels, dels)) end
- end;
-
-
-(** exeuctable content **)
-
-datatype thmproc = Preproc of {
- inlines: thm list,
- inline_procs: (string * (serial * (theory -> cterm list -> thm list))) list,
- preprocs: (string * (serial * (theory -> thm list -> thm list))) list,
- posts: thm list
-};
-
-fun mk_thmproc (((inlines, inline_procs), preprocs), posts) =
- Preproc { inlines = inlines, inline_procs = inline_procs, preprocs = preprocs,
- posts = posts };
-fun map_thmproc f (Preproc { inlines, inline_procs, preprocs, posts }) =
- mk_thmproc (f (((inlines, inline_procs), preprocs), posts));
-fun merge_thmproc (Preproc { inlines = inlines1, inline_procs = inline_procs1,
- preprocs = preprocs1, posts = posts1 },
- Preproc { inlines = inlines2, inline_procs = inline_procs2,
- preprocs = preprocs2, posts= posts2 }) =
- let
- val (touched1, inlines) = merge_thms (inlines1, inlines2);
- val (touched2, inline_procs) = merge_alist (op =) (eq_fst (op =)) (inline_procs1, inline_procs2);
- val (touched3, preprocs) = merge_alist (op =) (eq_fst (op =)) (preprocs1, preprocs2);
- val (_, posts) = merge_thms (posts1, posts2);
- in (touched1 orelse touched2 orelse touched3,
- mk_thmproc (((inlines, inline_procs), preprocs), posts)) end;
-
-fun join_func_thms (tabs as (tab1, tab2)) =
- let
- val cs1 = Consttab.keys tab1;
- val cs2 = Consttab.keys tab2;
- val cs' = filter (member CodegenConsts.eq_const cs2) cs1;
- val cs'' = subtract (op =) cs' cs1 @ subtract (op =) cs' cs2;
- val cs''' = ref [] : CodegenConsts.const list ref;
- fun merge c x = let val (touched, thms') = merge_sdthms x in
- (if touched then cs''' := cons c (!cs''') else (); thms') end;
- in (cs'' @ !cs''', Consttab.join merge tabs) end;
-fun merge_funcs (thms1, thms2) =
- let
- val (consts, thms) = join_func_thms (thms1, thms2);
- in (SOME consts, thms) end;
-
-val eq_string = op = : string * string -> bool;
-val eq_co = op = : (string * typ list) * (string * typ list) -> bool;
-fun eq_dtyp ((vs1, cs1), (vs2, cs2)) =
- gen_eq_set (eq_pair eq_string (gen_eq_set eq_string)) (vs1, vs2)
- andalso gen_eq_set eq_co (cs1, cs2);
-fun merge_dtyps (tabs as (tab1, tab2)) =
- let
- val tycos1 = Symtab.keys tab1;
- val tycos2 = Symtab.keys tab2;
- val tycos' = filter (member eq_string tycos2) tycos1;
- val new_types = not (gen_eq_set (op =) (tycos1, tycos2));
- val diff_types = not (gen_eq_set (eq_pair (op =) eq_dtyp)
- (AList.make (the o Symtab.lookup tab1) tycos',
- AList.make (the o Symtab.lookup tab2) tycos'));
- fun join _ (cos as (_, cos2)) = if eq_dtyp cos
- then raise Symtab.SAME else cos2;
- in ((new_types, diff_types), Symtab.join join tabs) end;
-
-datatype spec = Spec of {
- funcs: sdthms Consttab.table,
- dtyps: ((string * sort) list * (string * typ list) list) Symtab.table
-};
-
-fun mk_spec (funcs, dtyps) =
- Spec { funcs = funcs, dtyps = dtyps };
-fun map_spec f (Spec { funcs = funcs, dtyps = dtyps }) =
- mk_spec (f (funcs, dtyps));
-fun merge_spec (Spec { funcs = funcs1, dtyps = dtyps1 },
- Spec { funcs = funcs2, dtyps = dtyps2 }) =
- let
- val (touched_cs, funcs) = merge_funcs (funcs1, funcs2);
- val ((new_types, diff_types), dtyps) = merge_dtyps (dtyps1, dtyps2);
- val touched = if new_types orelse diff_types then NONE else touched_cs;
- in (touched, mk_spec (funcs, dtyps)) end;
-
-datatype exec = Exec of {
- thmproc: thmproc,
- spec: spec
-};
-
-fun mk_exec (thmproc, spec) =
- Exec { thmproc = thmproc, spec = spec };
-fun map_exec f (Exec { thmproc = thmproc, spec = spec }) =
- mk_exec (f (thmproc, spec));
-fun merge_exec (Exec { thmproc = thmproc1, spec = spec1 },
- Exec { thmproc = thmproc2, spec = spec2 }) =
- let
- val (touched', thmproc) = merge_thmproc (thmproc1, thmproc2);
- val (touched_cs, spec) = merge_spec (spec1, spec2);
- val touched = if touched' then NONE else touched_cs;
- in (touched, mk_exec (thmproc, spec)) end;
-val empty_exec = mk_exec (mk_thmproc ((([], []), []), []),
- mk_spec (Consttab.empty, Symtab.empty));
-
-fun the_thmproc (Exec { thmproc = Preproc x, ...}) = x;
-fun the_spec (Exec { spec = Spec x, ...}) = x;
-val the_funcs = #funcs o the_spec;
-val the_dtyps = #dtyps o the_spec;
-val map_thmproc = map_exec o apfst o map_thmproc;
-val map_funcs = map_exec o apsnd o map_spec o apfst;
-val map_dtyps = map_exec o apsnd o map_spec o apsnd;
-
-
-(* data slots dependent on executable content *)
-
-(*private copy avoids potential conflict of table exceptions*)
-structure Datatab = TableFun(type key = int val ord = int_ord);
-
-local
-
-type kind = {
- empty: Object.T,
- merge: Pretty.pp -> Object.T * Object.T -> Object.T,
- purge: theory option -> CodegenConsts.const list option -> Object.T -> Object.T
-};
-
-val kinds = ref (Datatab.empty: kind Datatab.table);
-val kind_keys = ref ([]: serial list);
-
-fun invoke f k = case Datatab.lookup (! kinds) k
- of SOME kind => f kind
- | NONE => sys_error "Invalid code data identifier";
-
-in
-
-fun declare_data empty merge purge =
- let
- val k = serial ();
- val kind = {empty = empty, merge = merge, purge = purge};
- val _ = change kinds (Datatab.update (k, kind));
- val _ = change kind_keys (cons k);
- in k end;
-
-fun invoke_empty k = invoke (fn kind => #empty kind) k;
-
-fun invoke_merge_all pp = Datatab.join
- (invoke (fn kind => #merge kind pp));
-
-fun invoke_purge_all thy_opt cs =
- fold (fn k => Datatab.map_entry k
- (invoke (fn kind => #purge kind thy_opt cs) k)) (! kind_keys);
-
-end; (*local*)
-
-
-(* theory store *)
-
-local
-
-type data = Object.T Datatab.table;
-
-structure CodeData = TheoryDataFun
-(
- type T = exec * data ref;
- val empty = (empty_exec, ref Datatab.empty : data ref);
- fun copy (exec, data) = (exec, ref (! data));
- val extend = copy;
- fun merge pp ((exec1, data1), (exec2, data2)) =
- let
- val (touched, exec) = merge_exec (exec1, exec2);
- val data1' = invoke_purge_all NONE touched (! data1);
- val data2' = invoke_purge_all NONE touched (! data2);
- val data = invoke_merge_all pp (data1', data2');
- in (exec, ref data) end;
-);
-
-val _ = Context.add_setup CodeData.init;
-
-fun ch r f = let val x = f (! r) in (r := x; x) end;
-fun thy_data f thy = f ((snd o CodeData.get) thy);
-
-fun get_ensure_init kind data_ref =
- case Datatab.lookup (! data_ref) kind
- of SOME x => x
- | NONE => let val y = invoke_empty kind
- in (change data_ref (Datatab.update (kind, y)); y) end;
-
-in
-
-(* access to executable content *)
-
-val get_exec = fst o CodeData.get;
-
-fun map_exec_purge touched f thy =
- CodeData.map (fn (exec, data) =>
- (f exec, ref (invoke_purge_all (SOME thy) touched (! data)))) thy;
-
-
-(* access to data dependent on abstract executable content *)
-
-fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest);
-
-fun change_data (kind, mk, dest) =
- let
- fun chnge data_ref f =
- let
- val data = get_ensure_init kind data_ref;
- val data' = f (dest data);
- in (change data_ref (Datatab.update (kind, mk data')); data') end;
- in thy_data chnge end;
-
-fun change_yield_data (kind, mk, dest) =
- let
- fun chnge data_ref f =
- let
- val data = get_ensure_init kind data_ref;
- val (x, data') = f (dest data);
- in (x, (change data_ref (Datatab.update (kind, mk data')); data')) end;
- in thy_data chnge end;
-
-end; (*local*)
-
-
-(* print executable content *)
-
-fun print_codesetup thy =
- let
- val ctxt = ProofContext.init thy;
- val exec = get_exec thy;
- fun pretty_func (s, lthms) =
- (Pretty.block o Pretty.fbreaks) (
- Pretty.str s :: pretty_sdthms ctxt lthms
- );
- fun pretty_dtyp (s, []) =
- Pretty.str s
- | pretty_dtyp (s, cos) =
- (Pretty.block o Pretty.breaks) (
- Pretty.str s
- :: Pretty.str "="
- :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str c
- | (c, tys) =>
- (Pretty.block o Pretty.breaks)
- (Pretty.str c :: Pretty.str "of" :: map (Pretty.quote o Sign.pretty_typ thy) tys)) cos)
- );
- val inlines = (#inlines o the_thmproc) exec;
- val inline_procs = (map fst o #inline_procs o the_thmproc) exec;
- val preprocs = (map fst o #preprocs o the_thmproc) exec;
- val funs = the_funcs exec
- |> Consttab.dest
- |> (map o apfst) (CodegenConsts.string_of_const thy)
- |> sort (string_ord o pairself fst);
- val dtyps = the_dtyps exec
- |> Symtab.dest
- |> map (fn (dtco, (vs, cos)) => (Sign.string_of_typ thy (Type (dtco, map TFree vs)), cos))
- |> sort (string_ord o pairself fst)
- in
- (Pretty.writeln o Pretty.chunks) [
- Pretty.block (
- Pretty.str "defining equations:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map pretty_func) funs
- ),
- Pretty.block (
- Pretty.str "inlining theorems:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map (ProofContext.pretty_thm ctxt)) inlines
- ),
- Pretty.block (
- Pretty.str "inlining procedures:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map Pretty.str) inline_procs
- ),
- Pretty.block (
- Pretty.str "preprocessors:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map Pretty.str) preprocs
- ),
- Pretty.block (
- Pretty.str "datatypes:"
- :: Pretty.fbrk
- :: (Pretty.fbreaks o map pretty_dtyp) dtyps
- )
- ]
- end;
-
-
-
-(** theorem transformation and certification **)
-
-fun common_typ_funcs [] = []
- | common_typ_funcs [thm] = [thm]
- | common_typ_funcs (thms as thm :: _) =
- let
- val thy = Thm.theory_of_thm thm;
- fun incr_thm thm max =
- let
- val thm' = incr_indexes max thm;
- val max' = Thm.maxidx_of thm' + 1;
- in (thm', max') end;
- val (thms', maxidx) = fold_map incr_thm thms 0;
- val ty1 :: tys = map (snd o CodegenFunc.head_func) thms';
- fun unify ty env = Sign.typ_unify thy (ty1, ty) env
- handle Type.TUNIFY =>
- error ("Type unificaton failed, while unifying defining equations\n"
- ^ (cat_lines o map Display.string_of_thm) thms
- ^ "\nwith types\n"
- ^ (cat_lines o map (CodegenConsts.string_of_typ thy)) (ty1 :: tys));
- val (env, _) = fold unify tys (Vartab.empty, maxidx)
- val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
- cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
- in map (Thm.instantiate (instT, [])) thms' end;
-
-fun certify_const thy const thms =
- let
- fun cert thm = if CodegenConsts.eq_const (const, fst (CodegenFunc.head_func thm))
- then thm else error ("Wrong head of defining equation,\nexpected constant "
- ^ CodegenConsts.string_of_const thy const ^ "\n" ^ string_of_thm thm)
- in map cert thms end;
-
-
-
-(** operational sort algebra and class discipline **)
-
-local
-
-fun aggr_neutr f y [] = y
- | aggr_neutr f y (x::xs) = aggr_neutr f (f y x) xs;
-
-fun aggregate f [] = NONE
- | aggregate f (x::xs) = SOME (aggr_neutr f x xs);
-
-fun inter_sorts thy =
- let
- val algebra = Sign.classes_of thy;
- val inters = curry (Sorts.inter_sort algebra);
- in aggregate (map2 inters) end;
-
-fun specific_constraints thy (class, tyco) =
- let
- val vs = Name.invents Name.context "" (Sign.arity_number thy tyco);
- val clsops = (these o Option.map snd o try (AxClass.params_of_class thy)) class;
- val funcs = clsops
- |> map (fn (clsop, _) => (clsop, SOME tyco))
- |> map (Consttab.lookup ((the_funcs o get_exec) thy))
- |> (map o Option.map) (Susp.force o fst)
- |> maps these
- |> map (Thm.transfer thy);
- val sorts = map (map (snd o dest_TVar) o snd o dest_Type o the_single
- o Sign.const_typargs thy o (fn ((c, _), ty) => (c, ty)) o CodegenFunc.head_func) funcs;
- in sorts end;
-
-fun weakest_constraints thy (class, tyco) =
- let
- val all_superclasses = class :: Graph.all_succs ((#classes o Sorts.rep_algebra o Sign.classes_of) thy) [class];
- in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) all_superclasses)
- of SOME sorts => sorts
- | NONE => Sign.arity_sorts thy tyco [class]
- end;
-
-fun strongest_constraints thy (class, tyco) =
- let
- val algebra = Sign.classes_of thy;
- val all_subclasses = class :: Graph.all_preds ((#classes o Sorts.rep_algebra) algebra) [class];
- val inst_subclasses = filter (can (Sorts.mg_domain algebra tyco) o single) all_subclasses;
- in case inter_sorts thy (maps (fn class => specific_constraints thy (class, tyco)) inst_subclasses)
- of SOME sorts => sorts
- | NONE => replicate
- (Sign.arity_number thy tyco) (Sign.certify_sort thy (Sign.all_classes thy))
- end;
-
-fun gen_classop_typ constr thy class (c, tyco) =
- let
- val (var, cs) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
- val ty = (the o AList.lookup (op =) cs) c;
- val sort_args = Name.names (Name.declare var Name.context) "'a"
- (constr thy (class, tyco));
- val ty_inst = Type (tyco, map TFree sort_args);
- in Logic.varifyT (map_type_tfree (K ty_inst) ty) end;
-
-fun retrieve_algebra thy operational =
- Sorts.subalgebra (Sign.pp thy) operational
- (weakest_constraints thy)
- (Sign.classes_of thy);
-
-in
-
-fun coregular_algebra thy = retrieve_algebra thy (K true) |> snd;
-fun operational_algebra thy =
- let
- fun add_iff_operational class =
- can (AxClass.get_definition thy) class ? cons class;
- val operational_classes = fold add_iff_operational (Sign.all_classes thy) []
- in retrieve_algebra thy (member (op =) operational_classes) end;
-
-val classop_weakest_typ = gen_classop_typ weakest_constraints;
-val classop_strongest_typ = gen_classop_typ strongest_constraints;
-
-fun assert_func_typ thm =
- let
- val thy = Thm.theory_of_thm thm;
- fun check_typ_classop class (const as (c, SOME tyco), thm) =
- let
- val (_, ty) = CodegenFunc.head_func thm;
- val ty_decl = classop_weakest_typ thy class (c, tyco);
- val ty_strongest = classop_strongest_typ thy class (c, tyco);
- fun constrain thm =
- let
- val max = Thm.maxidx_of thm + 1;
- val ty_decl' = Logic.incr_tvar max ty_decl;
- val (_, ty') = CodegenFunc.head_func thm;
- val (env, _) = Sign.typ_unify thy (ty_decl', ty') (Vartab.empty, max);
- val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
- cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
- in Thm.instantiate (instT, []) thm end;
- in if Sign.typ_instance thy (ty_strongest, ty)
- then if Sign.typ_instance thy (ty, ty_decl)
- then thm
- else (warning ("Constraining type\n" ^ CodegenConsts.string_of_typ thy ty
- ^ "\nof defining equation\n"
- ^ string_of_thm thm
- ^ "\nto permitted most general type\n"
- ^ CodegenConsts.string_of_typ thy ty_decl);
- constrain thm)
- else CodegenFunc.bad_thm ("Type\n" ^ CodegenConsts.string_of_typ thy ty
- ^ "\nof defining equation\n"
- ^ string_of_thm thm
- ^ "\nis incompatible with permitted least general type\n"
- ^ CodegenConsts.string_of_typ thy ty_strongest)
- end
- | check_typ_classop class ((c, NONE), thm) =
- CodegenFunc.bad_thm ("Illegal type for class operation " ^ quote c
- ^ "\nin defining equation\n"
- ^ string_of_thm thm);
- fun check_typ_fun (const as (c, _), thm) =
- let
- val (_, ty) = CodegenFunc.head_func thm;
- val ty_decl = Sign.the_const_type thy c;
- in if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
- then thm
- else CodegenFunc.bad_thm ("Type\n" ^ CodegenConsts.string_of_typ thy ty
- ^ "\nof defining equation\n"
- ^ string_of_thm thm
- ^ "\nis incompatible with declared function type\n"
- ^ CodegenConsts.string_of_typ thy ty_decl)
- end;
- fun check_typ (const as (c, _), thm) =
- case AxClass.class_of_param thy c
- of SOME class => check_typ_classop class (const, thm)
- | NONE => check_typ_fun (const, thm);
- in check_typ (fst (CodegenFunc.head_func thm), thm) end;
-
-val mk_func = CodegenFunc.error_thm
- (assert_func_typ o CodegenFunc.mk_func);
-val mk_func_liberal = CodegenFunc.warning_thm
- (assert_func_typ o CodegenFunc.mk_func);
-
-end;
-
-
-
-(** interfaces **)
-
-fun add_func true thm thy =
- let
- val func = mk_func thm;
- val (const, _) = CodegenFunc.head_func func;
- in map_exec_purge (SOME [const]) (map_funcs
- (Consttab.map_default
- (const, (Susp.value [], [])) (add_thm func))) thy
- end
- | add_func false thm thy =
- case mk_func_liberal thm
- of SOME func => let
- val (const, _) = CodegenFunc.head_func func
- in map_exec_purge (SOME [const]) (map_funcs
- (Consttab.map_default
- (const, (Susp.value [], [])) (add_thm func))) thy
- end
- | NONE => thy;
-
-fun delete_force msg key xs =
- if AList.defined (op =) xs key then AList.delete (op =) key xs
- else error ("No such " ^ msg ^ ": " ^ quote key);
-
-fun del_func thm thy =
- let
- val func = mk_func thm;
- val (const, _) = CodegenFunc.head_func func;
- in map_exec_purge (SOME [const]) (map_funcs
- (Consttab.map_entry
- const (del_thm func))) thy
- end;
-
-fun add_funcl (const, lthms) thy =
- let
- val lthms' = certificate thy (fn thy => certify_const thy const) lthms;
- (*FIXME must check compatibility with sort algebra;
- alas, naive checking results in non-termination!*)
- in
- map_exec_purge (SOME [const]) (map_funcs (Consttab.map_default (const, (Susp.value [], []))
- (add_lthms lthms'))) thy
- end;
-
-fun add_func_attr strict = Attrib.internal (fn _ => Thm.declaration_attribute
- (fn thm => Context.mapping (add_func strict thm) I));
-
-local
-
-fun del_datatype tyco thy =
- case Symtab.lookup ((the_dtyps o get_exec) thy) tyco
- of SOME (vs, cos) => let
- val consts = CodegenConsts.consts_of_cos thy tyco vs cos;
- in map_exec_purge (if null consts then NONE else SOME consts)
- (map_dtyps (Symtab.delete tyco)) thy end
- | NONE => thy;
-
-in
-
-fun add_datatype (tyco, (vs_cos as (vs, cos))) thy =
- let
- val consts = CodegenConsts.consts_of_cos thy tyco vs cos;
- in
- thy
- |> del_datatype tyco
- |> map_exec_purge (SOME consts) (map_dtyps (Symtab.update_new (tyco, vs_cos)))
- end;
-
-fun add_datatype_consts consts thy =
- add_datatype (CodegenConsts.cos_of_consts thy consts) thy;
-
-fun add_datatype_consts_cmd raw_cs thy =
- add_datatype_consts (map (CodegenConsts.read_const thy) raw_cs) thy
-
-end; (*local*)
-
-fun add_inline thm thy =
- (map_exec_purge NONE o map_thmproc o apfst o apfst o apfst)
- (insert Thm.eq_thm_prop (CodegenFunc.error_thm CodegenFunc.mk_rew thm)) thy;
- (*fully applied in order to get right context for mk_rew!*)
-
-fun del_inline thm thy =
- (map_exec_purge NONE o map_thmproc o apfst o apfst o apfst)
- (remove Thm.eq_thm_prop (CodegenFunc.error_thm CodegenFunc.mk_rew thm)) thy;
- (*fully applied in order to get right context for mk_rew!*)
-
-fun add_inline_proc (name, f) =
- (map_exec_purge NONE o map_thmproc o apfst o apfst o apsnd)
- (AList.update (op =) (name, (serial (), f)));
-
-fun del_inline_proc name =
- (map_exec_purge NONE o map_thmproc o apfst o apfst o apsnd)
- (delete_force "inline procedure" name);
-
-fun add_preproc (name, f) =
- (map_exec_purge NONE o map_thmproc o apfst o apsnd)
- (AList.update (op =) (name, (serial (), f)));
-
-fun del_preproc name =
- (map_exec_purge NONE o map_thmproc o apfst o apsnd)
- (delete_force "preprocessor" name);
-
-fun add_post thm thy =
- (map_exec_purge NONE o map_thmproc o apsnd)
- (insert Thm.eq_thm_prop (CodegenFunc.error_thm CodegenFunc.mk_rew thm)) thy;
- (*fully applied in order to get right context for mk_rew!*)
-
-fun del_post thm thy =
- (map_exec_purge NONE o map_thmproc o apsnd)
- (remove Thm.eq_thm_prop (CodegenFunc.error_thm CodegenFunc.mk_rew thm)) thy;
- (*fully applied in order to get right context for mk_rew!*)
-
-
-(** , post- and preprocessing **)
-
-local
-
-fun gen_apply_inline_proc prep post thy f x =
- let
- val cts = prep x;
- val rews = map CodegenFunc.assert_rew (f thy cts);
- in post rews x end;
-
-val apply_inline_proc = gen_apply_inline_proc (maps
- ((fn [args, rhs] => rhs :: (snd o Drule.strip_comb) args) o snd o Drule.strip_comb o Thm.cprop_of))
- (fn rews => map (CodegenFunc.rewrite_func rews));
-val apply_inline_proc_cterm = gen_apply_inline_proc single
- (MetaSimplifier.rewrite false);
-
-fun apply_preproc thy f [] = []
- | apply_preproc thy f (thms as (thm :: _)) =
- let
- val (const, _) = CodegenFunc.head_func thm;
- val thms' = f thy thms;
- in certify_const thy const thms' end;
-
-fun rhs_conv conv thm =
- let
- val thm' = (conv o Thm.rhs_of) thm;
- in Thm.transitive thm thm' end
-
-in
-
-fun preprocess thy thms =
- thms
- |> fold (fn (_, (_, f)) => apply_preproc thy f) ((#preprocs o the_thmproc o get_exec) thy)
- |> map (CodegenFunc.rewrite_func ((#inlines o the_thmproc o get_exec) thy))
- |> fold (fn (_, (_, f)) => apply_inline_proc thy f) ((#inline_procs o the_thmproc o get_exec) thy)
-(*FIXME - must check: rewrite rule, defining equation, proper constant |> map (snd o check_func false thy) *)
- |> common_typ_funcs;
-
-fun preprocess_conv ct =
- let
- val thy = Thm.theory_of_cterm ct;
- in
- ct
- |> MetaSimplifier.rewrite false ((#inlines o the_thmproc o get_exec) thy)
- |> fold (fn (_, (_, f)) => rhs_conv (apply_inline_proc_cterm thy f))
- ((#inline_procs o the_thmproc o get_exec) thy)
- end;
-
-fun postprocess_conv ct =
- let
- val thy = Thm.theory_of_cterm ct;
- in
- ct
- |> MetaSimplifier.rewrite false ((#posts o the_thmproc o get_exec) thy)
- end;
-
-end; (*local*)
-
-fun get_datatype thy tyco =
- case Symtab.lookup ((the_dtyps o get_exec) thy) tyco
- of SOME spec => spec
- | NONE => Sign.arity_number thy tyco
- |> Name.invents Name.context "'a"
- |> map (rpair [])
- |> rpair [];
-
-fun get_datatype_of_constr thy const =
- case CodegenConsts.co_of_const' thy const
- of SOME (tyco, (_, co)) => if member eq_co
- (Symtab.lookup (((the_dtyps o get_exec) thy)) tyco
- |> Option.map snd
- |> the_default []) co then SOME tyco else NONE
- | NONE => NONE;
-
-fun get_constr_typ thy const =
- case get_datatype_of_constr thy const
- of SOME tyco => let
- val (vs, cos) = get_datatype thy tyco;
- val (_, (_, (co, tys))) = CodegenConsts.co_of_const thy const
- in (tys ---> Type (tyco, map TFree vs))
- |> map_atyps (fn TFree (v, _) => TFree (v, AList.lookup (op =) vs v |> the))
- |> Logic.varifyT
- |> SOME end
- | NONE => NONE;
-
-fun default_typ_proto thy (const as (c, SOME tyco)) = classop_weakest_typ thy
- ((the o AxClass.class_of_param thy) c) (c, tyco) |> SOME
- | default_typ_proto thy (const as (c, NONE)) = case AxClass.class_of_param thy c
- of SOME class => SOME (Term.map_type_tvar
- (K (TVar (("'a", 0), [class]))) (Sign.the_const_type thy c))
- | NONE => get_constr_typ thy const;
-
-local
-
-fun get_funcs thy const =
- Consttab.lookup ((the_funcs o get_exec) thy) const
- |> Option.map (Susp.force o fst)
- |> these
- |> map (Thm.transfer thy);
-
-in
-
-fun these_funcs thy const =
- let
- fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals
- o ObjectLogic.drop_judgment thy o Thm.plain_prop_of);
- in
- get_funcs thy const
- |> preprocess thy
- |> drop_refl thy
- end;
-
-fun default_typ thy (const as (c, _)) = case default_typ_proto thy const
- of SOME ty => ty
- | NONE => (case get_funcs thy const
- of thm :: _ => snd (CodegenFunc.head_func thm)
- | [] => Sign.the_const_type thy c);
-
-end; (*local*)
-
-end; (*struct*)
-
-
-(** type-safe interfaces for data depedent on executable content **)
-
-functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA =
-struct
-
-type T = Data.T;
-exception Data of T;
-fun dest (Data x) = x
-
-val kind = CodegenData.declare_data (Data Data.empty)
- (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)))
- (fn thy_opt => fn cs => fn Data x => Data (Data.purge thy_opt cs x));
-
-val data_op = (kind, Data, dest);
-
-val get = CodegenData.get_data data_op;
-val change = CodegenData.change_data data_op;
-fun change_yield thy = CodegenData.change_yield_data data_op thy;
-
-end;
-
-structure CodegenData : CODEGEN_DATA =
-struct
-
-open CodegenData;
-
-end;
--- a/src/Pure/Tools/codegen_func.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-(* Title: Pure/Tools/codegen_func.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Basic handling of defining equations ("func"s) for code generator framework.
-*)
-
-signature CODEGEN_FUNC =
-sig
- val assert_rew: thm -> thm
- val mk_rew: thm -> thm
- val mk_func: thm -> thm
- val head_func: thm -> CodegenConsts.const * typ
- val bad_thm: string -> 'a
- val error_thm: (thm -> thm) -> thm -> thm
- val warning_thm: (thm -> thm) -> thm -> thm option
-
- val inst_thm: sort Vartab.table -> thm -> thm
- val expand_eta: int -> thm -> thm
- val rewrite_func: thm list -> thm -> thm
- val norm_args: thm list -> thm list
- val norm_varnames: (string -> string) -> (string -> string) -> thm list -> thm list
-end;
-
-structure CodegenFunc : CODEGEN_FUNC =
-struct
-
-
-(* auxiliary *)
-
-exception BAD_THM of string;
-fun bad_thm msg = raise BAD_THM msg;
-fun error_thm f thm = f thm handle BAD_THM msg => error msg;
-fun warning_thm f thm = SOME (f thm) handle BAD_THM msg
- => (warning ("code generator: " ^ msg); NONE);
-
-
-(* making rewrite theorems *)
-
-fun assert_rew thm =
- let
- val thy = Thm.theory_of_thm thm;
- val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
- handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
- | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
- fun vars_of t = fold_aterms
- (fn Var (v, _) => insert (op =) v
- | Free _ => bad_thm ("Illegal free variable in rewrite theorem\n"
- ^ Display.string_of_thm thm)
- | _ => I) t [];
- fun tvars_of t = fold_term_types
- (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v
- | TFree _ => bad_thm
- ("Illegal free type variable in rewrite theorem\n" ^ Display.string_of_thm thm))) t [];
- val lhs_vs = vars_of lhs;
- val rhs_vs = vars_of rhs;
- val lhs_tvs = tvars_of lhs;
- val rhs_tvs = tvars_of lhs;
- val _ = if null (subtract (op =) lhs_vs rhs_vs)
- then ()
- else bad_thm ("Free variables on right hand side of rewrite theorem\n"
- ^ Display.string_of_thm thm);
- val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
- then ()
- else bad_thm ("Free type variables on right hand side of rewrite theorem\n"
- ^ Display.string_of_thm thm)
- in thm end;
-
-fun mk_rew thm =
- let
- val thy = Thm.theory_of_thm thm;
- val ctxt = ProofContext.init thy;
- in
- thm
- |> LocalDefs.meta_rewrite_rule ctxt
- |> assert_rew
- end;
-
-
-(* making defining equations *)
-
-fun assert_func thm =
- let
- val thy = Thm.theory_of_thm thm;
- val (head, args) = (strip_comb o fst o Logic.dest_equals
- o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm;
- val _ = case head of Const _ => () | _ =>
- bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
- val _ =
- if has_duplicates (op =)
- ((fold o fold_aterms) (fn Var (v, _) => cons v
- | _ => I
- ) args [])
- then bad_thm ("Duplicated variables on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- else ()
- fun check _ (Abs _) = bad_thm
- ("Abstraction on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- | check 0 (Var _) = ()
- | check _ (Var _) = bad_thm
- ("Variable with application on left hand side of defining equation\n"
- ^ Display.string_of_thm thm)
- | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
- | check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty
- then bad_thm
- ("Partially applied constant on left hand side of equation\n"
- ^ Display.string_of_thm thm)
- else ();
- val _ = map (check 0) args;
- in thm end;
-
-val mk_func = assert_func o mk_rew;
-
-fun head_func thm =
- let
- val thy = Thm.theory_of_thm thm;
- val (Const (c_ty as (_, ty))) = (fst o strip_comb o fst o Logic.dest_equals
- o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm;
- val const = CodegenConsts.const_of_cexpr thy c_ty;
- in (const, ty) end;
-
-
-(* utilities *)
-
-fun inst_thm tvars' thm =
- let
- val thy = Thm.theory_of_thm thm;
- val tvars = (Term.add_tvars o Thm.prop_of) thm [];
- fun mk_inst (tvar as (v, _)) = case Vartab.lookup tvars' v
- of SOME sort => SOME (pairself (Thm.ctyp_of thy o TVar) (tvar, (v, sort)))
- | NONE => NONE;
- val insts = map_filter mk_inst tvars;
- in Thm.instantiate (insts, []) thm end;
-
-fun expand_eta k thm =
- let
- val thy = Thm.theory_of_thm thm;
- val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
- val (head, args) = strip_comb lhs;
- val l = if k = ~1
- then (length o fst o strip_abs) rhs
- else Int.max (0, k - length args);
- val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
- fun get_name _ 0 used = ([], used)
- | get_name (Abs (v, ty, t)) k used =
- used
- |> Name.variants [v]
- ||>> get_name t (k - 1)
- |>> (fn ([v'], vs') => (v', ty) :: vs')
- | get_name t k used =
- let
- val (tys, _) = (strip_type o fastype_of) t
- in case tys
- of [] => raise TERM ("expand_eta", [t])
- | ty :: _ =>
- used
- |> Name.variants [""]
- |-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
- #>> (fn vs' => (v, ty) :: vs'))
- end;
- val (vs, _) = get_name rhs l used;
- val vs_refl = map (fn (v, ty) => Thm.reflexive (Thm.cterm_of thy (Var ((v, 0), ty)))) vs;
- in
- thm
- |> fold (fn refl => fn thm => Thm.combination thm refl) vs_refl
- |> Conv.fconv_rule Drule.beta_eta_conversion
- end;
-
-fun rewrite_func rewrites thm =
- let
- val rewrite = MetaSimplifier.rewrite false rewrites;
- val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o Thm.cprop_of) thm;
- val Const ("==", _) = Thm.term_of ct_eq;
- val (ct_f, ct_args) = Drule.strip_comb ct_lhs;
- val rhs' = rewrite ct_rhs;
- val args' = map rewrite ct_args;
- val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1)
- args' (Thm.reflexive ct_f));
- in Thm.transitive (Thm.transitive lhs' thm) rhs' end;
-
-fun norm_args thms =
- let
- val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
- val k = fold (curry Int.max o num_args_of o Thm.plain_prop_of) thms 0;
- in
- thms
- |> map (expand_eta k)
- |> map (Conv.fconv_rule Drule.beta_eta_conversion)
- end;
-
-fun canonical_tvars purify_tvar thm =
- let
- val ctyp = Thm.ctyp_of (Thm.theory_of_thm thm);
- fun tvars_subst_for thm = (fold_types o fold_atyps)
- (fn TVar (v_i as (v, _), sort) => let
- val v' = purify_tvar v
- in if v = v' then I
- else insert (op =) (v_i, (v', sort)) end
- | _ => I) (prop_of thm) [];
- fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
- let
- val ty = TVar (v_i, sort)
- in
- (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
- end;
- val maxidx = Thm.maxidx_of thm + 1;
- val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
- in Thm.instantiate (inst, []) thm end;
-
-fun canonical_vars purify_var thm =
- let
- val cterm = Thm.cterm_of (Thm.theory_of_thm thm);
- fun vars_subst_for thm = fold_aterms
- (fn Var (v_i as (v, _), ty) => let
- val v' = purify_var v
- in if v = v' then I
- else insert (op =) (v_i, (v', ty)) end
- | _ => I) (prop_of thm) [];
- fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
- let
- val t = Var (v_i, ty)
- in
- (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
- end;
- val maxidx = Thm.maxidx_of thm + 1;
- val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
- in Thm.instantiate ([], inst) thm end;
-
-fun canonical_absvars purify_var thm =
- let
- val t = Thm.plain_prop_of thm;
- val t' = Term.map_abs_vars purify_var t;
- in Thm.rename_boundvars t t' thm end;
-
-fun norm_varnames purify_tvar purify_var thms =
- let
- fun burrow_thms f [] = []
- | burrow_thms f thms =
- thms
- |> Conjunction.intr_balanced
- |> f
- |> Conjunction.elim_balanced (length thms)
- in
- thms
- |> burrow_thms (canonical_tvars purify_tvar)
- |> map (canonical_vars purify_var)
- |> map (canonical_absvars purify_var)
- |> map Drule.zero_var_indexes
- end;
-
-end;
--- a/src/Pure/Tools/codegen_funcgr.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,409 +0,0 @@
-(* Title: Pure/Tools/codegen_funcgr.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Retrieving, normalizing and structuring defining equations
-in graph with explicit dependencies.
-*)
-
-signature CODEGEN_FUNCGR =
-sig
- type T
- val timing: bool ref
- val funcs: T -> CodegenConsts.const -> thm list
- val typ: T -> CodegenConsts.const -> typ
- val deps: T -> CodegenConsts.const list -> CodegenConsts.const list list
- val all: T -> CodegenConsts.const list
- val pretty: theory -> T -> Pretty.T
- structure Constgraph : GRAPH
-end
-
-signature CODEGEN_FUNCGR_RETRIEVAL =
-sig
- type T (* = CODEGEN_FUNCGR.T *)
- val make: theory -> CodegenConsts.const list -> T
- val make_consts: theory -> CodegenConsts.const list -> CodegenConsts.const list * T
- val make_term: theory -> (T -> (thm -> thm) -> cterm -> thm -> 'a) -> cterm -> 'a * T
- (*FIXME drop make_term as soon as possible*)
- val eval_conv: theory -> (T -> cterm -> thm) -> cterm -> thm
- val intervene: theory -> T -> T
- (*FIXME drop intervene as soon as possible*)
-end;
-
-structure CodegenFuncgr = (*signature is added later*)
-struct
-
-(** the graph type **)
-
-structure Constgraph = GraphFun (
- type key = CodegenConsts.const;
- val ord = CodegenConsts.const_ord;
-);
-
-type T = (typ * thm list) Constgraph.T;
-
-fun funcs funcgr =
- these o Option.map snd o try (Constgraph.get_node funcgr);
-
-fun typ funcgr =
- fst o Constgraph.get_node funcgr;
-
-fun deps funcgr cs =
- let
- val conn = Constgraph.strong_conn funcgr;
- val order = rev conn;
- in
- (map o filter) (member (op =) (Constgraph.all_succs funcgr cs)) order
- |> filter_out null
- end;
-
-fun all funcgr = Constgraph.keys funcgr;
-
-fun pretty thy funcgr =
- AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr)
- |> (map o apfst) (CodegenConsts.string_of_const thy)
- |> sort (string_ord o pairself fst)
- |> map (fn (s, thms) =>
- (Pretty.block o Pretty.fbreaks) (
- Pretty.str s
- :: map Display.pretty_thm thms
- ))
- |> Pretty.chunks;
-
-
-(** generic combinators **)
-
-fun fold_consts f thms =
- thms
- |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of)
- |> (fold o fold_aterms) (fn Const c => f c | _ => I);
-
-fun consts_of (const, []) = []
- | consts_of (const, thms as thm :: _) =
- let
- val thy = Thm.theory_of_thm thm;
- val is_refl = curry CodegenConsts.eq_const const;
- fun the_const c = case try (CodegenConsts.const_of_cexpr thy) c
- of SOME const => if is_refl const then I else insert CodegenConsts.eq_const const
- | NONE => I
- in fold_consts the_const thms [] end;
-
-fun insts_of thy algebra c ty_decl ty =
- let
- val tys_decl = Sign.const_typargs thy (c, ty_decl);
- val tys = Sign.const_typargs thy (c, ty);
- fun class_relation (x, _) _ = x;
- fun type_constructor tyco xs class =
- (tyco, class) :: maps (maps fst) xs;
- fun type_variable (TVar (_, sort)) = map (pair []) sort
- | type_variable (TFree (_, sort)) = map (pair []) sort;
- fun mk_inst ty (TVar (_, sort)) = cons (ty, sort)
- | mk_inst ty (TFree (_, sort)) = cons (ty, sort)
- | mk_inst (Type (_, tys1)) (Type (_, tys2)) = fold2 mk_inst tys1 tys2;
- fun of_sort_deriv (ty, sort) =
- Sorts.of_sort_derivation (Sign.pp thy) algebra
- { class_relation = class_relation, type_constructor = type_constructor,
- type_variable = type_variable }
- (ty, sort)
- in
- flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl []))
- end;
-
-fun drop_classes thy tfrees thm =
- let
- val (_, thm') = Thm.varifyT' [] thm;
- val tvars = Term.add_tvars (Thm.prop_of thm') [];
- val unconstr = map (Thm.ctyp_of thy o TVar) tvars;
- val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy)
- (TVar (v_i, []), TFree (v, sort))) tvars tfrees;
- in
- thm'
- |> fold Thm.unconstrainT unconstr
- |> Thm.instantiate (instmap, [])
- |> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy))
- end;
-
-
-(** graph algorithm **)
-
-val timing = ref false;
-
-local
-
-exception INVALID of CodegenConsts.const list * string;
-
-fun resort_thms algebra tap_typ [] = []
- | resort_thms algebra tap_typ (thms as thm :: _) =
- let
- val thy = Thm.theory_of_thm thm;
- val cs = fold_consts (insert (op =)) thms [];
- fun match_const c (ty, ty_decl) =
- let
- val tys = CodegenConsts.typargs thy (c, ty);
- val sorts = map (snd o dest_TVar) (CodegenConsts.typargs thy (c, ty_decl));
- in fold2 (curry (CodegenConsts.typ_sort_inst algebra)) tys sorts end;
- fun match (c_ty as (c, ty)) =
- case tap_typ c_ty
- of SOME ty_decl => match_const c (ty, ty_decl)
- | NONE => I;
- val tvars = fold match cs Vartab.empty;
- in map (CodegenFunc.inst_thm tvars) thms end;
-
-fun resort_funcss thy algebra funcgr =
- let
- val typ_funcgr = try (fst o Constgraph.get_node funcgr o CodegenConsts.const_of_cexpr thy);
- fun resort_dep (const, thms) = (const, resort_thms algebra typ_funcgr thms)
- handle Sorts.CLASS_ERROR e => raise INVALID ([const], Sorts.msg_class_error (Sign.pp thy) e
- ^ ",\nfor constant " ^ CodegenConsts.string_of_const thy const
- ^ "\nin defining equations\n"
- ^ (cat_lines o map string_of_thm) thms)
- fun resort_rec tap_typ (const, []) = (true, (const, []))
- | resort_rec tap_typ (const, thms as thm :: _) =
- let
- val (_, ty) = CodegenFunc.head_func thm;
- val thms' as thm' :: _ = resort_thms algebra tap_typ thms
- val (_, ty') = CodegenFunc.head_func thm';
- in (Sign.typ_equiv thy (ty, ty'), (const, thms')) end;
- fun resort_recs funcss =
- let
- fun tap_typ c_ty = case try (CodegenConsts.const_of_cexpr thy) c_ty
- of SOME const => AList.lookup (CodegenConsts.eq_const) funcss const
- |> these
- |> try hd
- |> Option.map (snd o CodegenFunc.head_func)
- | NONE => NONE;
- val (unchangeds, funcss') = split_list (map (resort_rec tap_typ) funcss);
- val unchanged = fold (fn x => fn y => x andalso y) unchangeds true;
- in (unchanged, funcss') end;
- fun resort_rec_until funcss =
- let
- val (unchanged, funcss') = resort_recs funcss;
- in if unchanged then funcss' else resort_rec_until funcss' end;
- in map resort_dep #> resort_rec_until end;
-
-fun instances_of thy algebra insts =
- let
- val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy;
- fun all_classops tyco class =
- try (AxClass.params_of_class thy) class
- |> Option.map snd
- |> these
- |> map (fn (c, _) => (c, SOME tyco))
- in
- Symtab.empty
- |> fold (fn (tyco, class) =>
- Symtab.map_default (tyco, []) (insert (op =) class)) insts
- |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops tyco)
- (Graph.all_succs thy_classes classes))) tab [])
- end;
-
-fun instances_of_consts thy algebra funcgr consts =
- let
- fun inst (cexpr as (c, ty)) = insts_of thy algebra c
- ((fst o Constgraph.get_node funcgr o CodegenConsts.const_of_cexpr thy) cexpr)
- ty handle CLASS_ERROR => [];
- in
- []
- |> fold (fold (insert (op =)) o inst) consts
- |> instances_of thy algebra
- end;
-
-fun ensure_const' rewrites thy algebra funcgr const auxgr =
- if can (Constgraph.get_node funcgr) const
- then (NONE, auxgr)
- else if can (Constgraph.get_node auxgr) const
- then (SOME const, auxgr)
- else if is_some (CodegenData.get_datatype_of_constr thy const) then
- auxgr
- |> Constgraph.new_node (const, [])
- |> pair (SOME const)
- else let
- val thms = CodegenData.these_funcs thy const
- |> map (CodegenFunc.rewrite_func (rewrites thy))
- |> CodegenFunc.norm_args
- |> CodegenFunc.norm_varnames CodegenNames.purify_tvar CodegenNames.purify_var;
- val rhs = consts_of (const, thms);
- in
- auxgr
- |> Constgraph.new_node (const, thms)
- |> fold_map (ensure_const rewrites thy algebra funcgr) rhs
- |-> (fn rhs' => fold (fn SOME const' => Constgraph.add_edge (const, const')
- | NONE => I) rhs')
- |> pair (SOME const)
- end
-and ensure_const rewrites thy algebra funcgr const =
- let
- val timeap = if !timing
- then Output.timeap_msg ("time for " ^ CodegenConsts.string_of_const thy const)
- else I;
- in timeap (ensure_const' rewrites thy algebra funcgr const) end;
-
-fun merge_funcss rewrites thy algebra raw_funcss funcgr =
- let
- val funcss = raw_funcss
- |> resort_funcss thy algebra funcgr
- |> filter_out (can (Constgraph.get_node funcgr) o fst);
- fun typ_func const [] = CodegenData.default_typ thy const
- | typ_func (_, NONE) (thm :: _) = (snd o CodegenFunc.head_func) thm
- | typ_func (const as (c, SOME tyco)) (thms as (thm :: _)) =
- let
- val (_, ty) = CodegenFunc.head_func thm;
- val SOME class = AxClass.class_of_param thy c;
- val sorts_decl = Sorts.mg_domain algebra tyco [class];
- val tys = CodegenConsts.typargs thy (c, ty);
- val sorts = map (snd o dest_TVar) tys;
- in if sorts = sorts_decl then ty
- else raise INVALID ([const], "Illegal instantation for class operation "
- ^ CodegenConsts.string_of_const thy const
- ^ "\nin defining equations\n"
- ^ (cat_lines o map string_of_thm) thms)
- end;
- fun add_funcs (const, thms) =
- Constgraph.new_node (const, (typ_func const thms, thms));
- fun add_deps (funcs as (const, thms)) funcgr =
- let
- val deps = consts_of funcs;
- val insts = instances_of_consts thy algebra funcgr
- (fold_consts (insert (op =)) thms []);
- in
- funcgr
- |> ensure_consts' rewrites thy algebra insts
- |> fold (curry Constgraph.add_edge const) deps
- |> fold (curry Constgraph.add_edge const) insts
- end;
- in
- funcgr
- |> fold add_funcs funcss
- |> fold add_deps funcss
- end
-and ensure_consts' rewrites thy algebra cs funcgr =
- let
- val auxgr = Constgraph.empty
- |> fold (snd oo ensure_const rewrites thy algebra funcgr) cs;
- in
- funcgr
- |> fold (merge_funcss rewrites thy algebra)
- (map (AList.make (Constgraph.get_node auxgr))
- (rev (Constgraph.strong_conn auxgr)))
- end handle INVALID (cs', msg)
- => raise INVALID (fold (insert CodegenConsts.eq_const) cs' cs, msg);
-
-fun ensure_consts rewrites thy consts funcgr =
- let
- val algebra = CodegenData.coregular_algebra thy
- in ensure_consts' rewrites thy algebra consts funcgr
- handle INVALID (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) "
- ^ commas (map (CodegenConsts.string_of_const thy) cs'))
- end;
-
-in
-
-(** retrieval interfaces **)
-
-val ensure_consts = ensure_consts;
-
-fun check_consts rewrites thy consts funcgr =
- let
- val algebra = CodegenData.coregular_algebra thy;
- fun try_const const funcgr =
- (SOME const, ensure_consts' rewrites thy algebra [const] funcgr)
- handle INVALID (cs', msg) => (NONE, funcgr);
- val (consts', funcgr') = fold_map try_const consts funcgr;
- in (map_filter I consts', funcgr') end;
-
-fun ensure_consts_term rewrites thy f ct funcgr =
- let
- fun consts_of thy t =
- fold_aterms (fn Const c => cons (CodegenConsts.const_of_cexpr thy c) | _ => I) t []
- fun rhs_conv conv thm =
- let
- val thm' = (conv o Thm.rhs_of) thm;
- in Thm.transitive thm thm' end
- val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct);
- val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) ();
- val thm1 = CodegenData.preprocess_conv ct
- |> fold (rhs_conv o MetaSimplifier.rewrite false o single) (rewrites thy);
- val ct' = Thm.rhs_of thm1;
- val consts = consts_of thy (Thm.term_of ct');
- val funcgr' = ensure_consts rewrites thy consts funcgr;
- val algebra = CodegenData.coregular_algebra thy;
- val (_, thm2) = Thm.varifyT' [] thm1;
- val thm3 = Thm.reflexive (Thm.rhs_of thm2);
- val typ_funcgr = try (fst o Constgraph.get_node funcgr' o CodegenConsts.const_of_cexpr thy);
- val [thm4] = resort_thms algebra typ_funcgr [thm3];
- val tfrees = Term.add_tfrees (Thm.prop_of thm1) [];
- fun inst thm =
- let
- val tvars = Term.add_tvars (Thm.prop_of thm) [];
- val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy)
- (TVar (v_i, sort), TFree (v, sort))) tvars tfrees;
- in Thm.instantiate (instmap, []) thm end;
- val thm5 = inst thm2;
- val thm6 = inst thm4;
- val ct'' = Thm.rhs_of thm6;
- val cs = fold_aterms (fn Const c => cons c | _ => I) (Thm.term_of ct'') [];
- val drop = drop_classes thy tfrees;
- val instdefs = instances_of_consts thy algebra funcgr' cs;
- val funcgr'' = ensure_consts rewrites thy instdefs funcgr';
- in (f funcgr'' drop ct'' thm5, funcgr'') end;
-
-fun ensure_consts_eval rewrites thy conv =
- let
- fun conv' funcgr drop_classes ct thm1 =
- let
- val thm2 = conv funcgr ct;
- val thm3 = CodegenData.postprocess_conv (Thm.rhs_of thm2);
- val thm23 = drop_classes (Thm.transitive thm2 thm3);
- in
- Thm.transitive thm1 thm23 handle THM _ =>
- error ("eval_conv - could not construct proof:\n"
- ^ (cat_lines o map string_of_thm) [thm1, thm2, thm3])
- end;
- in ensure_consts_term rewrites thy conv' end;
-
-end; (*local*)
-
-end; (*struct*)
-
-functor CodegenFuncgrRetrieval (val rewrites: theory -> thm list) : CODEGEN_FUNCGR_RETRIEVAL =
-struct
-
-(** code data **)
-
-type T = CodegenFuncgr.T;
-
-structure Funcgr = CodeDataFun
-(struct
- type T = T;
- val empty = CodegenFuncgr.Constgraph.empty;
- fun merge _ _ = CodegenFuncgr.Constgraph.empty;
- fun purge _ NONE _ = CodegenFuncgr.Constgraph.empty
- | purge _ (SOME cs) funcgr =
- CodegenFuncgr.Constgraph.del_nodes ((CodegenFuncgr.Constgraph.all_preds funcgr
- o filter (can (CodegenFuncgr.Constgraph.get_node funcgr))) cs) funcgr;
-end);
-
-fun make thy =
- Funcgr.change thy o CodegenFuncgr.ensure_consts rewrites thy;
-
-fun make_consts thy =
- Funcgr.change_yield thy o CodegenFuncgr.check_consts rewrites thy;
-
-fun make_term thy f =
- Funcgr.change_yield thy o CodegenFuncgr.ensure_consts_term rewrites thy f;
-
-fun eval_conv thy f =
- fst o Funcgr.change_yield thy o CodegenFuncgr.ensure_consts_eval rewrites thy f;
-
-(*FIXME*)
-fun intervene thy funcgr =
- Funcgr.change thy (K funcgr);
-
-end; (*functor*)
-
-structure CodegenFuncgr : CODEGEN_FUNCGR =
-struct
-
-open CodegenFuncgr;
-
-end; (*struct*)
--- a/src/Pure/Tools/codegen_names.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,426 +0,0 @@
-(* Title: Pure/Tools/codegen_names.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Naming policies for code generation: prefixing any name by corresponding theory name,
-conversion to alphanumeric representation.
-Mappings are incrementally cached.
-*)
-
-signature CODEGEN_NAMES =
-sig
- val purify_var: string -> string
- val purify_tvar: string -> string
- val check_modulename: string -> string
- type var_ctxt;
- val make_vars: string list -> var_ctxt;
- val intro_vars: string list -> var_ctxt -> var_ctxt;
- val lookup_var: var_ctxt -> string -> string;
-
- type tyco = string
- type const = string
- val class: theory -> class -> class
- val class_rev: theory -> class -> class option
- val classrel: theory -> class * class -> string
- val classrel_rev: theory -> string -> (class * class) option
- val tyco: theory -> tyco -> tyco
- val tyco_rev: theory -> tyco -> tyco option
- val instance: theory -> class * tyco -> string
- val instance_rev: theory -> string -> (class * tyco) option
- val const: theory -> CodegenConsts.const -> const
- val const_rev: theory -> const -> CodegenConsts.const option
- val labelled_name: theory -> string -> string
-end;
-
-structure CodegenNames: CODEGEN_NAMES =
-struct
-
-(** purification **)
-
-fun purify_name upper_else_lower =
- let
- fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse s = "'";
- val is_junk = not o is_valid andf Symbol.is_regular;
- val junk = Scan.many is_junk;
- val scan_valids = Symbol.scanner "Malformed input"
- ((junk |--
- (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
- --| junk))
- -- Scan.repeat ((Scan.many1 is_valid >> implode) --| junk) >> op ::);
- fun upper_lower cs = if upper_else_lower then nth_map 0 Symbol.to_ascii_upper cs
- else (if forall Symbol.is_ascii_upper cs
- then map else nth_map 0) Symbol.to_ascii_lower cs;
- in
- explode
- #> scan_valids
- #> space_implode "_"
- #> explode
- #> upper_lower
- #> implode
- end;
-
-fun purify_var "" = "x"
- | purify_var v = purify_name false v;
-
-fun purify_tvar "" = "'a"
- | purify_tvar v =
- (unprefix "'" #> explode #> filter Symbol.is_ascii_letter #> cons "'" #> implode) v;
-
-fun check_modulename mn =
- let
- val mns = NameSpace.explode mn;
- val mns' = map (purify_name true) mns;
- in
- if mns' = mns then mn else error ("Invalid module name: " ^ quote mn ^ "\n"
- ^ "perhaps try " ^ quote (NameSpace.implode mns'))
- end;
-
-
-(** variable name contexts **)
-
-type var_ctxt = string Symtab.table * Name.context;
-
-fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
- Name.make_context names);
-
-fun intro_vars names (namemap, namectxt) =
- let
- val (names', namectxt') = Name.variants names namectxt;
- val namemap' = fold2 (curry Symtab.update) names names' namemap;
- in (namemap', namectxt') end;
-
-fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
- of SOME name' => name'
- | NONE => error ("Invalid name in context: " ^ quote name);
-
-
-
-(** global names (identifiers) **)
-
-(* identifier categories *)
-
-val idf_class = "class";
-val idf_classrel = "clsrel"
-val idf_tyco = "tyco";
-val idf_instance = "inst";
-val idf_const = "const";
-
-fun string_of_classrel (class, superclass) = class ^ " < " ^ superclass;
-fun string_of_instance (class, tyco) = tyco ^ " :: " ^ class;
-
-fun add_idf nsp name =
- NameSpace.append name nsp;
-
-fun dest_idf nsp name =
- if NameSpace.base name = nsp
- then SOME (NameSpace.qualifier name)
- else NONE;
-
-local
-
-val name_mapping = [
- (idf_class, "class"),
- (idf_classrel, "subclass relation"),
- (idf_tyco, "type constructor"),
- (idf_instance, "instance"),
- (idf_const, "constant")
-]
-
-in
-
-val category_of = the o AList.lookup (op =) name_mapping o NameSpace.base;
-
-end;
-
-
-(* theory name lookup *)
-
-fun thyname_of thy f errmsg x =
- let
- fun thy_of thy =
- if f thy x then case get_first thy_of (Theory.parents_of thy)
- of NONE => SOME thy
- | thy => thy
- else NONE;
- in case thy_of thy
- of SOME thy => Context.theory_name thy
- | NONE => error (errmsg x) end;
-
-fun thyname_of_class thy =
- thyname_of thy (fn thy => member (op =) (Sign.all_classes thy))
- (fn class => "thyname_of_class: no such class: " ^ quote class);
-
-fun thyname_of_classrel thy =
- thyname_of thy (fn thy => fn (class1, class2) => Sign.subsort thy ([class1], [class2]))
- (fn (class1, class2) => "thyname_of_classrel: no such subclass relation: "
- ^ (quote o string_of_classrel) (class1, class2));
-
-fun thyname_of_tyco thy =
- thyname_of thy Sign.declared_tyname
- (fn tyco => "thyname_of_tyco: no such type constructor: " ^ quote tyco);
-
-fun thyname_of_instance thy =
- let
- fun test_instance thy (class, tyco) =
- can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
- in thyname_of thy test_instance
- (fn (class, tyco) => "thyname_of_instance: no such instance: "
- ^ (quote o string_of_instance) (class, tyco))
- end;
-
-fun thyname_of_const thy =
- thyname_of thy Sign.declared_const
- (fn c => "thyname_of_const: no such constant: " ^ quote c);
-
-
-(* naming policies *)
-
-val purify_prefix =
- explode
- (*should disappear as soon as hierarchical theory name spaces are available*)
- #> Symbol.scanner "Malformed name"
- (Scan.repeat ($$ "_" |-- $$ "_" >> (fn _ => ".") || Scan.one Symbol.is_regular))
- #> implode
- #> NameSpace.explode
- #> map (purify_name true);
-
-fun purify_base "op =" = "eq"
- | purify_base "op &" = "and"
- | purify_base "op |" = "or"
- | purify_base "op -->" = "implies"
- | purify_base "{}" = "empty"
- | purify_base "op :" = "member"
- | purify_base "op Int" = "intersect"
- | purify_base "op Un" = "union"
- | purify_base "*" = "product"
- | purify_base "+" = "sum"
- | purify_base s = purify_name false s;
-
-fun default_policy thy get_basename get_thyname name =
- let
- val prefix = (purify_prefix o get_thyname thy) name;
- val base = (purify_base o get_basename) name;
- in NameSpace.implode (prefix @ [base]) end;
-
-fun class_policy thy = default_policy thy NameSpace.base thyname_of_class;
-fun classrel_policy thy = default_policy thy (fn (class1, class2) =>
- NameSpace.base class2 ^ "_" ^ NameSpace.base class1) thyname_of_classrel;
- (*order fits nicely with composed projections*)
-fun tyco_policy thy = default_policy thy NameSpace.base thyname_of_tyco;
-fun instance_policy thy = default_policy thy (fn (class, tyco) =>
- NameSpace.base class ^ "_" ^ NameSpace.base tyco) thyname_of_instance;
-
-fun force_thyname thy (const as (c, opt_tyco)) =
- case CodegenData.get_datatype_of_constr thy const
- of SOME dtco => SOME (thyname_of_tyco thy dtco)
- | NONE => (case AxClass.class_of_param thy c
- of SOME class => (case opt_tyco
- of SOME tyco => SOME (thyname_of_instance thy (class, tyco))
- | NONE => SOME (thyname_of_class thy class))
- | NONE => NONE);
-
-fun const_policy thy (const as (c, opt_tyco)) =
- case force_thyname thy const
- of NONE => default_policy thy NameSpace.base thyname_of_const c
- | SOME thyname => let
- val prefix = purify_prefix thyname;
- val tycos = the_list opt_tyco;
- val base = map (purify_base o NameSpace.base) (c :: tycos);
- in NameSpace.implode (prefix @ [space_implode "_" base]) end;
-
-
-(* theory and code data *)
-
-type tyco = string;
-type const = string;
-structure Consttab = CodegenConsts.Consttab;
-
-structure StringPairTab =
- TableFun(
- type key = string * string;
- val ord = prod_ord fast_string_ord fast_string_ord;
- );
-
-datatype names = Names of {
- class: class Symtab.table * class Symtab.table,
- classrel: string StringPairTab.table * (class * class) Symtab.table,
- tyco: tyco Symtab.table * tyco Symtab.table,
- instance: string StringPairTab.table * (class * tyco) Symtab.table
-}
-
-val empty_names = Names {
- class = (Symtab.empty, Symtab.empty),
- classrel = (StringPairTab.empty, Symtab.empty),
- tyco = (Symtab.empty, Symtab.empty),
- instance = (StringPairTab.empty, Symtab.empty)
-};
-
-local
- fun mk_names (class, classrel, tyco, instance) =
- Names { class = class, classrel = classrel, tyco = tyco, instance = instance };
- fun map_names f (Names { class, classrel, tyco, instance }) =
- mk_names (f (class, classrel, tyco, instance));
-in
- fun merge_names (Names { class = (class1, classrev1),
- classrel = (classrel1, classrelrev1), tyco = (tyco1, tycorev1),
- instance = (instance1, instancerev1) },
- Names { class = (class2, classrev2),
- classrel = (classrel2, classrelrev2), tyco = (tyco2, tycorev2),
- instance = (instance2, instancerev2) }) =
- mk_names ((Symtab.merge (op =) (class1, class2), Symtab.merge (op =) (classrev1, classrev2)),
- (StringPairTab.merge (op =) (classrel1, classrel2), Symtab.merge (op =) (classrelrev1, classrelrev2)),
- (Symtab.merge (op =) (tyco1, tyco2), Symtab.merge (op =) (tycorev1, tycorev2)),
- (StringPairTab.merge (op =) (instance1, instance2), Symtab.merge (op =) (instancerev1, instancerev2)));
- fun map_class f = map_names
- (fn (class, classrel, tyco, inst) => (f class, classrel, tyco, inst));
- fun map_classrel f = map_names
- (fn (class, classrel, tyco, inst) => (class, f classrel, tyco, inst));
- fun map_tyco f = map_names
- (fn (class, classrel, tyco, inst) => (class, classrel, f tyco, inst));
- fun map_instance f = map_names
- (fn (class, classrel, tyco, inst) => (class, classrel, tyco, f inst));
-end; (*local*)
-
-structure CodeName = TheoryDataFun
-(
- type T = names ref;
- val empty = ref empty_names;
- fun copy (ref names) = ref names;
- val extend = copy;
- fun merge _ (ref names1, ref names2) = ref (merge_names (names1, names2));
-);
-
-structure ConstName = CodeDataFun
-(
- type T = const Consttab.table * (string * string option) Symtab.table;
- val empty = (Consttab.empty, Symtab.empty);
- fun merge _ ((const1, constrev1), (const2, constrev2)) : T =
- (Consttab.merge (op =) (const1, const2),
- Symtab.merge CodegenConsts.eq_const (constrev1, constrev2));
- fun purge _ NONE _ = empty
- | purge _ (SOME cs) (const, constrev) = (fold Consttab.delete_safe cs const,
- fold Symtab.delete (map_filter (Consttab.lookup const) cs) constrev);
-);
-
-val _ = Context.add_setup (CodeName.init);
-
-
-(* forward lookup with cache update *)
-
-fun get thy get_tabs get upd_names upd policy x =
- let
- val names_ref = CodeName.get thy;
- val (Names names) = ! names_ref;
- val tabs = get_tabs names;
- fun declare name =
- let
- val names' = upd_names (K (upd (x, name) (fst tabs),
- Symtab.update_new (name, x) (snd tabs))) (Names names)
- in (names_ref := names'; name) end;
- in case get (fst tabs) x
- of SOME name => name
- | NONE =>
- x
- |> policy thy
- |> Name.variant (Symtab.keys (snd tabs))
- |> declare
- end;
-
-fun get_const thy const =
- let
- val tabs = ConstName.get thy;
- fun declare name =
- let
- val names' = (Consttab.update (const, name) (fst tabs),
- Symtab.update_new (name, const) (snd tabs))
- in (ConstName.change thy (K names'); name) end;
- in case Consttab.lookup (fst tabs) const
- of SOME name => name
- | NONE =>
- const
- |> const_policy thy
- |> Name.variant (Symtab.keys (snd tabs))
- |> declare
- end;
-
-
-(* backward lookup *)
-
-fun rev thy get_tabs name =
- let
- val names_ref = CodeName.get thy
- val (Names names) = ! names_ref;
- val tab = (snd o get_tabs) names;
- in case Symtab.lookup tab name
- of SOME x => x
- | NONE => error ("No such " ^ category_of name ^ ": " ^ quote name)
- end;
-
-fun rev_const thy name =
- let
- val tab = snd (ConstName.get thy);
- in case Symtab.lookup tab name
- of SOME const => const
- | NONE => error ("No such " ^ category_of name ^ ": " ^ quote name)
- end;
-
-
-(* external interfaces *)
-
-fun class thy =
- get thy #class Symtab.lookup map_class Symtab.update class_policy
- #> add_idf idf_class;
-fun classrel thy =
- get thy #classrel StringPairTab.lookup map_classrel StringPairTab.update classrel_policy
- #> add_idf idf_classrel;
-fun tyco thy =
- get thy #tyco Symtab.lookup map_tyco Symtab.update tyco_policy
- #> add_idf idf_tyco;
-fun instance thy =
- get thy #instance StringPairTab.lookup map_instance StringPairTab.update instance_policy
- #> add_idf idf_instance;
-fun const thy =
- get_const thy
- #> add_idf idf_const;
-
-fun class_rev thy =
- dest_idf idf_class
- #> Option.map (rev thy #class);
-fun classrel_rev thy =
- dest_idf idf_classrel
- #> Option.map (rev thy #classrel);
-fun tyco_rev thy =
- dest_idf idf_tyco
- #> Option.map (rev thy #tyco);
-fun instance_rev thy =
- dest_idf idf_instance
- #> Option.map (rev thy #instance);
-fun const_rev thy =
- dest_idf idf_const
- #> Option.map (rev_const thy);
-
-local
-
-val f_mapping = [
- (idf_class, class_rev),
- (idf_classrel, Option.map string_of_classrel oo classrel_rev),
- (idf_tyco, tyco_rev),
- (idf_instance, Option.map string_of_instance oo instance_rev),
- (idf_const, fn thy => Option.map (CodegenConsts.string_of_const thy) o const_rev thy)
-];
-
-in
-
-fun labelled_name thy idf =
- let
- val category = category_of idf;
- val name = NameSpace.qualifier idf;
- val f = (the o AList.lookup (op =) f_mapping o NameSpace.base) idf
- in case f thy idf
- of SOME thing => category ^ " " ^ quote thing
- | NONE => error ("Unknown name: " ^ quote name)
- end;
-
-end;
-
-end;
--- a/src/Pure/Tools/codegen_package.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,673 +0,0 @@
-(* Title: Pure/Tools/codegen_package.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Code generator translation kernel. Code generator Isar setup.
-*)
-
-signature CODEGEN_PACKAGE =
-sig
- (* interfaces *)
- val eval_conv: theory
- -> (CodegenThingol.code -> CodegenThingol.iterm -> cterm -> thm) -> cterm -> thm;
- val codegen_command: theory -> string -> unit;
-
- (* primitive interfaces *)
- val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a;
- val satisfies_ref: bool option ref;
- val satisfies: theory -> term -> string list -> bool;
-
- (* axiomatic interfaces *)
- type appgen;
- val add_appconst: string * appgen -> theory -> theory;
- val appgen_let: appgen;
- val appgen_if: appgen;
- val appgen_case: (theory -> term
- -> ((string * typ) list * ((term * typ) * (term * term) list)) option)
- -> appgen;
-
- val timing: bool ref;
-end;
-
-structure CodegenPackage : CODEGEN_PACKAGE =
-struct
-
-open BasicCodegenThingol;
-val tracing = CodegenThingol.tracing;
-val succeed = CodegenThingol.succeed;
-val fail = CodegenThingol.fail;
-
-(** code translation **)
-
-(* theory data *)
-
-type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T
- -> CodegenFuncgr.T
- -> (string * typ) * term list -> CodegenThingol.transact -> iterm * CodegenThingol.transact;
-
-type appgens = (int * (appgen * stamp)) Symtab.table;
-val merge_appgens : appgens * appgens -> appgens =
- Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) =>
- bounds1 = bounds2 andalso stamp1 = stamp2);
-
-structure Consttab = CodegenConsts.Consttab;
-type abstypes = typ Symtab.table * CodegenConsts.const Consttab.table;
-fun merge_abstypes ((typs1, consts1) : abstypes, (typs2, consts2) : abstypes) =
- (Symtab.merge (Type.eq_type Vartab.empty) (typs1, typs2),
- Consttab.merge CodegenConsts.eq_const (consts1, consts2));
-
-structure Translation = TheoryDataFun
-(
- type T = appgens * abstypes;
- val empty = (Symtab.empty, (Symtab.empty, Consttab.empty));
- val copy = I;
- val extend = I;
- fun merge _ ((appgens1, abstypes1), (appgens2, abstypes2)) =
- (merge_appgens (appgens1, appgens2), merge_abstypes (abstypes1, abstypes2));
-);
-
-structure Funcgr = CodegenFuncgrRetrieval (fun rewrites thy = []);
-
-fun code_depgr thy [] = Funcgr.make thy []
- | code_depgr thy consts =
- let
- val gr = Funcgr.make thy consts;
- val select = CodegenFuncgr.Constgraph.all_succs gr consts;
- in
- CodegenFuncgr.Constgraph.subgraph
- (member CodegenConsts.eq_const select) gr
- end;
-
-fun code_thms thy =
- Pretty.writeln o CodegenFuncgr.pretty thy o code_depgr thy;
-
-fun code_deps thy consts =
- let
- val gr = code_depgr thy consts;
- fun mk_entry (const, (_, (_, parents))) =
- let
- val name = CodegenConsts.string_of_const thy const;
- val nameparents = map (CodegenConsts.string_of_const thy) parents;
- in { name = name, ID = name, dir = "", unfold = true,
- path = "", parents = nameparents }
- end;
- val prgr = CodegenFuncgr.Constgraph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr [];
- in Present.display_graph prgr end;
-
-structure Code = CodeDataFun
-(
- type T = CodegenThingol.code;
- val empty = CodegenThingol.empty_code;
- fun merge _ = CodegenThingol.merge_code;
- fun purge _ NONE _ = CodegenThingol.empty_code
- | purge NONE _ _ = CodegenThingol.empty_code
- | purge (SOME thy) (SOME cs) code =
- let
- val cs_exisiting =
- map_filter (CodegenNames.const_rev thy) (Graph.keys code);
- val dels = (Graph.all_preds code
- o map (CodegenNames.const thy)
- o filter (member CodegenConsts.eq_const cs_exisiting)
- ) cs;
- in Graph.del_nodes dels code end;
-);
-
-
-(* translation kernel *)
-
-fun get_abstype thy (tyco, tys) = case Symtab.lookup ((fst o snd o Translation.get) thy) tyco
- of SOME ty => SOME ((map_atyps (fn TFree (n, _) => nth tys (the (Int.fromString n)))) ty)
- | NONE => NONE;
-
-fun ensure_def thy = CodegenThingol.ensure_def (CodegenNames.labelled_name thy);
-
-fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class =
- let
- val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
- val (v, cs) = AxClass.params_of_class thy class;
- val class' = CodegenNames.class thy class;
- val classrels' = map (curry (CodegenNames.classrel thy) class) superclasses;
- val classops' = map (CodegenNames.const thy o CodegenConsts.const_of_cexpr thy) cs;
- val defgen_class =
- fold_map (ensure_def_class thy algbr funcgr) superclasses
- ##>> (fold_map (exprgen_typ thy algbr funcgr) o map snd) cs
- #-> (fn (superclasses, classoptyps) => succeed
- (CodegenThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps))))
- in
- tracing (fn _ => "generating class " ^ quote class)
- #> ensure_def thy defgen_class ("generating class " ^ quote class) class'
- #> pair class'
- end
-and ensure_def_classrel thy algbr funcgr (subclass, superclass) =
- ensure_def_class thy algbr funcgr subclass
- #>> (fn _ => CodegenNames.classrel thy (subclass, superclass))
-and ensure_def_tyco thy algbr funcgr "fun" trns =
- trns
- |> pair "fun"
- | ensure_def_tyco thy algbr funcgr tyco trns =
- let
- fun defgen_datatype trns =
- let
- val (vs, cos) = CodegenData.get_datatype thy tyco;
- in
- trns
- |> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
- ||>> fold_map (fn (c, tys) =>
- fold_map (exprgen_typ thy algbr funcgr) tys
- #-> (fn tys' =>
- pair ((CodegenNames.const thy o CodegenConsts.const_of_cexpr thy)
- (c, tys ---> Type (tyco, map TFree vs)), tys'))) cos
- |-> (fn (vs, cos) => succeed (CodegenThingol.Datatype (vs, cos)))
- end;
- val tyco' = CodegenNames.tyco thy tyco;
- in
- trns
- |> tracing (fn _ => "generating type constructor " ^ quote tyco)
- |> ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco'
- |> pair tyco'
- end
-and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) trns =
- trns
- |> fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort)
- |>> (fn sort => (unprefix "'" v, sort))
-and exprgen_typ thy algbr funcgr (TFree vs) trns =
- trns
- |> exprgen_tyvar_sort thy algbr funcgr vs
- |>> (fn (v, sort) => ITyVar v)
- | exprgen_typ thy algbr funcgr (Type (tyco, tys)) trns =
- case get_abstype thy (tyco, tys)
- of SOME ty =>
- trns
- |> exprgen_typ thy algbr funcgr ty
- | NONE =>
- trns
- |> ensure_def_tyco thy algbr funcgr tyco
- ||>> fold_map (exprgen_typ thy algbr funcgr) tys
- |>> (fn (tyco, tys) => tyco `%% tys);
-
-exception CONSTRAIN of (string * typ) * typ;
-val timing = ref false;
-
-fun exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) =
- let
- val pp = Sign.pp thy;
- datatype typarg =
- Global of (class * string) * typarg list list
- | Local of (class * class) list * (string * (int * sort));
- fun class_relation (Global ((_, tyco), yss), _) class =
- Global ((class, tyco), yss)
- | class_relation (Local (classrels, v), subclass) superclass =
- Local ((subclass, superclass) :: classrels, v);
- fun type_constructor tyco yss class =
- Global ((class, tyco), (map o map) fst yss);
- fun type_variable (TFree (v, sort)) =
- let
- val sort' = proj_sort sort;
- in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
- val typargs = Sorts.of_sort_derivation pp algebra
- {class_relation = class_relation, type_constructor = type_constructor,
- type_variable = type_variable}
- (ty_ctxt, proj_sort sort_decl);
- fun mk_dict (Global (inst, yss)) =
- ensure_def_inst thy algbr funcgr inst
- ##>> (fold_map o fold_map) mk_dict yss
- #>> (fn (inst, dss) => DictConst (inst, dss))
- | mk_dict (Local (classrels, (v, (k, sort)))) =
- fold_map (ensure_def_classrel thy algbr funcgr) classrels
- #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
- in
- fold_map mk_dict typargs
- end
-and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) =
- let
- val c' = CodegenConsts.const_of_cexpr thy (c, ty_ctxt)
- val idf = CodegenNames.const thy c';
- val ty_decl = Consts.the_declaration consts idf;
- val (tys, tys_decl) = pairself (curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl);
- val sorts = map (snd o dest_TVar) tys_decl;
- in
- fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts)
- end
-and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) trns =
- let
- val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
- val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
- val vs = Name.names (Name.declare var Name.context) "'a" (Sorts.mg_domain algebra tyco [class]);
- val arity_typ = Type (tyco, map TFree vs);
- fun gen_superarity superclass trns =
- trns
- |> ensure_def_class thy algbr funcgr superclass
- ||>> ensure_def_classrel thy algbr funcgr (class, superclass)
- ||>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass])
- |>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
- (superclass, (classrel, (inst, dss))));
- fun gen_classop_def (classop as (c, ty)) trns =
- trns
- |> ensure_def_const thy algbr funcgr (CodegenConsts.const_of_cexpr thy classop)
- ||>> exprgen_term thy algbr funcgr (Const (c, map_type_tfree (K arity_typ) ty));
- fun defgen_inst trns =
- trns
- |> ensure_def_class thy algbr funcgr class
- ||>> ensure_def_tyco thy algbr funcgr tyco
- ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
- ||>> fold_map gen_superarity superclasses
- ||>> fold_map gen_classop_def classops
- |-> (fn ((((class, tyco), arity), superarities), classops) =>
- succeed (CodegenThingol.Classinst ((class, (tyco, arity)), (superarities, classops))));
- val inst = CodegenNames.instance thy (class, tyco);
- in
- trns
- |> tracing (fn _ => "generating instance " ^ quote class ^ " / " ^ quote tyco)
- |> ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst
- |> pair inst
- end
-and ensure_def_const thy (algbr as (_, consts)) funcgr (const as (c, opt_tyco)) trns =
- let
- val c' = CodegenNames.const thy const;
- fun defgen_datatypecons trns =
- trns
- |> ensure_def_tyco thy algbr funcgr
- ((the o CodegenData.get_datatype_of_constr thy) const)
- |-> (fn _ => succeed CodegenThingol.Bot);
- fun defgen_classop trns =
- trns
- |> ensure_def_class thy algbr funcgr ((the o AxClass.class_of_param thy) c)
- |-> (fn _ => succeed CodegenThingol.Bot);
- fun defgen_fun trns =
- let
- val const' = perhaps (Consttab.lookup ((snd o snd o Translation.get) thy)) const;
- val raw_thms = CodegenFuncgr.funcs funcgr const';
- val ty = (Logic.unvarifyT o CodegenFuncgr.typ funcgr) const';
- val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty
- then raw_thms
- else map (CodegenFunc.expand_eta 1) raw_thms;
- val timeap = if !timing then Output.timeap_msg ("time for " ^ c')
- else I;
- val msg = cat_lines ("generating code for theorems " :: map string_of_thm thms);
- val vs = (map dest_TFree o Consts.typargs consts) (c', ty);
- val dest_eqthm =
- apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify o prop_of;
- fun exprgen_eq (args, rhs) trns =
- trns
- |> fold_map (exprgen_term thy algbr funcgr) args
- ||>> exprgen_term thy algbr funcgr rhs;
- in
- trns
- |> CodegenThingol.message msg (fn trns => trns
- |> timeap (fold_map (exprgen_eq o dest_eqthm) thms)
- ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
- ||>> exprgen_typ thy algbr funcgr ty
- |-> (fn ((eqs, vs), ty) => succeed (CodegenThingol.Fun (eqs, (vs, ty)))))
- end;
- val defgen = if (is_some o CodegenData.get_datatype_of_constr thy) const
- then defgen_datatypecons
- else if is_some opt_tyco
- orelse (not o is_some o AxClass.class_of_param thy) c
- then defgen_fun
- else defgen_classop
- in
- trns
- |> tracing (fn _ => "generating constant "
- ^ (quote o CodegenConsts.string_of_const thy) const)
- |> ensure_def thy defgen ("generating constant " ^ CodegenConsts.string_of_const thy const) c'
- |> pair c'
- end
-and exprgen_term thy algbr funcgr (Const (c, ty)) trns =
- trns
- |> select_appgen thy algbr funcgr ((c, ty), [])
- | exprgen_term thy algbr funcgr (Free (v, ty)) trns =
- trns
- |> exprgen_typ thy algbr funcgr ty
- |>> (fn _ => IVar v)
- | exprgen_term thy algbr funcgr (Abs (raw_v, ty, raw_t)) trns =
- let
- val (v, t) = Syntax.variant_abs (raw_v, ty, raw_t);
- in
- trns
- |> exprgen_typ thy algbr funcgr ty
- ||>> exprgen_term thy algbr funcgr t
- |>> (fn (ty, t) => (v, ty) `|-> t)
- end
- | exprgen_term thy algbr funcgr (t as _ $ _) trns =
- case strip_comb t
- of (Const (c, ty), ts) =>
- trns
- |> select_appgen thy algbr funcgr ((c, ty), ts)
- | (t', ts) =>
- trns
- |> exprgen_term thy algbr funcgr t'
- ||>> fold_map (exprgen_term thy algbr funcgr) ts
- |>> (fn (t, ts) => t `$$ ts)
-and appgen_default thy algbr funcgr ((c, ty), ts) trns =
- trns
- |> ensure_def_const thy algbr funcgr (CodegenConsts.const_of_cexpr thy (c, ty))
- ||>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty)
- ||>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)
- ||>> exprgen_dict_parms thy algbr funcgr (c, ty)
- ||>> fold_map (exprgen_term thy algbr funcgr) ts
- |>> (fn ((((c, tys), ty), iss), ts) => IConst (c, (iss, tys)) `$$ ts)
-and select_appgen thy algbr funcgr ((c, ty), ts) trns =
- case Symtab.lookup (fst (Translation.get thy)) c
- of SOME (i, (appgen, _)) =>
- if length ts < i then
- let
- val k = length ts;
- val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty;
- val ctxt = (fold o fold_aterms)
- (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
- val vs = Name.names ctxt "a" tys;
- in
- trns
- |> fold_map (exprgen_typ thy algbr funcgr) tys
- ||>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs)
- |>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t)
- end
- else if length ts > i then
- trns
- |> appgen thy algbr funcgr ((c, ty), Library.take (i, ts))
- ||>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts))
- |>> (fn (t, ts) => t `$$ ts)
- else
- trns
- |> appgen thy algbr funcgr ((c, ty), ts)
- | NONE =>
- trns
- |> appgen_default thy algbr funcgr ((c, ty), ts);
-
-
-(* entrance points into translation kernel *)
-
-fun ensure_def_const' thy algbr funcgr c trns =
- ensure_def_const thy algbr funcgr c trns
- handle CONSTRAIN ((c, ty), ty_decl) => error (
- "Constant " ^ c ^ " with most general type\n"
- ^ CodegenConsts.string_of_typ thy ty
- ^ "\noccurs with type\n"
- ^ CodegenConsts.string_of_typ thy ty_decl);
-
-fun perhaps_def_const thy algbr funcgr c trns =
- case try (ensure_def_const thy algbr funcgr c) trns
- of SOME (c, trns) => (SOME c, trns)
- | NONE => (NONE, trns);
-
-fun exprgen_term' thy algbr funcgr t trns =
- exprgen_term thy algbr funcgr t trns
- handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t
- ^ ",\nconstant " ^ c ^ " with most general type\n"
- ^ CodegenConsts.string_of_typ thy ty
- ^ "\noccurs with type\n"
- ^ CodegenConsts.string_of_typ thy ty_decl);
-
-
-(* parametrized application generators, for instantiation in object logic *)
-(* (axiomatic extensions of translation kernel) *)
-
-fun appgen_case dest_case_expr thy algbr funcgr (app as (c_ty, ts)) =
- let
- val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts));
- fun clause_gen (dt, bt) =
- exprgen_term thy algbr funcgr dt
- ##>> exprgen_term thy algbr funcgr bt;
- in
- exprgen_term thy algbr funcgr st
- ##>> exprgen_typ thy algbr funcgr sty
- ##>> fold_map clause_gen ds
- ##>> appgen_default thy algbr funcgr app
- #>> (fn (((se, sty), ds), t0) => ICase (((se, sty), ds), t0))
- end;
-
-fun appgen_let thy algbr funcgr (app as (_, [st, ct])) =
- exprgen_term thy algbr funcgr ct
- ##>> exprgen_term thy algbr funcgr st
- ##>> appgen_default thy algbr funcgr app
- #>> (fn (((v, ty) `|-> be, se), t0) =>
- ICase (CodegenThingol.collapse_let (((v, ty), se), be), t0)
- | (_, t0) => t0);
-
-fun appgen_if thy algbr funcgr (app as (_, [tb, tt, tf])) =
- exprgen_term thy algbr funcgr tb
- ##>> exprgen_typ thy algbr funcgr (Type ("bool", []))
- ##>> exprgen_term thy algbr funcgr (Const ("True", Type ("bool", [])))
- ##>> exprgen_term thy algbr funcgr tt
- ##>> exprgen_term thy algbr funcgr (Const ("False", Type ("bool", [])))
- ##>> exprgen_term thy algbr funcgr tf
- ##>> appgen_default thy algbr funcgr app
- #>> (fn ((((((tb, B), T), tt), F), tf), t0) => ICase (((tb, B), [(T, tt), (F, tf)]), t0));
-
-fun add_appconst (c, appgen) thy =
- let
- val i = (length o fst o strip_type o Sign.the_const_type thy) c;
- val _ = Code.change thy (K CodegenThingol.empty_code);
- in
- (Translation.map o apfst)
- (Symtab.update (c, (i, (appgen, stamp ())))) thy
- end;
-
-
-
-(** abstype and constsubst interface **)
-
-local
-
-fun add_consts thy f (c1, c2 as (c, opt_tyco)) =
- let
- val _ = if
- is_some (AxClass.class_of_param thy c) andalso is_none opt_tyco
- orelse is_some (CodegenData.get_datatype_of_constr thy c2)
- then error ("Not a function: " ^ CodegenConsts.string_of_const thy c2)
- else ();
- val funcgr = Funcgr.make thy [c1, c2];
- val ty1 = (f o CodegenFuncgr.typ funcgr) c1;
- val ty2 = CodegenFuncgr.typ funcgr c2;
- val _ = if Sign.typ_equiv thy (ty1, ty2) then () else
- error ("Incompatiable type signatures of " ^ CodegenConsts.string_of_const thy c1
- ^ " and " ^ CodegenConsts.string_of_const thy c2 ^ ":\n"
- ^ CodegenConsts.string_of_typ thy ty1 ^ "\n" ^ CodegenConsts.string_of_typ thy ty2);
- in Consttab.update (c1, c2) end;
-
-fun gen_abstyp prep_const prep_typ (raw_abstyp, raw_substtyp) raw_absconsts thy =
- let
- val abstyp = Type.no_tvars (prep_typ thy raw_abstyp);
- val substtyp = Type.no_tvars (prep_typ thy raw_substtyp);
- val absconsts = (map o pairself) (prep_const thy) raw_absconsts;
- val Type (abstyco, tys) = abstyp handle BIND => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
- val typarms = map (fst o dest_TFree) tys handle MATCH => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
- fun mk_index v =
- let
- val k = find_index (fn w => v = w) typarms;
- in if k = ~1
- then error ("Free type variable: " ^ quote v)
- else TFree (string_of_int k, [])
- end;
- val typpat = map_atyps (fn TFree (v, _) => mk_index v) substtyp;
- fun apply_typpat (Type (tyco, tys)) =
- let
- val tys' = map apply_typpat tys;
- in if tyco = abstyco then
- (map_atyps (fn TFree (n, _) => nth tys' (the (Int.fromString n)))) typpat
- else
- Type (tyco, tys')
- end
- | apply_typpat ty = ty;
- val _ = Code.change thy (K CodegenThingol.empty_code);
- in
- thy
- |> (Translation.map o apsnd) (fn (abstypes, abscs) =>
- (abstypes
- |> Symtab.update (abstyco, typpat),
- abscs
- |> fold (add_consts thy apply_typpat) absconsts)
- )
- end;
-
-fun gen_constsubst prep_const raw_constsubsts thy =
- let
- val constsubsts = (map o pairself) (prep_const thy) raw_constsubsts;
- val _ = Code.change thy (K CodegenThingol.empty_code);
- in
- thy
- |> (Translation.map o apsnd o apsnd) (fold (add_consts thy I) constsubsts)
- end;
-
-in
-
-val abstyp = gen_abstyp (K I) Sign.certify_typ;
-val abstyp_e = gen_abstyp CodegenConsts.read_const Sign.read_typ;
-
-val constsubst = gen_constsubst (K I);
-val constsubst_e = gen_constsubst CodegenConsts.read_const;
-
-end; (*local*)
-
-
-(** code generation interfaces **)
-
-(* generic generation combinators *)
-
-fun generate thy funcgr gen it =
- let
- (*FIXME*)
- val _ = Funcgr.intervene thy funcgr;
- val cs = map_filter (Consttab.lookup ((snd o snd o Translation.get) thy))
- (CodegenFuncgr.all funcgr);
- val funcgr' = Funcgr.make thy cs;
- val naming = NameSpace.qualified_names NameSpace.default_naming;
- val consttab = Consts.empty
- |> fold (fn c => Consts.declare naming
- ((CodegenNames.const thy c, CodegenFuncgr.typ funcgr' c), true))
- (CodegenFuncgr.all funcgr');
- val algbr = (CodegenData.operational_algebra thy, consttab);
- in
- Code.change_yield thy
- (CodegenThingol.start_transact (gen thy algbr funcgr' it))
- |> fst
- end;
-
-fun eval_conv thy conv =
- let
- fun conv' funcgr ct =
- let
- val t = generate thy funcgr exprgen_term' (Thm.term_of ct);
- val consts = CodegenThingol.fold_constnames (insert (op =)) t [];
- val code = Code.get thy
- |> CodegenThingol.project_code true [] (SOME consts)
- in conv code t ct end;
- in Funcgr.eval_conv thy conv' end;
-
-fun codegen_term thy t =
- let
- val ct = Thm.cterm_of thy t;
- val (ct', funcgr) = Funcgr.make_term thy (K (K K)) ct;
- val t' = Thm.term_of ct';
- in generate thy funcgr exprgen_term' t' end;
-
-fun raw_eval_term thy (ref_spec, t) args =
- let
- val _ = (Term.map_types o Term.map_atyps) (fn _ =>
- error ("Term " ^ Sign.string_of_term thy t ^ " contains polymorphic type"))
- t;
- val t' = codegen_term thy t;
- in
- CodegenSerializer.eval_term thy CodegenNames.labelled_name
- (Code.get thy) (ref_spec, t') args
- end;
-
-val satisfies_ref : bool option ref = ref NONE;
-
-fun eval_term thy t = raw_eval_term thy t [];
-fun satisfies thy t witnesses = raw_eval_term thy
- (("CodegenPackage.satisfies_ref", satisfies_ref), t) witnesses;
-
-fun filter_generatable thy consts =
- let
- val (consts', funcgr) = Funcgr.make_consts thy consts;
- val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts';
- val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE)
- (consts' ~~ consts'');
- in consts''' end;
-
-
-(** toplevel interface and setup **)
-
-local
-
-structure P = OuterParse
-and K = OuterKeyword
-
-fun code raw_cs seris thy =
- let
- val (perm1, cs) = CodegenConsts.read_const_exprs thy
- (filter_generatable thy) raw_cs;
- val (perm2, cs') = case generate thy (Funcgr.make thy cs) (fold_map ooo ensure_def_const') cs
- of [] => (true, NONE)
- | cs => (false, SOME cs);
- val code = Code.get thy;
- val seris' = map (fn (((target, module), file), args) =>
- CodegenSerializer.get_serializer thy target (perm1 orelse perm2) module file args
- CodegenNames.labelled_name cs') seris;
- in
- (map (fn f => f code) seris' : unit list; ())
- end;
-
-fun code_thms_cmd thy =
- code_thms thy o snd o CodegenConsts.read_const_exprs thy (fst o Funcgr.make_consts thy);
-
-fun code_deps_cmd thy =
- code_deps thy o snd o CodegenConsts.read_const_exprs thy (fst o Funcgr.make_consts thy);
-
-val (inK, toK, fileK) = ("in", "to", "file");
-
-val code_exprP =
- (Scan.repeat P.term
- -- Scan.repeat (P.$$$ inK |-- P.name
- -- Scan.option (P.$$$ toK |-- P.name)
- -- Scan.option (P.$$$ fileK |-- P.name)
- -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") []
- ) >> (fn (raw_cs, seris) => code raw_cs seris));
-
-val _ = OuterSyntax.add_keywords [inK, toK, fileK];
-
-val (codeK, code_abstypeK, code_axiomsK, code_thmsK, code_depsK) =
- ("code_gen", "code_abstype", "code_axioms", "code_thms", "code_deps");
-
-in
-
-val codeP =
- OuterSyntax.improper_command codeK "generate executable code for constants"
- K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
-
-fun codegen_command thy cmd =
- case Scan.read OuterLex.stopper (P.!!! code_exprP) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd)
- of SOME f => (writeln "Now generating code..."; f thy)
- | NONE => error ("Bad directive " ^ quote cmd);
-
-val code_abstypeP =
- OuterSyntax.command code_abstypeK "axiomatic abstypes for code generation" K.thy_decl (
- (P.typ -- P.typ -- Scan.optional (P.$$$ "where" |-- Scan.repeat1
- (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)) [])
- >> (Toplevel.theory o uncurry abstyp_e)
- );
-
-val code_axiomsP =
- OuterSyntax.command code_axiomsK "axiomatic constant equalities for code generation" K.thy_decl (
- Scan.repeat1 (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)
- >> (Toplevel.theory o constsubst_e)
- );
-
-val code_thmsP =
- OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag
- (Scan.repeat P.term
- >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
- o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
-
-val code_depsP =
- OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag
- (Scan.repeat P.term
- >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
- o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
-
-val _ = OuterSyntax.add_parsers [codeP, code_abstypeP, code_axiomsP, code_thmsP, code_depsP];
-
-end; (* local *)
-
-end; (* struct *)
--- a/src/Pure/Tools/codegen_serializer.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2215 +0,0 @@
-(* Title: Pure/Tools/codegen_serializer.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Serializer from intermediate language ("Thin-gol") to
-target languages (like SML or Haskell).
-*)
-
-signature CODEGEN_SERIALIZER =
-sig
- include BASIC_CODEGEN_THINGOL;
-
- val add_syntax_class: string -> class
- -> (string * (CodegenConsts.const * string) list) option -> theory -> theory;
- val add_syntax_inst: string -> string * class -> bool -> theory -> theory;
- val add_syntax_tycoP: string -> string -> OuterParse.token list
- -> (theory -> theory) * OuterParse.token list;
- val add_syntax_constP: string -> string -> OuterParse.token list
- -> (theory -> theory) * OuterParse.token list;
-
- val add_undefined: string -> string -> string -> theory -> theory;
- val add_pretty_list: string -> string -> string -> theory -> theory;
- val add_pretty_list_string: string -> string -> string
- -> string -> string list -> theory -> theory;
- val add_pretty_char: string -> string -> string list -> theory -> theory
- val add_pretty_numeral: string -> bool -> string * typ -> string -> string -> string
- -> string -> string -> theory -> theory;
- val add_pretty_ml_string: string -> string -> string list -> string
- -> string -> string -> theory -> theory;
- val add_pretty_imperative_monad_bind: string -> string -> theory -> theory;
-
- type serializer;
- val add_serializer: string * serializer -> theory -> theory;
- val get_serializer: theory -> string -> bool -> string option -> string option -> Args.T list
- -> (theory -> string -> string) -> string list option -> CodegenThingol.code -> unit;
- val assert_serializer: theory -> string -> string;
-
- val eval_verbose: bool ref;
- val eval_term: theory -> (theory -> string -> string) -> CodegenThingol.code
- -> (string * 'a option ref) * CodegenThingol.iterm -> string list -> 'a;
- val code_width: int ref;
-end;
-
-structure CodegenSerializer : CODEGEN_SERIALIZER =
-struct
-
-open BasicCodegenThingol;
-val tracing = CodegenThingol.tracing;
-
-(** basics **)
-
-infixr 5 @@;
-infixr 5 @|;
-fun x @@ y = [x, y];
-fun xs @| y = xs @ [y];
-val str = PrintMode.with_default Pretty.str;
-val concat = Pretty.block o Pretty.breaks;
-val brackets = Pretty.enclose "(" ")" o Pretty.breaks;
-fun semicolon ps = Pretty.block [concat ps, str ";"];
-
-
-(** syntax **)
-
-datatype lrx = L | R | X;
-
-datatype fixity =
- BR
- | NOBR
- | INFX of (int * lrx);
-
-val APP = INFX (~1, L);
-
-fun eval_lrx L L = false
- | eval_lrx R R = false
- | eval_lrx _ _ = true;
-
-fun eval_fxy NOBR NOBR = false
- | eval_fxy BR NOBR = false
- | eval_fxy NOBR BR = false
- | eval_fxy (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) =
- pr < pr_ctxt
- orelse pr = pr_ctxt
- andalso eval_lrx lr lr_ctxt
- orelse pr_ctxt = ~1
- | eval_fxy _ (INFX _) = false
- | eval_fxy (INFX _) NOBR = false
- | eval_fxy _ _ = true;
-
-fun gen_brackify _ [p] = p
- | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
- | gen_brackify false (ps as _::_) = Pretty.block ps;
-
-fun brackify fxy_ctxt ps =
- gen_brackify (eval_fxy BR fxy_ctxt) (Pretty.breaks ps);
-
-fun brackify_infix infx fxy_ctxt ps =
- gen_brackify (eval_fxy (INFX infx) fxy_ctxt) (Pretty.breaks ps);
-
-type class_syntax = string * (string -> string option);
-type typ_syntax = int * ((fixity -> itype -> Pretty.T)
- -> fixity -> itype list -> Pretty.T);
-type term_syntax = int * ((CodegenNames.var_ctxt -> fixity -> iterm -> Pretty.T)
- -> CodegenNames.var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
-
-
-(* user-defined syntax *)
-
-datatype 'a mixfix =
- Arg of fixity
- | Pretty of Pretty.T;
-
-fun mk_mixfix prep_arg (fixity_this, mfx) =
- let
- fun is_arg (Arg _) = true
- | is_arg _ = false;
- val i = (length o filter is_arg) mfx;
- fun fillin _ [] [] =
- []
- | fillin pr (Arg fxy :: mfx) (a :: args) =
- (pr fxy o prep_arg) a :: fillin pr mfx args
- | fillin pr (Pretty p :: mfx) args =
- p :: fillin pr mfx args
- | fillin _ [] _ =
- error ("Inconsistent mixfix: too many arguments")
- | fillin _ _ [] =
- error ("Inconsistent mixfix: too less arguments");
- in
- (i, fn pr => fn fixity_ctxt => fn args =>
- gen_brackify (eval_fxy fixity_this fixity_ctxt) (fillin pr mfx args))
- end;
-
-fun parse_infix prep_arg (x, i) s =
- let
- val l = case x of L => INFX (i, L) | _ => INFX (i, X);
- val r = case x of R => INFX (i, R) | _ => INFX (i, X);
- in
- mk_mixfix prep_arg (INFX (i, x), [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r])
- end;
-
-fun parse_mixfix prep_arg s =
- let
- val sym_any = Scan.one Symbol.is_regular;
- val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat (
- ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
- || ($$ "_" >> K (Arg BR))
- || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length))
- || (Scan.repeat1
- ( $$ "'" |-- sym_any
- || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
- sym_any) >> (Pretty o str o implode)));
- in case Scan.finite Symbol.stopper parse (Symbol.explode s)
- of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p)
- | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p)
- | _ => Scan.!! (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail ()
- end;
-
-fun parse_args f args =
- case Scan.read Args.stopper f args
- of SOME x => x
- | NONE => error "Bad serializer arguments";
-
-
-(* generic serializer combinators *)
-
-fun gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons
- lhs vars fxy (app as ((c, (_, tys)), ts)) =
- case const_syntax c
- of NONE => if lhs andalso not (is_cons c) then
- error ("non-constructor on left hand side of equation: " ^ labelled_name c)
- else brackify fxy (pr_app' lhs vars app)
- | SOME (i, pr) =>
- let
- val k = if i < 0 then length tys else i;
- fun pr' fxy ts = pr (pr_term lhs) vars fxy (ts ~~ curry Library.take k tys);
- in if k = length ts
- then pr' fxy ts
- else if k < length ts
- then case chop k ts of (ts1, ts2) =>
- brackify fxy (pr' APP ts1 :: map (pr_term lhs vars BR) ts2)
- else pr_term lhs vars fxy (CodegenThingol.eta_expand app k)
- end;
-
-fun gen_pr_bind pr_bind' pr_term fxy ((v, pat), ty) vars =
- let
- val vs = case pat
- of SOME pat => CodegenThingol.fold_varnames (insert (op =)) pat []
- | NONE => [];
- val vars' = CodegenNames.intro_vars (the_list v) vars;
- val vars'' = CodegenNames.intro_vars vs vars';
- val v' = Option.map (CodegenNames.lookup_var vars') v;
- val pat' = Option.map (pr_term vars'' fxy) pat;
- in (pr_bind' ((v', pat'), ty), vars'') end;
-
-
-(* list, char, string, numeral and monad abstract syntax transformations *)
-
-fun implode_list c_nil c_cons t =
- let
- fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
- if c = c_cons
- then SOME (t1, t2)
- else NONE
- | dest_cons _ = NONE;
- val (ts, t') = CodegenThingol.unfoldr dest_cons t;
- in case t'
- of IConst (c, _) => if c = c_nil then SOME ts else NONE
- | _ => NONE
- end;
-
-fun decode_char c_nibbles (IConst (c1, _), IConst (c2, _)) =
- let
- fun idx c = find_index (curry (op =) c) 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 c_char c_nibbles mk_char mk_string ts =
- let
- fun implode_char (IConst (c, _) `$ t1 `$ t2) =
- if c = c_char then decode_char c_nibbles (t1, t2) else NONE
- | implode_char _ = NONE;
- val ts' = map implode_char ts;
- in if forall is_some ts'
- then (SOME o str o mk_string o implode o map_filter I) ts'
- else NONE
- end;
-
-fun implode_numeral c_bit0 c_bit1 c_pls c_min c_bit =
- let
- fun dest_bit (IConst (c, _)) = if c = c_bit0 then SOME 0
- else if c = c_bit1 then SOME 1
- else NONE
- | dest_bit _ = NONE;
- fun dest_numeral (IConst (c, _)) = if c = c_pls then SOME (IntInf.fromInt 0)
- else if c = c_min then SOME (IntInf.fromInt ~1)
- else NONE
- | dest_numeral (IConst (c, _) `$ t1 `$ t2) =
- if c = c_bit then case (dest_numeral t1, dest_bit t2)
- of (SOME n, SOME b) => SOME (IntInf.fromInt 2 * n + IntInf.fromInt b)
- | _ => NONE
- else NONE
- | dest_numeral _ = NONE;
- in dest_numeral end;
-
-fun implode_monad c_mbind c_kbind t =
- let
- fun dest_monad (IConst (c, _) `$ t1 `$ t2) =
- if c = c_mbind
- then case CodegenThingol.split_abs t2
- of SOME (((v, pat), ty), t') => SOME ((SOME (((SOME v, pat), ty), true), t1), t')
- | NONE => NONE
- else if c = c_kbind
- then SOME ((NONE, t1), t2)
- else NONE
- | dest_monad t = case CodegenThingol.split_let t
- of SOME (((pat, ty), tbind), t') => SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t')
- | NONE => NONE;
- in CodegenThingol.unfoldr dest_monad t end;
-
-
-(** name auxiliary **)
-
-val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
-val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
-
-val dest_name =
- apfst NameSpace.implode o split_last o fst o split_last o NameSpace.explode;
-
-fun mk_modl_name_tab init_names prefix module_alias code =
- let
- fun nsp_map f = NameSpace.explode #> map f #> NameSpace.implode;
- fun mk_alias name =
- case module_alias name
- of SOME name' => name'
- | NONE => nsp_map (fn name => (the_single o fst)
- (Name.variants [name] init_names)) name;
- fun mk_prefix name =
- case prefix
- of SOME prefix => NameSpace.append prefix name
- | NONE => name;
- val tab =
- Symtab.empty
- |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
- o fst o dest_name o fst)
- code
- in fn name => (the o Symtab.lookup tab) name end;
-
-
-
-(** SML/OCaml serializer **)
-
-datatype ml_def =
- MLFuns of (string * ((iterm list * iterm) list * typscheme)) list
- | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list
- | MLClass of string * ((class * string) list * (vname * (string * itype) list))
- | MLClassinst of string * ((class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * (string * iterm) list));
-
-fun pr_sml tyco_syntax const_syntax labelled_name init_syms deresolv is_cons ml_def =
- let
- val pr_label_classrel = translate_string (fn "." => "__" | c => c) o NameSpace.qualifier;
- val pr_label_classop = NameSpace.base o NameSpace.qualifier;
- fun pr_dicts fxy ds =
- let
- fun pr_dictvar (v, (_, 1)) = first_upper v ^ "_"
- | pr_dictvar (v, (i, _)) = first_upper v ^ string_of_int (i+1) ^ "_";
- fun pr_proj [] p =
- p
- | pr_proj [p'] p =
- brackets [p', p]
- | pr_proj (ps as _ :: _) p =
- brackets [Pretty.enum " o" "(" ")" ps, p];
- fun pr_dictc fxy (DictConst (inst, dss)) =
- brackify fxy ((str o deresolv) inst :: map (pr_dicts BR) dss)
- | pr_dictc fxy (DictVar (classrels, v)) =
- pr_proj (map (str o deresolv) classrels) ((str o pr_dictvar) v)
- in case ds
- of [] => str "()"
- | [d] => pr_dictc fxy d
- | _ :: _ => (Pretty.list "(" ")" o map (pr_dictc NOBR)) ds
- end;
- fun pr_tyvars vs =
- vs
- |> map (fn (v, sort) => map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort)
- |> map (pr_dicts BR);
- fun pr_tycoexpr fxy (tyco, tys) =
- let
- val tyco' = (str o deresolv) tyco
- in case map (pr_typ BR) tys
- of [] => tyco'
- | [p] => Pretty.block [p, Pretty.brk 1, tyco']
- | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
- end
- and pr_typ fxy (tyco `%% tys) =
- (case tyco_syntax tyco
- of NONE => pr_tycoexpr fxy (tyco, tys)
- | SOME (i, pr) =>
- if not (i = length tys)
- then error ("Number of argument mismatch in customary serialization: "
- ^ (string_of_int o length) tys ^ " given, "
- ^ string_of_int i ^ " expected")
- else pr pr_typ fxy tys)
- | pr_typ fxy (ITyVar v) =
- str ("'" ^ v);
- fun pr_term lhs vars fxy (IConst c) =
- pr_app lhs vars fxy (c, [])
- | pr_term lhs vars fxy (IVar v) =
- str (CodegenNames.lookup_var vars v)
- | pr_term lhs vars fxy (t as t1 `$ t2) =
- (case CodegenThingol.unfold_const_app t
- of SOME c_ts => pr_app lhs vars fxy c_ts
- | NONE =>
- brackify fxy [pr_term lhs vars NOBR t1, pr_term lhs vars BR t2])
- | pr_term lhs vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = CodegenThingol.unfold_abs t;
- fun pr ((v, pat), ty) =
- pr_bind NOBR ((SOME v, pat), ty)
- #>> (fn p => concat [str "fn", p, str "=>"]);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (ps @ [pr_term lhs vars' NOBR t']) end
- | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodegenThingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
- then pr_case vars fxy cases
- else pr_app lhs vars fxy c_ts
- | NONE => pr_case vars fxy cases)
- and pr_app' lhs vars (app as ((c, (iss, tys)), ts)) =
- if is_cons c then let
- val k = length tys
- in if k < 2 then
- (str o deresolv) c :: map (pr_term lhs vars BR) ts
- else if k = length ts then
- [(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term lhs vars NOBR) ts)]
- else [pr_term lhs vars BR (CodegenThingol.eta_expand app k)] end else
- (str o deresolv) c
- :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term lhs vars BR) ts
- and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
- and pr_bind' ((NONE, NONE), _) = str "_"
- | pr_bind' ((SOME v, NONE), _) = str v
- | pr_bind' ((NONE, SOME p), _) = p
- | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p]
- and pr_bind fxy = gen_pr_bind pr_bind' (pr_term false) fxy
- and pr_case vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, t') = CodegenThingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind NOBR ((NONE, SOME pat), ty)
- |>> (fn p => semicolon [str "val", p, str "=", pr_term false vars NOBR t])
- val (ps, vars') = fold_map pr binds vars;
- in
- Pretty.chunks [
- [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block,
- [str ("in"), Pretty.fbrk, pr_term false vars' NOBR t'] |> Pretty.block,
- str ("end")
- ]
- end
- | pr_case vars fxy (((td, ty), b::bs), _) =
- let
- fun pr delim (pat, t) =
- let
- val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
- in
- concat [str delim, p, str "=>", pr_term false vars' NOBR t]
- end;
- in
- (Pretty.enclose "(" ")" o single o brackify fxy) (
- str "case"
- :: pr_term false vars NOBR td
- :: pr "of" b
- :: map (pr "|") bs
- )
- end
- | pr_case vars fxy ((_, []), _) = str "raise Fail \"empty case\""
- fun pr_def (MLFuns (funns as (funn :: funns'))) =
- let
- val definer =
- let
- fun mk [] [] = "val"
- | mk (_::_) _ = "fun"
- | mk [] vs = if (null o filter_out (null o snd)) vs then "val" else "fun";
- fun chk (_, ((ts, _) :: _, (vs, _))) NONE = SOME (mk ts vs)
- | chk (_, ((ts, _) :: _, (vs, _))) (SOME defi) =
- if defi = mk ts vs then SOME defi
- else error ("Mixing simultaneous vals and funs not implemented: "
- ^ commas (map (labelled_name o fst) funns));
- in the (fold chk funns NONE) end;
- fun pr_funn definer (name, (eqs as eq::eqs', (raw_vs, ty))) =
- let
- val vs = filter_out (null o snd) raw_vs;
- val shift = if null eqs' then I else
- map (Pretty.block o single o Pretty.block o single);
- fun pr_eq definer (ts, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- ((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts
- |> CodegenNames.intro_vars ((fold o CodegenThingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- concat (
- [str definer, (str o deresolv) name]
- @ (if null ts andalso null vs
- andalso not (ty = ITyVar "_")(*for evaluation*)
- then [str ":", pr_typ NOBR ty]
- else
- pr_tyvars vs
- @ map (pr_term true vars BR) ts)
- @ [str "=", pr_term false vars NOBR t]
- )
- end
- in
- (Pretty.block o Pretty.fbreaks o shift) (
- pr_eq definer eq
- :: map (pr_eq "|") eqs'
- )
- end;
- val (ps, p) = split_last (pr_funn definer funn :: map (pr_funn "and") funns');
- in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end
- | pr_def (MLDatas (datas as (data :: datas'))) =
- let
- fun pr_co (co, []) =
- str (deresolv co)
- | pr_co (co, tys) =
- concat [
- str (deresolv co),
- str "of",
- Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
- ];
- fun pr_data definer (tyco, (vs, [])) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- @@ str "EMPTY__"
- )
- | pr_data definer (tyco, (vs, cos)) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- :: separate (str "|") (map pr_co cos)
- );
- val (ps, p) = split_last (pr_data "datatype" data :: map (pr_data "and") datas');
- in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end
- | pr_def (MLClass (class, (superclasses, (v, classops)))) =
- let
- val w = first_upper v ^ "_";
- fun pr_superclass_field (class, classrel) =
- (concat o map str) [
- pr_label_classrel classrel, ":", "'" ^ v, deresolv class
- ];
- fun pr_classop_field (classop, ty) =
- concat [
- (str o pr_label_classop) classop, str ":", pr_typ NOBR ty
- ];
- fun pr_classop_proj (classop, _) =
- semicolon [
- str "fun",
- (str o deresolv) classop,
- Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
- str "=",
- str ("#" ^ pr_label_classop classop),
- str w
- ];
- fun pr_superclass_proj (_, classrel) =
- semicolon [
- str "fun",
- (str o deresolv) classrel,
- Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
- str "=",
- str ("#" ^ pr_label_classrel classrel),
- str w
- ];
- in
- Pretty.chunks (
- concat [
- str ("type '" ^ v),
- (str o deresolv) class,
- str "=",
- Pretty.enum "," "{" "};" (
- map pr_superclass_field superclasses @ map pr_classop_field classops
- )
- ]
- :: map pr_superclass_proj superclasses
- @ map pr_classop_proj classops
- )
- end
- | pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) =
- let
- fun pr_superclass (_, (classrel, dss)) =
- concat [
- (str o pr_label_classrel) classrel,
- str "=",
- pr_dicts NOBR [DictConst dss]
- ];
- fun pr_classop (classop, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- (CodegenThingol.fold_constnames (insert (op =)) t []);
- val vars = CodegenNames.intro_vars consts init_syms;
- in
- concat [
- (str o pr_label_classop) classop,
- str "=",
- pr_term false vars NOBR t
- ]
- end;
- in
- semicolon ([
- str (if null arity then "val" else "fun"),
- (str o deresolv) inst ] @
- pr_tyvars arity @ [
- str "=",
- Pretty.enum "," "{" "}" (map pr_superclass superarities @ map pr_classop classop_defs),
- str ":",
- pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
- ])
- end;
- in pr_def ml_def end;
-
-fun pr_sml_modl name content =
- Pretty.chunks ([
- str ("structure " ^ name ^ " = "),
- str "struct",
- str ""
- ] @ content @ [
- str "",
- str ("end; (*struct " ^ name ^ "*)")
- ]);
-
-fun pr_ocaml tyco_syntax const_syntax labelled_name init_syms deresolv is_cons ml_def =
- let
- fun pr_dicts fxy ds =
- let
- fun pr_dictvar (v, (_, 1)) = "_" ^ first_upper v
- | pr_dictvar (v, (i, _)) = "_" ^ first_upper v ^ string_of_int (i+1);
- fun pr_proj ps p =
- fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p
- fun pr_dictc fxy (DictConst (inst, dss)) =
- brackify fxy ((str o deresolv) inst :: map (pr_dicts BR) dss)
- | pr_dictc fxy (DictVar (classrels, v)) =
- pr_proj (map deresolv classrels) ((str o pr_dictvar) v)
- in case ds
- of [] => str "()"
- | [d] => pr_dictc fxy d
- | _ :: _ => (Pretty.list "(" ")" o map (pr_dictc NOBR)) ds
- end;
- fun pr_tyvars vs =
- vs
- |> map (fn (v, sort) => map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort)
- |> map (pr_dicts BR);
- fun pr_tycoexpr fxy (tyco, tys) =
- let
- val tyco' = (str o deresolv) tyco
- in case map (pr_typ BR) tys
- of [] => tyco'
- | [p] => Pretty.block [p, Pretty.brk 1, tyco']
- | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
- end
- and pr_typ fxy (tyco `%% tys) =
- (case tyco_syntax tyco
- of NONE => pr_tycoexpr fxy (tyco, tys)
- | SOME (i, pr) =>
- if not (i = length tys)
- then error ("Number of argument mismatch in customary serialization: "
- ^ (string_of_int o length) tys ^ " given, "
- ^ string_of_int i ^ " expected")
- else pr pr_typ fxy tys)
- | pr_typ fxy (ITyVar v) =
- str ("'" ^ v);
- fun pr_term lhs vars fxy (IConst c) =
- pr_app lhs vars fxy (c, [])
- | pr_term lhs vars fxy (IVar v) =
- str (CodegenNames.lookup_var vars v)
- | pr_term lhs vars fxy (t as t1 `$ t2) =
- (case CodegenThingol.unfold_const_app t
- of SOME c_ts => pr_app lhs vars fxy c_ts
- | NONE =>
- brackify fxy [pr_term lhs vars NOBR t1, pr_term lhs vars BR t2])
- | pr_term lhs vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = CodegenThingol.unfold_abs t;
- fun pr ((v, pat), ty) = pr_bind BR ((SOME v, pat), ty);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (str "fun" :: ps @ str "->" @@ pr_term lhs vars' NOBR t') end
- | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodegenThingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
- then pr_case vars fxy cases
- else pr_app lhs vars fxy c_ts
- | NONE => pr_case vars fxy cases)
- and pr_app' lhs vars (app as ((c, (iss, tys)), ts)) =
- if is_cons c then
- if length tys = length ts
- then case ts
- of [] => [(str o deresolv) c]
- | [t] => [(str o deresolv) c, pr_term lhs vars BR t]
- | _ => [(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term lhs vars NOBR) ts)]
- else [pr_term lhs vars BR (CodegenThingol.eta_expand app (length tys))]
- else (str o deresolv) c
- :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term lhs vars BR) ts)
- and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
- and pr_bind' ((NONE, NONE), _) = str "_"
- | pr_bind' ((SOME v, NONE), _) = str v
- | pr_bind' ((NONE, SOME p), _) = p
- | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v]
- and pr_bind fxy = gen_pr_bind pr_bind' (pr_term false) fxy
- and pr_case vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, t') = CodegenThingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind NOBR ((NONE, SOME pat), ty)
- |>> (fn p => concat [str "let", p, str "=", pr_term false vars NOBR t, str "in"])
- val (ps, vars') = fold_map pr binds vars;
- in Pretty.chunks (ps @| pr_term false vars' NOBR t') end
- | pr_case vars fxy (((td, ty), b::bs), _) =
- let
- fun pr delim (pat, t) =
- let
- val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
- in concat [str delim, p, str "->", pr_term false vars' NOBR t] end;
- in
- (Pretty.enclose "(" ")" o single o brackify fxy) (
- str "match"
- :: pr_term false vars NOBR td
- :: pr "with" b
- :: map (pr "|") bs
- )
- end
- | pr_case vars fxy ((_, []), _) = str "failwith \"empty case\"";
- fun pr_def (MLFuns (funns as funn :: funns')) =
- let
- fun fish_parm _ (w as SOME _) = w
- | fish_parm (IVar v) NONE = SOME v
- | fish_parm _ NONE = NONE;
- fun fillup_parm _ (_, SOME v) = v
- | fillup_parm x (i, NONE) = x ^ string_of_int i;
- fun fish_parms vars eqs =
- let
- val raw_fished = fold (map2 fish_parm) eqs (replicate (length (hd eqs)) NONE);
- val x = Name.variant (map_filter I raw_fished) "x";
- val fished = map_index (fillup_parm x) raw_fished;
- val vars' = CodegenNames.intro_vars fished vars;
- in map (CodegenNames.lookup_var vars') fished end;
- fun pr_eq (ts, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- ((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts
- |> CodegenNames.intro_vars ((fold o CodegenThingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in concat [
- (Pretty.block o Pretty.commas) (map (pr_term true vars NOBR) ts),
- str "->",
- pr_term false vars NOBR t
- ] end;
- fun pr_eqs [(ts, t)] =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- ((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts
- |> CodegenNames.intro_vars ((fold o CodegenThingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- concat (
- map (pr_term true vars BR) ts
- @ str "="
- @@ pr_term false vars NOBR t
- )
- end
- | pr_eqs (eqs as (eq as ([_], _)) :: eqs') =
- Pretty.block (
- str "="
- :: Pretty.brk 1
- :: str "function"
- :: Pretty.brk 1
- :: pr_eq eq
- :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs'
- )
- | pr_eqs (eqs as eq :: eqs') =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- ((fold o CodegenThingol.fold_constnames) (insert (op =)) (map snd eqs) []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts;
- val dummy_parms = (map str o fish_parms vars o map fst) eqs;
- in
- Pretty.block (
- Pretty.breaks dummy_parms
- @ Pretty.brk 1
- :: str "="
- :: Pretty.brk 1
- :: str "match"
- :: Pretty.brk 1
- :: (Pretty.block o Pretty.commas) dummy_parms
- :: Pretty.brk 1
- :: str "with"
- :: Pretty.brk 1
- :: pr_eq eq
- :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs'
- )
- end;
- fun pr_funn definer (name, (eqs, (vs, ty))) =
- concat (
- str definer
- :: (str o deresolv) name
- :: pr_tyvars (filter_out (null o snd) vs)
- @| pr_eqs eqs
- );
- val (ps, p) = split_last (pr_funn "let rec" funn :: map (pr_funn "and") funns');
- in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end
- | pr_def (MLDatas (datas as (data :: datas'))) =
- let
- fun pr_co (co, []) =
- str (deresolv co)
- | pr_co (co, tys) =
- concat [
- str (deresolv co),
- str "of",
- Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
- ];
- fun pr_data definer (tyco, (vs, [])) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- @@ str "EMPTY_"
- )
- | pr_data definer (tyco, (vs, cos)) =
- concat (
- str definer
- :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
- :: str "="
- :: separate (str "|") (map pr_co cos)
- );
- val (ps, p) = split_last (pr_data "type" data :: map (pr_data "and") datas');
- in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end
- | pr_def (MLClass (class, (superclasses, (v, classops)))) =
- let
- val w = "_" ^ first_upper v;
- fun pr_superclass_field (class, classrel) =
- (concat o map str) [
- deresolv classrel, ":", "'" ^ v, deresolv class
- ];
- fun pr_classop_field (classop, ty) =
- concat [
- (str o deresolv) classop, str ":", pr_typ NOBR ty
- ];
- fun pr_classop_proj (classop, _) =
- concat [
- str "let",
- (str o deresolv) classop,
- str w,
- str "=",
- str (w ^ "." ^ deresolv classop ^ ";;")
- ];
- in Pretty.chunks (
- concat [
- str ("type '" ^ v),
- (str o deresolv) class,
- str "=",
- Pretty.enum ";" "{" "};;" (
- map pr_superclass_field superclasses @ map pr_classop_field classops
- )
- ]
- :: map pr_classop_proj classops
- ) end
- | pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) =
- let
- fun pr_superclass (_, (classrel, dss)) =
- concat [
- (str o deresolv) classrel,
- str "=",
- pr_dicts NOBR [DictConst dss]
- ];
- fun pr_classop_def (classop, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- (CodegenThingol.fold_constnames (insert (op =)) t []);
- val vars = CodegenNames.intro_vars consts init_syms;
- in
- concat [
- (str o deresolv) classop,
- str "=",
- pr_term false vars NOBR t
- ]
- end;
- in
- concat (
- str "let"
- :: (str o deresolv) inst
- :: pr_tyvars arity
- @ str "="
- @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [
- Pretty.enum ";" "{" "}" (map pr_superclass superarities @ map pr_classop_def classop_defs),
- str ":",
- pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
- ]
- )
- end;
- in pr_def ml_def end;
-
-fun pr_ocaml_modl name content =
- Pretty.chunks ([
- str ("module " ^ name ^ " = "),
- str "struct",
- str ""
- ] @ content @ [
- str "",
- str ("end;; (*struct " ^ name ^ "*)")
- ]);
-
-val code_width = ref 80;
-fun code_output p = Pretty.setmp_margin (!code_width) Pretty.output p ^ "\n";
-
-fun seri_ml pr_def pr_modl module output labelled_name reserved_syms raw_module_alias module_prolog
- (_ : string -> class_syntax option) tyco_syntax const_syntax code =
- let
- val module_alias = if is_some module then K module else raw_module_alias;
- val is_cons = CodegenThingol.is_cons code;
- datatype node =
- Def of string * ml_def option
- | Module of string * ((Name.context * Name.context) * node Graph.T);
- val init_names = Name.make_context reserved_syms;
- val init_module = ((init_names, init_names), Graph.empty);
- fun map_node [] f = f
- | map_node (m::ms) f =
- Graph.default_node (m, Module (m, init_module))
- #> Graph.map_node m (fn (Module (dmodlname, (nsp, nodes))) => Module (dmodlname, (nsp, map_node ms f nodes)));
- fun map_nsp_yield [] f (nsp, nodes) =
- let
- val (x, nsp') = f nsp
- in (x, (nsp', nodes)) end
- | map_nsp_yield (m::ms) f (nsp, nodes) =
- let
- val (x, nodes') =
- nodes
- |> Graph.default_node (m, Module (m, init_module))
- |> Graph.map_node_yield m (fn Module (dmodlname, nsp_nodes) =>
- let
- val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes
- in (x, Module (dmodlname, nsp_nodes')) end)
- in (x, (nsp, nodes')) end;
- val init_syms = CodegenNames.make_vars reserved_syms;
- val name_modl = mk_modl_name_tab init_names NONE module_alias code;
- fun name_def upper name nsp =
- let
- val (_, base) = dest_name name;
- val base' = if upper then first_upper base else base;
- val ([base''], nsp') = Name.variants [base'] nsp;
- in (base'', nsp') end;
- fun map_nsp_fun f (nsp_fun, nsp_typ) =
- let
- val (x, nsp_fun') = f nsp_fun
- in (x, (nsp_fun', nsp_typ)) end;
- fun map_nsp_typ f (nsp_fun, nsp_typ) =
- let
- val (x, nsp_typ') = f nsp_typ
- in (x, (nsp_fun, nsp_typ')) end;
- fun mk_funs defs =
- fold_map
- (fn (name, CodegenThingol.Fun info) =>
- map_nsp_fun (name_def false name) >> (fn base => (base, (name, info)))
- | (name, def) => error ("Function block containing illegal definition: " ^ labelled_name name)
- ) defs
- >> (split_list #> apsnd MLFuns);
- fun mk_datatype defs =
- fold_map
- (fn (name, CodegenThingol.Datatype info) =>
- map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info)))
- | (name, CodegenThingol.Datatypecons _) =>
- map_nsp_fun (name_def true name) >> (fn base => (base, NONE))
- | (name, def) => error ("Datatype block containing illegal definition: " ^ labelled_name name)
- ) defs
- >> (split_list #> apsnd (map_filter I
- #> (fn [] => error ("Datatype block without data definition: " ^ (commas o map (labelled_name o fst)) defs)
- | infos => MLDatas infos)));
- fun mk_class defs =
- fold_map
- (fn (name, CodegenThingol.Class info) =>
- map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info)))
- | (name, CodegenThingol.Classrel _) =>
- map_nsp_fun (name_def false name) >> (fn base => (base, NONE))
- | (name, CodegenThingol.Classop _) =>
- map_nsp_fun (name_def false name) >> (fn base => (base, NONE))
- | (name, def) => error ("Class block containing illegal definition: " ^ labelled_name name)
- ) defs
- >> (split_list #> apsnd (map_filter I
- #> (fn [] => error ("Class block without class definition: " ^ (commas o map (labelled_name o fst)) defs)
- | [info] => MLClass info)));
- fun mk_inst [(name, CodegenThingol.Classinst info)] =
- map_nsp_fun (name_def false name)
- >> (fn base => ([base], MLClassinst (name, info)));
- fun add_group mk defs nsp_nodes =
- let
- val names as (name :: names') = map fst defs;
- val deps =
- []
- |> fold (fold (insert (op =)) o Graph.imm_succs code) names
- |> subtract (op =) names;
- val (modls, _) = (split_list o map dest_name) names;
- val modl' = (the_single o distinct (op =) o map name_modl) modls
- handle Empty =>
- error ("Illegal mutual dependencies: " ^ commas (map labelled_name names));
- val modl_explode = NameSpace.explode modl';
- fun add_dep name name'' =
- let
- val modl'' = (name_modl o fst o dest_name) name'';
- in if modl' = modl'' then
- map_node modl_explode
- (Graph.add_edge (name, name''))
- else let
- val (common, (diff1::_, diff2::_)) = chop_prefix (op =)
- (modl_explode, NameSpace.explode modl'');
- in
- map_node common
- (fn gr => Graph.add_edge_acyclic (diff1, diff2) gr
- handle Graph.CYCLES _ => error ("Dependency "
- ^ quote name ^ " -> " ^ quote name''
- ^ " would result in module dependency cycle"))
- end end;
- in
- nsp_nodes
- |> map_nsp_yield modl_explode (mk defs)
- |-> (fn (base' :: bases', def') =>
- apsnd (map_node modl_explode (Graph.new_node (name, (Def (base', SOME def')))
- #> fold2 (fn name' => fn base' => Graph.new_node (name', (Def (base', NONE)))) names' bases')))
- |> apsnd (fold (fn name => fold (add_dep name) deps) names)
- |> apsnd (fold (map_node modl_explode o Graph.add_edge) (product names names))
- end;
- fun group_defs [(_, CodegenThingol.Bot)] =
- I
- | group_defs ((defs as (_, CodegenThingol.Fun _)::_)) =
- add_group mk_funs defs
- | group_defs ((defs as (_, CodegenThingol.Datatypecons _)::_)) =
- add_group mk_datatype defs
- | group_defs ((defs as (_, CodegenThingol.Datatype _)::_)) =
- add_group mk_datatype defs
- | group_defs ((defs as (_, CodegenThingol.Class _)::_)) =
- add_group mk_class defs
- | group_defs ((defs as (_, CodegenThingol.Classrel _)::_)) =
- add_group mk_class defs
- | group_defs ((defs as (_, CodegenThingol.Classop _)::_)) =
- add_group mk_class defs
- | group_defs ((defs as [(_, CodegenThingol.Classinst _)])) =
- add_group mk_inst defs
- | group_defs defs = error ("Illegal mutual dependencies: " ^
- (commas o map (labelled_name o fst)) defs)
- val (_, nodes) =
- init_module
- |> fold group_defs (map (AList.make (Graph.get_node code))
- (rev (Graph.strong_conn code)))
- fun deresolver prefix name =
- let
- val modl = (fst o dest_name) name;
- val modl' = (NameSpace.explode o name_modl) modl;
- val (_, (_, remainder)) = chop_prefix (op =) (prefix, modl');
- val defname' =
- nodes
- |> fold (fn m => fn g => case Graph.get_node g m
- of Module (_, (_, g)) => g) modl'
- |> (fn g => case Graph.get_node g name of Def (defname, _) => defname);
- in
- NameSpace.implode (remainder @ [defname'])
- end handle Graph.UNDEF _ =>
- error ("Unknown definition name: " ^ labelled_name name);
- fun the_prolog modlname = case module_prolog modlname
- of NONE => []
- | SOME p => [p, str ""];
- fun pr_node prefix (Def (_, NONE)) =
- NONE
- | pr_node prefix (Def (_, SOME def)) =
- SOME (pr_def tyco_syntax const_syntax labelled_name init_syms
- (deresolver prefix) is_cons def)
- | pr_node prefix (Module (dmodlname, (_, nodes))) =
- SOME (pr_modl dmodlname (the_prolog (NameSpace.implode (prefix @ [dmodlname]))
- @ separate (str "") ((map_filter (pr_node (prefix @ [dmodlname]) o Graph.get_node nodes)
- o rev o flat o Graph.strong_conn) nodes)));
- val p = Pretty.chunks (the_prolog "" @ separate (str "") ((map_filter
- (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes))
- in output p end;
-
-val eval_verbose = ref false;
-
-fun isar_seri_sml module file =
- let
- val output = case file
- of NONE => use_text "generated code" Output.ml_output (!eval_verbose) o code_output
- | SOME "-" => writeln o code_output
- | SOME file => File.write (Path.explode file) o code_output;
- in
- parse_args (Scan.succeed ())
- #> (fn () => seri_ml pr_sml pr_sml_modl module output)
- end;
-
-fun isar_seri_ocaml module file =
- let
- val output = case file
- of NONE => error "OCaml: no internal compilation"
- | SOME "-" => writeln o code_output
- | SOME file => File.write (Path.explode file) o code_output;
- fun output_file file = File.write (Path.explode file) o code_output;
- val output_diag = writeln o code_output;
- in
- parse_args (Scan.succeed ())
- #> (fn () => seri_ml pr_ocaml pr_ocaml_modl module output)
- end;
-
-
-(** Haskell serializer **)
-
-local
-
-fun pr_bind' ((NONE, NONE), _) = str "_"
- | pr_bind' ((SOME v, NONE), _) = str v
- | pr_bind' ((NONE, SOME p), _) = p
- | pr_bind' ((SOME v, SOME p), _) = brackets [str v, str "@", p]
-
-val pr_bind_haskell = gen_pr_bind pr_bind';
-
-in
-
-fun pr_haskell class_syntax tyco_syntax const_syntax labelled_name init_syms
- deresolv_here deresolv is_cons deriving_show def =
- let
- fun class_name class = case class_syntax class
- of NONE => deresolv class
- | SOME (class, _) => class;
- fun classop_name class classop = case class_syntax class
- of NONE => deresolv_here classop
- | SOME (_, classop_syntax) => case classop_syntax classop
- of NONE => (snd o dest_name) classop
- | SOME classop => classop
- fun pr_typparms tyvars vs =
- case maps (fn (v, sort) => map (pair v) sort) vs
- of [] => str ""
- | xs => Pretty.block [
- Pretty.enum "," "(" ")" (
- map (fn (v, class) => str
- (class_name class ^ " " ^ CodegenNames.lookup_var tyvars v)) xs
- ),
- str " => "
- ];
- fun pr_tycoexpr tyvars fxy (tyco, tys) =
- brackify fxy (str tyco :: map (pr_typ tyvars BR) tys)
- and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) =
- (case tyco_syntax tyco
- of NONE =>
- pr_tycoexpr tyvars fxy (deresolv tyco, tys)
- | SOME (i, pr) =>
- if not (i = length tys)
- then error ("Number of argument mismatch in customary serialization: "
- ^ (string_of_int o length) tys ^ " given, "
- ^ string_of_int i ^ " expected")
- else pr (pr_typ tyvars) fxy tys)
- | pr_typ tyvars fxy (ITyVar v) =
- (str o CodegenNames.lookup_var tyvars) v;
- fun pr_typscheme_expr tyvars (vs, tycoexpr) =
- Pretty.block (pr_typparms tyvars vs @@ pr_tycoexpr tyvars NOBR tycoexpr);
- fun pr_typscheme tyvars (vs, ty) =
- Pretty.block (pr_typparms tyvars vs @@ pr_typ tyvars NOBR ty);
- fun pr_term lhs vars fxy (IConst c) =
- pr_app lhs vars fxy (c, [])
- | pr_term lhs vars fxy (t as (t1 `$ t2)) =
- (case CodegenThingol.unfold_const_app t
- of SOME app => pr_app lhs vars fxy app
- | _ =>
- brackify fxy [
- pr_term lhs vars NOBR t1,
- pr_term lhs vars BR t2
- ])
- | pr_term lhs vars fxy (IVar v) =
- (str o CodegenNames.lookup_var vars) v
- | pr_term lhs vars fxy (t as _ `|-> _) =
- let
- val (binds, t') = CodegenThingol.unfold_abs t;
- fun pr ((v, pat), ty) = pr_bind BR ((SOME v, pat), ty);
- val (ps, vars') = fold_map pr binds vars;
- in brackets (str "\\" :: ps @ str "->" @@ pr_term lhs vars' NOBR t') end
- | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodegenThingol.unfold_const_app t0
- of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
- then pr_case vars fxy cases
- else pr_app lhs vars fxy c_ts
- | NONE => pr_case vars fxy cases)
- and pr_app' lhs vars ((c, _), ts) =
- (str o deresolv) c :: map (pr_term lhs vars BR) ts
- and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
- and pr_bind fxy = pr_bind_haskell (pr_term false) fxy
- and pr_case vars fxy (cases as ((_, [_]), _)) =
- let
- val (binds, t) = CodegenThingol.unfold_let (ICase cases);
- fun pr ((pat, ty), t) vars =
- vars
- |> pr_bind BR ((NONE, SOME pat), ty)
- |>> (fn p => semicolon [p, str "=", pr_term false vars NOBR t])
- val (ps, vars') = fold_map pr binds vars;
- in
- Pretty.block_enclose (
- str "let {",
- concat [str "}", str "in", pr_term false vars' NOBR t]
- ) ps
- end
- | pr_case vars fxy (((td, ty), bs as _ :: _), _) =
- let
- fun pr (pat, t) =
- let
- val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
- in semicolon [p, str "->", pr_term false vars' NOBR t] end;
- in
- Pretty.block_enclose (
- concat [str "(case", pr_term false vars NOBR td, str "of", str "{"],
- str "})"
- ) (map pr bs)
- end
- | pr_case vars fxy ((_, []), _) = str "error \"empty case\"";
- fun pr_def (name, CodegenThingol.Fun (eqs, (vs, ty))) =
- let
- val tyvars = CodegenNames.intro_vars (map fst vs) init_syms;
- fun pr_eq (ts, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- ((fold o CodegenThingol.fold_constnames) (insert (op =)) (t :: ts) []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts
- |> CodegenNames.intro_vars ((fold o CodegenThingol.fold_unbound_varnames)
- (insert (op =)) ts []);
- in
- semicolon (
- (str o deresolv_here) name
- :: map (pr_term true vars BR) ts
- @ str "="
- @@ pr_term false vars NOBR t
- )
- end;
- in
- Pretty.chunks (
- Pretty.block [
- (str o suffix " ::" o deresolv_here) name,
- Pretty.brk 1,
- pr_typscheme tyvars (vs, ty),
- str ";"
- ]
- :: map pr_eq eqs
- )
- end
- | pr_def (name, CodegenThingol.Datatype (vs, [])) =
- let
- val tyvars = CodegenNames.intro_vars (map fst vs) init_syms;
- in
- semicolon [
- str "data",
- pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
- ]
- end
- | pr_def (name, CodegenThingol.Datatype (vs, [(co, [ty])])) =
- let
- val tyvars = CodegenNames.intro_vars (map fst vs) init_syms;
- in
- semicolon (
- str "newtype"
- :: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
- :: str "="
- :: (str o deresolv_here) co
- :: pr_typ tyvars BR ty
- :: (if deriving_show name then [str "deriving (Read, Show)"] else [])
- )
- end
- | pr_def (name, CodegenThingol.Datatype (vs, co :: cos)) =
- let
- val tyvars = CodegenNames.intro_vars (map fst vs) init_syms;
- fun pr_co (co, tys) =
- concat (
- (str o deresolv_here) co
- :: map (pr_typ tyvars BR) tys
- )
- in
- semicolon (
- str "data"
- :: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
- :: str "="
- :: pr_co co
- :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos
- @ (if deriving_show name then [str "deriving (Read, Show)"] else [])
- )
- end
- | pr_def (name, CodegenThingol.Class (superclasss, (v, classops))) =
- let
- val tyvars = CodegenNames.intro_vars [v] init_syms;
- fun pr_classop (classop, ty) =
- semicolon [
- (str o classop_name name) classop,
- str "::",
- pr_typ tyvars NOBR ty
- ]
- in
- Pretty.block_enclose (
- Pretty.block [
- str "class ",
- pr_typparms tyvars [(v, map fst superclasss)],
- str (deresolv_here name ^ " " ^ CodegenNames.lookup_var tyvars v),
- str " where {"
- ],
- str "};"
- ) (map pr_classop classops)
- end
- | pr_def (_, CodegenThingol.Classinst ((class, (tyco, vs)), (_, classop_defs))) =
- let
- val tyvars = CodegenNames.intro_vars (map fst vs) init_syms;
- fun pr_instdef (classop, t) =
- let
- val consts = map_filter
- (fn c => if (is_some o const_syntax) c
- then NONE else (SOME o NameSpace.base o deresolv) c)
- (CodegenThingol.fold_constnames (insert (op =)) t []);
- val vars = init_syms
- |> CodegenNames.intro_vars consts;
- in
- semicolon [
- (str o classop_name class) classop,
- str "=",
- pr_term false vars NOBR t
- ]
- end;
- in
- Pretty.block_enclose (
- Pretty.block [
- str "instance ",
- pr_typparms tyvars vs,
- str (class_name class ^ " "),
- pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
- str " where {"
- ],
- str "};"
- ) (map pr_instdef classop_defs)
- end;
- in pr_def def end;
-
-fun pretty_haskell_monad c_mbind c_kbind =
- let
- fun pretty pr vars fxy [(t, _)] =
- let
- val pr_bind = pr_bind_haskell pr;
- fun pr_mbind (NONE, t) vars =
- (semicolon [pr vars NOBR t], vars)
- | pr_mbind (SOME (bind, true), t) vars = vars
- |> pr_bind NOBR bind
- |>> (fn p => semicolon [p, str "<-", pr vars NOBR t])
- | pr_mbind (SOME (bind, false), t) vars = vars
- |> pr_bind NOBR bind
- |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]);
- val (binds, t) = implode_monad c_mbind c_kbind t;
- val (ps, vars') = fold_map pr_mbind binds vars;
- fun brack p = if eval_fxy BR fxy then Pretty.block [str "(", p, str ")"] else p;
- in (brack o Pretty.block_enclose (str "do {", str "}")) (ps @| pr vars' NOBR t) end;
- in (1, pretty) end;
-
-end; (*local*)
-
-fun seri_haskell module_prefix module destination string_classes labelled_name
- reserved_syms raw_module_alias module_prolog class_syntax tyco_syntax const_syntax code =
- let
- val _ = Option.map File.check destination;
- val is_cons = CodegenThingol.is_cons code;
- val module_alias = if is_some module then K module else raw_module_alias;
- val init_names = Name.make_context reserved_syms;
- val name_modl = mk_modl_name_tab init_names module_prefix module_alias code;
- fun add_def (name, (def, deps)) =
- let
- val (modl, base) = dest_name name;
- fun name_def base = Name.variants [base] #>> the_single;
- fun add_fun upper (nsp_fun, nsp_typ) =
- let
- val (base', nsp_fun') = name_def (if upper then first_upper base else base) nsp_fun
- in (base', (nsp_fun', nsp_typ)) end;
- fun add_typ (nsp_fun, nsp_typ) =
- let
- val (base', nsp_typ') = name_def (first_upper base) nsp_typ
- in (base', (nsp_fun, nsp_typ')) end;
- val add_name =
- case def
- of CodegenThingol.Bot => pair base
- | CodegenThingol.Fun _ => add_fun false
- | CodegenThingol.Datatype _ => add_typ
- | CodegenThingol.Datatypecons _ => add_fun true
- | CodegenThingol.Class _ => add_typ
- | CodegenThingol.Classrel _ => pair base
- | CodegenThingol.Classop _ => add_fun false
- | CodegenThingol.Classinst _ => pair base;
- val modlname' = name_modl modl;
- fun add_def base' =
- case def
- of CodegenThingol.Bot => I
- | CodegenThingol.Datatypecons _ =>
- cons (name, ((NameSpace.append modlname' base', base'), NONE))
- | CodegenThingol.Classrel _ => I
- | CodegenThingol.Classop _ =>
- cons (name, ((NameSpace.append modlname' base', base'), NONE))
- | _ => cons (name, ((NameSpace.append modlname' base', base'), SOME def));
- in
- Symtab.map_default (modlname', ([], ([], (init_names, init_names))))
- (apfst (fold (insert (op = : string * string -> bool)) deps))
- #> `(fn code => add_name ((snd o snd o the o Symtab.lookup code) modlname'))
- #-> (fn (base', names) =>
- (Symtab.map_entry modlname' o apsnd) (fn (defs, _) =>
- (add_def base' defs, names)))
- end;
- val code' =
- fold add_def (AList.make (fn name => (Graph.get_node code name, Graph.imm_succs code name))
- (Graph.strong_conn code |> flat)) Symtab.empty;
- val init_syms = CodegenNames.make_vars reserved_syms;
- fun deresolv name =
- (fst o fst o the o AList.lookup (op =) ((fst o snd o the
- o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name
- handle Option => error ("Unknown definition name: " ^ labelled_name name);
- fun deresolv_here name =
- (snd o fst o the o AList.lookup (op =) ((fst o snd o the
- o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name
- handle Option => error ("Unknown definition name: " ^ labelled_name name);
- fun deriving_show tyco =
- let
- fun deriv _ "fun" = false
- | deriv tycos tyco = member (op =) tycos tyco orelse
- case the_default CodegenThingol.Bot (try (Graph.get_node code) tyco)
- of CodegenThingol.Bot => true
- | CodegenThingol.Datatype (_, cs) => forall (deriv' (tyco :: tycos))
- (maps snd cs)
- and deriv' tycos (tyco `%% tys) = deriv tycos tyco
- andalso forall (deriv' tycos) tys
- | deriv' _ (ITyVar _) = true
- in deriv [] tyco end;
- fun seri_def qualified = pr_haskell class_syntax tyco_syntax const_syntax labelled_name init_syms
- deresolv_here (if qualified then deresolv else deresolv_here) is_cons
- (if string_classes then deriving_show else K false);
- fun write_module (SOME destination) modlname =
- let
- val filename = case modlname
- of "" => Path.explode "Main.hs"
- | _ => (Path.ext "hs" o Path.explode o implode o separate "/" o NameSpace.explode) modlname;
- val pathname = Path.append destination filename;
- val _ = File.mkdir (Path.dir pathname);
- in File.write pathname end
- | write_module NONE _ = writeln;
- fun seri_module (modlname', (imports, (defs, _))) =
- let
- val imports' =
- imports
- |> map (name_modl o fst o dest_name)
- |> distinct (op =)
- |> remove (op =) modlname';
- val qualified =
- imports
- |> map_filter (try deresolv)
- |> map NameSpace.base
- |> has_duplicates (op =);
- val mk_import = str o (if qualified
- then prefix "import qualified "
- else prefix "import ") o suffix ";";
- in
- Pretty.chunks (
- str ("module " ^ modlname' ^ " where {")
- :: str ""
- :: map mk_import imports'
- @ str ""
- :: separate (str "") ((case module_prolog modlname'
- of SOME prolog => [prolog]
- | NONE => [])
- @ map_filter
- (fn (name, (_, SOME def)) => SOME (seri_def qualified (name, def))
- | (_, (_, NONE)) => NONE) defs)
- @ str ""
- @@ str "}"
- )
- |> code_output
- |> write_module destination modlname'
- end;
- in Symtab.fold (fn modl => fn () => seri_module modl) code' () end;
-
-fun isar_seri_haskell module file =
- let
- val destination = case file
- of NONE => error ("Haskell: no internal compilation")
- | SOME "-" => NONE
- | SOME file => SOME (Path.explode file)
- in
- parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
- -- Scan.optional (Args.$$$ "string_classes" >> K true) false
- >> (fn (module_prefix, string_classes) =>
- seri_haskell module_prefix module destination string_classes))
- end;
-
-
-(** diagnosis serializer **)
-
-fun seri_diagnosis labelled_name _ _ _ _ _ _ code =
- let
- val init_names = CodegenNames.make_vars [];
- fun pr_fun "fun" = SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ])
- | pr_fun _ = NONE
- val pr = pr_haskell (K NONE) pr_fun (K NONE) labelled_name init_names I I (K false) (K false);
- in
- []
- |> Graph.fold (fn (name, (def, _)) => case try pr (name, def) of SOME p => cons p | NONE => I) code
- |> Pretty.chunks2
- |> code_output
- |> writeln
- end;
-
-
-
-(** theory data **)
-
-datatype syntax_expr = SyntaxExpr of {
- class: (string * (string -> string option)) Symtab.table,
- inst: unit Symtab.table,
- tyco: typ_syntax Symtab.table,
- const: term_syntax Symtab.table
-};
-
-fun mk_syntax_expr ((class, inst), (tyco, const)) =
- SyntaxExpr { class = class, inst = inst, tyco = tyco, const = const };
-fun map_syntax_expr f (SyntaxExpr { class, inst, tyco, const }) =
- mk_syntax_expr (f ((class, inst), (tyco, const)));
-fun merge_syntax_expr (SyntaxExpr { class = class1, inst = inst1, tyco = tyco1, const = const1 },
- SyntaxExpr { class = class2, inst = inst2, tyco = tyco2, const = const2 }) =
- mk_syntax_expr (
- (Symtab.join (K snd) (class1, class2),
- Symtab.join (K snd) (inst1, inst2)),
- (Symtab.join (K snd) (tyco1, tyco2),
- Symtab.join (K snd) (const1, const2))
- );
-
-datatype syntax_modl = SyntaxModl of {
- alias: string Symtab.table,
- prolog: Pretty.T Symtab.table
-};
-
-fun mk_syntax_modl (alias, prolog) =
- SyntaxModl { alias = alias, prolog = prolog };
-fun map_syntax_modl f (SyntaxModl { alias, prolog }) =
- mk_syntax_modl (f (alias, prolog));
-fun merge_syntax_modl (SyntaxModl { alias = alias1, prolog = prolog1 },
- SyntaxModl { alias = alias2, prolog = prolog2 }) =
- mk_syntax_modl (
- Symtab.join (K snd) (alias1, alias2),
- Symtab.join (K snd) (prolog1, prolog2)
- );
-
-type serializer =
- string option
- -> string option
- -> Args.T list
- -> (string -> string)
- -> string list
- -> (string -> string option)
- -> (string -> Pretty.T option)
- -> (string -> class_syntax option)
- -> (string -> typ_syntax option)
- -> (string -> term_syntax option)
- -> CodegenThingol.code -> unit;
-
-datatype target = Target of {
- serial: serial,
- serializer: serializer,
- syntax_expr: syntax_expr,
- syntax_modl: syntax_modl,
- reserved: string list
-};
-
-fun mk_target (serial, ((serializer, reserved), (syntax_expr, syntax_modl))) =
- Target { serial = serial, reserved = reserved, serializer = serializer, syntax_expr = syntax_expr, syntax_modl = syntax_modl };
-fun map_target f ( Target { serial, serializer, reserved, syntax_expr, syntax_modl } ) =
- mk_target (f (serial, ((serializer, reserved), (syntax_expr, syntax_modl))));
-fun merge_target target (Target { serial = serial1, serializer = serializer, reserved = reserved1,
- syntax_expr = syntax_expr1, syntax_modl = syntax_modl1 },
- Target { serial = serial2, serializer = _, reserved = reserved2,
- syntax_expr = syntax_expr2, syntax_modl = syntax_modl2 }) =
- if serial1 = serial2 then
- mk_target (serial1, ((serializer, merge (op =) (reserved1, reserved2)),
- (merge_syntax_expr (syntax_expr1, syntax_expr2),
- merge_syntax_modl (syntax_modl1, syntax_modl2))
- ))
- else
- error ("Incompatible serializers: " ^ quote target);
-
-structure CodegenSerializerData = TheoryDataFun
-(
- type T = target Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ = Symtab.join merge_target;
-);
-
-fun the_serializer (Target { serializer, ... }) = serializer;
-fun the_reserved (Target { reserved, ... }) = reserved;
-fun the_syntax_expr (Target { syntax_expr = SyntaxExpr x, ... }) = x;
-fun the_syntax_modl (Target { syntax_modl = SyntaxModl x, ... }) = x;
-
-fun assert_serializer thy target =
- case Symtab.lookup (CodegenSerializerData.get thy) target
- of SOME data => target
- | NONE => error ("Unknown code target language: " ^ quote target);
-
-fun add_serializer (target, seri) thy =
- let
- val _ = case Symtab.lookup (CodegenSerializerData.get thy) target
- of SOME _ => warning ("overwriting existing serializer " ^ quote target)
- | NONE => ();
- in
- thy
- |> (CodegenSerializerData.map oo Symtab.map_default)
- (target, mk_target (serial (), ((seri, []),
- (mk_syntax_expr ((Symtab.empty, Symtab.empty), (Symtab.empty, Symtab.empty)),
- mk_syntax_modl (Symtab.empty, Symtab.empty)))))
- (map_target (fn (serial, ((_, keywords), syntax)) => (serial, ((seri, keywords), syntax))))
- end;
-
-fun map_seri_data target f thy =
- let
- val _ = assert_serializer thy target;
- in
- thy
- |> (CodegenSerializerData.map o Symtab.map_entry target o map_target) f
- end;
-
-val target_SML = "SML";
-val target_OCaml = "OCaml";
-val target_Haskell = "Haskell";
-val target_diag = "diag";
-
-val _ = Context.add_setup (
- add_serializer (target_SML, isar_seri_sml)
- #> add_serializer (target_OCaml, isar_seri_ocaml)
- #> add_serializer (target_Haskell, isar_seri_haskell)
- #> add_serializer (target_diag, (fn _ => fn _ => fn _ => seri_diagnosis))
-);
-
-fun get_serializer thy target permissive module file args labelled_name = fn cs =>
- let
- val data = case Symtab.lookup (CodegenSerializerData.get thy) target
- of SOME data => data
- | NONE => error ("Unknown code target language: " ^ quote target);
- val seri = the_serializer data;
- val reserved = the_reserved data;
- val { alias, prolog } = the_syntax_modl data;
- val { class, inst, tyco, const } = the_syntax_expr data;
- val project = if target = target_diag then I
- else CodegenThingol.project_code permissive
- (Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const) cs;
- fun check_empty_funs code = case CodegenThingol.empty_funs code
- of [] => code
- | names => error ("No defining equations for " ^ commas (map (labelled_name thy) names));
- in
- project
- #> check_empty_funs
- #> seri module file args (labelled_name thy) reserved (Symtab.lookup alias) (Symtab.lookup prolog)
- (Symtab.lookup class) (Symtab.lookup tyco) (Symtab.lookup const)
- end;
-
-fun eval_term thy labelled_name code ((ref_name, reff), t) args =
- let
- val val_name = "Isabelle_Eval.EVAL.EVAL";
- val val_name' = "Isabelle_Eval.EVAL";
- val val_name'_args = space_implode " " (val_name' :: map (enclose "(" ")") args);
- val seri = get_serializer thy "SML" false (SOME "Isabelle_Eval") NONE [] labelled_name;
- fun eval code = (
- reff := NONE;
- seri (SOME [val_name]) code;
- use_text "generated code for evaluation" Output.ml_output (!eval_verbose)
- ("val _ = (" ^ ref_name ^ " := SOME (" ^ val_name'_args ^ "))");
- case !reff
- of NONE => error ("Could not retrieve value of ML reference " ^ quote ref_name
- ^ " (reference probably has been shadowed)")
- | SOME value => value
- );
- in
- code
- |> CodegenThingol.add_eval_def (val_name, t)
- |> eval
- end;
-
-
-
-(** optional pretty serialization **)
-
-local
-
-val pretty : (string * {
- pretty_char: string -> string,
- pretty_string: string -> string,
- pretty_numeral: bool -> IntInf.int -> string,
- pretty_list: Pretty.T list -> Pretty.T,
- infix_cons: int * string
- }) list = [
- ("SML", { pretty_char = prefix "#" o quote o ML_Syntax.print_char,
- pretty_string = ML_Syntax.print_string,
- pretty_numeral = fn unbounded => fn k =>
- if unbounded then "(" ^ IntInf.toString k ^ " : IntInf.int)"
- else IntInf.toString k,
- pretty_list = Pretty.enum "," "[" "]",
- infix_cons = (7, "::")}),
- ("OCaml", { pretty_char = fn c => enclose "'" "'"
- (let val i = ord c
- in if i < 32 orelse i = 39 orelse i = 92
- then prefix "\\" (string_of_int i)
- else c
- end),
- pretty_string = (fn _ => error "OCaml: no pretty strings"),
- pretty_numeral = fn unbounded => fn k => if k >= IntInf.fromInt 0 then
- if unbounded then
- "(Big_int.big_int_of_int " ^ IntInf.toString k ^ ")"
- else IntInf.toString k
- else
- if unbounded then
- "(Big_int.big_int_of_int " ^ (enclose "(" ")" o prefix "-"
- o IntInf.toString o op ~) k ^ ")"
- else (enclose "(" ")" o prefix "-" o IntInf.toString o op ~) k,
- pretty_list = Pretty.enum ";" "[" "]",
- infix_cons = (6, "::")}),
- ("Haskell", { pretty_char = fn c => enclose "'" "'"
- (let val i = ord c
- in if i < 32 orelse i = 39 orelse i = 92
- then Library.prefix "\\" (string_of_int i)
- else c
- end),
- pretty_string = ML_Syntax.print_string,
- pretty_numeral = fn unbounded => fn k => if k >= IntInf.fromInt 0 then
- IntInf.toString k
- else
- (enclose "(" ")" o Library.prefix "-" o IntInf.toString o IntInf.~) k,
- pretty_list = Pretty.enum "," "[" "]",
- infix_cons = (5, ":")})
-];
-
-in
-
-fun pr_pretty target = case AList.lookup (op =) pretty target
- of SOME x => x
- | NONE => error ("Unknown code target language: " ^ quote target);
-
-fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
- brackify_infix (target_fxy, R) fxy [
- pr (INFX (target_fxy, X)) t1,
- str target_cons,
- pr (INFX (target_fxy, R)) t2
- ];
-
-fun pretty_list c_nil c_cons target =
- let
- val pretty_ops = pr_pretty target;
- val mk_list = #pretty_list pretty_ops;
- fun pretty pr vars fxy [(t1, _), (t2, _)] =
- case Option.map (cons t1) (implode_list c_nil c_cons t2)
- of SOME ts => mk_list (map (pr vars NOBR) ts)
- | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2;
- in (2, pretty) end;
-
-fun pretty_list_string c_nil c_cons c_char c_nibbles target =
- let
- val pretty_ops = pr_pretty target;
- val mk_list = #pretty_list pretty_ops;
- val mk_char = #pretty_char pretty_ops;
- val mk_string = #pretty_string pretty_ops;
- fun pretty pr vars fxy [(t1, _), (t2, _)] =
- case Option.map (cons t1) (implode_list c_nil c_cons t2)
- of SOME ts => case implode_string c_char c_nibbles mk_char mk_string ts
- of SOME p => p
- | NONE => mk_list (map (pr vars NOBR) ts)
- | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2;
- in (2, pretty) end;
-
-fun pretty_char c_char c_nibbles target =
- let
- val mk_char = #pretty_char (pr_pretty target);
- fun pretty _ _ _ [(t1, _), (t2, _)] =
- case decode_char c_nibbles (t1, t2)
- of SOME c => (str o mk_char) c
- | NONE => error "Illegal character expression";
- in (2, pretty) end;
-
-fun pretty_numeral unbounded c_bit0 c_bit1 c_pls c_min c_bit target =
- let
- val mk_numeral = #pretty_numeral (pr_pretty target);
- fun pretty _ _ _ [(t, _)] =
- case implode_numeral c_bit0 c_bit1 c_pls c_min c_bit t
- of SOME k => (str o mk_numeral unbounded) k
- | NONE => error "Illegal numeral expression";
- in (1, pretty) end;
-
-fun pretty_ml_string c_char c_nibbles c_nil c_cons target =
- let
- val pretty_ops = pr_pretty target;
- val mk_char = #pretty_char pretty_ops;
- val mk_string = #pretty_string pretty_ops;
- fun pretty pr vars fxy [(t, _)] =
- case implode_list c_nil c_cons t
- of SOME ts => (case implode_string c_char c_nibbles mk_char mk_string ts
- of SOME p => p
- | NONE => error "Illegal ml_string expression")
- | NONE => error "Illegal ml_string expression";
- in (1, pretty) end;
-
-val pretty_imperative_monad_bind =
- let
- fun pretty (pr : CodegenNames.var_ctxt -> fixity -> iterm -> Pretty.T)
- vars fxy [(t1, _), ((v, ty) `|-> t2, _)] =
- pr vars fxy (ICase (((t1, ty), ([(IVar v, t2)])), IVar ""))
- | pretty pr vars fxy [(t1, _), (t2, ty2)] =
- let
- (*this code suffers from the lack of a proper concept for bindings*)
- val vs = CodegenThingol.fold_varnames cons t2 [];
- val v = Name.variant vs "x";
- val vars' = CodegenNames.intro_vars [v] vars;
- val var = IVar v;
- val ty = (hd o fst o CodegenThingol.unfold_fun) ty2;
- in pr vars' fxy (ICase (((t1, ty), ([(var, t2 `$ var)])), IVar "")) end;
- in (2, pretty) end;
-
-end; (*local*)
-
-(** ML and Isar interface **)
-
-local
-
-fun map_syntax_exprs target =
- map_seri_data target o apsnd o apsnd o apfst o map_syntax_expr;
-fun map_syntax_modls target =
- map_seri_data target o apsnd o apsnd o apsnd o map_syntax_modl;
-fun map_reserveds target =
- map_seri_data target o apsnd o apfst o apsnd;
-
-fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy =
- let
- val cls = prep_class thy raw_class;
- val class = CodegenNames.class thy cls;
- fun mk_classop (const as (c, _)) = case AxClass.class_of_param thy c
- of SOME class' => if cls = class' then CodegenNames.const thy const
- else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c)
- | NONE => error ("Not a class operation: " ^ quote c);
- fun mk_syntax_ops raw_ops = AList.lookup (op =)
- ((map o apfst) (mk_classop o prep_const thy) raw_ops);
- in case raw_syn
- of SOME (syntax, raw_ops) =>
- thy
- |> (map_syntax_exprs target o apfst o apfst)
- (Symtab.update (class, (syntax, mk_syntax_ops raw_ops)))
- | NONE =>
- thy
- |> (map_syntax_exprs target o apfst o apfst)
- (Symtab.delete_safe class)
- end;
-
-fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy =
- let
- val inst = CodegenNames.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco);
- in if add_del then
- thy
- |> (map_syntax_exprs target o apfst o apsnd)
- (Symtab.update (inst, ()))
- else
- thy
- |> (map_syntax_exprs target o apfst o apsnd)
- (Symtab.delete_safe inst)
- end;
-
-fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy =
- let
- val tyco = prep_tyco thy raw_tyco;
- val tyco' = if tyco = "fun" then "fun" else CodegenNames.tyco thy tyco;
- fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco
- then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
- else syntax
- in case raw_syn
- of SOME syntax =>
- thy
- |> (map_syntax_exprs target o apsnd o apfst)
- (Symtab.update (tyco', check_args syntax))
- | NONE =>
- thy
- |> (map_syntax_exprs target o apsnd o apfst)
- (Symtab.delete_safe tyco')
- end;
-
-fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
- let
- val c = prep_const thy raw_c;
- val c' = CodegenNames.const thy c;
- fun check_args (syntax as (n, _)) = if n > (length o fst o strip_type o Sign.the_const_type thy o fst) c
- then error ("Too many arguments in syntax for constant " ^ (quote o fst) c)
- else syntax;
- in case raw_syn
- of SOME syntax =>
- thy
- |> (map_syntax_exprs target o apsnd o apsnd)
- (Symtab.update (c', check_args syntax))
- | NONE =>
- thy
- |> (map_syntax_exprs target o apsnd o apsnd)
- (Symtab.delete_safe c')
- end;
-
-fun cert_class thy class =
- let
- val _ = AxClass.get_definition thy class;
- in class end;
-
-fun read_class thy raw_class =
- let
- val class = Sign.intern_class thy raw_class;
- val _ = AxClass.get_definition thy class;
- in class end;
-
-fun cert_tyco thy tyco =
- let
- val _ = if Sign.declared_tyname thy tyco then ()
- else error ("No such type constructor: " ^ quote tyco);
- in tyco end;
-
-fun read_tyco thy raw_tyco =
- let
- val tyco = Sign.intern_type thy raw_tyco;
- val _ = if Sign.declared_tyname thy tyco then ()
- else error ("No such type constructor: " ^ quote raw_tyco);
- in tyco end;
-
-fun idfs_of_const thy c =
- let
- val c' = (c, Sign.the_const_type thy c);
- val c'' = CodegenConsts.const_of_cexpr thy c';
- in (c'', CodegenNames.const thy c'') end;
-
-fun no_bindings x = (Option.map o apsnd)
- (fn pretty => fn pr => fn vars => pretty (pr vars)) x;
-
-fun gen_add_haskell_monad prep_const c_run c_mbind c_kbind thy =
- let
- val c_run' = prep_const thy c_run;
- val c_mbind' = prep_const thy c_mbind;
- val c_mbind'' = CodegenNames.const thy c_mbind';
- val c_kbind' = prep_const thy c_kbind;
- val c_kbind'' = CodegenNames.const thy c_kbind';
- val pr = pretty_haskell_monad c_mbind'' c_kbind''
- in
- thy
- |> gen_add_syntax_const (K I) target_Haskell c_run' (SOME pr)
- |> gen_add_syntax_const (K I) target_Haskell c_mbind'
- (no_bindings (SOME (parse_infix fst (L, 1) ">>=")))
- |> gen_add_syntax_const (K I) target_Haskell c_kbind'
- (no_bindings (SOME (parse_infix fst (L, 1) ">>")))
- end;
-
-fun add_reserved target =
- let
- fun add sym syms = if member (op =) syms sym
- then error ("Reserved symbol " ^ quote sym ^ " already declared")
- else insert (op =) sym syms
- in map_reserveds target o add end;
-
-fun add_modl_alias target =
- map_syntax_modls target o apfst o Symtab.update o apsnd CodegenNames.check_modulename;
-
-fun add_modl_prolog target =
- map_syntax_modls target o apsnd o
- (fn (modl, NONE) => Symtab.delete modl | (modl, SOME prolog) =>
- Symtab.update (modl, Pretty.str prolog));
-
-fun zip_list (x::xs) f g =
- f
- #-> (fn y =>
- fold_map (fn x => g |-- f >> pair x) xs
- #-> (fn xys => pair ((x, y) :: xys)));
-
-structure P = OuterParse
-and K = OuterKeyword
-
-fun parse_multi_syntax parse_thing parse_syntax =
- P.and_list1 parse_thing
- #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
- (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")"));
-
-val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr");
-
-fun parse_syntax prep_arg xs =
- Scan.option ((
- ((P.$$$ infixK >> K X)
- || (P.$$$ infixlK >> K L)
- || (P.$$$ infixrK >> K R))
- -- P.nat >> parse_infix prep_arg
- || Scan.succeed (parse_mixfix prep_arg))
- -- P.string
- >> (fn (parse, s) => parse s)) xs;
-
-val (code_classK, code_instanceK, code_typeK, code_constK, code_monadK,
- code_reservedK, code_modulenameK, code_moduleprologK) =
- ("code_class", "code_instance", "code_type", "code_const", "code_monad",
- "code_reserved", "code_modulename", "code_moduleprolog");
-
-in
-
-val parse_syntax = parse_syntax;
-
-val add_syntax_class = gen_add_syntax_class cert_class (K I);
-val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco;
-val add_syntax_tyco = gen_add_syntax_tyco cert_tyco;
-val add_syntax_const = gen_add_syntax_const (K I);
-
-val add_syntax_class_cmd = gen_add_syntax_class read_class CodegenConsts.read_const;
-val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
-val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
-val add_syntax_const_cmd = gen_add_syntax_const CodegenConsts.read_const;
-
-fun add_syntax_tycoP target tyco = parse_syntax I >> add_syntax_tyco_cmd target tyco;
-fun add_syntax_constP target c = parse_syntax fst >> (add_syntax_const_cmd target c o no_bindings);
-
-fun add_undefined target undef target_undefined thy =
- let
- val (undef', _) = idfs_of_const thy undef;
- fun pr _ _ _ _ = str target_undefined;
- in
- thy
- |> add_syntax_const target undef' (SOME (~1, pr))
- end;
-
-fun add_pretty_list target nill cons thy =
- let
- val (_, nil'') = idfs_of_const thy nill;
- val (cons', cons'') = idfs_of_const thy cons;
- val pr = pretty_list nil'' cons'' target;
- in
- thy
- |> add_syntax_const target cons' (SOME pr)
- end;
-
-fun add_pretty_list_string target nill cons charr nibbles thy =
- let
- val (_, nil'') = idfs_of_const thy nill;
- val (cons', cons'') = idfs_of_const thy cons;
- val (_, charr'') = idfs_of_const thy charr;
- val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
- val pr = pretty_list_string nil'' cons'' charr'' nibbles'' target;
- in
- thy
- |> add_syntax_const target cons' (SOME pr)
- end;
-
-fun add_pretty_char target charr nibbles thy =
- let
- val (charr', charr'') = idfs_of_const thy charr;
- val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
- val pr = pretty_char charr'' nibbles'' target;
- in
- thy
- |> add_syntax_const target charr' (SOME pr)
- end;
-
-fun add_pretty_numeral target unbounded number_of b0 b1 pls min bit thy =
- let
- val number_of' = CodegenConsts.const_of_cexpr thy number_of;
- val (_, b0'') = idfs_of_const thy b0;
- val (_, b1'') = idfs_of_const thy b1;
- val (_, pls'') = idfs_of_const thy pls;
- val (_, min'') = idfs_of_const thy min;
- val (_, bit'') = idfs_of_const thy bit;
- val pr = pretty_numeral unbounded b0'' b1'' pls'' min'' bit'' target;
- in
- thy
- |> add_syntax_const target number_of' (SOME pr)
- end;
-
-fun add_pretty_ml_string target charr nibbles nill cons str thy =
- let
- val (_, charr'') = idfs_of_const thy charr;
- val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
- val (_, nil'') = idfs_of_const thy nill;
- val (_, cons'') = idfs_of_const thy cons;
- val (str', _) = idfs_of_const thy str;
- val pr = pretty_ml_string charr'' nibbles'' nil'' cons'' target;
- in
- thy
- |> add_syntax_const target str' (SOME pr)
- end;
-
-fun add_pretty_imperative_monad_bind target bind thy =
- let
- val (bind', _) = idfs_of_const thy bind;
- val pr = pretty_imperative_monad_bind
- in
- thy
- |> add_syntax_const target bind' (SOME pr)
- end;
-
-val add_haskell_monad = gen_add_haskell_monad CodegenConsts.read_const;
-
-val code_classP =
- OuterSyntax.command code_classK "define code syntax for class" K.thy_decl (
- parse_multi_syntax P.xname
- (Scan.option (P.string -- Scan.optional (P.$$$ "where" |-- Scan.repeat1
- (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.string)) []))
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns)
- );
-
-val code_instanceP =
- OuterSyntax.command code_instanceK "define code syntax for instance" K.thy_decl (
- parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname)
- ((P.minus >> K true) || Scan.succeed false)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns)
- );
-
-val code_typeP =
- OuterSyntax.command code_typeK "define code syntax for type constructor" K.thy_decl (
- parse_multi_syntax P.xname (parse_syntax I)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns)
- );
-
-val code_constP =
- OuterSyntax.command code_constK "define code syntax for constant" K.thy_decl (
- parse_multi_syntax P.term (parse_syntax fst)
- >> (Toplevel.theory oo fold) (fn (target, syns) =>
- fold (fn (raw_const, syn) => add_syntax_const_cmd target raw_const (no_bindings syn)) syns)
- );
-
-val code_monadP =
- OuterSyntax.command code_monadK "define code syntax for Haskell monads" K.thy_decl (
- P.term -- P.term -- P.term
- >> (fn ((raw_run, raw_mbind), raw_kbind) => Toplevel.theory
- (add_haskell_monad raw_run raw_mbind raw_kbind))
- );
-
-val code_reservedP =
- OuterSyntax.command code_reservedK "declare words as reserved for target language" K.thy_decl (
- P.name -- Scan.repeat1 P.name
- >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)
- )
-
-val code_modulenameP =
- OuterSyntax.command code_modulenameK "alias module to other name" K.thy_decl (
- P.name -- Scan.repeat1 (P.name -- P.name)
- >> (fn (target, modlnames) => (Toplevel.theory o fold (add_modl_alias target)) modlnames)
- )
-
-val code_moduleprologP =
- OuterSyntax.command code_moduleprologK "add prolog to module" K.thy_decl (
- P.name -- Scan.repeat1 (P.name -- (P.text >> (fn "-" => NONE | s => SOME s)))
- >> (fn (target, prologs) => (Toplevel.theory o fold (add_modl_prolog target)) prologs)
- )
-
-val _ = OuterSyntax.add_keywords [infixK, infixlK, infixrK];
-
-val _ = OuterSyntax.add_parsers [code_classP, code_instanceP, code_typeP, code_constP,
- code_reservedP, code_modulenameP, code_moduleprologP, code_monadP];
-
-
-(*including serializer defaults*)
-val _ = Context.add_setup (
- add_syntax_tyco "SML" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- (gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> add_syntax_tyco "OCaml" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- (gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> add_syntax_tyco "Haskell" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy [
- pr_typ (INFX (1, X)) ty1,
- str "->",
- pr_typ (INFX (1, R)) ty2
- ]))
- #> fold (add_reserved "SML") ML_Syntax.reserved_names
- #> fold (add_reserved "SML")
- ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)]
- #> fold (add_reserved "OCaml") [
- "and", "as", "assert", "begin", "class",
- "constraint", "do", "done", "downto", "else", "end", "exception",
- "external", "false", "for", "fun", "function", "functor", "if",
- "in", "include", "inherit", "initializer", "lazy", "let", "match", "method",
- "module", "mutable", "new", "object", "of", "open", "or", "private", "rec",
- "sig", "struct", "then", "to", "true", "try", "type", "val",
- "virtual", "when", "while", "with"
- ]
- #> fold (add_reserved "OCaml") ["failwith", "mod"]
- #> fold (add_reserved "Haskell") [
- "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
- "import", "default", "forall", "let", "in", "class", "qualified", "data",
- "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
- ]
- #> fold (add_reserved "Haskell") [
- "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int",
- "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded",
- "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor",
- "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException",
- "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException",
- "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure",
- "ExitSuccess", "False", "GT", "HeapOverflow",
- "IOError", "IOException", "IllegalOperation",
- "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError",
- "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow",
- "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError",
- "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow",
- "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow",
- "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all",
- "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan",
- "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen",
- "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break",
- "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const",
- "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod",
- "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem",
- "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo",
- "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip",
- "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational",
- "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble",
- "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj",
- "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head",
- "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha",
- "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite",
- "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'",
- "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log",
- "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum",
- "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem",
- "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo",
- "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred",
- "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn",
- "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat",
- "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile",
- "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn",
- "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational",
- "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse",
- "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence",
- "sequence_", "show", "showChar", "showException", "showField", "showList",
- "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand",
- "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract",
- "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult",
- "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry",
- "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords",
- "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3"
- ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*)
-
-)
-
-end; (*local*)
-
-end; (*struct*)
--- a/src/Pure/Tools/codegen_thingol.ML Fri Aug 10 17:04:24 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,442 +0,0 @@
-(* Title: Pure/Tools/codegen_thingol.ML
- ID: $Id$
- Author: Florian Haftmann, TU Muenchen
-
-Intermediate language ("Thin-gol") representing executable code.
-*)
-
-infix 8 `%%;
-infixr 6 `->;
-infixr 6 `-->;
-infix 4 `$;
-infix 4 `$$;
-infixr 3 `|->;
-infixr 3 `|-->;
-
-signature BASIC_CODEGEN_THINGOL =
-sig
- type vname = string;
- datatype dict =
- DictConst of string * dict list list
- | DictVar of string list * (vname * (int * int));
- datatype itype =
- `%% of string * itype list
- | ITyVar of vname;
- datatype iterm =
- IConst of string * (dict list list * itype list (*types of arguments*))
- | IVar of vname
- | `$ of iterm * iterm
- | `|-> of (vname * itype) * iterm
- | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
- (*((term, type), [(selector pattern, body term )]), primitive term)*)
- val `-> : itype * itype -> itype;
- val `--> : itype list * itype -> itype;
- val `$$ : iterm * iterm list -> iterm;
- val `|--> : (vname * itype) list * iterm -> iterm;
- type typscheme = (vname * sort) list * itype;
-end;
-
-signature CODEGEN_THINGOL =
-sig
- include BASIC_CODEGEN_THINGOL;
- val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list;
- val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a;
- val unfold_fun: itype -> itype list * itype;
- val unfold_app: iterm -> iterm * iterm list;
- val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option;
- val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm;
- val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option;
- val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm;
- val unfold_const_app: iterm ->
- ((string * (dict list list * itype list)) * iterm list) option;
- val collapse_let: ((vname * itype) * iterm) * iterm
- -> (iterm * itype) * (iterm * iterm) list;
- val eta_expand: (string * (dict list list * itype list)) * iterm list -> int -> iterm;
- val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
- val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
- val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
-
- datatype def =
- Bot
- | Fun of (iterm list * iterm) list * typscheme
- | Datatype of (vname * sort) list * (string * itype list) list
- | Datatypecons of string
- | Class of (class * string) list * (vname * (string * itype) list)
- | Classop of class
- | Classrel of class * class
- | Classinst of (class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * (string * iterm) list);
- type code = def Graph.T;
- type transact;
- val empty_code: code;
- val merge_code: code * code -> code;
- val project_code: bool (*delete empty funs*)
- -> string list (*hidden*) -> string list option (*selected*)
- -> code -> code;
- val empty_funs: code -> string list;
- val is_cons: code -> string -> bool;
- val add_eval_def: string (*bind name*) * iterm -> code -> code;
-
- val ensure_def: (string -> string) -> (transact -> def * code) -> string
- -> string -> transact -> transact;
- val succeed: 'a -> transact -> 'a * code;
- val fail: string -> transact -> 'a * code;
- val message: string -> (transact -> 'a) -> transact -> 'a;
- val start_transact: (transact -> 'a * transact) -> code -> 'a * code;
-
- val trace: bool ref;
- val tracing: ('a -> string) -> 'a -> 'a;
-end;
-
-structure CodegenThingol: CODEGEN_THINGOL =
-struct
-
-(** auxiliary **)
-
-val trace = ref false;
-fun tracing f x = (if !trace then Output.tracing (f x) else (); x);
-
-fun unfoldl dest x =
- case dest x
- of NONE => (x, [])
- | SOME (x1, x2) =>
- let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end;
-
-fun unfoldr dest x =
- case dest x
- of NONE => ([], x)
- | SOME (x1, x2) =>
- let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end;
-
-
-(** language core - types, pattern, expressions **)
-
-(* language representation *)
-
-type vname = string;
-
-datatype dict =
- DictConst of string * dict list list
- | DictVar of string list * (vname * (int * int));
-
-datatype itype =
- `%% of string * itype list
- | ITyVar of vname;
-
-datatype iterm =
- IConst of string * (dict list list * itype list)
- | IVar of vname
- | `$ of iterm * iterm
- | `|-> of (vname * itype) * iterm
- | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
- (*see also signature*)
-
-(*
- variable naming conventions
-
- bare names:
- variable names v
- class names class
- type constructor names tyco
- datatype names dtco
- const names (general) c
- constructor names co
- class operation names clsop (op)
- arbitrary name s
-
- v, c, co, clsop also annotated with types etc.
-
- constructs:
- sort sort
- type parameters vs
- type ty
- type schemes tysm
- term t
- (term as pattern) p
- instance (class, tyco) inst
- *)
-
-fun ty1 `-> ty2 = "fun" `%% [ty1, ty2];
-val op `--> = Library.foldr (op `->);
-val op `$$ = Library.foldl (op `$);
-val op `|--> = Library.foldr (op `|->);
-
-val unfold_fun = unfoldr
- (fn "fun" `%% [ty1, ty2] => SOME (ty1, ty2)
- | _ => NONE);
-
-val unfold_app = unfoldl
- (fn op `$ t => SOME t
- | _ => NONE);
-
-val split_abs =
- (fn (v, ty) `|-> (t as ICase (((IVar w, _), [(p, t')]), _)) =>
- if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t)
- | (v, ty) `|-> t => SOME (((v, NONE), ty), t)
- | _ => NONE);
-
-val unfold_abs = unfoldr split_abs;
-
-val split_let =
- (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)
- | _ => NONE);
-
-val unfold_let = unfoldr split_let;
-
-fun unfold_const_app t =
- case unfold_app t
- of (IConst c, ts) => SOME (c, ts)
- | _ => NONE;
-
-fun fold_aiterms f (t as IConst _) = f t
- | fold_aiterms f (t as IVar _) = f t
- | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2
- | fold_aiterms f (t as _ `|-> t') = f t #> fold_aiterms f t'
- | fold_aiterms f (ICase (_, t)) = fold_aiterms f t;
-
-fun fold_constnames f =
- let
- fun add (IConst (c, _)) = f c
- | add _ = I;
- in fold_aiterms add end;
-
-fun fold_varnames f =
- let
- fun add (IVar v) = f v
- | add ((v, _) `|-> _) = f v
- | add _ = I;
- in fold_aiterms add end;
-
-fun fold_unbound_varnames f =
- let
- fun add _ (IConst _) = I
- | add vs (IVar v) = if not (member (op =) vs v) then f v else I
- | add vs (t1 `$ t2) = add vs t1 #> add vs t2
- | add vs ((v, _) `|-> t) = add (insert (op =) v vs) t
- | add vs (ICase (_, t)) = add vs t;
- in add [] end;
-
-fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) =
- let
- fun exists_v t = fold_unbound_varnames (fn w => fn b =>
- b orelse v = w) t false;
- in if v = w andalso forall (fn (t1, t2) =>
- exists_v t1 orelse not (exists_v t2)) ds
- then ((se, ty), ds)
- else ((se, ty), [(IVar v, be)])
- end
- | collapse_let (((v, ty), se), be) =
- ((se, ty), [(IVar v, be)])
-
-fun eta_expand (c as (_, (_, tys)), ts) k =
- let
- val j = length ts;
- val l = k - j;
- val ctxt = (fold o fold_varnames) Name.declare ts Name.context;
- val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys);
- in vs_tys `|--> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end;
-
-
-(** definitions, transactions **)
-
-(* type definitions *)
-
-type typscheme = (vname * sort) list * itype;
-datatype def =
- Bot
- | Fun of (iterm list * iterm) list * typscheme
- | Datatype of (vname * sort) list * (string * itype list) list
- | Datatypecons of string
- | Class of (class * string) list * (vname * (string * itype) list)
- | Classop of class
- | Classrel of class * class
- | Classinst of (class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * (string * iterm) list);
-
-type code = def Graph.T;
-type transact = Graph.key option * code;
-exception FAIL of string list;
-
-
-(* abstract code *)
-
-val empty_code = Graph.empty : code; (*read: "depends on"*)
-
-fun ensure_bot name = Graph.default_node (name, Bot);
-
-fun add_def_incr (name, Bot) code =
- (case the_default Bot (try (Graph.get_node code) name)
- of Bot => error "Attempted to add Bot to code"
- | _ => code)
- | add_def_incr (name, def) code =
- (case try (Graph.get_node code) name
- of NONE => Graph.new_node (name, def) code
- | SOME Bot => Graph.map_node name (K def) code
- | SOME _ => error ("Tried to overwrite definition " ^ quote name));
-
-fun add_dep (dep as (name1, name2)) =
- if name1 = name2 then I else Graph.add_edge dep;
-
-val merge_code : code * code -> code = Graph.merge (K true);
-
-fun project_code delete_empty_funs hidden raw_selected code =
- let
- fun is_empty_fun name = case Graph.get_node code name
- of Fun ([], _) => true
- | _ => false;
- val names = subtract (op =) hidden (Graph.keys code);
- val deleted = Graph.all_preds code (filter is_empty_fun names);
- val selected = case raw_selected
- of NONE => names |> subtract (op =) deleted
- | SOME sel => sel
- |> delete_empty_funs ? subtract (op =) deleted
- |> subtract (op =) hidden
- |> Graph.all_succs code
- |> delete_empty_funs ? subtract (op =) deleted
- |> subtract (op =) hidden;
- in
- code
- |> Graph.subgraph (member (op =) selected)
- end;
-
-fun empty_funs code =
- Graph.fold (fn (name, (Fun ([], _), _)) => cons name
- | _ => I) code [];
-
-fun is_cons code name = case Graph.get_node code name
- of Datatypecons _ => true
- | _ => false;
-
-fun check_samemodule names =
- fold (fn name =>
- let
- val module_name = (NameSpace.qualifier o NameSpace.qualifier) name
- in
- fn NONE => SOME module_name
- | SOME module_name' => if module_name = module_name' then SOME module_name
- else error ("Inconsistent name prefix for simultanous names: " ^ commas_quote names)
- end
- ) names NONE;
-
-fun check_funeqs eqs =
- (fold (fn (pats, _) =>
- let
- val l = length pats
- in
- fn NONE => SOME l
- | SOME l' => if l = l' then SOME l
- else error "Function definition with different number of arguments"
- end
- ) eqs NONE; eqs);
-
-fun check_prep_def code Bot =
- Bot
- | check_prep_def code (Fun (eqs, d)) =
- Fun (check_funeqs eqs, d)
- | check_prep_def code (d as Datatype _) =
- d
- | check_prep_def code (Datatypecons dtco) =
- error "Attempted to add bare term constructor"
- | check_prep_def code (d as Class _) =
- d
- | check_prep_def code (Classop _) =
- error "Attempted to add bare class operation"
- | check_prep_def code (Classrel _) =
- error "Attempted to add bare class relation"
- | check_prep_def code (d as Classinst ((class, (tyco, arity)), (_, inst_classops))) =
- let
- val Class (_, (_, classops)) = Graph.get_node code class;
- val _ = if length inst_classops > length classops
- then error "Too many class operations given"
- else ();
- fun check_classop (f, _) =
- if AList.defined (op =) inst_classops f
- then () else error ("Missing definition for class operation " ^ quote f);
- val _ = map check_classop classops;
- in d end;
-
-fun postprocess_def (name, Datatype (_, constrs)) =
- tap (fn _ => check_samemodule (name :: map fst constrs))
- #> fold (fn (co, _) =>
- add_def_incr (co, Datatypecons name)
- #> add_dep (co, name)
- #> add_dep (name, co)
- ) constrs
- | postprocess_def (name, Class (classrels, (_, classops))) =
- tap (fn _ => check_samemodule (name :: map fst classops @ map snd classrels))
- #> fold (fn (f, _) =>
- add_def_incr (f, Classop name)
- #> add_dep (f, name)
- #> add_dep (name, f)
- ) classops
- #> fold (fn (superclass, classrel) =>
- add_def_incr (classrel, Classrel (name, superclass))
- #> add_dep (classrel, name)
- #> add_dep (name, classrel)
- ) classrels
- | postprocess_def _ =
- I;
-
-
-(* transaction protocol *)
-
-fun ensure_def labelled_name defgen msg name (dep, code) =
- let
- val msg' = (case dep
- of NONE => msg
- | SOME dep => msg ^ ", required for " ^ labelled_name dep);
- fun add_dp NONE = I
- | add_dp (SOME dep) =
- tracing (fn _ => "adding dependency " ^ labelled_name dep
- ^ " -> " ^ labelled_name name)
- #> add_dep (dep, name);
- fun prep_def def code =
- (check_prep_def code def, code);
- fun invoke_generator name defgen code =
- defgen (SOME name, code) handle FAIL msgs => raise FAIL (msg' :: msgs);
- fun add_def false =
- ensure_bot name
- #> add_dp dep
- #> invoke_generator name defgen
- #-> (fn def => prep_def def)
- #-> (fn def => add_def_incr (name, def)
- #> postprocess_def (name, def))
- | add_def true =
- add_dp dep;
- in
- code
- |> add_def (can (Graph.get_node code) name)
- |> pair dep
- end;
-
-fun succeed some (_, code) = (some, code);
-
-fun fail msg (_, code) = raise FAIL [msg];
-
-fun message msg f trns =
- f trns handle FAIL msgs =>
- raise FAIL (msg :: msgs);
-
-fun start_transact f code =
- let
- fun handle_fail f x =
- (f x
- handle FAIL msgs =>
- (error o cat_lines) ("Code generation failed, while:" :: msgs))
- in
- (NONE, code)
- |> handle_fail f
- |-> (fn x => fn (_, code) => (x, code))
- end;
-
-fun add_eval_def (name, t) code =
- code
- |> Graph.new_node (name, Fun ([([], t)], ([("_", [])], ITyVar "_")))
- |> fold (curry Graph.add_edge name) (Graph.keys code);
-
-end; (*struct*)
-
-
-structure BasicCodegenThingol: BASIC_CODEGEN_THINGOL = CodegenThingol;
--- a/src/Pure/Tools/nbe.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Tools/nbe.ML Fri Aug 10 17:04:34 2007 +0200
@@ -65,7 +65,7 @@
(* theorem store *)
-structure Funcgr = CodegenFuncgrRetrieval (val rewrites = the_pres);
+structure Funcgr = CodeFuncgrRetrieval (val rewrites = the_pres);
(* code store *)
@@ -110,14 +110,14 @@
fun ensure_funs thy funcgr t =
let
fun consts_of thy t =
- fold_aterms (fn Const c => cons (CodegenConsts.const_of_cexpr thy c) | _ => I) t []
+ fold_aterms (fn Const c => cons (CodeUnit.const_of_cexpr thy c) | _ => I) t []
val consts = consts_of thy t;
val nbe_tab = NBE_Data.get thy;
in
- CodegenFuncgr.deps funcgr consts
- |> (map o filter_out) (Symtab.defined nbe_tab o CodegenNames.const thy)
+ CodeFuncgr.deps funcgr consts
+ |> (map o filter_out) (Symtab.defined nbe_tab o CodeNames.const thy)
|> filter_out null
- |> (map o map) (fn c => (CodegenNames.const thy c, CodegenFuncgr.funcs funcgr c))
+ |> (map o map) (fn c => (CodeNames.const thy c, CodeFuncgr.funcs funcgr c))
|> generate thy
end;
@@ -156,7 +156,7 @@
(* evaluation oracle *)
-exception Normalization of CodegenFuncgr.T * term;
+exception Normalization of CodeFuncgr.T * term;
fun normalization_oracle (thy, Normalization (funcgr, t)) =
Logic.mk_equals (t, eval_term thy funcgr t);
--- a/src/Pure/Tools/nbe_codegen.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Tools/nbe_codegen.ML Fri Aug 10 17:04:34 2007 +0200
@@ -149,7 +149,7 @@
let
fun to_term bounds (C s) tcount =
let
- val SOME (const as (c, _)) = CodegenNames.const_rev thy s;
+ val SOME (const as (c, _)) = CodeNames.const_rev thy s;
val ty = CodegenData.default_typ thy const;
val ty' = map_type_tvar (fn ((s,i),S) => TypeInfer.param (tcount + i) (s,S)) ty;
val tcount' = tcount + maxidx_of_typ ty + 1;
--- a/src/Pure/Tools/nbe_eval.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/Tools/nbe_eval.ML Fri Aug 10 17:04:34 2007 +0200
@@ -107,8 +107,8 @@
(* ------------------ evaluation with greetings to Tarski ------------------ *)
-fun prep_term thy (Const c) = Const (CodegenNames.const thy
- (CodegenConsts.const_of_cexpr thy c), dummyT)
+fun prep_term thy (Const c) = Const (CodeNames.const thy
+ (CodeUnit.const_of_cexpr thy c), dummyT)
| prep_term thy (Free v_ty) = Free v_ty
| prep_term thy (s $ t) = prep_term thy s $ prep_term thy t
| prep_term thy (Abs (raw_v, ty, raw_t)) =
--- a/src/Pure/codegen.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Pure/codegen.ML Fri Aug 10 17:04:34 2007 +0200
@@ -26,7 +26,6 @@
val add_codegen: string -> term codegen -> theory -> theory
val add_tycodegen: string -> typ codegen -> theory -> theory
- val add_attribute: string -> (Args.T list -> attribute * Args.T list) -> theory -> theory
val add_preprocessor: (theory -> thm list -> thm list) -> theory -> theory
val preprocess: theory -> thm list -> thm list
val preprocess_term: theory -> term -> term
@@ -202,8 +201,6 @@
(* theory data *)
-structure CodeData = CodegenData;
-
structure CodegenData = TheoryDataFun
(
type T =
@@ -211,29 +208,27 @@
tycodegens : (string * typ codegen) list,
consts : ((string * typ) * (term mixfix list * (string * string) list)) list,
types : (string * (typ mixfix list * (string * string) list)) list,
- attrs: (string * (Args.T list -> attribute * Args.T list)) list,
preprocs: (stamp * (theory -> thm list -> thm list)) list,
modules: codegr Symtab.table,
test_params: test_params};
val empty =
- {codegens = [], tycodegens = [], consts = [], types = [], attrs = [],
+ {codegens = [], tycodegens = [], consts = [], types = [],
preprocs = [], modules = Symtab.empty, test_params = default_test_params};
val copy = I;
val extend = I;
fun merge _
({codegens = codegens1, tycodegens = tycodegens1,
- consts = consts1, types = types1, attrs = attrs1,
+ consts = consts1, types = types1,
preprocs = preprocs1, modules = modules1, test_params = test_params1},
{codegens = codegens2, tycodegens = tycodegens2,
- consts = consts2, types = types2, attrs = attrs2,
+ consts = consts2, types = types2,
preprocs = preprocs2, modules = modules2, test_params = test_params2}) =
{codegens = AList.merge (op =) (K true) (codegens1, codegens2),
tycodegens = AList.merge (op =) (K true) (tycodegens1, tycodegens2),
consts = AList.merge (op =) (K true) (consts1, consts2),
types = AList.merge (op =) (K true) (types1, types2),
- attrs = AList.merge (op =) (K true) (attrs1, attrs2),
preprocs = AList.merge (op =) (K true) (preprocs1, preprocs2),
modules = Symtab.merge (K true) (modules1, modules2),
test_params = merge_test_params test_params1 test_params2};
@@ -253,10 +248,10 @@
fun get_test_params thy = #test_params (CodegenData.get thy);
fun map_test_params f thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ let val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy;
in CodegenData.put {codegens = codegens, tycodegens = tycodegens,
- consts = consts, types = types, attrs = attrs, preprocs = preprocs,
+ consts = consts, types = types, preprocs = preprocs,
modules = modules, test_params = f test_params} thy
end;
@@ -266,10 +261,10 @@
fun get_modules thy = #modules (CodegenData.get thy);
fun map_modules f thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ let val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy;
in CodegenData.put {codegens = codegens, tycodegens = tycodegens,
- consts = consts, types = types, attrs = attrs, preprocs = preprocs,
+ consts = consts, types = types, preprocs = preprocs,
modules = f modules, test_params = test_params} thy
end;
@@ -277,23 +272,23 @@
(**** add new code generators to theory ****)
fun add_codegen name f thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ let val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy
in (case AList.lookup (op =) codegens name of
NONE => CodegenData.put {codegens = (name, f) :: codegens,
tycodegens = tycodegens, consts = consts, types = types,
- attrs = attrs, preprocs = preprocs, modules = modules,
+ preprocs = preprocs, modules = modules,
test_params = test_params} thy
| SOME _ => error ("Code generator " ^ name ^ " already declared"))
end;
fun add_tycodegen name f thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ let val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy
in (case AList.lookup (op =) tycodegens name of
NONE => CodegenData.put {tycodegens = (name, f) :: tycodegens,
codegens = codegens, consts = consts, types = types,
- attrs = attrs, preprocs = preprocs, modules = modules,
+ preprocs = preprocs, modules = modules,
test_params = test_params} thy
| SOME _ => error ("Code generator " ^ name ^ " already declared"))
end;
@@ -302,11 +297,11 @@
(**** preprocessors ****)
fun add_preprocessor p thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ let val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy
in CodegenData.put {tycodegens = tycodegens,
codegens = codegens, consts = consts, types = types,
- attrs = attrs, preprocs = (stamp (), p) :: preprocs,
+ preprocs = (stamp (), p) :: preprocs,
modules = modules, test_params = test_params} thy
end;
@@ -341,53 +336,20 @@
end)
in add_preprocessor prep end;
-
-(**** code attribute ****)
-
-fun add_attribute name att thy =
- let val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
- CodegenData.get thy
- in (case AList.lookup (op =) attrs name of
- NONE => CodegenData.put {tycodegens = tycodegens,
- codegens = codegens, consts = consts, types = types,
- attrs = if name = "" then attrs @ [(name, att)] else (name, att) :: attrs,
- preprocs = preprocs, modules = modules,
- test_params = test_params} thy
- | SOME _ => error ("Code attribute " ^ name ^ " already declared"))
- end;
-
-fun mk_parser (a, p) = (if a = "" then Scan.succeed "" else Args.$$$ a) |-- p;
-
-val code_attr =
- Attrib.syntax (Scan.peek (fn context => List.foldr op || Scan.fail (map mk_parser
- (#attrs (CodegenData.get (Context.theory_of context))))));
-
val _ = Context.add_setup
- (Attrib.add_attributes [("code", code_attr, "declare theorems for code generation")]);
-
-local
- fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
- fun add_simple_attribute (name, f) =
- add_attribute name (Scan.succeed (mk_attribute f));
- fun add_del_attribute (name, (add, del)) =
- add_attribute name (Args.del |-- Scan.succeed (mk_attribute del)
- || Scan.succeed (mk_attribute add))
-in
- val _ = Context.add_setup (add_simple_attribute ("unfold",
- fn thm => add_unfold thm #> CodeData.add_inline thm));
- val _ = map (Context.add_setup o add_del_attribute) [
- ("func", (CodeData.add_func true, CodeData.del_func)),
- ("inline", (CodeData.add_inline, CodeData.del_inline)),
- ("post", (CodeData.add_post, CodeData.del_post))
- ];
-end; (*local*)
+ (let
+ fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+ in
+ Code.add_attribute ("unfold", Scan.succeed (mk_attribute
+ (fn thm => add_unfold thm #> Code.add_inline thm)))
+ end);
(**** associate constants with target language code ****)
fun gen_assoc_const prep_const (raw_const, syn) thy =
let
- val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy;
val (cname, T) = prep_const thy raw_const;
in
@@ -397,20 +359,20 @@
NONE => CodegenData.put {codegens = codegens,
tycodegens = tycodegens,
consts = ((cname, T), syn) :: consts,
- types = types, attrs = attrs, preprocs = preprocs,
+ types = types, preprocs = preprocs,
modules = modules, test_params = test_params} thy
| SOME _ => error ("Constant " ^ cname ^ " already associated with code")
end;
val assoc_const_i = gen_assoc_const (K I);
-val assoc_const = gen_assoc_const CodegenConsts.read_bare_const;
+val assoc_const = gen_assoc_const CodeUnit.read_bare_const;
(**** associate types with target language types ****)
fun assoc_type (s, syn) thy =
let
- val {codegens, tycodegens, consts, types, attrs, preprocs, modules, test_params} =
+ val {codegens, tycodegens, consts, types, preprocs, modules, test_params} =
CodegenData.get thy;
val tc = Sign.intern_type thy s;
in
@@ -421,7 +383,7 @@
else (case AList.lookup (op =) types tc of
NONE => CodegenData.put {codegens = codegens,
tycodegens = tycodegens, consts = consts,
- types = (tc, syn) :: types, attrs = attrs,
+ types = (tc, syn) :: types,
preprocs = preprocs, modules = modules, test_params = test_params} thy
| SOME _ => error ("Type " ^ tc ^ " already associated with code"))
| _ => error ("Not a type constructor: " ^ s)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_funcgr.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,407 @@
+(* Title: Tools/code/code_funcgr.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Retrieving, normalizing and structuring defining equations in graph
+with explicit dependencies.
+*)
+
+signature CODE_FUNCGR =
+sig
+ type T
+ val timing: bool ref
+ val funcs: T -> CodeUnit.const -> thm list
+ val typ: T -> CodeUnit.const -> typ
+ val deps: T -> CodeUnit.const list -> CodeUnit.const list list
+ val all: T -> CodeUnit.const list
+ val pretty: theory -> T -> Pretty.T
+ structure Constgraph : GRAPH
+end
+
+signature CODE_FUNCGR_RETRIEVAL =
+sig
+ type T (* = CODE_FUNCGR.T *)
+ val make: theory -> CodeUnit.const list -> T
+ val make_consts: theory -> CodeUnit.const list -> CodeUnit.const list * T
+ val make_term: theory -> (T -> (thm -> thm) -> cterm -> thm -> 'a) -> cterm -> 'a * T
+ (*FIXME drop make_term as soon as possible*)
+ val eval_conv: theory -> (T -> cterm -> thm) -> cterm -> thm
+ val intervene: theory -> T -> T
+ (*FIXME drop intervene as soon as possible*)
+end;
+
+structure CodeFuncgr = (*signature is added later*)
+struct
+
+(** the graph type **)
+
+structure Constgraph = GraphFun (
+ type key = CodeUnit.const;
+ val ord = CodeUnit.const_ord;
+);
+
+type T = (typ * thm list) Constgraph.T;
+
+fun funcs funcgr =
+ these o Option.map snd o try (Constgraph.get_node funcgr);
+
+fun typ funcgr =
+ fst o Constgraph.get_node funcgr;
+
+fun deps funcgr cs =
+ let
+ val conn = Constgraph.strong_conn funcgr;
+ val order = rev conn;
+ in
+ (map o filter) (member (op =) (Constgraph.all_succs funcgr cs)) order
+ |> filter_out null
+ end;
+
+fun all funcgr = Constgraph.keys funcgr;
+
+fun pretty thy funcgr =
+ AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr)
+ |> (map o apfst) (CodeUnit.string_of_const thy)
+ |> sort (string_ord o pairself fst)
+ |> map (fn (s, thms) =>
+ (Pretty.block o Pretty.fbreaks) (
+ Pretty.str s
+ :: map Display.pretty_thm thms
+ ))
+ |> Pretty.chunks;
+
+
+(** generic combinators **)
+
+fun fold_consts f thms =
+ thms
+ |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of)
+ |> (fold o fold_aterms) (fn Const c => f c | _ => I);
+
+fun consts_of (const, []) = []
+ | consts_of (const, thms as thm :: _) =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val is_refl = curry CodeUnit.eq_const const;
+ fun the_const c = case try (CodeUnit.const_of_cexpr thy) c
+ of SOME const => if is_refl const then I else insert CodeUnit.eq_const const
+ | NONE => I
+ in fold_consts the_const thms [] end;
+
+fun insts_of thy algebra c ty_decl ty =
+ let
+ val tys_decl = Sign.const_typargs thy (c, ty_decl);
+ val tys = Sign.const_typargs thy (c, ty);
+ fun class_relation (x, _) _ = x;
+ fun type_constructor tyco xs class =
+ (tyco, class) :: maps (maps fst) xs;
+ fun type_variable (TVar (_, sort)) = map (pair []) sort
+ | type_variable (TFree (_, sort)) = map (pair []) sort;
+ fun mk_inst ty (TVar (_, sort)) = cons (ty, sort)
+ | mk_inst ty (TFree (_, sort)) = cons (ty, sort)
+ | mk_inst (Type (_, tys1)) (Type (_, tys2)) = fold2 mk_inst tys1 tys2;
+ fun of_sort_deriv (ty, sort) =
+ Sorts.of_sort_derivation (Sign.pp thy) algebra
+ { class_relation = class_relation, type_constructor = type_constructor,
+ type_variable = type_variable }
+ (ty, sort)
+ in
+ flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl []))
+ end;
+
+fun drop_classes thy tfrees thm =
+ let
+ val (_, thm') = Thm.varifyT' [] thm;
+ val tvars = Term.add_tvars (Thm.prop_of thm') [];
+ val unconstr = map (Thm.ctyp_of thy o TVar) tvars;
+ val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy)
+ (TVar (v_i, []), TFree (v, sort))) tvars tfrees;
+ in
+ thm'
+ |> fold Thm.unconstrainT unconstr
+ |> Thm.instantiate (instmap, [])
+ |> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy))
+ end;
+
+
+(** graph algorithm **)
+
+val timing = ref false;
+
+local
+
+exception INVALID of CodeUnit.const list * string;
+
+fun resort_thms algebra tap_typ [] = []
+ | resort_thms algebra tap_typ (thms as thm :: _) =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val cs = fold_consts (insert (op =)) thms [];
+ fun match_const c (ty, ty_decl) =
+ let
+ val tys = CodeUnit.typargs thy (c, ty);
+ val sorts = map (snd o dest_TVar) (CodeUnit.typargs thy (c, ty_decl));
+ in fold2 (curry (CodeUnit.typ_sort_inst algebra)) tys sorts end;
+ fun match (c_ty as (c, ty)) =
+ case tap_typ c_ty
+ of SOME ty_decl => match_const c (ty, ty_decl)
+ | NONE => I;
+ val tvars = fold match cs Vartab.empty;
+ in map (CodeUnit.inst_thm tvars) thms end;
+
+fun resort_funcss thy algebra funcgr =
+ let
+ val typ_funcgr = try (fst o Constgraph.get_node funcgr o CodeUnit.const_of_cexpr thy);
+ fun resort_dep (const, thms) = (const, resort_thms algebra typ_funcgr thms)
+ handle Sorts.CLASS_ERROR e => raise INVALID ([const], Sorts.msg_class_error (Sign.pp thy) e
+ ^ ",\nfor constant " ^ CodeUnit.string_of_const thy const
+ ^ "\nin defining equations\n"
+ ^ (cat_lines o map string_of_thm) thms)
+ fun resort_rec tap_typ (const, []) = (true, (const, []))
+ | resort_rec tap_typ (const, thms as thm :: _) =
+ let
+ val (_, ty) = CodeUnit.head_func thm;
+ val thms' as thm' :: _ = resort_thms algebra tap_typ thms
+ val (_, ty') = CodeUnit.head_func thm';
+ in (Sign.typ_equiv thy (ty, ty'), (const, thms')) end;
+ fun resort_recs funcss =
+ let
+ fun tap_typ c_ty = case try (CodeUnit.const_of_cexpr thy) c_ty
+ of SOME const => AList.lookup (CodeUnit.eq_const) funcss const
+ |> these
+ |> try hd
+ |> Option.map (snd o CodeUnit.head_func)
+ | NONE => NONE;
+ val (unchangeds, funcss') = split_list (map (resort_rec tap_typ) funcss);
+ val unchanged = fold (fn x => fn y => x andalso y) unchangeds true;
+ in (unchanged, funcss') end;
+ fun resort_rec_until funcss =
+ let
+ val (unchanged, funcss') = resort_recs funcss;
+ in if unchanged then funcss' else resort_rec_until funcss' end;
+ in map resort_dep #> resort_rec_until end;
+
+fun instances_of thy algebra insts =
+ let
+ val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy;
+ fun all_classops tyco class =
+ try (AxClass.params_of_class thy) class
+ |> Option.map snd
+ |> these
+ |> map (fn (c, _) => (c, SOME tyco))
+ in
+ Symtab.empty
+ |> fold (fn (tyco, class) =>
+ Symtab.map_default (tyco, []) (insert (op =) class)) insts
+ |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops tyco)
+ (Graph.all_succs thy_classes classes))) tab [])
+ end;
+
+fun instances_of_consts thy algebra funcgr consts =
+ let
+ fun inst (cexpr as (c, ty)) = insts_of thy algebra c
+ ((fst o Constgraph.get_node funcgr o CodeUnit.const_of_cexpr thy) cexpr)
+ ty handle CLASS_ERROR => [];
+ in
+ []
+ |> fold (fold (insert (op =)) o inst) consts
+ |> instances_of thy algebra
+ end;
+
+fun ensure_const' rewrites thy algebra funcgr const auxgr =
+ if can (Constgraph.get_node funcgr) const
+ then (NONE, auxgr)
+ else if can (Constgraph.get_node auxgr) const
+ then (SOME const, auxgr)
+ else if is_some (Code.get_datatype_of_constr thy const) then
+ auxgr
+ |> Constgraph.new_node (const, [])
+ |> pair (SOME const)
+ else let
+ val thms = Code.these_funcs thy const
+ |> map (CodeUnit.rewrite_func (rewrites thy))
+ |> CodeUnit.norm_args
+ |> CodeUnit.norm_varnames CodeName.purify_tvar CodeName.purify_var;
+ val rhs = consts_of (const, thms);
+ in
+ auxgr
+ |> Constgraph.new_node (const, thms)
+ |> fold_map (ensure_const rewrites thy algebra funcgr) rhs
+ |-> (fn rhs' => fold (fn SOME const' => Constgraph.add_edge (const, const')
+ | NONE => I) rhs')
+ |> pair (SOME const)
+ end
+and ensure_const rewrites thy algebra funcgr const =
+ let
+ val timeap = if !timing
+ then Output.timeap_msg ("time for " ^ CodeUnit.string_of_const thy const)
+ else I;
+ in timeap (ensure_const' rewrites thy algebra funcgr const) end;
+
+fun merge_funcss rewrites thy algebra raw_funcss funcgr =
+ let
+ val funcss = raw_funcss
+ |> resort_funcss thy algebra funcgr
+ |> filter_out (can (Constgraph.get_node funcgr) o fst);
+ fun typ_func const [] = Code.default_typ thy const
+ | typ_func (_, NONE) (thm :: _) = (snd o CodeUnit.head_func) thm
+ | typ_func (const as (c, SOME tyco)) (thms as (thm :: _)) =
+ let
+ val (_, ty) = CodeUnit.head_func thm;
+ val SOME class = AxClass.class_of_param thy c;
+ val sorts_decl = Sorts.mg_domain algebra tyco [class];
+ val tys = CodeUnit.typargs thy (c, ty);
+ val sorts = map (snd o dest_TVar) tys;
+ in if sorts = sorts_decl then ty
+ else raise INVALID ([const], "Illegal instantation for class operation "
+ ^ CodeUnit.string_of_const thy const
+ ^ "\nin defining equations\n"
+ ^ (cat_lines o map string_of_thm) thms)
+ end;
+ fun add_funcs (const, thms) =
+ Constgraph.new_node (const, (typ_func const thms, thms));
+ fun add_deps (funcs as (const, thms)) funcgr =
+ let
+ val deps = consts_of funcs;
+ val insts = instances_of_consts thy algebra funcgr
+ (fold_consts (insert (op =)) thms []);
+ in
+ funcgr
+ |> ensure_consts' rewrites thy algebra insts
+ |> fold (curry Constgraph.add_edge const) deps
+ |> fold (curry Constgraph.add_edge const) insts
+ end;
+ in
+ funcgr
+ |> fold add_funcs funcss
+ |> fold add_deps funcss
+ end
+and ensure_consts' rewrites thy algebra cs funcgr =
+ let
+ val auxgr = Constgraph.empty
+ |> fold (snd oo ensure_const rewrites thy algebra funcgr) cs;
+ in
+ funcgr
+ |> fold (merge_funcss rewrites thy algebra)
+ (map (AList.make (Constgraph.get_node auxgr))
+ (rev (Constgraph.strong_conn auxgr)))
+ end handle INVALID (cs', msg)
+ => raise INVALID (fold (insert CodeUnit.eq_const) cs' cs, msg);
+
+fun ensure_consts rewrites thy consts funcgr =
+ let
+ val algebra = Code.coregular_algebra thy
+ in ensure_consts' rewrites thy algebra consts funcgr
+ handle INVALID (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) "
+ ^ commas (map (CodeUnit.string_of_const thy) cs'))
+ end;
+
+in
+
+(** retrieval interfaces **)
+
+val ensure_consts = ensure_consts;
+
+fun check_consts rewrites thy consts funcgr =
+ let
+ val algebra = Code.coregular_algebra thy;
+ fun try_const const funcgr =
+ (SOME const, ensure_consts' rewrites thy algebra [const] funcgr)
+ handle INVALID (cs', msg) => (NONE, funcgr);
+ val (consts', funcgr') = fold_map try_const consts funcgr;
+ in (map_filter I consts', funcgr') end;
+
+fun ensure_consts_term rewrites thy f ct funcgr =
+ let
+ fun consts_of thy t =
+ fold_aterms (fn Const c => cons (CodeUnit.const_of_cexpr thy c) | _ => I) t []
+ fun rhs_conv conv thm =
+ let
+ val thm' = (conv o Thm.rhs_of) thm;
+ in Thm.transitive thm thm' end
+ val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct);
+ val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) ();
+ val thm1 = Code.preprocess_conv ct
+ |> fold (rhs_conv o MetaSimplifier.rewrite false o single) (rewrites thy);
+ val ct' = Thm.rhs_of thm1;
+ val consts = consts_of thy (Thm.term_of ct');
+ val funcgr' = ensure_consts rewrites thy consts funcgr;
+ val algebra = Code.coregular_algebra thy;
+ val (_, thm2) = Thm.varifyT' [] thm1;
+ val thm3 = Thm.reflexive (Thm.rhs_of thm2);
+ val typ_funcgr = try (fst o Constgraph.get_node funcgr' o CodeUnit.const_of_cexpr thy);
+ val [thm4] = resort_thms algebra typ_funcgr [thm3];
+ val tfrees = Term.add_tfrees (Thm.prop_of thm1) [];
+ fun inst thm =
+ let
+ val tvars = Term.add_tvars (Thm.prop_of thm) [];
+ val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy)
+ (TVar (v_i, sort), TFree (v, sort))) tvars tfrees;
+ in Thm.instantiate (instmap, []) thm end;
+ val thm5 = inst thm2;
+ val thm6 = inst thm4;
+ val ct'' = Thm.rhs_of thm6;
+ val cs = fold_aterms (fn Const c => cons c | _ => I) (Thm.term_of ct'') [];
+ val drop = drop_classes thy tfrees;
+ val instdefs = instances_of_consts thy algebra funcgr' cs;
+ val funcgr'' = ensure_consts rewrites thy instdefs funcgr';
+ in (f funcgr'' drop ct'' thm5, funcgr'') end;
+
+fun ensure_consts_eval rewrites thy conv =
+ let
+ fun conv' funcgr drop_classes ct thm1 =
+ let
+ val thm2 = conv funcgr ct;
+ val thm3 = Code.postprocess_conv (Thm.rhs_of thm2);
+ val thm23 = drop_classes (Thm.transitive thm2 thm3);
+ in
+ Thm.transitive thm1 thm23 handle THM _ =>
+ error ("eval_conv - could not construct proof:\n"
+ ^ (cat_lines o map string_of_thm) [thm1, thm2, thm3])
+ end;
+ in ensure_consts_term rewrites thy conv' end;
+
+end; (*local*)
+
+end; (*struct*)
+
+functor CodeFuncgrRetrieval (val rewrites: theory -> thm list) : CODE_FUNCGR_RETRIEVAL =
+struct
+
+(** code data **)
+
+type T = CodeFuncgr.T;
+
+structure Funcgr = CodeDataFun
+(struct
+ type T = T;
+ val empty = CodeFuncgr.Constgraph.empty;
+ fun merge _ _ = CodeFuncgr.Constgraph.empty;
+ fun purge _ NONE _ = CodeFuncgr.Constgraph.empty
+ | purge _ (SOME cs) funcgr =
+ CodeFuncgr.Constgraph.del_nodes ((CodeFuncgr.Constgraph.all_preds funcgr
+ o filter (can (CodeFuncgr.Constgraph.get_node funcgr))) cs) funcgr;
+end);
+
+fun make thy =
+ Funcgr.change thy o CodeFuncgr.ensure_consts rewrites thy;
+
+fun make_consts thy =
+ Funcgr.change_yield thy o CodeFuncgr.check_consts rewrites thy;
+
+fun make_term thy f =
+ Funcgr.change_yield thy o CodeFuncgr.ensure_consts_term rewrites thy f;
+
+fun eval_conv thy f =
+ fst o Funcgr.change_yield thy o CodeFuncgr.ensure_consts_eval rewrites thy f;
+
+fun intervene thy funcgr = Funcgr.change thy (K funcgr);
+
+end; (*functor*)
+
+structure CodeFuncgr : CODE_FUNCGR =
+struct
+
+open CodeFuncgr;
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_name.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,428 @@
+(* Title: Tools/code/code_name.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Naming policies for code generation: prefixing any name by corresponding
+theory name, conversion to alphanumeric representation.
+Mappings are incrementally cached.
+*)
+
+signature CODE_NAME =
+sig
+ val purify_var: string -> string
+ val purify_tvar: string -> string
+ val check_modulename: string -> string
+ type var_ctxt;
+ val make_vars: string list -> var_ctxt;
+ val intro_vars: string list -> var_ctxt -> var_ctxt;
+ val lookup_var: var_ctxt -> string -> string;
+
+ type tyco = string
+ type const = string
+ val class: theory -> class -> class
+ val class_rev: theory -> class -> class option
+ val classrel: theory -> class * class -> string
+ val classrel_rev: theory -> string -> (class * class) option
+ val tyco: theory -> tyco -> tyco
+ val tyco_rev: theory -> tyco -> tyco option
+ val instance: theory -> class * tyco -> string
+ val instance_rev: theory -> string -> (class * tyco) option
+ val const: theory -> CodeUnit.const -> const
+ val const_rev: theory -> const -> CodeUnit.const option
+ val labelled_name: theory -> string -> string
+
+ val setup: theory -> theory
+end;
+
+structure CodeName: CODE_NAME =
+struct
+
+(** purification **)
+
+fun purify_name upper_else_lower =
+ let
+ fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse s = "'";
+ val is_junk = not o is_valid andf Symbol.is_regular;
+ val junk = Scan.many is_junk;
+ val scan_valids = Symbol.scanner "Malformed input"
+ ((junk |--
+ (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
+ --| junk))
+ -- Scan.repeat ((Scan.many1 is_valid >> implode) --| junk) >> op ::);
+ fun upper_lower cs = if upper_else_lower then nth_map 0 Symbol.to_ascii_upper cs
+ else (if forall Symbol.is_ascii_upper cs
+ then map else nth_map 0) Symbol.to_ascii_lower cs;
+ in
+ explode
+ #> scan_valids
+ #> space_implode "_"
+ #> explode
+ #> upper_lower
+ #> implode
+ end;
+
+fun purify_var "" = "x"
+ | purify_var v = purify_name false v;
+
+fun purify_tvar "" = "'a"
+ | purify_tvar v =
+ (unprefix "'" #> explode #> filter Symbol.is_ascii_letter #> cons "'" #> implode) v;
+
+fun check_modulename mn =
+ let
+ val mns = NameSpace.explode mn;
+ val mns' = map (purify_name true) mns;
+ in
+ if mns' = mns then mn else error ("Invalid module name: " ^ quote mn ^ "\n"
+ ^ "perhaps try " ^ quote (NameSpace.implode mns'))
+ end;
+
+
+(** variable name contexts **)
+
+type var_ctxt = string Symtab.table * Name.context;
+
+fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty,
+ Name.make_context names);
+
+fun intro_vars names (namemap, namectxt) =
+ let
+ val (names', namectxt') = Name.variants names namectxt;
+ val namemap' = fold2 (curry Symtab.update) names names' namemap;
+ in (namemap', namectxt') end;
+
+fun lookup_var (namemap, _) name = case Symtab.lookup namemap name
+ of SOME name' => name'
+ | NONE => error ("Invalid name in context: " ^ quote name);
+
+
+
+(** global names (identifiers) **)
+
+(* identifier categories *)
+
+val idf_class = "class";
+val idf_classrel = "clsrel"
+val idf_tyco = "tyco";
+val idf_instance = "inst";
+val idf_const = "const";
+
+fun string_of_classrel (class, superclass) = class ^ " < " ^ superclass;
+fun string_of_instance (class, tyco) = tyco ^ " :: " ^ class;
+
+fun add_idf nsp name =
+ NameSpace.append name nsp;
+
+fun dest_idf nsp name =
+ if NameSpace.base name = nsp
+ then SOME (NameSpace.qualifier name)
+ else NONE;
+
+local
+
+val name_mapping = [
+ (idf_class, "class"),
+ (idf_classrel, "subclass relation"),
+ (idf_tyco, "type constructor"),
+ (idf_instance, "instance"),
+ (idf_const, "constant")
+]
+
+in
+
+val category_of = the o AList.lookup (op =) name_mapping o NameSpace.base;
+
+end;
+
+
+(* theory name lookup *)
+
+fun thyname_of thy f errmsg x =
+ let
+ fun thy_of thy =
+ if f thy x then case get_first thy_of (Theory.parents_of thy)
+ of NONE => SOME thy
+ | thy => thy
+ else NONE;
+ in case thy_of thy
+ of SOME thy => Context.theory_name thy
+ | NONE => error (errmsg x) end;
+
+fun thyname_of_class thy =
+ thyname_of thy (fn thy => member (op =) (Sign.all_classes thy))
+ (fn class => "thyname_of_class: no such class: " ^ quote class);
+
+fun thyname_of_classrel thy =
+ thyname_of thy (fn thy => fn (class1, class2) => Sign.subsort thy ([class1], [class2]))
+ (fn (class1, class2) => "thyname_of_classrel: no such subclass relation: "
+ ^ (quote o string_of_classrel) (class1, class2));
+
+fun thyname_of_tyco thy =
+ thyname_of thy Sign.declared_tyname
+ (fn tyco => "thyname_of_tyco: no such type constructor: " ^ quote tyco);
+
+fun thyname_of_instance thy =
+ let
+ fun test_instance thy (class, tyco) =
+ can (Sorts.mg_domain (Sign.classes_of thy) tyco) [class]
+ in thyname_of thy test_instance
+ (fn (class, tyco) => "thyname_of_instance: no such instance: "
+ ^ (quote o string_of_instance) (class, tyco))
+ end;
+
+fun thyname_of_const thy =
+ thyname_of thy Sign.declared_const
+ (fn c => "thyname_of_const: no such constant: " ^ quote c);
+
+
+(* naming policies *)
+
+val purify_prefix =
+ explode
+ (*should disappear as soon as hierarchical theory name spaces are available*)
+ #> Symbol.scanner "Malformed name"
+ (Scan.repeat ($$ "_" |-- $$ "_" >> (fn _ => ".") || Scan.one Symbol.is_regular))
+ #> implode
+ #> NameSpace.explode
+ #> map (purify_name true);
+
+fun purify_base "op =" = "eq"
+ | purify_base "op &" = "and"
+ | purify_base "op |" = "or"
+ | purify_base "op -->" = "implies"
+ | purify_base "{}" = "empty"
+ | purify_base "op :" = "member"
+ | purify_base "op Int" = "intersect"
+ | purify_base "op Un" = "union"
+ | purify_base "*" = "product"
+ | purify_base "+" = "sum"
+ | purify_base s = purify_name false s;
+
+fun default_policy thy get_basename get_thyname name =
+ let
+ val prefix = (purify_prefix o get_thyname thy) name;
+ val base = (purify_base o get_basename) name;
+ in NameSpace.implode (prefix @ [base]) end;
+
+fun class_policy thy = default_policy thy NameSpace.base thyname_of_class;
+fun classrel_policy thy = default_policy thy (fn (class1, class2) =>
+ NameSpace.base class2 ^ "_" ^ NameSpace.base class1) thyname_of_classrel;
+ (*order fits nicely with composed projections*)
+fun tyco_policy thy = default_policy thy NameSpace.base thyname_of_tyco;
+fun instance_policy thy = default_policy thy (fn (class, tyco) =>
+ NameSpace.base class ^ "_" ^ NameSpace.base tyco) thyname_of_instance;
+
+fun force_thyname thy (const as (c, opt_tyco)) =
+ case Code.get_datatype_of_constr thy const
+ of SOME dtco => SOME (thyname_of_tyco thy dtco)
+ | NONE => (case AxClass.class_of_param thy c
+ of SOME class => (case opt_tyco
+ of SOME tyco => SOME (thyname_of_instance thy (class, tyco))
+ | NONE => SOME (thyname_of_class thy class))
+ | NONE => NONE);
+
+fun const_policy thy (const as (c, opt_tyco)) =
+ case force_thyname thy const
+ of NONE => default_policy thy NameSpace.base thyname_of_const c
+ | SOME thyname => let
+ val prefix = purify_prefix thyname;
+ val tycos = the_list opt_tyco;
+ val base = map (purify_base o NameSpace.base) (c :: tycos);
+ in NameSpace.implode (prefix @ [space_implode "_" base]) end;
+
+
+(* theory and code data *)
+
+type tyco = string;
+type const = string;
+structure Consttab = CodeUnit.Consttab;
+
+structure StringPairTab =
+ TableFun(
+ type key = string * string;
+ val ord = prod_ord fast_string_ord fast_string_ord;
+ );
+
+datatype names = Names of {
+ class: class Symtab.table * class Symtab.table,
+ classrel: string StringPairTab.table * (class * class) Symtab.table,
+ tyco: tyco Symtab.table * tyco Symtab.table,
+ instance: string StringPairTab.table * (class * tyco) Symtab.table
+}
+
+val empty_names = Names {
+ class = (Symtab.empty, Symtab.empty),
+ classrel = (StringPairTab.empty, Symtab.empty),
+ tyco = (Symtab.empty, Symtab.empty),
+ instance = (StringPairTab.empty, Symtab.empty)
+};
+
+local
+ fun mk_names (class, classrel, tyco, instance) =
+ Names { class = class, classrel = classrel, tyco = tyco, instance = instance };
+ fun map_names f (Names { class, classrel, tyco, instance }) =
+ mk_names (f (class, classrel, tyco, instance));
+in
+ fun merge_names (Names { class = (class1, classrev1),
+ classrel = (classrel1, classrelrev1), tyco = (tyco1, tycorev1),
+ instance = (instance1, instancerev1) },
+ Names { class = (class2, classrev2),
+ classrel = (classrel2, classrelrev2), tyco = (tyco2, tycorev2),
+ instance = (instance2, instancerev2) }) =
+ mk_names ((Symtab.merge (op =) (class1, class2), Symtab.merge (op =) (classrev1, classrev2)),
+ (StringPairTab.merge (op =) (classrel1, classrel2), Symtab.merge (op =) (classrelrev1, classrelrev2)),
+ (Symtab.merge (op =) (tyco1, tyco2), Symtab.merge (op =) (tycorev1, tycorev2)),
+ (StringPairTab.merge (op =) (instance1, instance2), Symtab.merge (op =) (instancerev1, instancerev2)));
+ fun map_class f = map_names
+ (fn (class, classrel, tyco, inst) => (f class, classrel, tyco, inst));
+ fun map_classrel f = map_names
+ (fn (class, classrel, tyco, inst) => (class, f classrel, tyco, inst));
+ fun map_tyco f = map_names
+ (fn (class, classrel, tyco, inst) => (class, classrel, f tyco, inst));
+ fun map_instance f = map_names
+ (fn (class, classrel, tyco, inst) => (class, classrel, tyco, f inst));
+end; (*local*)
+
+structure CodeName = TheoryDataFun
+(
+ type T = names ref;
+ val empty = ref empty_names;
+ fun copy (ref names) = ref names;
+ val extend = copy;
+ fun merge _ (ref names1, ref names2) = ref (merge_names (names1, names2));
+);
+
+structure ConstName = CodeDataFun
+(
+ type T = const Consttab.table * (string * string option) Symtab.table;
+ val empty = (Consttab.empty, Symtab.empty);
+ fun merge _ ((const1, constrev1), (const2, constrev2)) : T =
+ (Consttab.merge (op =) (const1, const2),
+ Symtab.merge CodeUnit.eq_const (constrev1, constrev2));
+ fun purge _ NONE _ = empty
+ | purge _ (SOME cs) (const, constrev) = (fold Consttab.delete_safe cs const,
+ fold Symtab.delete (map_filter (Consttab.lookup const) cs) constrev);
+);
+
+val setup = CodeName.init;
+
+
+(* forward lookup with cache update *)
+
+fun get thy get_tabs get upd_names upd policy x =
+ let
+ val names_ref = CodeName.get thy;
+ val (Names names) = ! names_ref;
+ val tabs = get_tabs names;
+ fun declare name =
+ let
+ val names' = upd_names (K (upd (x, name) (fst tabs),
+ Symtab.update_new (name, x) (snd tabs))) (Names names)
+ in (names_ref := names'; name) end;
+ in case get (fst tabs) x
+ of SOME name => name
+ | NONE =>
+ x
+ |> policy thy
+ |> Name.variant (Symtab.keys (snd tabs))
+ |> declare
+ end;
+
+fun get_const thy const =
+ let
+ val tabs = ConstName.get thy;
+ fun declare name =
+ let
+ val names' = (Consttab.update (const, name) (fst tabs),
+ Symtab.update_new (name, const) (snd tabs))
+ in (ConstName.change thy (K names'); name) end;
+ in case Consttab.lookup (fst tabs) const
+ of SOME name => name
+ | NONE =>
+ const
+ |> const_policy thy
+ |> Name.variant (Symtab.keys (snd tabs))
+ |> declare
+ end;
+
+
+(* backward lookup *)
+
+fun rev thy get_tabs name =
+ let
+ val names_ref = CodeName.get thy
+ val (Names names) = ! names_ref;
+ val tab = (snd o get_tabs) names;
+ in case Symtab.lookup tab name
+ of SOME x => x
+ | NONE => error ("No such " ^ category_of name ^ ": " ^ quote name)
+ end;
+
+fun rev_const thy name =
+ let
+ val tab = snd (ConstName.get thy);
+ in case Symtab.lookup tab name
+ of SOME const => const
+ | NONE => error ("No such " ^ category_of name ^ ": " ^ quote name)
+ end;
+
+
+(* external interfaces *)
+
+fun class thy =
+ get thy #class Symtab.lookup map_class Symtab.update class_policy
+ #> add_idf idf_class;
+fun classrel thy =
+ get thy #classrel StringPairTab.lookup map_classrel StringPairTab.update classrel_policy
+ #> add_idf idf_classrel;
+fun tyco thy =
+ get thy #tyco Symtab.lookup map_tyco Symtab.update tyco_policy
+ #> add_idf idf_tyco;
+fun instance thy =
+ get thy #instance StringPairTab.lookup map_instance StringPairTab.update instance_policy
+ #> add_idf idf_instance;
+fun const thy =
+ get_const thy
+ #> add_idf idf_const;
+
+fun class_rev thy =
+ dest_idf idf_class
+ #> Option.map (rev thy #class);
+fun classrel_rev thy =
+ dest_idf idf_classrel
+ #> Option.map (rev thy #classrel);
+fun tyco_rev thy =
+ dest_idf idf_tyco
+ #> Option.map (rev thy #tyco);
+fun instance_rev thy =
+ dest_idf idf_instance
+ #> Option.map (rev thy #instance);
+fun const_rev thy =
+ dest_idf idf_const
+ #> Option.map (rev_const thy);
+
+local
+
+val f_mapping = [
+ (idf_class, class_rev),
+ (idf_classrel, Option.map string_of_classrel oo classrel_rev),
+ (idf_tyco, tyco_rev),
+ (idf_instance, Option.map string_of_instance oo instance_rev),
+ (idf_const, fn thy => Option.map (CodeUnit.string_of_const thy) o const_rev thy)
+];
+
+in
+
+fun labelled_name thy idf =
+ let
+ val category = category_of idf;
+ val name = NameSpace.qualifier idf;
+ val f = (the o AList.lookup (op =) f_mapping o NameSpace.base) idf
+ in case f thy idf
+ of SOME thing => category ^ " " ^ quote thing
+ | NONE => error ("Unknown name: " ^ quote name)
+ end;
+
+end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_package.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,667 @@
+(* Title: Tools/code/code_package.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Code generator translation kernel. Code generator Isar setup.
+*)
+
+signature CODE_PACKAGE =
+sig
+ (* interfaces *)
+ val eval_conv: theory
+ -> (CodeThingol.code -> CodeThingol.iterm -> cterm -> thm) -> cterm -> thm;
+ val codegen_command: theory -> string -> unit;
+
+ (* primitive interfaces *)
+ val eval_term: theory -> (string (*reference name!*) * 'a option ref) * term -> 'a;
+ val satisfies_ref: bool option ref;
+ val satisfies: theory -> term -> string list -> bool;
+
+ (* axiomatic interfaces *)
+ type appgen;
+ val add_appconst: string * appgen -> theory -> theory;
+ val appgen_let: appgen;
+ val appgen_if: appgen;
+ val appgen_case: (theory -> term
+ -> ((string * typ) list * ((term * typ) * (term * term) list)) option)
+ -> appgen;
+
+ val timing: bool ref;
+end;
+
+structure CodePackage : CODE_PACKAGE =
+struct
+
+open BasicCodeThingol;
+val succeed = CodeThingol.succeed;
+val fail = CodeThingol.fail;
+
+(** code translation **)
+
+(* theory data *)
+
+type appgen = theory -> ((sort -> sort) * Sorts.algebra) * Consts.T
+ -> CodeFuncgr.T
+ -> (string * typ) * term list -> CodeThingol.transact -> iterm * CodeThingol.transact;
+
+type appgens = (int * (appgen * stamp)) Symtab.table;
+val merge_appgens : appgens * appgens -> appgens =
+ Symtab.merge (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) =>
+ bounds1 = bounds2 andalso stamp1 = stamp2);
+
+structure Consttab = CodeUnit.Consttab;
+type abstypes = typ Symtab.table * CodeUnit.const Consttab.table;
+fun merge_abstypes ((typs1, consts1) : abstypes, (typs2, consts2) : abstypes) =
+ (Symtab.merge (Type.eq_type Vartab.empty) (typs1, typs2),
+ Consttab.merge CodeUnit.eq_const (consts1, consts2));
+
+structure Translation = TheoryDataFun
+(
+ type T = appgens * abstypes;
+ val empty = (Symtab.empty, (Symtab.empty, Consttab.empty));
+ val copy = I;
+ val extend = I;
+ fun merge _ ((appgens1, abstypes1), (appgens2, abstypes2)) =
+ (merge_appgens (appgens1, appgens2), merge_abstypes (abstypes1, abstypes2));
+);
+
+structure Funcgr = CodeFuncgrRetrieval (fun rewrites thy = []);
+
+fun code_depgr thy [] = Funcgr.make thy []
+ | code_depgr thy consts =
+ let
+ val gr = Funcgr.make thy consts;
+ val select = CodeFuncgr.Constgraph.all_succs gr consts;
+ in
+ CodeFuncgr.Constgraph.subgraph
+ (member CodeUnit.eq_const select) gr
+ end;
+
+fun code_thms thy =
+ Pretty.writeln o CodeFuncgr.pretty thy o code_depgr thy;
+
+fun code_deps thy consts =
+ let
+ val gr = code_depgr thy consts;
+ fun mk_entry (const, (_, (_, parents))) =
+ let
+ val name = CodeUnit.string_of_const thy const;
+ val nameparents = map (CodeUnit.string_of_const thy) parents;
+ in { name = name, ID = name, dir = "", unfold = true,
+ path = "", parents = nameparents }
+ end;
+ val prgr = CodeFuncgr.Constgraph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr [];
+ in Present.display_graph prgr end;
+
+structure Program = CodeDataFun
+(
+ type T = CodeThingol.code;
+ val empty = CodeThingol.empty_code;
+ fun merge _ = CodeThingol.merge_code;
+ fun purge _ NONE _ = CodeThingol.empty_code
+ | purge NONE _ _ = CodeThingol.empty_code
+ | purge (SOME thy) (SOME cs) code =
+ let
+ val cs_exisiting =
+ map_filter (CodeName.const_rev thy) (Graph.keys code);
+ val dels = (Graph.all_preds code
+ o map (CodeName.const thy)
+ o filter (member CodeUnit.eq_const cs_exisiting)
+ ) cs;
+ in Graph.del_nodes dels code end;
+);
+
+
+(* translation kernel *)
+
+fun get_abstype thy (tyco, tys) = case Symtab.lookup ((fst o snd o Translation.get) thy) tyco
+ of SOME ty => SOME ((map_atyps (fn TFree (n, _) => nth tys (the (Int.fromString n)))) ty)
+ | NONE => NONE;
+
+fun ensure_def thy = CodeThingol.ensure_def (CodeName.labelled_name thy);
+
+fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class =
+ let
+ val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
+ val (v, cs) = AxClass.params_of_class thy class;
+ val class' = CodeName.class thy class;
+ val classrels' = map (curry (CodeName.classrel thy) class) superclasses;
+ val classops' = map (CodeName.const thy o CodeUnit.const_of_cexpr thy) cs;
+ val defgen_class =
+ fold_map (ensure_def_class thy algbr funcgr) superclasses
+ ##>> (fold_map (exprgen_typ thy algbr funcgr) o map snd) cs
+ #-> (fn (superclasses, classoptyps) => succeed
+ (CodeThingol.Class (superclasses ~~ classrels', (unprefix "'" v, classops' ~~ classoptyps))))
+ in
+ ensure_def thy defgen_class ("generating class " ^ quote class) class'
+ #> pair class'
+ end
+and ensure_def_classrel thy algbr funcgr (subclass, superclass) =
+ ensure_def_class thy algbr funcgr subclass
+ #>> (fn _ => CodeName.classrel thy (subclass, superclass))
+and ensure_def_tyco thy algbr funcgr "fun" trns =
+ trns
+ |> pair "fun"
+ | ensure_def_tyco thy algbr funcgr tyco trns =
+ let
+ fun defgen_datatype trns =
+ let
+ val (vs, cos) = Code.get_datatype thy tyco;
+ in
+ trns
+ |> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
+ ||>> fold_map (fn (c, tys) =>
+ fold_map (exprgen_typ thy algbr funcgr) tys
+ #-> (fn tys' =>
+ pair ((CodeName.const thy o CodeUnit.const_of_cexpr thy)
+ (c, tys ---> Type (tyco, map TFree vs)), tys'))) cos
+ |-> (fn (vs, cos) => succeed (CodeThingol.Datatype (vs, cos)))
+ end;
+ val tyco' = CodeName.tyco thy tyco;
+ in
+ trns
+ |> ensure_def thy defgen_datatype ("generating type constructor " ^ quote tyco) tyco'
+ |> pair tyco'
+ end
+and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) trns =
+ trns
+ |> fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort)
+ |>> (fn sort => (unprefix "'" v, sort))
+and exprgen_typ thy algbr funcgr (TFree vs) trns =
+ trns
+ |> exprgen_tyvar_sort thy algbr funcgr vs
+ |>> (fn (v, sort) => ITyVar v)
+ | exprgen_typ thy algbr funcgr (Type (tyco, tys)) trns =
+ case get_abstype thy (tyco, tys)
+ of SOME ty =>
+ trns
+ |> exprgen_typ thy algbr funcgr ty
+ | NONE =>
+ trns
+ |> ensure_def_tyco thy algbr funcgr tyco
+ ||>> fold_map (exprgen_typ thy algbr funcgr) tys
+ |>> (fn (tyco, tys) => tyco `%% tys);
+
+exception CONSTRAIN of (string * typ) * typ;
+val timing = ref false;
+
+fun exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) =
+ let
+ val pp = Sign.pp thy;
+ datatype typarg =
+ Global of (class * string) * typarg list list
+ | Local of (class * class) list * (string * (int * sort));
+ fun class_relation (Global ((_, tyco), yss), _) class =
+ Global ((class, tyco), yss)
+ | class_relation (Local (classrels, v), subclass) superclass =
+ Local ((subclass, superclass) :: classrels, v);
+ fun type_constructor tyco yss class =
+ Global ((class, tyco), (map o map) fst yss);
+ fun type_variable (TFree (v, sort)) =
+ let
+ val sort' = proj_sort sort;
+ in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
+ val typargs = Sorts.of_sort_derivation pp algebra
+ {class_relation = class_relation, type_constructor = type_constructor,
+ type_variable = type_variable}
+ (ty_ctxt, proj_sort sort_decl);
+ fun mk_dict (Global (inst, yss)) =
+ ensure_def_inst thy algbr funcgr inst
+ ##>> (fold_map o fold_map) mk_dict yss
+ #>> (fn (inst, dss) => DictConst (inst, dss))
+ | mk_dict (Local (classrels, (v, (k, sort)))) =
+ fold_map (ensure_def_classrel thy algbr funcgr) classrels
+ #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort))))
+ in
+ fold_map mk_dict typargs
+ end
+and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) =
+ let
+ val c' = CodeUnit.const_of_cexpr thy (c, ty_ctxt)
+ val idf = CodeName.const thy c';
+ val ty_decl = Consts.the_declaration consts idf;
+ val (tys, tys_decl) = pairself (curry (Consts.typargs consts) idf) (ty_ctxt, ty_decl);
+ val sorts = map (snd o dest_TVar) tys_decl;
+ in
+ fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts)
+ end
+and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) trns =
+ let
+ val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class;
+ val (var, classops) = try (AxClass.params_of_class thy) class |> the_default ("'a", [])
+ val vs = Name.names (Name.declare var Name.context) "'a" (Sorts.mg_domain algebra tyco [class]);
+ val arity_typ = Type (tyco, map TFree vs);
+ fun gen_superarity superclass trns =
+ trns
+ |> ensure_def_class thy algbr funcgr superclass
+ ||>> ensure_def_classrel thy algbr funcgr (class, superclass)
+ ||>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass])
+ |>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
+ (superclass, (classrel, (inst, dss))));
+ fun gen_classop_def (classop as (c, ty)) trns =
+ trns
+ |> ensure_def_const thy algbr funcgr (CodeUnit.const_of_cexpr thy classop)
+ ||>> exprgen_term thy algbr funcgr (Const (c, map_type_tfree (K arity_typ) ty));
+ fun defgen_inst trns =
+ trns
+ |> ensure_def_class thy algbr funcgr class
+ ||>> ensure_def_tyco thy algbr funcgr tyco
+ ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
+ ||>> fold_map gen_superarity superclasses
+ ||>> fold_map gen_classop_def classops
+ |-> (fn ((((class, tyco), arity), superarities), classops) =>
+ succeed (CodeThingol.Classinst ((class, (tyco, arity)), (superarities, classops))));
+ val inst = CodeName.instance thy (class, tyco);
+ in
+ trns
+ |> ensure_def thy defgen_inst ("generating instance " ^ quote class ^ " / " ^ quote tyco) inst
+ |> pair inst
+ end
+and ensure_def_const thy (algbr as (_, consts)) funcgr (const as (c, opt_tyco)) trns =
+ let
+ val c' = CodeName.const thy const;
+ fun defgen_datatypecons trns =
+ trns
+ |> ensure_def_tyco thy algbr funcgr
+ ((the o Code.get_datatype_of_constr thy) const)
+ |-> (fn _ => succeed CodeThingol.Bot);
+ fun defgen_classop trns =
+ trns
+ |> ensure_def_class thy algbr funcgr ((the o AxClass.class_of_param thy) c)
+ |-> (fn _ => succeed CodeThingol.Bot);
+ fun defgen_fun trns =
+ let
+ val const' = perhaps (Consttab.lookup ((snd o snd o Translation.get) thy)) const;
+ val raw_thms = CodeFuncgr.funcs funcgr const';
+ val ty = (Logic.unvarifyT o CodeFuncgr.typ funcgr) const';
+ val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty
+ then raw_thms
+ else map (CodeUnit.expand_eta 1) raw_thms;
+ val timeap = if !timing then Output.timeap_msg ("time for " ^ c')
+ else I;
+ val msg = cat_lines ("generating code for theorems " :: map string_of_thm thms);
+ val vs = (map dest_TFree o Consts.typargs consts) (c', ty);
+ val dest_eqthm =
+ apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify o prop_of;
+ fun exprgen_eq (args, rhs) trns =
+ trns
+ |> fold_map (exprgen_term thy algbr funcgr) args
+ ||>> exprgen_term thy algbr funcgr rhs;
+ in
+ trns
+ |> CodeThingol.message msg (fn trns => trns
+ |> timeap (fold_map (exprgen_eq o dest_eqthm) thms)
+ ||>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs
+ ||>> exprgen_typ thy algbr funcgr ty
+ |-> (fn ((eqs, vs), ty) => succeed (CodeThingol.Fun (eqs, (vs, ty)))))
+ end;
+ val defgen = if (is_some o Code.get_datatype_of_constr thy) const
+ then defgen_datatypecons
+ else if is_some opt_tyco
+ orelse (not o is_some o AxClass.class_of_param thy) c
+ then defgen_fun
+ else defgen_classop
+ in
+ trns
+ |> ensure_def thy defgen ("generating constant " ^ CodeUnit.string_of_const thy const) c'
+ |> pair c'
+ end
+and exprgen_term thy algbr funcgr (Const (c, ty)) trns =
+ trns
+ |> select_appgen thy algbr funcgr ((c, ty), [])
+ | exprgen_term thy algbr funcgr (Free (v, ty)) trns =
+ trns
+ |> exprgen_typ thy algbr funcgr ty
+ |>> (fn _ => IVar v)
+ | exprgen_term thy algbr funcgr (Abs (raw_v, ty, raw_t)) trns =
+ let
+ val (v, t) = Syntax.variant_abs (raw_v, ty, raw_t);
+ in
+ trns
+ |> exprgen_typ thy algbr funcgr ty
+ ||>> exprgen_term thy algbr funcgr t
+ |>> (fn (ty, t) => (v, ty) `|-> t)
+ end
+ | exprgen_term thy algbr funcgr (t as _ $ _) trns =
+ case strip_comb t
+ of (Const (c, ty), ts) =>
+ trns
+ |> select_appgen thy algbr funcgr ((c, ty), ts)
+ | (t', ts) =>
+ trns
+ |> exprgen_term thy algbr funcgr t'
+ ||>> fold_map (exprgen_term thy algbr funcgr) ts
+ |>> (fn (t, ts) => t `$$ ts)
+and appgen_default thy algbr funcgr ((c, ty), ts) trns =
+ trns
+ |> ensure_def_const thy algbr funcgr (CodeUnit.const_of_cexpr thy (c, ty))
+ ||>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty)
+ ||>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)
+ ||>> exprgen_dict_parms thy algbr funcgr (c, ty)
+ ||>> fold_map (exprgen_term thy algbr funcgr) ts
+ |>> (fn ((((c, tys), ty), iss), ts) => IConst (c, (iss, tys)) `$$ ts)
+and select_appgen thy algbr funcgr ((c, ty), ts) trns =
+ case Symtab.lookup (fst (Translation.get thy)) c
+ of SOME (i, (appgen, _)) =>
+ if length ts < i then
+ let
+ val k = length ts;
+ val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty;
+ val ctxt = (fold o fold_aterms)
+ (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
+ val vs = Name.names ctxt "a" tys;
+ in
+ trns
+ |> fold_map (exprgen_typ thy algbr funcgr) tys
+ ||>> appgen thy algbr funcgr ((c, ty), ts @ map Free vs)
+ |>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t)
+ end
+ else if length ts > i then
+ trns
+ |> appgen thy algbr funcgr ((c, ty), Library.take (i, ts))
+ ||>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts))
+ |>> (fn (t, ts) => t `$$ ts)
+ else
+ trns
+ |> appgen thy algbr funcgr ((c, ty), ts)
+ | NONE =>
+ trns
+ |> appgen_default thy algbr funcgr ((c, ty), ts);
+
+
+(* entrance points into translation kernel *)
+
+fun ensure_def_const' thy algbr funcgr c trns =
+ ensure_def_const thy algbr funcgr c trns
+ handle CONSTRAIN ((c, ty), ty_decl) => error (
+ "Constant " ^ c ^ " with most general type\n"
+ ^ CodeUnit.string_of_typ thy ty
+ ^ "\noccurs with type\n"
+ ^ CodeUnit.string_of_typ thy ty_decl);
+
+fun perhaps_def_const thy algbr funcgr c trns =
+ case try (ensure_def_const thy algbr funcgr c) trns
+ of SOME (c, trns) => (SOME c, trns)
+ | NONE => (NONE, trns);
+
+fun exprgen_term' thy algbr funcgr t trns =
+ exprgen_term thy algbr funcgr t trns
+ handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t
+ ^ ",\nconstant " ^ c ^ " with most general type\n"
+ ^ CodeUnit.string_of_typ thy ty
+ ^ "\noccurs with type\n"
+ ^ CodeUnit.string_of_typ thy ty_decl);
+
+
+(* parametrized application generators, for instantiation in object logic *)
+(* (axiomatic extensions of translation kernel) *)
+
+fun appgen_case dest_case_expr thy algbr funcgr (app as (c_ty, ts)) =
+ let
+ val SOME ([], ((st, sty), ds)) = dest_case_expr thy (list_comb (Const c_ty, ts));
+ fun clause_gen (dt, bt) =
+ exprgen_term thy algbr funcgr dt
+ ##>> exprgen_term thy algbr funcgr bt;
+ in
+ exprgen_term thy algbr funcgr st
+ ##>> exprgen_typ thy algbr funcgr sty
+ ##>> fold_map clause_gen ds
+ ##>> appgen_default thy algbr funcgr app
+ #>> (fn (((se, sty), ds), t0) => ICase (((se, sty), ds), t0))
+ end;
+
+fun appgen_let thy algbr funcgr (app as (_, [st, ct])) =
+ exprgen_term thy algbr funcgr ct
+ ##>> exprgen_term thy algbr funcgr st
+ ##>> appgen_default thy algbr funcgr app
+ #>> (fn (((v, ty) `|-> be, se), t0) =>
+ ICase (CodeThingol.collapse_let (((v, ty), se), be), t0)
+ | (_, t0) => t0);
+
+fun appgen_if thy algbr funcgr (app as (_, [tb, tt, tf])) =
+ exprgen_term thy algbr funcgr tb
+ ##>> exprgen_typ thy algbr funcgr (Type ("bool", []))
+ ##>> exprgen_term thy algbr funcgr (Const ("True", Type ("bool", [])))
+ ##>> exprgen_term thy algbr funcgr tt
+ ##>> exprgen_term thy algbr funcgr (Const ("False", Type ("bool", [])))
+ ##>> exprgen_term thy algbr funcgr tf
+ ##>> appgen_default thy algbr funcgr app
+ #>> (fn ((((((tb, B), T), tt), F), tf), t0) => ICase (((tb, B), [(T, tt), (F, tf)]), t0));
+
+fun add_appconst (c, appgen) thy =
+ let
+ val i = (length o fst o strip_type o Sign.the_const_type thy) c;
+ val _ = Program.change thy (K CodeThingol.empty_code);
+ in
+ (Translation.map o apfst)
+ (Symtab.update (c, (i, (appgen, stamp ())))) thy
+ end;
+
+
+
+(** abstype and constsubst interface **)
+
+local
+
+fun add_consts thy f (c1, c2 as (c, opt_tyco)) =
+ let
+ val _ = if
+ is_some (AxClass.class_of_param thy c) andalso is_none opt_tyco
+ orelse is_some (Code.get_datatype_of_constr thy c2)
+ then error ("Not a function: " ^ CodeUnit.string_of_const thy c2)
+ else ();
+ val funcgr = Funcgr.make thy [c1, c2];
+ val ty1 = (f o CodeFuncgr.typ funcgr) c1;
+ val ty2 = CodeFuncgr.typ funcgr c2;
+ val _ = if Sign.typ_equiv thy (ty1, ty2) then () else
+ error ("Incompatiable type signatures of " ^ CodeUnit.string_of_const thy c1
+ ^ " and " ^ CodeUnit.string_of_const thy c2 ^ ":\n"
+ ^ CodeUnit.string_of_typ thy ty1 ^ "\n" ^ CodeUnit.string_of_typ thy ty2);
+ in Consttab.update (c1, c2) end;
+
+fun gen_abstyp prep_const prep_typ (raw_abstyp, raw_substtyp) raw_absconsts thy =
+ let
+ val abstyp = Type.no_tvars (prep_typ thy raw_abstyp);
+ val substtyp = Type.no_tvars (prep_typ thy raw_substtyp);
+ val absconsts = (map o pairself) (prep_const thy) raw_absconsts;
+ val Type (abstyco, tys) = abstyp handle BIND => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
+ val typarms = map (fst o dest_TFree) tys handle MATCH => error ("Bad type: " ^ Sign.string_of_typ thy abstyp);
+ fun mk_index v =
+ let
+ val k = find_index (fn w => v = w) typarms;
+ in if k = ~1
+ then error ("Free type variable: " ^ quote v)
+ else TFree (string_of_int k, [])
+ end;
+ val typpat = map_atyps (fn TFree (v, _) => mk_index v) substtyp;
+ fun apply_typpat (Type (tyco, tys)) =
+ let
+ val tys' = map apply_typpat tys;
+ in if tyco = abstyco then
+ (map_atyps (fn TFree (n, _) => nth tys' (the (Int.fromString n)))) typpat
+ else
+ Type (tyco, tys')
+ end
+ | apply_typpat ty = ty;
+ val _ = Program.change thy (K CodeThingol.empty_code);
+ in
+ thy
+ |> (Translation.map o apsnd) (fn (abstypes, abscs) =>
+ (abstypes
+ |> Symtab.update (abstyco, typpat),
+ abscs
+ |> fold (add_consts thy apply_typpat) absconsts)
+ )
+ end;
+
+fun gen_constsubst prep_const raw_constsubsts thy =
+ let
+ val constsubsts = (map o pairself) (prep_const thy) raw_constsubsts;
+ val _ = Program.change thy (K CodeThingol.empty_code);
+ in
+ thy
+ |> (Translation.map o apsnd o apsnd) (fold (add_consts thy I) constsubsts)
+ end;
+
+in
+
+val abstyp = gen_abstyp (K I) Sign.certify_typ;
+val abstyp_e = gen_abstyp CodeUnit.read_const Sign.read_typ;
+
+val constsubst = gen_constsubst (K I);
+val constsubst_e = gen_constsubst CodeUnit.read_const;
+
+end; (*local*)
+
+
+(** code generation interfaces **)
+
+(* generic generation combinators *)
+
+fun generate thy funcgr gen it =
+ let
+ (*FIXME*)
+ val _ = Funcgr.intervene thy funcgr;
+ val cs = map_filter (Consttab.lookup ((snd o snd o Translation.get) thy))
+ (CodeFuncgr.all funcgr);
+ val funcgr' = Funcgr.make thy cs;
+ val naming = NameSpace.qualified_names NameSpace.default_naming;
+ val consttab = Consts.empty
+ |> fold (fn c => Consts.declare naming
+ ((CodeName.const thy c, CodeFuncgr.typ funcgr' c), true))
+ (CodeFuncgr.all funcgr');
+ val algbr = (Code.operational_algebra thy, consttab);
+ in
+ Program.change_yield thy
+ (CodeThingol.start_transact (gen thy algbr funcgr' it))
+ |> fst
+ end;
+
+fun eval_conv thy conv =
+ let
+ fun conv' funcgr ct =
+ let
+ val t = generate thy funcgr exprgen_term' (Thm.term_of ct);
+ val consts = CodeThingol.fold_constnames (insert (op =)) t [];
+ val code = Program.get thy
+ |> CodeThingol.project_code true [] (SOME consts)
+ in conv code t ct end;
+ in Funcgr.eval_conv thy conv' end;
+
+fun codegen_term thy t =
+ let
+ val ct = Thm.cterm_of thy t;
+ val (ct', funcgr) = Funcgr.make_term thy (K (K K)) ct;
+ val t' = Thm.term_of ct';
+ in generate thy funcgr exprgen_term' t' end;
+
+fun raw_eval_term thy (ref_spec, t) args =
+ let
+ val _ = (Term.map_types o Term.map_atyps) (fn _ =>
+ error ("Term " ^ Sign.string_of_term thy t ^ " contains polymorphic type"))
+ t;
+ val t' = codegen_term thy t;
+ in
+ CodeTarget.eval_term thy CodeName.labelled_name
+ (Program.get thy) (ref_spec, t') args
+ end;
+
+val satisfies_ref : bool option ref = ref NONE;
+
+fun eval_term thy t = raw_eval_term thy t [];
+fun satisfies thy t witnesses = raw_eval_term thy
+ (("CodePackage.satisfies_ref", satisfies_ref), t) witnesses;
+
+fun filter_generatable thy consts =
+ let
+ val (consts', funcgr) = Funcgr.make_consts thy consts;
+ val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts';
+ val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE)
+ (consts' ~~ consts'');
+ in consts''' end;
+
+
+(** toplevel interface and setup **)
+
+local
+
+structure P = OuterParse
+and K = OuterKeyword
+
+fun code raw_cs seris thy =
+ let
+ val (perm1, cs) = CodeUnit.read_const_exprs thy
+ (filter_generatable thy) raw_cs;
+ val (perm2, cs') = case generate thy (Funcgr.make thy cs) (fold_map ooo ensure_def_const') cs
+ of [] => (true, NONE)
+ | cs => (false, SOME cs);
+ val code = Program.get thy;
+ val seris' = map (fn (((target, module), file), args) =>
+ CodeTarget.get_serializer thy target (perm1 orelse perm2) module file args
+ CodeName.labelled_name cs') seris;
+ in
+ (map (fn f => f code) seris' : unit list; ())
+ end;
+
+fun code_thms_cmd thy =
+ code_thms thy o snd o CodeUnit.read_const_exprs thy (fst o Funcgr.make_consts thy);
+
+fun code_deps_cmd thy =
+ code_deps thy o snd o CodeUnit.read_const_exprs thy (fst o Funcgr.make_consts thy);
+
+val (inK, toK, fileK) = ("in", "to", "file");
+
+val code_exprP =
+ (Scan.repeat P.term
+ -- Scan.repeat (P.$$$ inK |-- P.name
+ -- Scan.option (P.$$$ toK |-- P.name)
+ -- Scan.option (P.$$$ fileK |-- P.name)
+ -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") []
+ ) >> (fn (raw_cs, seris) => code raw_cs seris));
+
+val _ = OuterSyntax.add_keywords [inK, toK, fileK];
+
+val (codeK, code_abstypeK, code_axiomsK, code_thmsK, code_depsK) =
+ ("code_gen", "code_abstype", "code_axioms", "code_thms", "code_deps");
+
+in
+
+val codeP =
+ OuterSyntax.improper_command codeK "generate executable code for constants"
+ K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
+
+fun codegen_command thy cmd =
+ case Scan.read OuterLex.stopper (P.!!! code_exprP) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd)
+ of SOME f => (writeln "Now generating code..."; f thy)
+ | NONE => error ("Bad directive " ^ quote cmd);
+
+val code_abstypeP =
+ OuterSyntax.command code_abstypeK "axiomatic abstypes for code generation" K.thy_decl (
+ (P.typ -- P.typ -- Scan.optional (P.$$$ "where" |-- Scan.repeat1
+ (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)) [])
+ >> (Toplevel.theory o uncurry abstyp_e)
+ );
+
+val code_axiomsP =
+ OuterSyntax.command code_axiomsK "axiomatic constant equalities for code generation" K.thy_decl (
+ Scan.repeat1 (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term)
+ >> (Toplevel.theory o constsubst_e)
+ );
+
+val code_thmsP =
+ OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag
+ (Scan.repeat P.term
+ >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
+ o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
+
+val code_depsP =
+ OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag
+ (Scan.repeat P.term
+ >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
+ o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
+
+val _ = OuterSyntax.add_parsers [codeP, code_abstypeP, code_axiomsP, code_thmsP, code_depsP];
+
+end; (* local *)
+
+end; (* struct *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_target.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,2211 @@
+(* Title: Tools/code/code_target.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Serializer from intermediate language ("Thin-gol")
+to target languages (like SML or Haskell).
+*)
+
+signature CODE_TARGET =
+sig
+ include BASIC_CODE_THINGOL;
+
+ val add_syntax_class: string -> class
+ -> (string * (CodeUnit.const * string) list) option -> theory -> theory;
+ val add_syntax_inst: string -> string * class -> bool -> theory -> theory;
+ val add_syntax_tycoP: string -> string -> OuterParse.token list
+ -> (theory -> theory) * OuterParse.token list;
+ val add_syntax_constP: string -> string -> OuterParse.token list
+ -> (theory -> theory) * OuterParse.token list;
+
+ val add_undefined: string -> string -> string -> theory -> theory;
+ val add_pretty_list: string -> string -> string -> theory -> theory;
+ val add_pretty_list_string: string -> string -> string
+ -> string -> string list -> theory -> theory;
+ val add_pretty_char: string -> string -> string list -> theory -> theory
+ val add_pretty_numeral: string -> bool -> string * typ -> string -> string -> string
+ -> string -> string -> theory -> theory;
+ val add_pretty_ml_string: string -> string -> string list -> string
+ -> string -> string -> theory -> theory;
+ val add_pretty_imperative_monad_bind: string -> string -> theory -> theory;
+
+ type serializer;
+ val add_serializer: string * serializer -> theory -> theory;
+ val get_serializer: theory -> string -> bool -> string option -> string option -> Args.T list
+ -> (theory -> string -> string) -> string list option -> CodeThingol.code -> unit;
+ val assert_serializer: theory -> string -> string;
+
+ val eval_verbose: bool ref;
+ val eval_term: theory -> (theory -> string -> string) -> CodeThingol.code
+ -> (string * 'a option ref) * CodeThingol.iterm -> string list -> 'a;
+ val code_width: int ref;
+
+ val setup: theory -> theory;
+end;
+
+structure CodeTarget : CODE_TARGET =
+struct
+
+open BasicCodeThingol;
+
+(** basics **)
+
+infixr 5 @@;
+infixr 5 @|;
+fun x @@ y = [x, y];
+fun xs @| y = xs @ [y];
+val str = PrintMode.with_default Pretty.str;
+val concat = Pretty.block o Pretty.breaks;
+val brackets = Pretty.enclose "(" ")" o Pretty.breaks;
+fun semicolon ps = Pretty.block [concat ps, str ";"];
+
+
+(** syntax **)
+
+datatype lrx = L | R | X;
+
+datatype fixity =
+ BR
+ | NOBR
+ | INFX of (int * lrx);
+
+val APP = INFX (~1, L);
+
+fun eval_lrx L L = false
+ | eval_lrx R R = false
+ | eval_lrx _ _ = true;
+
+fun eval_fxy NOBR NOBR = false
+ | eval_fxy BR NOBR = false
+ | eval_fxy NOBR BR = false
+ | eval_fxy (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) =
+ pr < pr_ctxt
+ orelse pr = pr_ctxt
+ andalso eval_lrx lr lr_ctxt
+ orelse pr_ctxt = ~1
+ | eval_fxy _ (INFX _) = false
+ | eval_fxy (INFX _) NOBR = false
+ | eval_fxy _ _ = true;
+
+fun gen_brackify _ [p] = p
+ | gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps
+ | gen_brackify false (ps as _::_) = Pretty.block ps;
+
+fun brackify fxy_ctxt ps =
+ gen_brackify (eval_fxy BR fxy_ctxt) (Pretty.breaks ps);
+
+fun brackify_infix infx fxy_ctxt ps =
+ gen_brackify (eval_fxy (INFX infx) fxy_ctxt) (Pretty.breaks ps);
+
+type class_syntax = string * (string -> string option);
+type typ_syntax = int * ((fixity -> itype -> Pretty.T)
+ -> fixity -> itype list -> Pretty.T);
+type term_syntax = int * ((CodeName.var_ctxt -> fixity -> iterm -> Pretty.T)
+ -> CodeName.var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T);
+
+
+(* user-defined syntax *)
+
+datatype 'a mixfix =
+ Arg of fixity
+ | Pretty of Pretty.T;
+
+fun mk_mixfix prep_arg (fixity_this, mfx) =
+ let
+ fun is_arg (Arg _) = true
+ | is_arg _ = false;
+ val i = (length o filter is_arg) mfx;
+ fun fillin _ [] [] =
+ []
+ | fillin pr (Arg fxy :: mfx) (a :: args) =
+ (pr fxy o prep_arg) a :: fillin pr mfx args
+ | fillin pr (Pretty p :: mfx) args =
+ p :: fillin pr mfx args
+ | fillin _ [] _ =
+ error ("Inconsistent mixfix: too many arguments")
+ | fillin _ _ [] =
+ error ("Inconsistent mixfix: too less arguments");
+ in
+ (i, fn pr => fn fixity_ctxt => fn args =>
+ gen_brackify (eval_fxy fixity_this fixity_ctxt) (fillin pr mfx args))
+ end;
+
+fun parse_infix prep_arg (x, i) s =
+ let
+ val l = case x of L => INFX (i, L) | _ => INFX (i, X);
+ val r = case x of R => INFX (i, R) | _ => INFX (i, X);
+ in
+ mk_mixfix prep_arg (INFX (i, x), [Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r])
+ end;
+
+fun parse_mixfix prep_arg s =
+ let
+ val sym_any = Scan.one Symbol.is_regular;
+ val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat (
+ ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR))
+ || ($$ "_" >> K (Arg BR))
+ || ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length))
+ || (Scan.repeat1
+ ( $$ "'" |-- sym_any
+ || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")")
+ sym_any) >> (Pretty o str o implode)));
+ in case Scan.finite Symbol.stopper parse (Symbol.explode s)
+ of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p)
+ | ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p)
+ | _ => Scan.!! (the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail ()
+ end;
+
+fun parse_args f args =
+ case Scan.read Args.stopper f args
+ of SOME x => x
+ | NONE => error "Bad serializer arguments";
+
+
+(* generic serializer combinators *)
+
+fun gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons
+ lhs vars fxy (app as ((c, (_, tys)), ts)) =
+ case const_syntax c
+ of NONE => if lhs andalso not (is_cons c) then
+ error ("non-constructor on left hand side of equation: " ^ labelled_name c)
+ else brackify fxy (pr_app' lhs vars app)
+ | SOME (i, pr) =>
+ let
+ val k = if i < 0 then length tys else i;
+ fun pr' fxy ts = pr (pr_term lhs) vars fxy (ts ~~ curry Library.take k tys);
+ in if k = length ts
+ then pr' fxy ts
+ else if k < length ts
+ then case chop k ts of (ts1, ts2) =>
+ brackify fxy (pr' APP ts1 :: map (pr_term lhs vars BR) ts2)
+ else pr_term lhs vars fxy (CodeThingol.eta_expand app k)
+ end;
+
+fun gen_pr_bind pr_bind' pr_term fxy ((v, pat), ty) vars =
+ let
+ val vs = case pat
+ of SOME pat => CodeThingol.fold_varnames (insert (op =)) pat []
+ | NONE => [];
+ val vars' = CodeName.intro_vars (the_list v) vars;
+ val vars'' = CodeName.intro_vars vs vars';
+ val v' = Option.map (CodeName.lookup_var vars') v;
+ val pat' = Option.map (pr_term vars'' fxy) pat;
+ in (pr_bind' ((v', pat'), ty), vars'') end;
+
+
+(* list, char, string, numeral and monad abstract syntax transformations *)
+
+fun implode_list c_nil c_cons t =
+ let
+ fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
+ if c = c_cons
+ then SOME (t1, t2)
+ else NONE
+ | dest_cons _ = NONE;
+ val (ts, t') = CodeThingol.unfoldr dest_cons t;
+ in case t'
+ of IConst (c, _) => if c = c_nil then SOME ts else NONE
+ | _ => NONE
+ end;
+
+fun decode_char c_nibbles (IConst (c1, _), IConst (c2, _)) =
+ let
+ fun idx c = find_index (curry (op =) c) 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 c_char c_nibbles mk_char mk_string ts =
+ let
+ fun implode_char (IConst (c, _) `$ t1 `$ t2) =
+ if c = c_char then decode_char c_nibbles (t1, t2) else NONE
+ | implode_char _ = NONE;
+ val ts' = map implode_char ts;
+ in if forall is_some ts'
+ then (SOME o str o mk_string o implode o map_filter I) ts'
+ else NONE
+ end;
+
+fun implode_numeral c_bit0 c_bit1 c_pls c_min c_bit =
+ let
+ fun dest_bit (IConst (c, _)) = if c = c_bit0 then SOME 0
+ else if c = c_bit1 then SOME 1
+ else NONE
+ | dest_bit _ = NONE;
+ fun dest_numeral (IConst (c, _)) = if c = c_pls then SOME (IntInf.fromInt 0)
+ else if c = c_min then SOME (IntInf.fromInt ~1)
+ else NONE
+ | dest_numeral (IConst (c, _) `$ t1 `$ t2) =
+ if c = c_bit then case (dest_numeral t1, dest_bit t2)
+ of (SOME n, SOME b) => SOME (IntInf.fromInt 2 * n + IntInf.fromInt b)
+ | _ => NONE
+ else NONE
+ | dest_numeral _ = NONE;
+ in dest_numeral end;
+
+fun implode_monad c_mbind c_kbind t =
+ let
+ fun dest_monad (IConst (c, _) `$ t1 `$ t2) =
+ if c = c_mbind
+ then case CodeThingol.split_abs t2
+ of SOME (((v, pat), ty), t') => SOME ((SOME (((SOME v, pat), ty), true), t1), t')
+ | NONE => NONE
+ else if c = c_kbind
+ then SOME ((NONE, t1), t2)
+ else NONE
+ | dest_monad t = case CodeThingol.split_let t
+ of SOME (((pat, ty), tbind), t') => SOME ((SOME (((NONE, SOME pat), ty), false), tbind), t')
+ | NONE => NONE;
+ in CodeThingol.unfoldr dest_monad t end;
+
+
+(** name auxiliary **)
+
+val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
+val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
+
+val dest_name =
+ apfst NameSpace.implode o split_last o fst o split_last o NameSpace.explode;
+
+fun mk_modl_name_tab init_names prefix module_alias code =
+ let
+ fun nsp_map f = NameSpace.explode #> map f #> NameSpace.implode;
+ fun mk_alias name =
+ case module_alias name
+ of SOME name' => name'
+ | NONE => nsp_map (fn name => (the_single o fst)
+ (Name.variants [name] init_names)) name;
+ fun mk_prefix name =
+ case prefix
+ of SOME prefix => NameSpace.append prefix name
+ | NONE => name;
+ val tab =
+ Symtab.empty
+ |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name))
+ o fst o dest_name o fst)
+ code
+ in fn name => (the o Symtab.lookup tab) name end;
+
+
+
+(** SML/OCaml serializer **)
+
+datatype ml_def =
+ MLFuns of (string * ((iterm list * iterm) list * typscheme)) list
+ | MLDatas of (string * ((vname * sort) list * (string * itype list) list)) list
+ | MLClass of string * ((class * string) list * (vname * (string * itype) list))
+ | MLClassinst of string * ((class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * (string * iterm) list));
+
+fun pr_sml tyco_syntax const_syntax labelled_name init_syms deresolv is_cons ml_def =
+ let
+ val pr_label_classrel = translate_string (fn "." => "__" | c => c) o NameSpace.qualifier;
+ val pr_label_classop = NameSpace.base o NameSpace.qualifier;
+ fun pr_dicts fxy ds =
+ let
+ fun pr_dictvar (v, (_, 1)) = first_upper v ^ "_"
+ | pr_dictvar (v, (i, _)) = first_upper v ^ string_of_int (i+1) ^ "_";
+ fun pr_proj [] p =
+ p
+ | pr_proj [p'] p =
+ brackets [p', p]
+ | pr_proj (ps as _ :: _) p =
+ brackets [Pretty.enum " o" "(" ")" ps, p];
+ fun pr_dictc fxy (DictConst (inst, dss)) =
+ brackify fxy ((str o deresolv) inst :: map (pr_dicts BR) dss)
+ | pr_dictc fxy (DictVar (classrels, v)) =
+ pr_proj (map (str o deresolv) classrels) ((str o pr_dictvar) v)
+ in case ds
+ of [] => str "()"
+ | [d] => pr_dictc fxy d
+ | _ :: _ => (Pretty.list "(" ")" o map (pr_dictc NOBR)) ds
+ end;
+ fun pr_tyvars vs =
+ vs
+ |> map (fn (v, sort) => map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort)
+ |> map (pr_dicts BR);
+ fun pr_tycoexpr fxy (tyco, tys) =
+ let
+ val tyco' = (str o deresolv) tyco
+ in case map (pr_typ BR) tys
+ of [] => tyco'
+ | [p] => Pretty.block [p, Pretty.brk 1, tyco']
+ | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
+ end
+ and pr_typ fxy (tyco `%% tys) =
+ (case tyco_syntax tyco
+ of NONE => pr_tycoexpr fxy (tyco, tys)
+ | SOME (i, pr) =>
+ if not (i = length tys)
+ then error ("Number of argument mismatch in customary serialization: "
+ ^ (string_of_int o length) tys ^ " given, "
+ ^ string_of_int i ^ " expected")
+ else pr pr_typ fxy tys)
+ | pr_typ fxy (ITyVar v) =
+ str ("'" ^ v);
+ fun pr_term lhs vars fxy (IConst c) =
+ pr_app lhs vars fxy (c, [])
+ | pr_term lhs vars fxy (IVar v) =
+ str (CodeName.lookup_var vars v)
+ | pr_term lhs vars fxy (t as t1 `$ t2) =
+ (case CodeThingol.unfold_const_app t
+ of SOME c_ts => pr_app lhs vars fxy c_ts
+ | NONE =>
+ brackify fxy [pr_term lhs vars NOBR t1, pr_term lhs vars BR t2])
+ | pr_term lhs vars fxy (t as _ `|-> _) =
+ let
+ val (binds, t') = CodeThingol.unfold_abs t;
+ fun pr ((v, pat), ty) =
+ pr_bind NOBR ((SOME v, pat), ty)
+ #>> (fn p => concat [str "fn", p, str "=>"]);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (ps @ [pr_term lhs vars' NOBR t']) end
+ | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodeThingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
+ then pr_case vars fxy cases
+ else pr_app lhs vars fxy c_ts
+ | NONE => pr_case vars fxy cases)
+ and pr_app' lhs vars (app as ((c, (iss, tys)), ts)) =
+ if is_cons c then let
+ val k = length tys
+ in if k < 2 then
+ (str o deresolv) c :: map (pr_term lhs vars BR) ts
+ else if k = length ts then
+ [(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term lhs vars NOBR) ts)]
+ else [pr_term lhs vars BR (CodeThingol.eta_expand app k)] end else
+ (str o deresolv) c
+ :: (map (pr_dicts BR) o filter_out null) iss @ map (pr_term lhs vars BR) ts
+ and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
+ and pr_bind' ((NONE, NONE), _) = str "_"
+ | pr_bind' ((SOME v, NONE), _) = str v
+ | pr_bind' ((NONE, SOME p), _) = p
+ | pr_bind' ((SOME v, SOME p), _) = concat [str v, str "as", p]
+ and pr_bind fxy = gen_pr_bind pr_bind' (pr_term false) fxy
+ and pr_case vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, t') = CodeThingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind NOBR ((NONE, SOME pat), ty)
+ |>> (fn p => semicolon [str "val", p, str "=", pr_term false vars NOBR t])
+ val (ps, vars') = fold_map pr binds vars;
+ in
+ Pretty.chunks [
+ [str ("let"), Pretty.fbrk, Pretty.chunks ps] |> Pretty.block,
+ [str ("in"), Pretty.fbrk, pr_term false vars' NOBR t'] |> Pretty.block,
+ str ("end")
+ ]
+ end
+ | pr_case vars fxy (((td, ty), b::bs), _) =
+ let
+ fun pr delim (pat, t) =
+ let
+ val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
+ in
+ concat [str delim, p, str "=>", pr_term false vars' NOBR t]
+ end;
+ in
+ (Pretty.enclose "(" ")" o single o brackify fxy) (
+ str "case"
+ :: pr_term false vars NOBR td
+ :: pr "of" b
+ :: map (pr "|") bs
+ )
+ end
+ | pr_case vars fxy ((_, []), _) = str "raise Fail \"empty case\""
+ fun pr_def (MLFuns (funns as (funn :: funns'))) =
+ let
+ val definer =
+ let
+ fun mk [] [] = "val"
+ | mk (_::_) _ = "fun"
+ | mk [] vs = if (null o filter_out (null o snd)) vs then "val" else "fun";
+ fun chk (_, ((ts, _) :: _, (vs, _))) NONE = SOME (mk ts vs)
+ | chk (_, ((ts, _) :: _, (vs, _))) (SOME defi) =
+ if defi = mk ts vs then SOME defi
+ else error ("Mixing simultaneous vals and funs not implemented: "
+ ^ commas (map (labelled_name o fst) funns));
+ in the (fold chk funns NONE) end;
+ fun pr_funn definer (name, (eqs as eq::eqs', (raw_vs, ty))) =
+ let
+ val vs = filter_out (null o snd) raw_vs;
+ val shift = if null eqs' then I else
+ map (Pretty.block o single o Pretty.block o single);
+ fun pr_eq definer (ts, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts
+ |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ concat (
+ [str definer, (str o deresolv) name]
+ @ (if null ts andalso null vs
+ andalso not (ty = ITyVar "_")(*for evaluation*)
+ then [str ":", pr_typ NOBR ty]
+ else
+ pr_tyvars vs
+ @ map (pr_term true vars BR) ts)
+ @ [str "=", pr_term false vars NOBR t]
+ )
+ end
+ in
+ (Pretty.block o Pretty.fbreaks o shift) (
+ pr_eq definer eq
+ :: map (pr_eq "|") eqs'
+ )
+ end;
+ val (ps, p) = split_last (pr_funn definer funn :: map (pr_funn "and") funns');
+ in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end
+ | pr_def (MLDatas (datas as (data :: datas'))) =
+ let
+ fun pr_co (co, []) =
+ str (deresolv co)
+ | pr_co (co, tys) =
+ concat [
+ str (deresolv co),
+ str "of",
+ Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
+ ];
+ fun pr_data definer (tyco, (vs, [])) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ @@ str "EMPTY__"
+ )
+ | pr_data definer (tyco, (vs, cos)) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ :: separate (str "|") (map pr_co cos)
+ );
+ val (ps, p) = split_last (pr_data "datatype" data :: map (pr_data "and") datas');
+ in Pretty.chunks (ps @ [Pretty.block ([p, str ";"])]) end
+ | pr_def (MLClass (class, (superclasses, (v, classops)))) =
+ let
+ val w = first_upper v ^ "_";
+ fun pr_superclass_field (class, classrel) =
+ (concat o map str) [
+ pr_label_classrel classrel, ":", "'" ^ v, deresolv class
+ ];
+ fun pr_classop_field (classop, ty) =
+ concat [
+ (str o pr_label_classop) classop, str ":", pr_typ NOBR ty
+ ];
+ fun pr_classop_proj (classop, _) =
+ semicolon [
+ str "fun",
+ (str o deresolv) classop,
+ Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
+ str "=",
+ str ("#" ^ pr_label_classop classop),
+ str w
+ ];
+ fun pr_superclass_proj (_, classrel) =
+ semicolon [
+ str "fun",
+ (str o deresolv) classrel,
+ Pretty.enclose "(" ")" [str (w ^ ":'" ^ v ^ " " ^ deresolv class)],
+ str "=",
+ str ("#" ^ pr_label_classrel classrel),
+ str w
+ ];
+ in
+ Pretty.chunks (
+ concat [
+ str ("type '" ^ v),
+ (str o deresolv) class,
+ str "=",
+ Pretty.enum "," "{" "};" (
+ map pr_superclass_field superclasses @ map pr_classop_field classops
+ )
+ ]
+ :: map pr_superclass_proj superclasses
+ @ map pr_classop_proj classops
+ )
+ end
+ | pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) =
+ let
+ fun pr_superclass (_, (classrel, dss)) =
+ concat [
+ (str o pr_label_classrel) classrel,
+ str "=",
+ pr_dicts NOBR [DictConst dss]
+ ];
+ fun pr_classop (classop, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ (CodeThingol.fold_constnames (insert (op =)) t []);
+ val vars = CodeName.intro_vars consts init_syms;
+ in
+ concat [
+ (str o pr_label_classop) classop,
+ str "=",
+ pr_term false vars NOBR t
+ ]
+ end;
+ in
+ semicolon ([
+ str (if null arity then "val" else "fun"),
+ (str o deresolv) inst ] @
+ pr_tyvars arity @ [
+ str "=",
+ Pretty.enum "," "{" "}" (map pr_superclass superarities @ map pr_classop classop_defs),
+ str ":",
+ pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
+ ])
+ end;
+ in pr_def ml_def end;
+
+fun pr_sml_modl name content =
+ Pretty.chunks ([
+ str ("structure " ^ name ^ " = "),
+ str "struct",
+ str ""
+ ] @ content @ [
+ str "",
+ str ("end; (*struct " ^ name ^ "*)")
+ ]);
+
+fun pr_ocaml tyco_syntax const_syntax labelled_name init_syms deresolv is_cons ml_def =
+ let
+ fun pr_dicts fxy ds =
+ let
+ fun pr_dictvar (v, (_, 1)) = "_" ^ first_upper v
+ | pr_dictvar (v, (i, _)) = "_" ^ first_upper v ^ string_of_int (i+1);
+ fun pr_proj ps p =
+ fold_rev (fn p2 => fn p1 => Pretty.block [p1, str ".", str p2]) ps p
+ fun pr_dictc fxy (DictConst (inst, dss)) =
+ brackify fxy ((str o deresolv) inst :: map (pr_dicts BR) dss)
+ | pr_dictc fxy (DictVar (classrels, v)) =
+ pr_proj (map deresolv classrels) ((str o pr_dictvar) v)
+ in case ds
+ of [] => str "()"
+ | [d] => pr_dictc fxy d
+ | _ :: _ => (Pretty.list "(" ")" o map (pr_dictc NOBR)) ds
+ end;
+ fun pr_tyvars vs =
+ vs
+ |> map (fn (v, sort) => map_index (fn (i, _) => DictVar ([], (v, (i, length sort)))) sort)
+ |> map (pr_dicts BR);
+ fun pr_tycoexpr fxy (tyco, tys) =
+ let
+ val tyco' = (str o deresolv) tyco
+ in case map (pr_typ BR) tys
+ of [] => tyco'
+ | [p] => Pretty.block [p, Pretty.brk 1, tyco']
+ | (ps as _::_) => Pretty.block [Pretty.list "(" ")" ps, Pretty.brk 1, tyco']
+ end
+ and pr_typ fxy (tyco `%% tys) =
+ (case tyco_syntax tyco
+ of NONE => pr_tycoexpr fxy (tyco, tys)
+ | SOME (i, pr) =>
+ if not (i = length tys)
+ then error ("Number of argument mismatch in customary serialization: "
+ ^ (string_of_int o length) tys ^ " given, "
+ ^ string_of_int i ^ " expected")
+ else pr pr_typ fxy tys)
+ | pr_typ fxy (ITyVar v) =
+ str ("'" ^ v);
+ fun pr_term lhs vars fxy (IConst c) =
+ pr_app lhs vars fxy (c, [])
+ | pr_term lhs vars fxy (IVar v) =
+ str (CodeName.lookup_var vars v)
+ | pr_term lhs vars fxy (t as t1 `$ t2) =
+ (case CodeThingol.unfold_const_app t
+ of SOME c_ts => pr_app lhs vars fxy c_ts
+ | NONE =>
+ brackify fxy [pr_term lhs vars NOBR t1, pr_term lhs vars BR t2])
+ | pr_term lhs vars fxy (t as _ `|-> _) =
+ let
+ val (binds, t') = CodeThingol.unfold_abs t;
+ fun pr ((v, pat), ty) = pr_bind BR ((SOME v, pat), ty);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (str "fun" :: ps @ str "->" @@ pr_term lhs vars' NOBR t') end
+ | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodeThingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
+ then pr_case vars fxy cases
+ else pr_app lhs vars fxy c_ts
+ | NONE => pr_case vars fxy cases)
+ and pr_app' lhs vars (app as ((c, (iss, tys)), ts)) =
+ if is_cons c then
+ if length tys = length ts
+ then case ts
+ of [] => [(str o deresolv) c]
+ | [t] => [(str o deresolv) c, pr_term lhs vars BR t]
+ | _ => [(str o deresolv) c, Pretty.enum "," "(" ")" (map (pr_term lhs vars NOBR) ts)]
+ else [pr_term lhs vars BR (CodeThingol.eta_expand app (length tys))]
+ else (str o deresolv) c
+ :: ((map (pr_dicts BR) o filter_out null) iss @ map (pr_term lhs vars BR) ts)
+ and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
+ and pr_bind' ((NONE, NONE), _) = str "_"
+ | pr_bind' ((SOME v, NONE), _) = str v
+ | pr_bind' ((NONE, SOME p), _) = p
+ | pr_bind' ((SOME v, SOME p), _) = brackets [p, str "as", str v]
+ and pr_bind fxy = gen_pr_bind pr_bind' (pr_term false) fxy
+ and pr_case vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, t') = CodeThingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind NOBR ((NONE, SOME pat), ty)
+ |>> (fn p => concat [str "let", p, str "=", pr_term false vars NOBR t, str "in"])
+ val (ps, vars') = fold_map pr binds vars;
+ in Pretty.chunks (ps @| pr_term false vars' NOBR t') end
+ | pr_case vars fxy (((td, ty), b::bs), _) =
+ let
+ fun pr delim (pat, t) =
+ let
+ val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
+ in concat [str delim, p, str "->", pr_term false vars' NOBR t] end;
+ in
+ (Pretty.enclose "(" ")" o single o brackify fxy) (
+ str "match"
+ :: pr_term false vars NOBR td
+ :: pr "with" b
+ :: map (pr "|") bs
+ )
+ end
+ | pr_case vars fxy ((_, []), _) = str "failwith \"empty case\"";
+ fun pr_def (MLFuns (funns as funn :: funns')) =
+ let
+ fun fish_parm _ (w as SOME _) = w
+ | fish_parm (IVar v) NONE = SOME v
+ | fish_parm _ NONE = NONE;
+ fun fillup_parm _ (_, SOME v) = v
+ | fillup_parm x (i, NONE) = x ^ string_of_int i;
+ fun fish_parms vars eqs =
+ let
+ val raw_fished = fold (map2 fish_parm) eqs (replicate (length (hd eqs)) NONE);
+ val x = Name.variant (map_filter I raw_fished) "x";
+ val fished = map_index (fillup_parm x) raw_fished;
+ val vars' = CodeName.intro_vars fished vars;
+ in map (CodeName.lookup_var vars') fished end;
+ fun pr_eq (ts, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts
+ |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in concat [
+ (Pretty.block o Pretty.commas) (map (pr_term true vars NOBR) ts),
+ str "->",
+ pr_term false vars NOBR t
+ ] end;
+ fun pr_eqs [(ts, t)] =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts
+ |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ concat (
+ map (pr_term true vars BR) ts
+ @ str "="
+ @@ pr_term false vars NOBR t
+ )
+ end
+ | pr_eqs (eqs as (eq as ([_], _)) :: eqs') =
+ Pretty.block (
+ str "="
+ :: Pretty.brk 1
+ :: str "function"
+ :: Pretty.brk 1
+ :: pr_eq eq
+ :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs'
+ )
+ | pr_eqs (eqs as eq :: eqs') =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ ((fold o CodeThingol.fold_constnames) (insert (op =)) (map snd eqs) []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts;
+ val dummy_parms = (map str o fish_parms vars o map fst) eqs;
+ in
+ Pretty.block (
+ Pretty.breaks dummy_parms
+ @ Pretty.brk 1
+ :: str "="
+ :: Pretty.brk 1
+ :: str "match"
+ :: Pretty.brk 1
+ :: (Pretty.block o Pretty.commas) dummy_parms
+ :: Pretty.brk 1
+ :: str "with"
+ :: Pretty.brk 1
+ :: pr_eq eq
+ :: maps (append [Pretty.fbrk, str "|", Pretty.brk 1] o single o pr_eq) eqs'
+ )
+ end;
+ fun pr_funn definer (name, (eqs, (vs, ty))) =
+ concat (
+ str definer
+ :: (str o deresolv) name
+ :: pr_tyvars (filter_out (null o snd) vs)
+ @| pr_eqs eqs
+ );
+ val (ps, p) = split_last (pr_funn "let rec" funn :: map (pr_funn "and") funns');
+ in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end
+ | pr_def (MLDatas (datas as (data :: datas'))) =
+ let
+ fun pr_co (co, []) =
+ str (deresolv co)
+ | pr_co (co, tys) =
+ concat [
+ str (deresolv co),
+ str "of",
+ Pretty.enum " *" "" "" (map (pr_typ (INFX (2, X))) tys)
+ ];
+ fun pr_data definer (tyco, (vs, [])) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ @@ str "EMPTY_"
+ )
+ | pr_data definer (tyco, (vs, cos)) =
+ concat (
+ str definer
+ :: pr_tycoexpr NOBR (tyco, map (ITyVar o fst) vs)
+ :: str "="
+ :: separate (str "|") (map pr_co cos)
+ );
+ val (ps, p) = split_last (pr_data "type" data :: map (pr_data "and") datas');
+ in Pretty.chunks (ps @ [Pretty.block ([p, str ";;"])]) end
+ | pr_def (MLClass (class, (superclasses, (v, classops)))) =
+ let
+ val w = "_" ^ first_upper v;
+ fun pr_superclass_field (class, classrel) =
+ (concat o map str) [
+ deresolv classrel, ":", "'" ^ v, deresolv class
+ ];
+ fun pr_classop_field (classop, ty) =
+ concat [
+ (str o deresolv) classop, str ":", pr_typ NOBR ty
+ ];
+ fun pr_classop_proj (classop, _) =
+ concat [
+ str "let",
+ (str o deresolv) classop,
+ str w,
+ str "=",
+ str (w ^ "." ^ deresolv classop ^ ";;")
+ ];
+ in Pretty.chunks (
+ concat [
+ str ("type '" ^ v),
+ (str o deresolv) class,
+ str "=",
+ Pretty.enum ";" "{" "};;" (
+ map pr_superclass_field superclasses @ map pr_classop_field classops
+ )
+ ]
+ :: map pr_classop_proj classops
+ ) end
+ | pr_def (MLClassinst (inst, ((class, (tyco, arity)), (superarities, classop_defs)))) =
+ let
+ fun pr_superclass (_, (classrel, dss)) =
+ concat [
+ (str o deresolv) classrel,
+ str "=",
+ pr_dicts NOBR [DictConst dss]
+ ];
+ fun pr_classop_def (classop, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ (CodeThingol.fold_constnames (insert (op =)) t []);
+ val vars = CodeName.intro_vars consts init_syms;
+ in
+ concat [
+ (str o deresolv) classop,
+ str "=",
+ pr_term false vars NOBR t
+ ]
+ end;
+ in
+ concat (
+ str "let"
+ :: (str o deresolv) inst
+ :: pr_tyvars arity
+ @ str "="
+ @@ (Pretty.enclose "(" ");;" o Pretty.breaks) [
+ Pretty.enum ";" "{" "}" (map pr_superclass superarities @ map pr_classop_def classop_defs),
+ str ":",
+ pr_tycoexpr NOBR (class, [tyco `%% map (ITyVar o fst) arity])
+ ]
+ )
+ end;
+ in pr_def ml_def end;
+
+fun pr_ocaml_modl name content =
+ Pretty.chunks ([
+ str ("module " ^ name ^ " = "),
+ str "struct",
+ str ""
+ ] @ content @ [
+ str "",
+ str ("end;; (*struct " ^ name ^ "*)")
+ ]);
+
+val code_width = ref 80;
+fun code_output p = Pretty.setmp_margin (!code_width) Pretty.output p ^ "\n";
+
+fun seri_ml pr_def pr_modl module output labelled_name reserved_syms raw_module_alias module_prolog
+ (_ : string -> class_syntax option) tyco_syntax const_syntax code =
+ let
+ val module_alias = if is_some module then K module else raw_module_alias;
+ val is_cons = CodeThingol.is_cons code;
+ datatype node =
+ Def of string * ml_def option
+ | Module of string * ((Name.context * Name.context) * node Graph.T);
+ val init_names = Name.make_context reserved_syms;
+ val init_module = ((init_names, init_names), Graph.empty);
+ fun map_node [] f = f
+ | map_node (m::ms) f =
+ Graph.default_node (m, Module (m, init_module))
+ #> Graph.map_node m (fn (Module (dmodlname, (nsp, nodes))) => Module (dmodlname, (nsp, map_node ms f nodes)));
+ fun map_nsp_yield [] f (nsp, nodes) =
+ let
+ val (x, nsp') = f nsp
+ in (x, (nsp', nodes)) end
+ | map_nsp_yield (m::ms) f (nsp, nodes) =
+ let
+ val (x, nodes') =
+ nodes
+ |> Graph.default_node (m, Module (m, init_module))
+ |> Graph.map_node_yield m (fn Module (dmodlname, nsp_nodes) =>
+ let
+ val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes
+ in (x, Module (dmodlname, nsp_nodes')) end)
+ in (x, (nsp, nodes')) end;
+ val init_syms = CodeName.make_vars reserved_syms;
+ val name_modl = mk_modl_name_tab init_names NONE module_alias code;
+ fun name_def upper name nsp =
+ let
+ val (_, base) = dest_name name;
+ val base' = if upper then first_upper base else base;
+ val ([base''], nsp') = Name.variants [base'] nsp;
+ in (base'', nsp') end;
+ fun map_nsp_fun f (nsp_fun, nsp_typ) =
+ let
+ val (x, nsp_fun') = f nsp_fun
+ in (x, (nsp_fun', nsp_typ)) end;
+ fun map_nsp_typ f (nsp_fun, nsp_typ) =
+ let
+ val (x, nsp_typ') = f nsp_typ
+ in (x, (nsp_fun, nsp_typ')) end;
+ fun mk_funs defs =
+ fold_map
+ (fn (name, CodeThingol.Fun info) =>
+ map_nsp_fun (name_def false name) >> (fn base => (base, (name, info)))
+ | (name, def) => error ("Function block containing illegal definition: " ^ labelled_name name)
+ ) defs
+ >> (split_list #> apsnd MLFuns);
+ fun mk_datatype defs =
+ fold_map
+ (fn (name, CodeThingol.Datatype info) =>
+ map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info)))
+ | (name, CodeThingol.Datatypecons _) =>
+ map_nsp_fun (name_def true name) >> (fn base => (base, NONE))
+ | (name, def) => error ("Datatype block containing illegal definition: " ^ labelled_name name)
+ ) defs
+ >> (split_list #> apsnd (map_filter I
+ #> (fn [] => error ("Datatype block without data definition: " ^ (commas o map (labelled_name o fst)) defs)
+ | infos => MLDatas infos)));
+ fun mk_class defs =
+ fold_map
+ (fn (name, CodeThingol.Class info) =>
+ map_nsp_typ (name_def false name) >> (fn base => (base, SOME (name, info)))
+ | (name, CodeThingol.Classrel _) =>
+ map_nsp_fun (name_def false name) >> (fn base => (base, NONE))
+ | (name, CodeThingol.Classop _) =>
+ map_nsp_fun (name_def false name) >> (fn base => (base, NONE))
+ | (name, def) => error ("Class block containing illegal definition: " ^ labelled_name name)
+ ) defs
+ >> (split_list #> apsnd (map_filter I
+ #> (fn [] => error ("Class block without class definition: " ^ (commas o map (labelled_name o fst)) defs)
+ | [info] => MLClass info)));
+ fun mk_inst [(name, CodeThingol.Classinst info)] =
+ map_nsp_fun (name_def false name)
+ >> (fn base => ([base], MLClassinst (name, info)));
+ fun add_group mk defs nsp_nodes =
+ let
+ val names as (name :: names') = map fst defs;
+ val deps =
+ []
+ |> fold (fold (insert (op =)) o Graph.imm_succs code) names
+ |> subtract (op =) names;
+ val (modls, _) = (split_list o map dest_name) names;
+ val modl' = (the_single o distinct (op =) o map name_modl) modls
+ handle Empty =>
+ error ("Illegal mutual dependencies: " ^ commas (map labelled_name names));
+ val modl_explode = NameSpace.explode modl';
+ fun add_dep name name'' =
+ let
+ val modl'' = (name_modl o fst o dest_name) name'';
+ in if modl' = modl'' then
+ map_node modl_explode
+ (Graph.add_edge (name, name''))
+ else let
+ val (common, (diff1::_, diff2::_)) = chop_prefix (op =)
+ (modl_explode, NameSpace.explode modl'');
+ in
+ map_node common
+ (fn gr => Graph.add_edge_acyclic (diff1, diff2) gr
+ handle Graph.CYCLES _ => error ("Dependency "
+ ^ quote name ^ " -> " ^ quote name''
+ ^ " would result in module dependency cycle"))
+ end end;
+ in
+ nsp_nodes
+ |> map_nsp_yield modl_explode (mk defs)
+ |-> (fn (base' :: bases', def') =>
+ apsnd (map_node modl_explode (Graph.new_node (name, (Def (base', SOME def')))
+ #> fold2 (fn name' => fn base' => Graph.new_node (name', (Def (base', NONE)))) names' bases')))
+ |> apsnd (fold (fn name => fold (add_dep name) deps) names)
+ |> apsnd (fold (map_node modl_explode o Graph.add_edge) (product names names))
+ end;
+ fun group_defs [(_, CodeThingol.Bot)] =
+ I
+ | group_defs ((defs as (_, CodeThingol.Fun _)::_)) =
+ add_group mk_funs defs
+ | group_defs ((defs as (_, CodeThingol.Datatypecons _)::_)) =
+ add_group mk_datatype defs
+ | group_defs ((defs as (_, CodeThingol.Datatype _)::_)) =
+ add_group mk_datatype defs
+ | group_defs ((defs as (_, CodeThingol.Class _)::_)) =
+ add_group mk_class defs
+ | group_defs ((defs as (_, CodeThingol.Classrel _)::_)) =
+ add_group mk_class defs
+ | group_defs ((defs as (_, CodeThingol.Classop _)::_)) =
+ add_group mk_class defs
+ | group_defs ((defs as [(_, CodeThingol.Classinst _)])) =
+ add_group mk_inst defs
+ | group_defs defs = error ("Illegal mutual dependencies: " ^
+ (commas o map (labelled_name o fst)) defs)
+ val (_, nodes) =
+ init_module
+ |> fold group_defs (map (AList.make (Graph.get_node code))
+ (rev (Graph.strong_conn code)))
+ fun deresolver prefix name =
+ let
+ val modl = (fst o dest_name) name;
+ val modl' = (NameSpace.explode o name_modl) modl;
+ val (_, (_, remainder)) = chop_prefix (op =) (prefix, modl');
+ val defname' =
+ nodes
+ |> fold (fn m => fn g => case Graph.get_node g m
+ of Module (_, (_, g)) => g) modl'
+ |> (fn g => case Graph.get_node g name of Def (defname, _) => defname);
+ in
+ NameSpace.implode (remainder @ [defname'])
+ end handle Graph.UNDEF _ =>
+ error ("Unknown definition name: " ^ labelled_name name);
+ fun the_prolog modlname = case module_prolog modlname
+ of NONE => []
+ | SOME p => [p, str ""];
+ fun pr_node prefix (Def (_, NONE)) =
+ NONE
+ | pr_node prefix (Def (_, SOME def)) =
+ SOME (pr_def tyco_syntax const_syntax labelled_name init_syms
+ (deresolver prefix) is_cons def)
+ | pr_node prefix (Module (dmodlname, (_, nodes))) =
+ SOME (pr_modl dmodlname (the_prolog (NameSpace.implode (prefix @ [dmodlname]))
+ @ separate (str "") ((map_filter (pr_node (prefix @ [dmodlname]) o Graph.get_node nodes)
+ o rev o flat o Graph.strong_conn) nodes)));
+ val p = Pretty.chunks (the_prolog "" @ separate (str "") ((map_filter
+ (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes))
+ in output p end;
+
+val eval_verbose = ref false;
+
+fun isar_seri_sml module file =
+ let
+ val output = case file
+ of NONE => use_text "generated code" Output.ml_output (!eval_verbose) o code_output
+ | SOME "-" => writeln o code_output
+ | SOME file => File.write (Path.explode file) o code_output;
+ in
+ parse_args (Scan.succeed ())
+ #> (fn () => seri_ml pr_sml pr_sml_modl module output)
+ end;
+
+fun isar_seri_ocaml module file =
+ let
+ val output = case file
+ of NONE => error "OCaml: no internal compilation"
+ | SOME "-" => writeln o code_output
+ | SOME file => File.write (Path.explode file) o code_output;
+ fun output_file file = File.write (Path.explode file) o code_output;
+ val output_diag = writeln o code_output;
+ in
+ parse_args (Scan.succeed ())
+ #> (fn () => seri_ml pr_ocaml pr_ocaml_modl module output)
+ end;
+
+
+(** Haskell serializer **)
+
+local
+
+fun pr_bind' ((NONE, NONE), _) = str "_"
+ | pr_bind' ((SOME v, NONE), _) = str v
+ | pr_bind' ((NONE, SOME p), _) = p
+ | pr_bind' ((SOME v, SOME p), _) = brackets [str v, str "@", p]
+
+val pr_bind_haskell = gen_pr_bind pr_bind';
+
+in
+
+fun pr_haskell class_syntax tyco_syntax const_syntax labelled_name init_syms
+ deresolv_here deresolv is_cons deriving_show def =
+ let
+ fun class_name class = case class_syntax class
+ of NONE => deresolv class
+ | SOME (class, _) => class;
+ fun classop_name class classop = case class_syntax class
+ of NONE => deresolv_here classop
+ | SOME (_, classop_syntax) => case classop_syntax classop
+ of NONE => (snd o dest_name) classop
+ | SOME classop => classop
+ fun pr_typparms tyvars vs =
+ case maps (fn (v, sort) => map (pair v) sort) vs
+ of [] => str ""
+ | xs => Pretty.block [
+ Pretty.enum "," "(" ")" (
+ map (fn (v, class) => str
+ (class_name class ^ " " ^ CodeName.lookup_var tyvars v)) xs
+ ),
+ str " => "
+ ];
+ fun pr_tycoexpr tyvars fxy (tyco, tys) =
+ brackify fxy (str tyco :: map (pr_typ tyvars BR) tys)
+ and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) =
+ (case tyco_syntax tyco
+ of NONE =>
+ pr_tycoexpr tyvars fxy (deresolv tyco, tys)
+ | SOME (i, pr) =>
+ if not (i = length tys)
+ then error ("Number of argument mismatch in customary serialization: "
+ ^ (string_of_int o length) tys ^ " given, "
+ ^ string_of_int i ^ " expected")
+ else pr (pr_typ tyvars) fxy tys)
+ | pr_typ tyvars fxy (ITyVar v) =
+ (str o CodeName.lookup_var tyvars) v;
+ fun pr_typscheme_expr tyvars (vs, tycoexpr) =
+ Pretty.block (pr_typparms tyvars vs @@ pr_tycoexpr tyvars NOBR tycoexpr);
+ fun pr_typscheme tyvars (vs, ty) =
+ Pretty.block (pr_typparms tyvars vs @@ pr_typ tyvars NOBR ty);
+ fun pr_term lhs vars fxy (IConst c) =
+ pr_app lhs vars fxy (c, [])
+ | pr_term lhs vars fxy (t as (t1 `$ t2)) =
+ (case CodeThingol.unfold_const_app t
+ of SOME app => pr_app lhs vars fxy app
+ | _ =>
+ brackify fxy [
+ pr_term lhs vars NOBR t1,
+ pr_term lhs vars BR t2
+ ])
+ | pr_term lhs vars fxy (IVar v) =
+ (str o CodeName.lookup_var vars) v
+ | pr_term lhs vars fxy (t as _ `|-> _) =
+ let
+ val (binds, t') = CodeThingol.unfold_abs t;
+ fun pr ((v, pat), ty) = pr_bind BR ((SOME v, pat), ty);
+ val (ps, vars') = fold_map pr binds vars;
+ in brackets (str "\\" :: ps @ str "->" @@ pr_term lhs vars' NOBR t') end
+ | pr_term lhs vars fxy (ICase (cases as (_, t0))) = (case CodeThingol.unfold_const_app t0
+ of SOME (c_ts as ((c, _), _)) => if is_none (const_syntax c)
+ then pr_case vars fxy cases
+ else pr_app lhs vars fxy c_ts
+ | NONE => pr_case vars fxy cases)
+ and pr_app' lhs vars ((c, _), ts) =
+ (str o deresolv) c :: map (pr_term lhs vars BR) ts
+ and pr_app lhs vars = gen_pr_app pr_app' pr_term const_syntax labelled_name is_cons lhs vars
+ and pr_bind fxy = pr_bind_haskell (pr_term false) fxy
+ and pr_case vars fxy (cases as ((_, [_]), _)) =
+ let
+ val (binds, t) = CodeThingol.unfold_let (ICase cases);
+ fun pr ((pat, ty), t) vars =
+ vars
+ |> pr_bind BR ((NONE, SOME pat), ty)
+ |>> (fn p => semicolon [p, str "=", pr_term false vars NOBR t])
+ val (ps, vars') = fold_map pr binds vars;
+ in
+ Pretty.block_enclose (
+ str "let {",
+ concat [str "}", str "in", pr_term false vars' NOBR t]
+ ) ps
+ end
+ | pr_case vars fxy (((td, ty), bs as _ :: _), _) =
+ let
+ fun pr (pat, t) =
+ let
+ val (p, vars') = pr_bind NOBR ((NONE, SOME pat), ty) vars;
+ in semicolon [p, str "->", pr_term false vars' NOBR t] end;
+ in
+ Pretty.block_enclose (
+ concat [str "(case", pr_term false vars NOBR td, str "of", str "{"],
+ str "})"
+ ) (map pr bs)
+ end
+ | pr_case vars fxy ((_, []), _) = str "error \"empty case\"";
+ fun pr_def (name, CodeThingol.Fun (eqs, (vs, ty))) =
+ let
+ val tyvars = CodeName.intro_vars (map fst vs) init_syms;
+ fun pr_eq (ts, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ ((fold o CodeThingol.fold_constnames) (insert (op =)) (t :: ts) []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts
+ |> CodeName.intro_vars ((fold o CodeThingol.fold_unbound_varnames)
+ (insert (op =)) ts []);
+ in
+ semicolon (
+ (str o deresolv_here) name
+ :: map (pr_term true vars BR) ts
+ @ str "="
+ @@ pr_term false vars NOBR t
+ )
+ end;
+ in
+ Pretty.chunks (
+ Pretty.block [
+ (str o suffix " ::" o deresolv_here) name,
+ Pretty.brk 1,
+ pr_typscheme tyvars (vs, ty),
+ str ";"
+ ]
+ :: map pr_eq eqs
+ )
+ end
+ | pr_def (name, CodeThingol.Datatype (vs, [])) =
+ let
+ val tyvars = CodeName.intro_vars (map fst vs) init_syms;
+ in
+ semicolon [
+ str "data",
+ pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
+ ]
+ end
+ | pr_def (name, CodeThingol.Datatype (vs, [(co, [ty])])) =
+ let
+ val tyvars = CodeName.intro_vars (map fst vs) init_syms;
+ in
+ semicolon (
+ str "newtype"
+ :: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
+ :: str "="
+ :: (str o deresolv_here) co
+ :: pr_typ tyvars BR ty
+ :: (if deriving_show name then [str "deriving (Read, Show)"] else [])
+ )
+ end
+ | pr_def (name, CodeThingol.Datatype (vs, co :: cos)) =
+ let
+ val tyvars = CodeName.intro_vars (map fst vs) init_syms;
+ fun pr_co (co, tys) =
+ concat (
+ (str o deresolv_here) co
+ :: map (pr_typ tyvars BR) tys
+ )
+ in
+ semicolon (
+ str "data"
+ :: pr_typscheme_expr tyvars (vs, (deresolv_here name, map (ITyVar o fst) vs))
+ :: str "="
+ :: pr_co co
+ :: map ((fn p => Pretty.block [str "| ", p]) o pr_co) cos
+ @ (if deriving_show name then [str "deriving (Read, Show)"] else [])
+ )
+ end
+ | pr_def (name, CodeThingol.Class (superclasss, (v, classops))) =
+ let
+ val tyvars = CodeName.intro_vars [v] init_syms;
+ fun pr_classop (classop, ty) =
+ semicolon [
+ (str o classop_name name) classop,
+ str "::",
+ pr_typ tyvars NOBR ty
+ ]
+ in
+ Pretty.block_enclose (
+ Pretty.block [
+ str "class ",
+ pr_typparms tyvars [(v, map fst superclasss)],
+ str (deresolv_here name ^ " " ^ CodeName.lookup_var tyvars v),
+ str " where {"
+ ],
+ str "};"
+ ) (map pr_classop classops)
+ end
+ | pr_def (_, CodeThingol.Classinst ((class, (tyco, vs)), (_, classop_defs))) =
+ let
+ val tyvars = CodeName.intro_vars (map fst vs) init_syms;
+ fun pr_instdef (classop, t) =
+ let
+ val consts = map_filter
+ (fn c => if (is_some o const_syntax) c
+ then NONE else (SOME o NameSpace.base o deresolv) c)
+ (CodeThingol.fold_constnames (insert (op =)) t []);
+ val vars = init_syms
+ |> CodeName.intro_vars consts;
+ in
+ semicolon [
+ (str o classop_name class) classop,
+ str "=",
+ pr_term false vars NOBR t
+ ]
+ end;
+ in
+ Pretty.block_enclose (
+ Pretty.block [
+ str "instance ",
+ pr_typparms tyvars vs,
+ str (class_name class ^ " "),
+ pr_typ tyvars BR (tyco `%% map (ITyVar o fst) vs),
+ str " where {"
+ ],
+ str "};"
+ ) (map pr_instdef classop_defs)
+ end;
+ in pr_def def end;
+
+fun pretty_haskell_monad c_mbind c_kbind =
+ let
+ fun pretty pr vars fxy [(t, _)] =
+ let
+ val pr_bind = pr_bind_haskell pr;
+ fun pr_mbind (NONE, t) vars =
+ (semicolon [pr vars NOBR t], vars)
+ | pr_mbind (SOME (bind, true), t) vars = vars
+ |> pr_bind NOBR bind
+ |>> (fn p => semicolon [p, str "<-", pr vars NOBR t])
+ | pr_mbind (SOME (bind, false), t) vars = vars
+ |> pr_bind NOBR bind
+ |>> (fn p => semicolon [str "let", p, str "=", pr vars NOBR t]);
+ val (binds, t) = implode_monad c_mbind c_kbind t;
+ val (ps, vars') = fold_map pr_mbind binds vars;
+ fun brack p = if eval_fxy BR fxy then Pretty.block [str "(", p, str ")"] else p;
+ in (brack o Pretty.block_enclose (str "do {", str "}")) (ps @| pr vars' NOBR t) end;
+ in (1, pretty) end;
+
+end; (*local*)
+
+fun seri_haskell module_prefix module destination string_classes labelled_name
+ reserved_syms raw_module_alias module_prolog class_syntax tyco_syntax const_syntax code =
+ let
+ val _ = Option.map File.check destination;
+ val is_cons = CodeThingol.is_cons code;
+ val module_alias = if is_some module then K module else raw_module_alias;
+ val init_names = Name.make_context reserved_syms;
+ val name_modl = mk_modl_name_tab init_names module_prefix module_alias code;
+ fun add_def (name, (def, deps)) =
+ let
+ val (modl, base) = dest_name name;
+ fun name_def base = Name.variants [base] #>> the_single;
+ fun add_fun upper (nsp_fun, nsp_typ) =
+ let
+ val (base', nsp_fun') = name_def (if upper then first_upper base else base) nsp_fun
+ in (base', (nsp_fun', nsp_typ)) end;
+ fun add_typ (nsp_fun, nsp_typ) =
+ let
+ val (base', nsp_typ') = name_def (first_upper base) nsp_typ
+ in (base', (nsp_fun, nsp_typ')) end;
+ val add_name =
+ case def
+ of CodeThingol.Bot => pair base
+ | CodeThingol.Fun _ => add_fun false
+ | CodeThingol.Datatype _ => add_typ
+ | CodeThingol.Datatypecons _ => add_fun true
+ | CodeThingol.Class _ => add_typ
+ | CodeThingol.Classrel _ => pair base
+ | CodeThingol.Classop _ => add_fun false
+ | CodeThingol.Classinst _ => pair base;
+ val modlname' = name_modl modl;
+ fun add_def base' =
+ case def
+ of CodeThingol.Bot => I
+ | CodeThingol.Datatypecons _ =>
+ cons (name, ((NameSpace.append modlname' base', base'), NONE))
+ | CodeThingol.Classrel _ => I
+ | CodeThingol.Classop _ =>
+ cons (name, ((NameSpace.append modlname' base', base'), NONE))
+ | _ => cons (name, ((NameSpace.append modlname' base', base'), SOME def));
+ in
+ Symtab.map_default (modlname', ([], ([], (init_names, init_names))))
+ (apfst (fold (insert (op = : string * string -> bool)) deps))
+ #> `(fn code => add_name ((snd o snd o the o Symtab.lookup code) modlname'))
+ #-> (fn (base', names) =>
+ (Symtab.map_entry modlname' o apsnd) (fn (defs, _) =>
+ (add_def base' defs, names)))
+ end;
+ val code' =
+ fold add_def (AList.make (fn name => (Graph.get_node code name, Graph.imm_succs code name))
+ (Graph.strong_conn code |> flat)) Symtab.empty;
+ val init_syms = CodeName.make_vars reserved_syms;
+ fun deresolv name =
+ (fst o fst o the o AList.lookup (op =) ((fst o snd o the
+ o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name
+ handle Option => error ("Unknown definition name: " ^ labelled_name name);
+ fun deresolv_here name =
+ (snd o fst o the o AList.lookup (op =) ((fst o snd o the
+ o Symtab.lookup code') ((name_modl o fst o dest_name) name))) name
+ handle Option => error ("Unknown definition name: " ^ labelled_name name);
+ fun deriving_show tyco =
+ let
+ fun deriv _ "fun" = false
+ | deriv tycos tyco = member (op =) tycos tyco orelse
+ case the_default CodeThingol.Bot (try (Graph.get_node code) tyco)
+ of CodeThingol.Bot => true
+ | CodeThingol.Datatype (_, cs) => forall (deriv' (tyco :: tycos))
+ (maps snd cs)
+ and deriv' tycos (tyco `%% tys) = deriv tycos tyco
+ andalso forall (deriv' tycos) tys
+ | deriv' _ (ITyVar _) = true
+ in deriv [] tyco end;
+ fun seri_def qualified = pr_haskell class_syntax tyco_syntax const_syntax labelled_name init_syms
+ deresolv_here (if qualified then deresolv else deresolv_here) is_cons
+ (if string_classes then deriving_show else K false);
+ fun write_module (SOME destination) modlname =
+ let
+ val filename = case modlname
+ of "" => Path.explode "Main.hs"
+ | _ => (Path.ext "hs" o Path.explode o implode o separate "/" o NameSpace.explode) modlname;
+ val pathname = Path.append destination filename;
+ val _ = File.mkdir (Path.dir pathname);
+ in File.write pathname end
+ | write_module NONE _ = writeln;
+ fun seri_module (modlname', (imports, (defs, _))) =
+ let
+ val imports' =
+ imports
+ |> map (name_modl o fst o dest_name)
+ |> distinct (op =)
+ |> remove (op =) modlname';
+ val qualified =
+ imports
+ |> map_filter (try deresolv)
+ |> map NameSpace.base
+ |> has_duplicates (op =);
+ val mk_import = str o (if qualified
+ then prefix "import qualified "
+ else prefix "import ") o suffix ";";
+ in
+ Pretty.chunks (
+ str ("module " ^ modlname' ^ " where {")
+ :: str ""
+ :: map mk_import imports'
+ @ str ""
+ :: separate (str "") ((case module_prolog modlname'
+ of SOME prolog => [prolog]
+ | NONE => [])
+ @ map_filter
+ (fn (name, (_, SOME def)) => SOME (seri_def qualified (name, def))
+ | (_, (_, NONE)) => NONE) defs)
+ @ str ""
+ @@ str "}"
+ )
+ |> code_output
+ |> write_module destination modlname'
+ end;
+ in Symtab.fold (fn modl => fn () => seri_module modl) code' () end;
+
+fun isar_seri_haskell module file =
+ let
+ val destination = case file
+ of NONE => error ("Haskell: no internal compilation")
+ | SOME "-" => NONE
+ | SOME file => SOME (Path.explode file)
+ in
+ parse_args (Scan.option (Args.$$$ "root" -- Args.colon |-- Args.name)
+ -- Scan.optional (Args.$$$ "string_classes" >> K true) false
+ >> (fn (module_prefix, string_classes) =>
+ seri_haskell module_prefix module destination string_classes))
+ end;
+
+
+(** diagnosis serializer **)
+
+fun seri_diagnosis labelled_name _ _ _ _ _ _ code =
+ let
+ val init_names = CodeName.make_vars [];
+ fun pr_fun "fun" = SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ])
+ | pr_fun _ = NONE
+ val pr = pr_haskell (K NONE) pr_fun (K NONE) labelled_name init_names I I (K false) (K false);
+ in
+ []
+ |> Graph.fold (fn (name, (def, _)) => case try pr (name, def) of SOME p => cons p | NONE => I) code
+ |> Pretty.chunks2
+ |> code_output
+ |> writeln
+ end;
+
+
+
+(** theory data **)
+
+datatype syntax_expr = SyntaxExpr of {
+ class: (string * (string -> string option)) Symtab.table,
+ inst: unit Symtab.table,
+ tyco: typ_syntax Symtab.table,
+ const: term_syntax Symtab.table
+};
+
+fun mk_syntax_expr ((class, inst), (tyco, const)) =
+ SyntaxExpr { class = class, inst = inst, tyco = tyco, const = const };
+fun map_syntax_expr f (SyntaxExpr { class, inst, tyco, const }) =
+ mk_syntax_expr (f ((class, inst), (tyco, const)));
+fun merge_syntax_expr (SyntaxExpr { class = class1, inst = inst1, tyco = tyco1, const = const1 },
+ SyntaxExpr { class = class2, inst = inst2, tyco = tyco2, const = const2 }) =
+ mk_syntax_expr (
+ (Symtab.join (K snd) (class1, class2),
+ Symtab.join (K snd) (inst1, inst2)),
+ (Symtab.join (K snd) (tyco1, tyco2),
+ Symtab.join (K snd) (const1, const2))
+ );
+
+datatype syntax_modl = SyntaxModl of {
+ alias: string Symtab.table,
+ prolog: Pretty.T Symtab.table
+};
+
+fun mk_syntax_modl (alias, prolog) =
+ SyntaxModl { alias = alias, prolog = prolog };
+fun map_syntax_modl f (SyntaxModl { alias, prolog }) =
+ mk_syntax_modl (f (alias, prolog));
+fun merge_syntax_modl (SyntaxModl { alias = alias1, prolog = prolog1 },
+ SyntaxModl { alias = alias2, prolog = prolog2 }) =
+ mk_syntax_modl (
+ Symtab.join (K snd) (alias1, alias2),
+ Symtab.join (K snd) (prolog1, prolog2)
+ );
+
+type serializer =
+ string option
+ -> string option
+ -> Args.T list
+ -> (string -> string)
+ -> string list
+ -> (string -> string option)
+ -> (string -> Pretty.T option)
+ -> (string -> class_syntax option)
+ -> (string -> typ_syntax option)
+ -> (string -> term_syntax option)
+ -> CodeThingol.code -> unit;
+
+datatype target = Target of {
+ serial: serial,
+ serializer: serializer,
+ syntax_expr: syntax_expr,
+ syntax_modl: syntax_modl,
+ reserved: string list
+};
+
+fun mk_target (serial, ((serializer, reserved), (syntax_expr, syntax_modl))) =
+ Target { serial = serial, reserved = reserved, serializer = serializer, syntax_expr = syntax_expr, syntax_modl = syntax_modl };
+fun map_target f ( Target { serial, serializer, reserved, syntax_expr, syntax_modl } ) =
+ mk_target (f (serial, ((serializer, reserved), (syntax_expr, syntax_modl))));
+fun merge_target target (Target { serial = serial1, serializer = serializer, reserved = reserved1,
+ syntax_expr = syntax_expr1, syntax_modl = syntax_modl1 },
+ Target { serial = serial2, serializer = _, reserved = reserved2,
+ syntax_expr = syntax_expr2, syntax_modl = syntax_modl2 }) =
+ if serial1 = serial2 then
+ mk_target (serial1, ((serializer, merge (op =) (reserved1, reserved2)),
+ (merge_syntax_expr (syntax_expr1, syntax_expr2),
+ merge_syntax_modl (syntax_modl1, syntax_modl2))
+ ))
+ else
+ error ("Incompatible serializers: " ^ quote target);
+
+structure CodeTargetData = TheoryDataFun
+(
+ type T = target Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ = Symtab.join merge_target;
+);
+
+fun the_serializer (Target { serializer, ... }) = serializer;
+fun the_reserved (Target { reserved, ... }) = reserved;
+fun the_syntax_expr (Target { syntax_expr = SyntaxExpr x, ... }) = x;
+fun the_syntax_modl (Target { syntax_modl = SyntaxModl x, ... }) = x;
+
+fun assert_serializer thy target =
+ case Symtab.lookup (CodeTargetData.get thy) target
+ of SOME data => target
+ | NONE => error ("Unknown code target language: " ^ quote target);
+
+fun add_serializer (target, seri) thy =
+ let
+ val _ = case Symtab.lookup (CodeTargetData.get thy) target
+ of SOME _ => warning ("overwriting existing serializer " ^ quote target)
+ | NONE => ();
+ in
+ thy
+ |> (CodeTargetData.map oo Symtab.map_default)
+ (target, mk_target (serial (), ((seri, []),
+ (mk_syntax_expr ((Symtab.empty, Symtab.empty), (Symtab.empty, Symtab.empty)),
+ mk_syntax_modl (Symtab.empty, Symtab.empty)))))
+ (map_target (fn (serial, ((_, keywords), syntax)) => (serial, ((seri, keywords), syntax))))
+ end;
+
+fun map_seri_data target f thy =
+ let
+ val _ = assert_serializer thy target;
+ in
+ thy
+ |> (CodeTargetData.map o Symtab.map_entry target o map_target) f
+ end;
+
+val target_SML = "SML";
+val target_OCaml = "OCaml";
+val target_Haskell = "Haskell";
+val target_diag = "diag";
+
+fun get_serializer thy target permissive module file args labelled_name = fn cs =>
+ let
+ val data = case Symtab.lookup (CodeTargetData.get thy) target
+ of SOME data => data
+ | NONE => error ("Unknown code target language: " ^ quote target);
+ val seri = the_serializer data;
+ val reserved = the_reserved data;
+ val { alias, prolog } = the_syntax_modl data;
+ val { class, inst, tyco, const } = the_syntax_expr data;
+ val project = if target = target_diag then I
+ else CodeThingol.project_code permissive
+ (Symtab.keys class @ Symtab.keys inst @ Symtab.keys tyco @ Symtab.keys const) cs;
+ fun check_empty_funs code = case CodeThingol.empty_funs code
+ of [] => code
+ | names => error ("No defining equations for " ^ commas (map (labelled_name thy) names));
+ in
+ project
+ #> check_empty_funs
+ #> seri module file args (labelled_name thy) reserved (Symtab.lookup alias) (Symtab.lookup prolog)
+ (Symtab.lookup class) (Symtab.lookup tyco) (Symtab.lookup const)
+ end;
+
+fun eval_term thy labelled_name code ((ref_name, reff), t) args =
+ let
+ val val_name = "Isabelle_Eval.EVAL.EVAL";
+ val val_name' = "Isabelle_Eval.EVAL";
+ val val_name'_args = space_implode " " (val_name' :: map (enclose "(" ")") args);
+ val seri = get_serializer thy "SML" false (SOME "Isabelle_Eval") NONE [] labelled_name;
+ fun eval code = (
+ reff := NONE;
+ seri (SOME [val_name]) code;
+ use_text "generated code for evaluation" Output.ml_output (!eval_verbose)
+ ("val _ = (" ^ ref_name ^ " := SOME (" ^ val_name'_args ^ "))");
+ case !reff
+ of NONE => error ("Could not retrieve value of ML reference " ^ quote ref_name
+ ^ " (reference probably has been shadowed)")
+ | SOME value => value
+ );
+ in
+ code
+ |> CodeThingol.add_eval_def (val_name, t)
+ |> eval
+ end;
+
+
+
+(** optional pretty serialization **)
+
+local
+
+val pretty : (string * {
+ pretty_char: string -> string,
+ pretty_string: string -> string,
+ pretty_numeral: bool -> IntInf.int -> string,
+ pretty_list: Pretty.T list -> Pretty.T,
+ infix_cons: int * string
+ }) list = [
+ ("SML", { pretty_char = prefix "#" o quote o ML_Syntax.print_char,
+ pretty_string = ML_Syntax.print_string,
+ pretty_numeral = fn unbounded => fn k =>
+ if unbounded then "(" ^ IntInf.toString k ^ " : IntInf.int)"
+ else IntInf.toString k,
+ pretty_list = Pretty.enum "," "[" "]",
+ infix_cons = (7, "::")}),
+ ("OCaml", { pretty_char = fn c => enclose "'" "'"
+ (let val i = ord c
+ in if i < 32 orelse i = 39 orelse i = 92
+ then prefix "\\" (string_of_int i)
+ else c
+ end),
+ pretty_string = (fn _ => error "OCaml: no pretty strings"),
+ pretty_numeral = fn unbounded => fn k => if k >= IntInf.fromInt 0 then
+ if unbounded then
+ "(Big_int.big_int_of_int " ^ IntInf.toString k ^ ")"
+ else IntInf.toString k
+ else
+ if unbounded then
+ "(Big_int.big_int_of_int " ^ (enclose "(" ")" o prefix "-"
+ o IntInf.toString o op ~) k ^ ")"
+ else (enclose "(" ")" o prefix "-" o IntInf.toString o op ~) k,
+ pretty_list = Pretty.enum ";" "[" "]",
+ infix_cons = (6, "::")}),
+ ("Haskell", { pretty_char = fn c => enclose "'" "'"
+ (let val i = ord c
+ in if i < 32 orelse i = 39 orelse i = 92
+ then Library.prefix "\\" (string_of_int i)
+ else c
+ end),
+ pretty_string = ML_Syntax.print_string,
+ pretty_numeral = fn unbounded => fn k => if k >= IntInf.fromInt 0 then
+ IntInf.toString k
+ else
+ (enclose "(" ")" o Library.prefix "-" o IntInf.toString o IntInf.~) k,
+ pretty_list = Pretty.enum "," "[" "]",
+ infix_cons = (5, ":")})
+];
+
+in
+
+fun pr_pretty target = case AList.lookup (op =) pretty target
+ of SOME x => x
+ | NONE => error ("Unknown code target language: " ^ quote target);
+
+fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
+ brackify_infix (target_fxy, R) fxy [
+ pr (INFX (target_fxy, X)) t1,
+ str target_cons,
+ pr (INFX (target_fxy, R)) t2
+ ];
+
+fun pretty_list c_nil c_cons target =
+ let
+ val pretty_ops = pr_pretty target;
+ val mk_list = #pretty_list pretty_ops;
+ fun pretty pr vars fxy [(t1, _), (t2, _)] =
+ case Option.map (cons t1) (implode_list c_nil c_cons t2)
+ of SOME ts => mk_list (map (pr vars NOBR) ts)
+ | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2;
+ in (2, pretty) end;
+
+fun pretty_list_string c_nil c_cons c_char c_nibbles target =
+ let
+ val pretty_ops = pr_pretty target;
+ val mk_list = #pretty_list pretty_ops;
+ val mk_char = #pretty_char pretty_ops;
+ val mk_string = #pretty_string pretty_ops;
+ fun pretty pr vars fxy [(t1, _), (t2, _)] =
+ case Option.map (cons t1) (implode_list c_nil c_cons t2)
+ of SOME ts => case implode_string c_char c_nibbles mk_char mk_string ts
+ of SOME p => p
+ | NONE => mk_list (map (pr vars NOBR) ts)
+ | NONE => default_list (#infix_cons pretty_ops) (pr vars) fxy t1 t2;
+ in (2, pretty) end;
+
+fun pretty_char c_char c_nibbles target =
+ let
+ val mk_char = #pretty_char (pr_pretty target);
+ fun pretty _ _ _ [(t1, _), (t2, _)] =
+ case decode_char c_nibbles (t1, t2)
+ of SOME c => (str o mk_char) c
+ | NONE => error "Illegal character expression";
+ in (2, pretty) end;
+
+fun pretty_numeral unbounded c_bit0 c_bit1 c_pls c_min c_bit target =
+ let
+ val mk_numeral = #pretty_numeral (pr_pretty target);
+ fun pretty _ _ _ [(t, _)] =
+ case implode_numeral c_bit0 c_bit1 c_pls c_min c_bit t
+ of SOME k => (str o mk_numeral unbounded) k
+ | NONE => error "Illegal numeral expression";
+ in (1, pretty) end;
+
+fun pretty_ml_string c_char c_nibbles c_nil c_cons target =
+ let
+ val pretty_ops = pr_pretty target;
+ val mk_char = #pretty_char pretty_ops;
+ val mk_string = #pretty_string pretty_ops;
+ fun pretty pr vars fxy [(t, _)] =
+ case implode_list c_nil c_cons t
+ of SOME ts => (case implode_string c_char c_nibbles mk_char mk_string ts
+ of SOME p => p
+ | NONE => error "Illegal ml_string expression")
+ | NONE => error "Illegal ml_string expression";
+ in (1, pretty) end;
+
+val pretty_imperative_monad_bind =
+ let
+ fun pretty (pr : CodeName.var_ctxt -> fixity -> iterm -> Pretty.T)
+ vars fxy [(t1, _), ((v, ty) `|-> t2, _)] =
+ pr vars fxy (ICase (((t1, ty), ([(IVar v, t2)])), IVar ""))
+ | pretty pr vars fxy [(t1, _), (t2, ty2)] =
+ let
+ (*this code suffers from the lack of a proper concept for bindings*)
+ val vs = CodeThingol.fold_varnames cons t2 [];
+ val v = Name.variant vs "x";
+ val vars' = CodeName.intro_vars [v] vars;
+ val var = IVar v;
+ val ty = (hd o fst o CodeThingol.unfold_fun) ty2;
+ in pr vars' fxy (ICase (((t1, ty), ([(var, t2 `$ var)])), IVar "")) end;
+ in (2, pretty) end;
+
+end; (*local*)
+
+(** ML and Isar interface **)
+
+local
+
+fun map_syntax_exprs target =
+ map_seri_data target o apsnd o apsnd o apfst o map_syntax_expr;
+fun map_syntax_modls target =
+ map_seri_data target o apsnd o apsnd o apsnd o map_syntax_modl;
+fun map_reserveds target =
+ map_seri_data target o apsnd o apfst o apsnd;
+
+fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy =
+ let
+ val cls = prep_class thy raw_class;
+ val class = CodeName.class thy cls;
+ fun mk_classop (const as (c, _)) = case AxClass.class_of_param thy c
+ of SOME class' => if cls = class' then CodeName.const thy const
+ else error ("Not a class operation for class " ^ quote class ^ ": " ^ quote c)
+ | NONE => error ("Not a class operation: " ^ quote c);
+ fun mk_syntax_ops raw_ops = AList.lookup (op =)
+ ((map o apfst) (mk_classop o prep_const thy) raw_ops);
+ in case raw_syn
+ of SOME (syntax, raw_ops) =>
+ thy
+ |> (map_syntax_exprs target o apfst o apfst)
+ (Symtab.update (class, (syntax, mk_syntax_ops raw_ops)))
+ | NONE =>
+ thy
+ |> (map_syntax_exprs target o apfst o apfst)
+ (Symtab.delete_safe class)
+ end;
+
+fun gen_add_syntax_inst prep_class prep_tyco target (raw_tyco, raw_class) add_del thy =
+ let
+ val inst = CodeName.instance thy (prep_class thy raw_class, prep_tyco thy raw_tyco);
+ in if add_del then
+ thy
+ |> (map_syntax_exprs target o apfst o apsnd)
+ (Symtab.update (inst, ()))
+ else
+ thy
+ |> (map_syntax_exprs target o apfst o apsnd)
+ (Symtab.delete_safe inst)
+ end;
+
+fun gen_add_syntax_tyco prep_tyco target raw_tyco raw_syn thy =
+ let
+ val tyco = prep_tyco thy raw_tyco;
+ val tyco' = if tyco = "fun" then "fun" else CodeName.tyco thy tyco;
+ fun check_args (syntax as (n, _)) = if n <> Sign.arity_number thy tyco
+ then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
+ else syntax
+ in case raw_syn
+ of SOME syntax =>
+ thy
+ |> (map_syntax_exprs target o apsnd o apfst)
+ (Symtab.update (tyco', check_args syntax))
+ | NONE =>
+ thy
+ |> (map_syntax_exprs target o apsnd o apfst)
+ (Symtab.delete_safe tyco')
+ end;
+
+fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
+ let
+ val c = prep_const thy raw_c;
+ val c' = CodeName.const thy c;
+ fun check_args (syntax as (n, _)) = if n > (length o fst o strip_type o Sign.the_const_type thy o fst) c
+ then error ("Too many arguments in syntax for constant " ^ (quote o fst) c)
+ else syntax;
+ in case raw_syn
+ of SOME syntax =>
+ thy
+ |> (map_syntax_exprs target o apsnd o apsnd)
+ (Symtab.update (c', check_args syntax))
+ | NONE =>
+ thy
+ |> (map_syntax_exprs target o apsnd o apsnd)
+ (Symtab.delete_safe c')
+ end;
+
+fun cert_class thy class =
+ let
+ val _ = AxClass.get_definition thy class;
+ in class end;
+
+fun read_class thy raw_class =
+ let
+ val class = Sign.intern_class thy raw_class;
+ val _ = AxClass.get_definition thy class;
+ in class end;
+
+fun cert_tyco thy tyco =
+ let
+ val _ = if Sign.declared_tyname thy tyco then ()
+ else error ("No such type constructor: " ^ quote tyco);
+ in tyco end;
+
+fun read_tyco thy raw_tyco =
+ let
+ val tyco = Sign.intern_type thy raw_tyco;
+ val _ = if Sign.declared_tyname thy tyco then ()
+ else error ("No such type constructor: " ^ quote raw_tyco);
+ in tyco end;
+
+fun idfs_of_const thy c =
+ let
+ val c' = (c, Sign.the_const_type thy c);
+ val c'' = CodeUnit.const_of_cexpr thy c';
+ in (c'', CodeName.const thy c'') end;
+
+fun no_bindings x = (Option.map o apsnd)
+ (fn pretty => fn pr => fn vars => pretty (pr vars)) x;
+
+fun gen_add_haskell_monad prep_const c_run c_mbind c_kbind thy =
+ let
+ val c_run' = prep_const thy c_run;
+ val c_mbind' = prep_const thy c_mbind;
+ val c_mbind'' = CodeName.const thy c_mbind';
+ val c_kbind' = prep_const thy c_kbind;
+ val c_kbind'' = CodeName.const thy c_kbind';
+ val pr = pretty_haskell_monad c_mbind'' c_kbind''
+ in
+ thy
+ |> gen_add_syntax_const (K I) target_Haskell c_run' (SOME pr)
+ |> gen_add_syntax_const (K I) target_Haskell c_mbind'
+ (no_bindings (SOME (parse_infix fst (L, 1) ">>=")))
+ |> gen_add_syntax_const (K I) target_Haskell c_kbind'
+ (no_bindings (SOME (parse_infix fst (L, 1) ">>")))
+ end;
+
+fun add_reserved target =
+ let
+ fun add sym syms = if member (op =) syms sym
+ then error ("Reserved symbol " ^ quote sym ^ " already declared")
+ else insert (op =) sym syms
+ in map_reserveds target o add end;
+
+fun add_modl_alias target =
+ map_syntax_modls target o apfst o Symtab.update o apsnd CodeName.check_modulename;
+
+fun add_modl_prolog target =
+ map_syntax_modls target o apsnd o
+ (fn (modl, NONE) => Symtab.delete modl | (modl, SOME prolog) =>
+ Symtab.update (modl, Pretty.str prolog));
+
+fun zip_list (x::xs) f g =
+ f
+ #-> (fn y =>
+ fold_map (fn x => g |-- f >> pair x) xs
+ #-> (fn xys => pair ((x, y) :: xys)));
+
+structure P = OuterParse
+and K = OuterKeyword
+
+fun parse_multi_syntax parse_thing parse_syntax =
+ P.and_list1 parse_thing
+ #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
+ (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")"));
+
+val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr");
+
+fun parse_syntax prep_arg xs =
+ Scan.option ((
+ ((P.$$$ infixK >> K X)
+ || (P.$$$ infixlK >> K L)
+ || (P.$$$ infixrK >> K R))
+ -- P.nat >> parse_infix prep_arg
+ || Scan.succeed (parse_mixfix prep_arg))
+ -- P.string
+ >> (fn (parse, s) => parse s)) xs;
+
+val (code_classK, code_instanceK, code_typeK, code_constK, code_monadK,
+ code_reservedK, code_modulenameK, code_moduleprologK) =
+ ("code_class", "code_instance", "code_type", "code_const", "code_monad",
+ "code_reserved", "code_modulename", "code_moduleprolog");
+
+in
+
+val parse_syntax = parse_syntax;
+
+val add_syntax_class = gen_add_syntax_class cert_class (K I);
+val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco;
+val add_syntax_tyco = gen_add_syntax_tyco cert_tyco;
+val add_syntax_const = gen_add_syntax_const (K I);
+
+val add_syntax_class_cmd = gen_add_syntax_class read_class CodeUnit.read_const;
+val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
+val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
+val add_syntax_const_cmd = gen_add_syntax_const CodeUnit.read_const;
+
+fun add_syntax_tycoP target tyco = parse_syntax I >> add_syntax_tyco_cmd target tyco;
+fun add_syntax_constP target c = parse_syntax fst >> (add_syntax_const_cmd target c o no_bindings);
+
+fun add_undefined target undef target_undefined thy =
+ let
+ val (undef', _) = idfs_of_const thy undef;
+ fun pr _ _ _ _ = str target_undefined;
+ in
+ thy
+ |> add_syntax_const target undef' (SOME (~1, pr))
+ end;
+
+fun add_pretty_list target nill cons thy =
+ let
+ val (_, nil'') = idfs_of_const thy nill;
+ val (cons', cons'') = idfs_of_const thy cons;
+ val pr = pretty_list nil'' cons'' target;
+ in
+ thy
+ |> add_syntax_const target cons' (SOME pr)
+ end;
+
+fun add_pretty_list_string target nill cons charr nibbles thy =
+ let
+ val (_, nil'') = idfs_of_const thy nill;
+ val (cons', cons'') = idfs_of_const thy cons;
+ val (_, charr'') = idfs_of_const thy charr;
+ val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
+ val pr = pretty_list_string nil'' cons'' charr'' nibbles'' target;
+ in
+ thy
+ |> add_syntax_const target cons' (SOME pr)
+ end;
+
+fun add_pretty_char target charr nibbles thy =
+ let
+ val (charr', charr'') = idfs_of_const thy charr;
+ val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
+ val pr = pretty_char charr'' nibbles'' target;
+ in
+ thy
+ |> add_syntax_const target charr' (SOME pr)
+ end;
+
+fun add_pretty_numeral target unbounded number_of b0 b1 pls min bit thy =
+ let
+ val number_of' = CodeUnit.const_of_cexpr thy number_of;
+ val (_, b0'') = idfs_of_const thy b0;
+ val (_, b1'') = idfs_of_const thy b1;
+ val (_, pls'') = idfs_of_const thy pls;
+ val (_, min'') = idfs_of_const thy min;
+ val (_, bit'') = idfs_of_const thy bit;
+ val pr = pretty_numeral unbounded b0'' b1'' pls'' min'' bit'' target;
+ in
+ thy
+ |> add_syntax_const target number_of' (SOME pr)
+ end;
+
+fun add_pretty_ml_string target charr nibbles nill cons str thy =
+ let
+ val (_, charr'') = idfs_of_const thy charr;
+ val (_, nibbles'') = split_list (map (idfs_of_const thy) nibbles);
+ val (_, nil'') = idfs_of_const thy nill;
+ val (_, cons'') = idfs_of_const thy cons;
+ val (str', _) = idfs_of_const thy str;
+ val pr = pretty_ml_string charr'' nibbles'' nil'' cons'' target;
+ in
+ thy
+ |> add_syntax_const target str' (SOME pr)
+ end;
+
+fun add_pretty_imperative_monad_bind target bind thy =
+ let
+ val (bind', _) = idfs_of_const thy bind;
+ val pr = pretty_imperative_monad_bind
+ in
+ thy
+ |> add_syntax_const target bind' (SOME pr)
+ end;
+
+val add_haskell_monad = gen_add_haskell_monad CodeUnit.read_const;
+
+val code_classP =
+ OuterSyntax.command code_classK "define code syntax for class" K.thy_decl (
+ parse_multi_syntax P.xname
+ (Scan.option (P.string -- Scan.optional (P.$$$ "where" |-- Scan.repeat1
+ (P.term --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.string)) []))
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns)
+ );
+
+val code_instanceP =
+ OuterSyntax.command code_instanceK "define code syntax for instance" K.thy_decl (
+ parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname)
+ ((P.minus >> K true) || Scan.succeed false)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns)
+ );
+
+val code_typeP =
+ OuterSyntax.command code_typeK "define code syntax for type constructor" K.thy_decl (
+ parse_multi_syntax P.xname (parse_syntax I)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns)
+ );
+
+val code_constP =
+ OuterSyntax.command code_constK "define code syntax for constant" K.thy_decl (
+ parse_multi_syntax P.term (parse_syntax fst)
+ >> (Toplevel.theory oo fold) (fn (target, syns) =>
+ fold (fn (raw_const, syn) => add_syntax_const_cmd target raw_const (no_bindings syn)) syns)
+ );
+
+val code_monadP =
+ OuterSyntax.command code_monadK "define code syntax for Haskell monads" K.thy_decl (
+ P.term -- P.term -- P.term
+ >> (fn ((raw_run, raw_mbind), raw_kbind) => Toplevel.theory
+ (add_haskell_monad raw_run raw_mbind raw_kbind))
+ );
+
+val code_reservedP =
+ OuterSyntax.command code_reservedK "declare words as reserved for target language" K.thy_decl (
+ P.name -- Scan.repeat1 P.name
+ >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)
+ )
+
+val code_modulenameP =
+ OuterSyntax.command code_modulenameK "alias module to other name" K.thy_decl (
+ P.name -- Scan.repeat1 (P.name -- P.name)
+ >> (fn (target, modlnames) => (Toplevel.theory o fold (add_modl_alias target)) modlnames)
+ )
+
+val code_moduleprologP =
+ OuterSyntax.command code_moduleprologK "add prolog to module" K.thy_decl (
+ P.name -- Scan.repeat1 (P.name -- (P.text >> (fn "-" => NONE | s => SOME s)))
+ >> (fn (target, prologs) => (Toplevel.theory o fold (add_modl_prolog target)) prologs)
+ )
+
+val _ = OuterSyntax.add_keywords [infixK, infixlK, infixrK];
+
+val _ = OuterSyntax.add_parsers [code_classP, code_instanceP, code_typeP, code_constP,
+ code_reservedP, code_modulenameP, code_moduleprologP, code_monadP];
+
+
+(*including serializer defaults*)
+val setup =
+ add_serializer (target_SML, isar_seri_sml)
+ #> add_serializer (target_OCaml, isar_seri_ocaml)
+ #> add_serializer (target_Haskell, isar_seri_haskell)
+ #> add_serializer (target_diag, fn _ => fn _ => fn _ => seri_diagnosis)
+ #> add_syntax_tyco "SML" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ (gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> add_syntax_tyco "OCaml" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ (gen_brackify (case fxy of NOBR => false | _ => eval_fxy (INFX (1, R)) fxy) o Pretty.breaks) [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> add_syntax_tyco "Haskell" "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy [
+ pr_typ (INFX (1, X)) ty1,
+ str "->",
+ pr_typ (INFX (1, R)) ty2
+ ]))
+ #> fold (add_reserved "SML") ML_Syntax.reserved_names
+ #> fold (add_reserved "SML")
+ ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)]
+ #> fold (add_reserved "OCaml") [
+ "and", "as", "assert", "begin", "class",
+ "constraint", "do", "done", "downto", "else", "end", "exception",
+ "external", "false", "for", "fun", "function", "functor", "if",
+ "in", "include", "inherit", "initializer", "lazy", "let", "match", "method",
+ "module", "mutable", "new", "object", "of", "open", "or", "private", "rec",
+ "sig", "struct", "then", "to", "true", "try", "type", "val",
+ "virtual", "when", "while", "with"
+ ]
+ #> fold (add_reserved "OCaml") ["failwith", "mod"]
+ #> fold (add_reserved "Haskell") [
+ "hiding", "deriving", "where", "case", "of", "infix", "infixl", "infixr",
+ "import", "default", "forall", "let", "in", "class", "qualified", "data",
+ "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
+ ]
+ #> fold (add_reserved "Haskell") [
+ "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int",
+ "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded",
+ "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor",
+ "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException",
+ "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException",
+ "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure",
+ "ExitSuccess", "False", "GT", "HeapOverflow",
+ "IOError", "IOException", "IllegalOperation",
+ "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError",
+ "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow",
+ "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError",
+ "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow",
+ "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow",
+ "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all",
+ "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan",
+ "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen",
+ "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break",
+ "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const",
+ "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod",
+ "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem",
+ "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo",
+ "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip",
+ "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational",
+ "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble",
+ "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj",
+ "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head",
+ "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha",
+ "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite",
+ "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'",
+ "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log",
+ "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum",
+ "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem",
+ "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo",
+ "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred",
+ "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn",
+ "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat",
+ "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile",
+ "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn",
+ "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational",
+ "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse",
+ "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence",
+ "sequence_", "show", "showChar", "showException", "showField", "showList",
+ "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand",
+ "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract",
+ "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult",
+ "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry",
+ "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords",
+ "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3"
+ ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*);
+
+end; (*local*)
+
+end; (*struct*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/code/code_thingol.ML Fri Aug 10 17:04:34 2007 +0200
@@ -0,0 +1,433 @@
+(* Title: Tools/code/code_thingol.ML
+ ID: $Id$
+ Author: Florian Haftmann, TU Muenchen
+
+Intermediate language ("Thin-gol") representing executable code.
+*)
+
+infix 8 `%%;
+infixr 6 `->;
+infixr 6 `-->;
+infix 4 `$;
+infix 4 `$$;
+infixr 3 `|->;
+infixr 3 `|-->;
+
+signature BASIC_CODE_THINGOL =
+sig
+ type vname = string;
+ datatype dict =
+ DictConst of string * dict list list
+ | DictVar of string list * (vname * (int * int));
+ datatype itype =
+ `%% of string * itype list
+ | ITyVar of vname;
+ datatype iterm =
+ IConst of string * (dict list list * itype list (*types of arguments*))
+ | IVar of vname
+ | `$ of iterm * iterm
+ | `|-> of (vname * itype) * iterm
+ | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
+ (*((term, type), [(selector pattern, body term )]), primitive term)*)
+ val `-> : itype * itype -> itype;
+ val `--> : itype list * itype -> itype;
+ val `$$ : iterm * iterm list -> iterm;
+ val `|--> : (vname * itype) list * iterm -> iterm;
+ type typscheme = (vname * sort) list * itype;
+end;
+
+signature CODE_THINGOL =
+sig
+ include BASIC_CODE_THINGOL;
+ val unfoldl: ('a -> ('a * 'b) option) -> 'a -> 'a * 'b list;
+ val unfoldr: ('a -> ('b * 'a) option) -> 'a -> 'b list * 'a;
+ val unfold_fun: itype -> itype list * itype;
+ val unfold_app: iterm -> iterm * iterm list;
+ val split_abs: iterm -> (((vname * iterm option) * itype) * iterm) option;
+ val unfold_abs: iterm -> ((vname * iterm option) * itype) list * iterm;
+ val split_let: iterm -> (((iterm * itype) * iterm) * iterm) option;
+ val unfold_let: iterm -> ((iterm * itype) * iterm) list * iterm;
+ val unfold_const_app: iterm ->
+ ((string * (dict list list * itype list)) * iterm list) option;
+ val collapse_let: ((vname * itype) * iterm) * iterm
+ -> (iterm * itype) * (iterm * iterm) list;
+ val eta_expand: (string * (dict list list * itype list)) * iterm list -> int -> iterm;
+ val fold_constnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
+ val fold_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
+ val fold_unbound_varnames: (string -> 'a -> 'a) -> iterm -> 'a -> 'a;
+
+ datatype def =
+ Bot
+ | Fun of (iterm list * iterm) list * typscheme
+ | Datatype of (vname * sort) list * (string * itype list) list
+ | Datatypecons of string
+ | Class of (class * string) list * (vname * (string * itype) list)
+ | Classop of class
+ | Classrel of class * class
+ | Classinst of (class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * (string * iterm) list);
+ type code = def Graph.T;
+ type transact;
+ val empty_code: code;
+ val merge_code: code * code -> code;
+ val project_code: bool (*delete empty funs*)
+ -> string list (*hidden*) -> string list option (*selected*)
+ -> code -> code;
+ val empty_funs: code -> string list;
+ val is_cons: code -> string -> bool;
+ val add_eval_def: string (*bind name*) * iterm -> code -> code;
+
+ val ensure_def: (string -> string) -> (transact -> def * code) -> string
+ -> string -> transact -> transact;
+ val succeed: 'a -> transact -> 'a * code;
+ val fail: string -> transact -> 'a * code;
+ val message: string -> (transact -> 'a) -> transact -> 'a;
+ val start_transact: (transact -> 'a * transact) -> code -> 'a * code;
+end;
+
+structure CodeThingol: CODE_THINGOL =
+struct
+
+(** auxiliary **)
+
+fun unfoldl dest x =
+ case dest x
+ of NONE => (x, [])
+ | SOME (x1, x2) =>
+ let val (x', xs') = unfoldl dest x1 in (x', xs' @ [x2]) end;
+
+fun unfoldr dest x =
+ case dest x
+ of NONE => ([], x)
+ | SOME (x1, x2) =>
+ let val (xs', x') = unfoldr dest x2 in (x1::xs', x') end;
+
+
+(** language core - types, pattern, expressions **)
+
+(* language representation *)
+
+type vname = string;
+
+datatype dict =
+ DictConst of string * dict list list
+ | DictVar of string list * (vname * (int * int));
+
+datatype itype =
+ `%% of string * itype list
+ | ITyVar of vname;
+
+datatype iterm =
+ IConst of string * (dict list list * itype list)
+ | IVar of vname
+ | `$ of iterm * iterm
+ | `|-> of (vname * itype) * iterm
+ | ICase of ((iterm * itype) * (iterm * iterm) list) * iterm;
+ (*see also signature*)
+
+(*
+ variable naming conventions
+
+ bare names:
+ variable names v
+ class names class
+ type constructor names tyco
+ datatype names dtco
+ const names (general) c
+ constructor names co
+ class operation names clsop (op)
+ arbitrary name s
+
+ v, c, co, clsop also annotated with types etc.
+
+ constructs:
+ sort sort
+ type parameters vs
+ type ty
+ type schemes tysm
+ term t
+ (term as pattern) p
+ instance (class, tyco) inst
+ *)
+
+fun ty1 `-> ty2 = "fun" `%% [ty1, ty2];
+val op `--> = Library.foldr (op `->);
+val op `$$ = Library.foldl (op `$);
+val op `|--> = Library.foldr (op `|->);
+
+val unfold_fun = unfoldr
+ (fn "fun" `%% [ty1, ty2] => SOME (ty1, ty2)
+ | _ => NONE);
+
+val unfold_app = unfoldl
+ (fn op `$ t => SOME t
+ | _ => NONE);
+
+val split_abs =
+ (fn (v, ty) `|-> (t as ICase (((IVar w, _), [(p, t')]), _)) =>
+ if v = w then SOME (((v, SOME p), ty), t') else SOME (((v, NONE), ty), t)
+ | (v, ty) `|-> t => SOME (((v, NONE), ty), t)
+ | _ => NONE);
+
+val unfold_abs = unfoldr split_abs;
+
+val split_let =
+ (fn ICase (((td, ty), [(p, t)]), _) => SOME (((p, ty), td), t)
+ | _ => NONE);
+
+val unfold_let = unfoldr split_let;
+
+fun unfold_const_app t =
+ case unfold_app t
+ of (IConst c, ts) => SOME (c, ts)
+ | _ => NONE;
+
+fun fold_aiterms f (t as IConst _) = f t
+ | fold_aiterms f (t as IVar _) = f t
+ | fold_aiterms f (t1 `$ t2) = fold_aiterms f t1 #> fold_aiterms f t2
+ | fold_aiterms f (t as _ `|-> t') = f t #> fold_aiterms f t'
+ | fold_aiterms f (ICase (_, t)) = fold_aiterms f t;
+
+fun fold_constnames f =
+ let
+ fun add (IConst (c, _)) = f c
+ | add _ = I;
+ in fold_aiterms add end;
+
+fun fold_varnames f =
+ let
+ fun add (IVar v) = f v
+ | add ((v, _) `|-> _) = f v
+ | add _ = I;
+ in fold_aiterms add end;
+
+fun fold_unbound_varnames f =
+ let
+ fun add _ (IConst _) = I
+ | add vs (IVar v) = if not (member (op =) vs v) then f v else I
+ | add vs (t1 `$ t2) = add vs t1 #> add vs t2
+ | add vs ((v, _) `|-> t) = add (insert (op =) v vs) t
+ | add vs (ICase (_, t)) = add vs t;
+ in add [] end;
+
+fun collapse_let (((v, ty), se), be as ICase (((IVar w, _), ds), _)) =
+ let
+ fun exists_v t = fold_unbound_varnames (fn w => fn b =>
+ b orelse v = w) t false;
+ in if v = w andalso forall (fn (t1, t2) =>
+ exists_v t1 orelse not (exists_v t2)) ds
+ then ((se, ty), ds)
+ else ((se, ty), [(IVar v, be)])
+ end
+ | collapse_let (((v, ty), se), be) =
+ ((se, ty), [(IVar v, be)])
+
+fun eta_expand (c as (_, (_, tys)), ts) k =
+ let
+ val j = length ts;
+ val l = k - j;
+ val ctxt = (fold o fold_varnames) Name.declare ts Name.context;
+ val vs_tys = Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys);
+ in vs_tys `|--> IConst c `$$ ts @ map (fn (v, _) => IVar v) vs_tys end;
+
+
+(** definitions, transactions **)
+
+(* type definitions *)
+
+type typscheme = (vname * sort) list * itype;
+datatype def =
+ Bot
+ | Fun of (iterm list * iterm) list * typscheme
+ | Datatype of (vname * sort) list * (string * itype list) list
+ | Datatypecons of string
+ | Class of (class * string) list * (vname * (string * itype) list)
+ | Classop of class
+ | Classrel of class * class
+ | Classinst of (class * (string * (vname * sort) list))
+ * ((class * (string * (string * dict list list))) list
+ * (string * iterm) list);
+
+type code = def Graph.T;
+type transact = Graph.key option * code;
+exception FAIL of string list;
+
+
+(* abstract code *)
+
+val empty_code = Graph.empty : code; (*read: "depends on"*)
+
+fun ensure_bot name = Graph.default_node (name, Bot);
+
+fun add_def_incr (name, Bot) code =
+ (case the_default Bot (try (Graph.get_node code) name)
+ of Bot => error "Attempted to add Bot to code"
+ | _ => code)
+ | add_def_incr (name, def) code =
+ (case try (Graph.get_node code) name
+ of NONE => Graph.new_node (name, def) code
+ | SOME Bot => Graph.map_node name (K def) code
+ | SOME _ => error ("Tried to overwrite definition " ^ quote name));
+
+fun add_dep (dep as (name1, name2)) =
+ if name1 = name2 then I else Graph.add_edge dep;
+
+val merge_code : code * code -> code = Graph.merge (K true);
+
+fun project_code delete_empty_funs hidden raw_selected code =
+ let
+ fun is_empty_fun name = case Graph.get_node code name
+ of Fun ([], _) => true
+ | _ => false;
+ val names = subtract (op =) hidden (Graph.keys code);
+ val deleted = Graph.all_preds code (filter is_empty_fun names);
+ val selected = case raw_selected
+ of NONE => names |> subtract (op =) deleted
+ | SOME sel => sel
+ |> delete_empty_funs ? subtract (op =) deleted
+ |> subtract (op =) hidden
+ |> Graph.all_succs code
+ |> delete_empty_funs ? subtract (op =) deleted
+ |> subtract (op =) hidden;
+ in
+ code
+ |> Graph.subgraph (member (op =) selected)
+ end;
+
+fun empty_funs code =
+ Graph.fold (fn (name, (Fun ([], _), _)) => cons name
+ | _ => I) code [];
+
+fun is_cons code name = case Graph.get_node code name
+ of Datatypecons _ => true
+ | _ => false;
+
+fun check_samemodule names =
+ fold (fn name =>
+ let
+ val module_name = (NameSpace.qualifier o NameSpace.qualifier) name
+ in
+ fn NONE => SOME module_name
+ | SOME module_name' => if module_name = module_name' then SOME module_name
+ else error ("Inconsistent name prefix for simultanous names: " ^ commas_quote names)
+ end
+ ) names NONE;
+
+fun check_funeqs eqs =
+ (fold (fn (pats, _) =>
+ let
+ val l = length pats
+ in
+ fn NONE => SOME l
+ | SOME l' => if l = l' then SOME l
+ else error "Function definition with different number of arguments"
+ end
+ ) eqs NONE; eqs);
+
+fun check_prep_def code Bot =
+ Bot
+ | check_prep_def code (Fun (eqs, d)) =
+ Fun (check_funeqs eqs, d)
+ | check_prep_def code (d as Datatype _) =
+ d
+ | check_prep_def code (Datatypecons dtco) =
+ error "Attempted to add bare term constructor"
+ | check_prep_def code (d as Class _) =
+ d
+ | check_prep_def code (Classop _) =
+ error "Attempted to add bare class operation"
+ | check_prep_def code (Classrel _) =
+ error "Attempted to add bare class relation"
+ | check_prep_def code (d as Classinst ((class, (tyco, arity)), (_, inst_classops))) =
+ let
+ val Class (_, (_, classops)) = Graph.get_node code class;
+ val _ = if length inst_classops > length classops
+ then error "Too many class operations given"
+ else ();
+ fun check_classop (f, _) =
+ if AList.defined (op =) inst_classops f
+ then () else error ("Missing definition for class operation " ^ quote f);
+ val _ = map check_classop classops;
+ in d end;
+
+fun postprocess_def (name, Datatype (_, constrs)) =
+ tap (fn _ => check_samemodule (name :: map fst constrs))
+ #> fold (fn (co, _) =>
+ add_def_incr (co, Datatypecons name)
+ #> add_dep (co, name)
+ #> add_dep (name, co)
+ ) constrs
+ | postprocess_def (name, Class (classrels, (_, classops))) =
+ tap (fn _ => check_samemodule (name :: map fst classops @ map snd classrels))
+ #> fold (fn (f, _) =>
+ add_def_incr (f, Classop name)
+ #> add_dep (f, name)
+ #> add_dep (name, f)
+ ) classops
+ #> fold (fn (superclass, classrel) =>
+ add_def_incr (classrel, Classrel (name, superclass))
+ #> add_dep (classrel, name)
+ #> add_dep (name, classrel)
+ ) classrels
+ | postprocess_def _ =
+ I;
+
+
+(* transaction protocol *)
+
+fun ensure_def labelled_name defgen msg name (dep, code) =
+ let
+ val msg' = (case dep
+ of NONE => msg
+ | SOME dep => msg ^ ", required for " ^ labelled_name dep);
+ fun add_dp NONE = I
+ | add_dp (SOME dep) = add_dep (dep, name);
+ fun prep_def def code =
+ (check_prep_def code def, code);
+ fun invoke_generator name defgen code =
+ defgen (SOME name, code) handle FAIL msgs => raise FAIL (msg' :: msgs);
+ fun add_def false =
+ ensure_bot name
+ #> add_dp dep
+ #> invoke_generator name defgen
+ #-> (fn def => prep_def def)
+ #-> (fn def => add_def_incr (name, def)
+ #> postprocess_def (name, def))
+ | add_def true =
+ add_dp dep;
+ in
+ code
+ |> add_def (can (Graph.get_node code) name)
+ |> pair dep
+ end;
+
+fun succeed some (_, code) = (some, code);
+
+fun fail msg (_, code) = raise FAIL [msg];
+
+fun message msg f trns =
+ f trns handle FAIL msgs =>
+ raise FAIL (msg :: msgs);
+
+fun start_transact f code =
+ let
+ fun handle_fail f x =
+ (f x
+ handle FAIL msgs =>
+ (error o cat_lines) ("Code generation failed, while:" :: msgs))
+ in
+ (NONE, code)
+ |> handle_fail f
+ |-> (fn x => fn (_, code) => (x, code))
+ end;
+
+fun add_eval_def (name, t) code =
+ code
+ |> Graph.new_node (name, Fun ([([], t)], ([("_", [])], ITyVar "_")))
+ |> fold (curry Graph.add_edge name) (Graph.keys code);
+
+end; (*struct*)
+
+
+structure BasicCodeThingol: BASIC_CODE_THINGOL = CodeThingol;
--- a/src/Tools/nbe.ML Fri Aug 10 17:04:24 2007 +0200
+++ b/src/Tools/nbe.ML Fri Aug 10 17:04:34 2007 +0200
@@ -25,7 +25,7 @@
val app: Univ -> Univ -> Univ (*explicit application*)
val univs_ref: Univ list ref
- val lookup_fun: CodegenNames.const -> Univ
+ val lookup_fun: CodeName.const -> Univ
val normalization_conv: cterm -> thm
@@ -183,7 +183,7 @@
end;
-open BasicCodegenThingol;
+open BasicCodeThingol;
(* greetings to Tarski *)
@@ -191,7 +191,7 @@
let
fun of_iterm t =
let
- val (t', ts) = CodegenThingol.unfold_app t
+ val (t', ts) = CodeThingol.unfold_app t
in of_itermapp t' (fold (cons o of_iterm) ts []) end
and of_itermapp (IConst (c, (dss, _))) ts =
(case num_args c
@@ -225,7 +225,7 @@
let
val cs = map fst eqnss;
val num_args = cs ~~ map (fn (_, (args, rhs) :: _) => length args) eqnss;
- val funs = fold (fold (CodegenThingol.fold_constnames
+ val funs = fold (fold (CodeThingol.fold_constnames
(insert (op =))) o map snd o snd) eqnss [];
val bind_funs = map nbe_lookup (filter is_fun funs);
val bind_locals = ml_fundefs (map nbe_fun cs ~~ map
@@ -235,29 +235,29 @@
fun assemble_eval thy is_fun t =
let
- val funs = CodegenThingol.fold_constnames (insert (op =)) t [];
- val frees = CodegenThingol.fold_unbound_varnames (insert (op =)) t [];
+ val funs = CodeThingol.fold_constnames (insert (op =)) t [];
+ val frees = CodeThingol.fold_unbound_varnames (insert (op =)) t [];
val bind_funs = map nbe_lookup (filter is_fun funs);
val bind_value = ml_fundefs [(nbe_value, [([ml_list (map nbe_bound frees)],
assemble_iterm thy is_fun (K NONE) t)])];
val result = ml_list [nbe_value `$` ml_list (map nbe_free frees)];
in ([nbe_value], ml_Let (bind_funs @ [bind_value]) result) end;
-fun eqns_of_stmt (name, CodegenThingol.Fun ([], _)) =
+fun eqns_of_stmt (name, CodeThingol.Fun ([], _)) =
NONE
- | eqns_of_stmt (name, CodegenThingol.Fun (eqns, _)) =
+ | eqns_of_stmt (name, CodeThingol.Fun (eqns, _)) =
SOME (name, eqns)
- | eqns_of_stmt (_, CodegenThingol.Datatypecons _) =
+ | eqns_of_stmt (_, CodeThingol.Datatypecons _) =
NONE
- | eqns_of_stmt (_, CodegenThingol.Datatype _) =
+ | eqns_of_stmt (_, CodeThingol.Datatype _) =
NONE
- | eqns_of_stmt (_, CodegenThingol.Class _) =
+ | eqns_of_stmt (_, CodeThingol.Class _) =
NONE
- | eqns_of_stmt (_, CodegenThingol.Classrel _) =
+ | eqns_of_stmt (_, CodeThingol.Classrel _) =
NONE
- | eqns_of_stmt (_, CodegenThingol.Classop _) =
+ | eqns_of_stmt (_, CodeThingol.Classop _) =
NONE
- | eqns_of_stmt (_, CodegenThingol.Classinst _) =
+ | eqns_of_stmt (_, CodeThingol.Classinst _) =
NONE;
fun compile_stmts thy is_fun =
@@ -297,8 +297,8 @@
#>> (fn ts' => list_comb (t, rev ts'))
and of_univ bounds (Const (name, ts)) typidx =
let
- val SOME (const as (c, _)) = CodegenNames.const_rev thy name;
- val T = CodegenData.default_typ thy const;
+ val SOME (const as (c, _)) = CodeName.const_rev thy name;
+ val T = Code.default_typ thy const;
val T' = map_type_tvar (fn ((v, i), S) => TypeInfer.param (typidx + i) (v, S)) T;
val typidx' = typidx + maxidx_of_typ T' + 1;
in of_apps bounds (Term.Const (c, T'), ts) typidx' end
@@ -345,7 +345,7 @@
(* evaluation oracle *)
-exception Normalization of CodegenThingol.code * term * CodegenThingol.iterm;
+exception Normalization of CodeThingol.code * term * CodeThingol.iterm;
fun normalization_oracle (thy, Normalization (code, t, t')) =
Logic.mk_equals (t, eval thy code t t');
@@ -361,7 +361,7 @@
let
val t = Thm.term_of ct;
in normalization_invoke thy code t t' end;
- in CodegenPackage.eval_conv thy conv ct end;
+ in CodePackage.eval_conv thy conv ct end;
(* evaluation command *)