Theory Nano_JSON_Manual
section‹The Nano JSON Manual›
theory
Nano_JSON_Manual
imports
Nano_JSON_Query
begin
subsection‹Isabelle/HOL›
text‹
On the Isabelle/HOL level, we provide three new commands for working with JSON formatted
data.
›
paragraph‹@{command "JSON"}.› text‹ For defining JSON data within HOL terms,
we provide the following command:
\<^rail>‹ @@{command "JSON"} json_data ›
where \emph{json-data} is a cartouche containing the JSON data.
›
paragraph‹@{command "JSON_file"}.› text‹ For reading JSON data from an external
file, we provide the command
\<^rail>‹ @@{command "JSON_file"} json_filename 'defining' json_def ›
where \emph{json-filename} is the path (file name) of the JSON file that should be read
and \emph{json-def} is the name the HOL definition representing the JSON data is bound to.›
paragraph‹@{command "JSON_export"}.› text‹ For serializing (exporting) JSON
encoded data, we provide the command
\<^rail>‹ @@{command "JSON_export"} json_def ('file' json_filename)? ›
where \emph{json-def} is the definition of the JSON data that should be exported. Optionally
a file name (\emph{json-filename}) can be provided. If a file name is provided, it is used
for storing the serialized data in Isabelle's virtual file system. If no file name is provided,
the serialized data is printed in Isabelle's output window.
›
paragraph‹Configuration.› text ‹We provide three configuration attributes:
▪ The attribute @{attribute "JSON_num_type"} (default @{type "int"} allows for configuring the
HOL-type used representing JSON numerals.
▪ The attribute @{attribute "JSON_string_type"} (default @{type "string"}) allows for configuring
the HOL-type used representing JSON string.
▪ The attribute @{attribute "JSON_verbose"} (default: false) allows for enabling warnings during the
JSON processing.
›
subsection‹Isabelle/ML›
text‹
We refer users who want to use or extend Nano JSON on the Isabelle/ML level to the ML signatures
shown in the implementation chapter, i.e., \autoref{ch:implementation}.
›
end