# HG changeset patch # User haftmann # Date 1260199668 -3600 # Node ID 1e620676303635d2acf9c26e40e0bdb504a68735 # Parent 7c2c38a5bca303f77b57b914c9da5ef95535488b split off evaluation mechanisms in separte module Code_Eval diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/Code_Evaluation.thy --- a/src/HOL/Code_Evaluation.thy Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/Code_Evaluation.thy Mon Dec 07 16:27:48 2009 +0100 @@ -279,7 +279,7 @@ val eval_ref = Unsynchronized.ref (NONE : (unit -> term) option); fun eval_term thy t = - Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) []; + Code_Eval.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) []; end; *} diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/HOL.thy Mon Dec 07 16:27:48 2009 +0100 @@ -1971,7 +1971,7 @@ val t = Thm.term_of ct; val dummy = @{cprop True}; in case try HOLogic.dest_Trueprop t - of SOME t' => if Code_ML.eval NONE + of SOME t' => if Code_Eval.eval NONE ("Eval_Method.eval_ref", Eval_Method.eval_ref) (K I) thy t' [] then Thm.capply (Thm.capply @{cterm "op \ \ prop \ prop \ prop"} ct) dummy else dummy diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/IsaMakefile Mon Dec 07 16:27:48 2009 +0100 @@ -103,6 +103,7 @@ $(SRC)/Provers/hypsubst.ML \ $(SRC)/Provers/quantifier1.ML \ $(SRC)/Provers/splitter.ML \ + $(SRC)/Tools/Code/code_eval.ML \ $(SRC)/Tools/Code/code_haskell.ML \ $(SRC)/Tools/Code/code_ml.ML \ $(SRC)/Tools/Code/code_preproc.ML \ diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/Library/Eval_Witness.thy --- a/src/HOL/Library/Eval_Witness.thy Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/Library/Eval_Witness.thy Mon Dec 07 16:27:48 2009 +0100 @@ -68,7 +68,7 @@ | dest_exs _ _ = sys_error "dest_exs"; val t = dest_exs (length ws) (HOLogic.dest_Trueprop goal); in - if Code_ML.eval NONE ("Eval_Witness_Method.eval_ref", Eval_Witness_Method.eval_ref) (K I) thy t ws + if Code_Eval.eval NONE ("Eval_Witness_Method.eval_ref", Eval_Witness_Method.eval_ref) (K I) thy t ws then Thm.cterm_of thy goal else @{cprop True} (*dummy*) end diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Dec 07 16:27:48 2009 +0100 @@ -2586,11 +2586,11 @@ val t' = mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t; val eval = if random then - Code_ML.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref) + Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref) (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' [] |> Random_Engine.run else - Code_ML.eval NONE ("Predicate_Compile_Core.eval_ref", eval_ref) Predicate.map thy t' [] + Code_Eval.eval NONE ("Predicate_Compile_Core.eval_ref", eval_ref) Predicate.map thy t' [] in (T, eval) end; fun values ctxt param_user_modes options k t_compr = diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Mon Dec 07 16:27:48 2009 +0100 @@ -106,7 +106,7 @@ mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term} (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))) val _ = tracing (Syntax.string_of_term ctxt' qc_term) - val compile = Code_ML.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref) + val compile = Code_Eval.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref) (fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc) thy'' qc_term [] in diff -r 7c2c38a5bca3 -r 1e6206763036 src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Mon Dec 07 14:54:28 2009 +0100 +++ b/src/HOL/Tools/quickcheck_generators.ML Mon Dec 07 16:27:48 2009 +0100 @@ -379,7 +379,7 @@ let val Ts = (map snd o fst o strip_abs) t; val t' = mk_generator_expr thy t Ts; - val compile = Code_ML.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref) + val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref) (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' []; in compile #> Random_Engine.run end; @@ -388,7 +388,7 @@ val setup = Typecopy.interpretation ensure_random_typecopy #> Datatype.interpretation ensure_random_datatype - #> Code_Target.extend_target (target, (Code_ML.target_Eval, K I)) + #> Code_Target.extend_target (target, (Code_Eval.target, K I)) #> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of); end; diff -r 7c2c38a5bca3 -r 1e6206763036 src/Tools/Code/code_eval.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Code/code_eval.ML Mon Dec 07 16:27:48 2009 +0100 @@ -0,0 +1,153 @@ +(* Title: Tools/code/code_eval.ML_ + Author: Florian Haftmann, TU Muenchen + +Runtime services building on code generation into implementation language SML. +*) + +signature CODE_EVAL = +sig + val target: string + val eval: string option -> string * (unit -> 'a) option Unsynchronized.ref + -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a + val evaluation_code: theory -> string list -> string list + -> string * ((string * string) list * (string * string) list) + val setup: theory -> theory +end; + +structure Code_Eval : CODE_EVAL = +struct + +(** generic **) + +val target = "Eval"; + +val eval_struct_name = "Code" + +fun evaluation_code thy tycos consts = + let + val (consts', (naming, program)) = Code_Thingol.consts_program thy consts; + val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; + val (ml_code, target_names) = Code_ML.evaluation_code_of thy target + (SOME eval_struct_name) naming program (consts' @ tycos'); + val (consts'', tycos'') = chop (length consts') target_names; + val consts_map = map2 (fn const => fn NONE => + error ("Constant " ^ (quote o Code.string_of_const thy) const + ^ "\nhas a user-defined serialization") + | SOME const'' => (const, const'')) consts consts'' + val tycos_map = map2 (fn tyco => fn NONE => + error ("Type " ^ (quote o Sign.extern_type thy) tyco + ^ "\nhas a user-defined serialization") + | SOME tyco'' => (tyco, tyco'')) tycos tycos''; + in (ml_code, (tycos_map, consts_map)) end; + + +(** evaluation **) + +fun eval some_target reff postproc thy t args = + let + val ctxt = ProofContext.init thy; + fun evaluator naming program ((_, (_, ty)), t) deps = + let + val _ = if Code_Thingol.contains_dictvar t then + error "Term to be evaluated contains free dictionaries" else (); + val value_name = "Value.VALUE.value" + val program' = program + |> Graph.new_node (value_name, + Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))]))) + |> fold (curry Graph.add_edge value_name) deps; + val (value_code, [SOME value_name']) = Code_ML.evaluation_code_of thy + (the_default target some_target) NONE naming program' [value_name]; + val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' + ^ space_implode " " (map (enclose "(" ")") args) ^ " end"; + in ML_Context.evaluate ctxt false reff sml_code end; + in Code_Thingol.eval thy postproc evaluator t end; + + +(** instrumentalization by antiquotation **) + +local + +structure CodeAntiqData = Proof_Data +( + type T = (string list * string list) * (bool * (string + * (string * ((string * string) list * (string * string) list)) lazy)); + fun init _ = (([], []), (true, ("", Lazy.value ("", ([], []))))); +); + +val is_first_occ = fst o snd o CodeAntiqData.get; + +fun register_code new_tycos new_consts ctxt = + let + val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt; + val tycos' = fold (insert (op =)) new_tycos tycos; + val consts' = fold (insert (op =)) new_consts consts; + val (struct_name', ctxt') = if struct_name = "" + then ML_Antiquote.variant eval_struct_name ctxt + else (struct_name, ctxt); + val acc_code = Lazy.lazy (fn () => evaluation_code (ProofContext.theory_of ctxt) tycos' consts'); + in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end; + +fun register_const const = register_code [] [const]; + +fun register_datatype tyco constrs = register_code [tyco] constrs; + +fun print_const const all_struct_name tycos_map consts_map = + (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; + +fun print_datatype tyco constrs all_struct_name tycos_map consts_map = + let + val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode; + fun check_base name name'' = + if upperize (Long_Name.base_name name) = upperize name'' + then () else error ("Name as printed " ^ quote name'' + ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry."); + val tyco'' = (the o AList.lookup (op =) tycos_map) tyco; + val constrs'' = map (the o AList.lookup (op =) consts_map) constrs; + val _ = check_base tyco tyco''; + val _ = map2 check_base constrs constrs''; + in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end; + +fun print_code struct_name is_first print_it ctxt = + let + val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; + val (ml_code, (tycos_map, consts_map)) = Lazy.force acc_code; + val ml_code = if is_first then ml_code + else ""; + val all_struct_name = Long_Name.append struct_name struct_code_name; + in (ml_code, print_it all_struct_name tycos_map consts_map) end; + +in + +fun ml_code_antiq raw_const {struct_name, background} = + let + val const = Code.check_const (ProofContext.theory_of background) raw_const; + val is_first = is_first_occ background; + val background' = register_const const background; + in (print_code struct_name is_first (print_const const), background') end; + +fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} = + let + val thy = ProofContext.theory_of background; + val tyco = Sign.intern_type thy raw_tyco; + val constrs = map (Code.check_const thy) raw_constrs; + val constrs' = (map fst o snd o Code.get_datatype thy) tyco; + val _ = if eq_set (op =) (constrs, constrs') then () + else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors") + val is_first = is_first_occ background; + val background' = register_datatype tyco constrs background; + in (print_code struct_name is_first (print_datatype tyco constrs), background') end; + +end; (*local*) + + +(** Isar setup **) + +val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); +val _ = ML_Context.add_antiq "code_datatype" (fn _ => + (Args.tyname --| Scan.lift (Args.$$$ "=") + -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term))) + >> ml_code_datatype_antiq); + +val setup = Code_Target.extend_target (target, (Code_ML.target_SML, K I)); + +end; (*struct*) diff -r 7c2c38a5bca3 -r 1e6206763036 src/Tools/Code/code_ml.ML --- a/src/Tools/Code/code_ml.ML Mon Dec 07 14:54:28 2009 +0100 +++ b/src/Tools/Code/code_ml.ML Mon Dec 07 16:27:48 2009 +0100 @@ -6,9 +6,9 @@ signature CODE_ML = sig - val eval: string option -> string * (unit -> 'a) option Unsynchronized.ref - -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a - val target_Eval: string + val target_SML: string + val evaluation_code_of: theory -> string -> string option + -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list val setup: theory -> theory end; @@ -26,7 +26,6 @@ val target_SML = "SML"; val target_OCaml = "OCaml"; -val target_Eval = "Eval"; datatype ml_binding = ML_Function of string * (typscheme * ((iterm list * iterm) * (thm * bool)) list) @@ -937,142 +936,19 @@ end; (*local*) -(** ML (system language) code for evaluation and instrumentalization **) +(** for instrumentalization **) -val eval_struct_name = "Code" - -fun eval_code_of some_target with_struct thy = +fun evaluation_code_of thy target some_struct_name = let - val target = the_default target_Eval some_target; - val serialize = if with_struct then fn _ => fn [] => - serialize_ml target_SML (SOME (K ())) print_sml_module print_sml_stmt (SOME eval_struct_name) true + val serialize = if is_some some_struct_name then fn _ => fn [] => + serialize_ml target (SOME (K ())) print_sml_module print_sml_stmt some_struct_name true else fn _ => fn [] => - serialize_ml target_SML (SOME (K ())) ((K o K) Pretty.chunks2) print_sml_stmt (SOME eval_struct_name) true; + serialize_ml target (SOME (K ())) ((K o K) Pretty.chunks2) print_sml_stmt some_struct_name true; in Code_Target.serialize_custom thy (target, (serialize, literals_sml)) end; -(* evaluation *) - -fun eval some_target reff postproc thy t args = - let - val ctxt = ProofContext.init thy; - fun evaluator naming program ((_, (_, ty)), t) deps = - let - val _ = if Code_Thingol.contains_dictvar t then - error "Term to be evaluated contains free dictionaries" else (); - val value_name = "Value.VALUE.value" - val program' = program - |> Graph.new_node (value_name, - Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))]))) - |> fold (curry Graph.add_edge value_name) deps; - val (value_code, [SOME value_name']) = eval_code_of some_target false thy naming program' [value_name]; - val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' - ^ space_implode " " (map (enclose "(" ")") args) ^ " end"; - in ML_Context.evaluate ctxt false reff sml_code end; - in Code_Thingol.eval thy postproc evaluator t end; - - -(* instrumentalization by antiquotation *) - -local - -structure CodeAntiqData = Proof_Data -( - type T = (string list * string list) * (bool * (string - * (string * ((string * string) list * (string * string) list)) lazy)); - fun init _ = (([], []), (true, ("", Lazy.value ("", ([], []))))); -); - -val is_first_occ = fst o snd o CodeAntiqData.get; - -fun delayed_code thy tycos consts () = - let - val (consts', (naming, program)) = Code_Thingol.consts_program thy consts; - val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; - val (ml_code, target_names) = eval_code_of NONE true thy naming program (consts' @ tycos'); - val (consts'', tycos'') = chop (length consts') target_names; - val consts_map = map2 (fn const => fn NONE => - error ("Constant " ^ (quote o Code.string_of_const thy) const - ^ "\nhas a user-defined serialization") - | SOME const'' => (const, const'')) consts consts'' - val tycos_map = map2 (fn tyco => fn NONE => - error ("Type " ^ (quote o Sign.extern_type thy) tyco - ^ "\nhas a user-defined serialization") - | SOME tyco'' => (tyco, tyco'')) tycos tycos''; - in (ml_code, (tycos_map, consts_map)) end; - -fun register_code new_tycos new_consts ctxt = - let - val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt; - val tycos' = fold (insert (op =)) new_tycos tycos; - val consts' = fold (insert (op =)) new_consts consts; - val (struct_name', ctxt') = if struct_name = "" - then ML_Antiquote.variant eval_struct_name ctxt - else (struct_name, ctxt); - val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) tycos' consts'); - in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end; - -fun register_const const = register_code [] [const]; - -fun register_datatype tyco constrs = register_code [tyco] constrs; - -fun print_const const all_struct_name tycos_map consts_map = - (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; - -fun print_datatype tyco constrs all_struct_name tycos_map consts_map = - let - val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode; - fun check_base name name'' = - if upperize (Long_Name.base_name name) = upperize name'' - then () else error ("Name as printed " ^ quote name'' - ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry."); - val tyco'' = (the o AList.lookup (op =) tycos_map) tyco; - val constrs'' = map (the o AList.lookup (op =) consts_map) constrs; - val _ = check_base tyco tyco''; - val _ = map2 check_base constrs constrs''; - in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end; - -fun print_code struct_name is_first print_it ctxt = - let - val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; - val (ml_code, (tycos_map, consts_map)) = Lazy.force acc_code; - val ml_code = if is_first then ml_code - else ""; - val all_struct_name = Long_Name.append struct_name struct_code_name; - in (ml_code, print_it all_struct_name tycos_map consts_map) end; - -in - -fun ml_code_antiq raw_const {struct_name, background} = - let - val const = Code.check_const (ProofContext.theory_of background) raw_const; - val is_first = is_first_occ background; - val background' = register_const const background; - in (print_code struct_name is_first (print_const const), background') end; - -fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} = - let - val thy = ProofContext.theory_of background; - val tyco = Sign.intern_type thy raw_tyco; - val constrs = map (Code.check_const thy) raw_constrs; - val constrs' = (map fst o snd o Code.get_datatype thy) tyco; - val _ = if eq_set (op =) (constrs, constrs') then () - else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors") - val is_first = is_first_occ background; - val background' = register_datatype tyco constrs background; - in (print_code struct_name is_first (print_datatype tyco constrs), background') end; - -end; (*local*) - - (** Isar setup **) -val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); -val _ = ML_Context.add_antiq "code_datatype" (fn _ => - (Args.tyname --| Scan.lift (Args.$$$ "=") - -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term))) - >> ml_code_datatype_antiq); - fun isar_seri_sml module_name = Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true >> (fn with_signatures => serialize_ml target_SML @@ -1087,7 +963,6 @@ val setup = Code_Target.add_target (target_SML, (isar_seri_sml, literals_sml)) #> Code_Target.add_target (target_OCaml, (isar_seri_ocaml, literals_ocaml)) - #> Code_Target.extend_target (target_Eval, (target_SML, K I)) #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] => brackify_infix (1, R) fxy [ print_typ (INFX (1, X)) ty1, diff -r 7c2c38a5bca3 -r 1e6206763036 src/Tools/Code_Generator.thy --- a/src/Tools/Code_Generator.thy Mon Dec 07 14:54:28 2009 +0100 +++ b/src/Tools/Code_Generator.thy Mon Dec 07 16:27:48 2009 +0100 @@ -16,6 +16,7 @@ "~~/src/Tools/Code/code_printer.ML" "~~/src/Tools/Code/code_target.ML" "~~/src/Tools/Code/code_ml.ML" + "~~/src/Tools/Code/code_eval.ML" "~~/src/Tools/Code/code_haskell.ML" "~~/src/Tools/nbe.ML" begin @@ -23,6 +24,7 @@ setup {* Code_Preproc.setup #> Code_ML.setup + #> Code_Eval.setup #> Code_Haskell.setup #> Nbe.setup *}