store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
integrity check of SHA1 digests produced in Scala vs. ML;
tuned signature;
(* Title: Pure/pure_syn.ML
Author: Makarius
Minimal outer syntax for bootstrapping Isabelle/Pure.
*)
structure Pure_Syn: sig end =
struct
val _ =
Outer_Syntax.command
(("theory", Keyword.tag_theory Keyword.thy_begin), Position.file "pure_syn.ML")
"begin theory context"
(Thy_Header.args >> (fn header =>
Toplevel.print o
Toplevel.init_theory
(fn () => Thy_Info.toplevel_begin_theory (! ProofGeneral.master_path) header)));
val _ =
Outer_Syntax.command
(("ML_file", Keyword.tag_ml Keyword.thy_load), Position.file "pure_syn.ML")
"ML text from file"
(Thy_Load.parse_files "ML_file" >> (fn files => Toplevel.generic_theory (fn gthy =>
let
val [{src_path, lines, digest, pos}] = files (Context.theory_of gthy);
val provide = Thy_Load.provide (src_path, digest);
in
gthy
|> ML_Context.exec (fn () => ML_Context.eval_text true pos (cat_lines lines))
|> Local_Theory.propagate_ml_env
|> Context.mapping provide (Local_Theory.background_theory provide)
end)));
end;