Theory Generator_static
text‹We present two solutions for obtaining an Isabelle file.›
section‹Static Meta Embedding with Exportation›
theory Generator_static
imports Printer "../../Antiquote_Setup"
begin
declare[[cartouche_type' = "abr_string"]]
text ‹In the ``static'' solution: the user manually generates
the Isabelle file after writing by hand a Toy input to translate.
The input is not written with the syntax of the Toy Language,
but with raw Isabelle constructors.›
subsection‹Giving an Input to Translate›
definition "Design =
(let n = λn1 n2. ToyTyObj (ToyTyCore_pre n1) (case n2 of None ⇒ [] | Some n2 ⇒ [[ToyTyCore_pre n2]])
; mk = λn l. toy_class_raw.make n l [] False in
[ mk (n ‹Galaxy› None) [(‹sound›, ToyTy_raw ‹unit›), (‹moving›, ToyTy_raw ‹bool›)]
, mk (n ‹Planet› (Some ‹Galaxy›)) [(‹weight›, ToyTy_raw ‹nat›)]
, mk (n ‹Person› (Some ‹Planet›)) [(‹salary›, ToyTy_raw ‹int›)] ])"
text ‹Since we are in a Isabelle session, at this time, it becomes possible to inspect with
the command @{command value} the result of the translations applied with @{term Design}.
A suitable environment should nevertheless be provided,
one can typically experiment this by copying-pasting the following environment
initialized in the above ‹main›:›
definition "main =
(let n = λn1. ToyTyObj (ToyTyCore_pre n1) []
; ToyMult = λm r. toy_multiplicity.make [m] r [Set] in
write_file
(compiler_env_config.extend
(compiler_env_config_empty True None (oidInit (Oid 0)) Gen_only_design (None, False)
⦇ D_output_disable_thy := False
, D_output_header_thy := Some (‹Design_generated›
,[‹../Toy_Library›]
,‹../embedding/Generator_dynamic_sequential›) ⦈)
( L.map (META_class_raw Floor1) Design
@@@@ [ META_association (toy_association.make
ToyAssTy_association
(ToyAssRel [ (n ‹Person›, ToyMult (Mult_star, None) None)
, (n ‹Person›, ToyMult (Mult_nat 0, Some (Mult_nat 1)) (Some ‹boss›))]))
, META_flush_all ToyFlushAll]
, None)))"
subsection‹Statically Executing the Exportation›
text‹
@{verbatim "apply_code_printing ()"} \\
@{verbatim "export_code main"} \\
@{verbatim " (* in Haskell *)"} \\
@{verbatim " (* in OCaml module_name M *)"} \\
@{verbatim " (* in Scala module_name M *)"} \\
@{verbatim " (* in SML module_name M *)"}
›
text‹After the exportation and executing the exported, we obtain an Isabelle \verb|.thy| file
containing the generated code associated to the above input.›
end