--- a/NEWS Wed May 19 14:17:40 2021 +0100
+++ b/NEWS Sun May 23 20:34:43 2021 +0200
@@ -34,6 +34,60 @@
*** Document preparation ***
+* Document antiquotations for ML text have been refined: "def" and "ref"
+variants support index entries, e.g. @{ML} (no entry) vs. @{ML_def}
+(bold entry) vs. @{ML_ref} (regular entry); @{ML_type} supports explicit
+type arguments for constructors (only relevant for index), e.g.
+@{ML_type \<open>'a list\<close>} vs. @{ML_type 'a \<open>list\<close>}; @{ML_op} has been renamed
+to @{ML_infix}. Minor INCOMPATIBILITY concerning name and syntax.
+
+* Option "document_logo" determines if an instance of the Isabelle logo
+should be created in the document output directory. The given string
+specifies the name of the logo variant, while "_" (underscore) refers to
+the unnamed variant. The output file name is always "isabelle_logo.pdf".
+
+* Option "document_preprocessor" specifies the name of an executable
+that is run within the document output directory, after preparing the
+document sources and before the actual build process. This allows to
+apply adhoc patches, without requiring a separate "build" script.
+
+* Option "document_build" determines the document build engine, as
+defined in Isabelle/Scala (as system service). The subsequent engines
+are provided by the Isabelle distribution:
+
+ - "lualatex" (default): use ISABELLE_LUALATEX for a standard LaTeX
+ build with optional ISABELLE_BIBTEX and ISABELLE_MAKEINDEX
+
+ - "pdflatex": as above, but use ISABELLE_PDFLATEX (legacy mode for
+ special LaTeX styles)
+
+ - "build": delegate to the executable "./build pdf"
+
+The presence of a "build" command within the document output directory
+explicitly requires document_build=build. Minor INCOMPATIBILITY, need to
+adjust session ROOT options.
+
+* The command-line tool "isabelle latex" has been discontinued,
+INCOMPATIBILITY for old document build scripts.
+
+ - Former "isabelle latex -o sty" has become obsolete: Isabelle .sty
+ files are automatically generated within the document output
+ directory.
+
+ - Former "isabelle latex -o pdf" should be replaced by
+ "$ISABELLE_LUALATEX root" or "$ISABELLE_PDFLATEX root" (without
+ quotes), according to the intended LaTeX engine.
+
+ - Former "isabelle latex -o bbl" should be replaced by
+ "$ISABELLE_BIBTEX root" (without quotes).
+
+ - Former "isabelle latex -o idx" should be replaced by
+ "$ISABELLE_MAKEINDEX root" (without quotes).
+
+* Option "document_bibliography" explicitly enables the use of bibtex;
+the default is to check the presence of root.bib, but it could have a
+different name.
+
* Improved LaTeX typesetting of \<open>...\<close> using \guilsinglleft ...
\guilsinglright. INCOMPATIBILITY, need to use \usepackage[T1]{fontenc}
(which is now also the default in "isabelle mkroot").
--- a/etc/options Wed May 19 14:17:40 2021 +0100
+++ b/etc/options Sun May 23 20:34:43 2021 +0200
@@ -13,6 +13,14 @@
-- "alternative document variants (separated by colons)"
option document_tags : string = ""
-- "default command tags (separated by commas)"
+option document_bibliography : bool = false
+ -- "explicitly enable use of bibtex (default: according to presence of root.bib)"
+option document_preprocessor : string = ""
+ -- "document preprocessor: executable relative to document output directory"
+option document_build : string = "lualatex"
+ -- "document build engine (e.g. lualatex, pdflatex, build)"
+option document_logo : string = ""
+ -- "generate named instance of Isabelle logo (underscore means unnamed variant)"
option thy_output_display : bool = false
-- "indicate output as multi-line display-style material"
--- a/etc/settings Wed May 19 14:17:40 2021 +0100
+++ b/etc/settings Sun May 23 20:34:43 2021 +0200
@@ -16,7 +16,7 @@
ISABELLE_TOOL_JAVA_OPTIONS="-Djava.awt.headless=true -Xms512m -Xmx4g -Xss16m"
-ISABELLE_SCALAC_OPTIONS="-encoding UTF-8 -Wconf:cat=other-match-analysis:silent -feature -deprecation -target:11 -J-Xms512m -J-Xmx4g -J-Xss16m"
+ISABELLE_SCALAC_OPTIONS="-encoding UTF-8 -Wconf:cat=other-match-analysis:silent -feature -deprecation -target:11 -Xsource:3 -J-Xms512m -J-Xmx4g -J-Xss16m"
classpath "$ISABELLE_HOME/lib/classes/Pure.jar"
@@ -34,6 +34,10 @@
isabelle_scala_service 'isabelle.Simplifier_Trace$Handler'
isabelle_scala_service 'isabelle.Server_Commands'
+isabelle_scala_service 'isabelle.Document_Build$LuaLaTeX_Engine'
+isabelle_scala_service 'isabelle.Document_Build$PDFLaTeX_Engine'
+isabelle_scala_service 'isabelle.Document_Build$Build_Engine'
+
#paranoia settings -- avoid intrusion of alien options
unset "_JAVA_OPTIONS"
unset "JAVA_TOOL_OPTIONS"
@@ -57,12 +61,18 @@
###
-### Document preparation (cf. isabelle latex/document)
+### Document preparation (cf. isabelle latex)
###
-ISABELLE_PDFLATEX="lualatex --file-line-error"
+if [ "$ISABELLE_PLATFORM_FAMILY" = "windows" ]; then
+ ISABELLE_PDFLATEX="pdflatex -interaction=nonstopmode -c-style-errors"
+else
+ ISABELLE_PDFLATEX="pdflatex -interaction=nonstopmode -file-line-error"
+fi
+
+ISABELLE_LUALATEX="lualatex --interaction=nonstopmode --file-line-error"
ISABELLE_BIBTEX="bibtex"
-ISABELLE_MAKEINDEX="makeindex"
+ISABELLE_MAKEINDEX="makeindex -c -q"
ISABELLE_EPSTOPDF="epstopdf"
--- a/lib/Tools/latex Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Markus Wenzel, TU Muenchen
-#
-# DESCRIPTION: run LaTeX (and related tools)
-
-
-PRG="$(basename "$0")"
-
-function usage()
-{
- echo
- echo "Usage: isabelle $PRG [OPTIONS] [FILE]"
- echo
- echo " Options are:"
- echo " -o FORMAT specify output format: pdf (default), bbl, idx, sty"
- echo
- echo " Run LaTeX (and related tools) on FILE (default root.tex),"
- echo " producing the specified output format."
- echo
- exit 1
-}
-
-function fail()
-{
- echo "$1" >&2
- exit 2
-}
-
-
-## process command line
-
-# options
-
-OUTFORMAT=pdf
-
-while getopts "o:" OPT
-do
- case "$OPT" in
- o)
- OUTFORMAT="$OPTARG"
- ;;
- \?)
- usage
- ;;
- esac
-done
-
-shift $(($OPTIND - 1))
-
-
-# args
-
-FILE="root.tex"
-[ "$#" -ge 1 ] && { FILE="$1"; shift; }
-
-[ "$#" -ne 0 ] && usage
-
-
-## main
-
-# root file
-
-DIR="$(dirname "$FILE")"
-FILEBASE="$(basename "$FILE" .tex)"
-[ "$DIR" = . ] || FILEBASE="$DIR/$FILEBASE"
-
-function check_root () { [ -f "$FILEBASE.tex" ] || fail "Bad file '$FILE'"; }
-
-
-# operations
-
-function run_pdflatex () { $ISABELLE_PDFLATEX "\\nonstopmode\\input{$FILEBASE.tex}"; }
-function run_bibtex () {
- $ISABELLE_BIBTEX </dev/null "$FILEBASE"
- RC="$?"
- if [ "$RC" -gt 0 -a -f "${FILEBASE}.blg" ]; then
- perl -n -e 'if (m/^I (found no.*$)/) { print "bibtex $1\n"; }' "${FILEBASE}.blg" >&2
- fi
- return "$RC"
-}
-function run_makeindex () { $ISABELLE_MAKEINDEX </dev/null "$FILEBASE"; }
-function copy_styles ()
-{
- for STYLEFILE in "$ISABELLE_HOME/lib/texinputs"/*.sty
- do
- TARGET="$DIR"/$(basename "$STYLEFILE")
- perl -p -e 's/\$[I]d:?(?:\s)*([^\$]*)\$//g' "$STYLEFILE" > "$TARGET"
- done
-}
-
-case "$OUTFORMAT" in
- pdf)
- check_root && \
- run_pdflatex
- RC="$?"
- ;;
- bbl)
- check_root && \
- run_bibtex
- RC="$?"
- ;;
- idx)
- check_root && \
- run_makeindex
- RC="$?"
- ;;
- sty)
- copy_styles
- RC="$?"
- ;;
- *)
- fail "Bad output format '$OUTFORMAT'"
- ;;
-esac
-
-exit "$RC"
--- a/lib/texinputs/isabelle.sty Wed May 19 14:17:40 2021 +0100
+++ b/lib/texinputs/isabelle.sty Sun May 23 20:34:43 2021 +0200
@@ -173,6 +173,12 @@
\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}
+% index entries
+
+\newcommand{\isaindexdef}[1]{\textbf{#1}}
+\newcommand{\isaindexref}[1]{#1}
+
+
% styles
\def\isabellestyle#1{\csname isabellestyle#1\endcsname}
--- a/src/Doc/Classes/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo Isar
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Classes/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Classes/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -10,7 +10,7 @@
\hyphenation{Isar}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_isar}
+\title{\includegraphics[scale=0.5]{isabelle_logo}
\\[4ex] Haskell-style type classes with Isabelle/Isar}
\author{\emph{Florian Haftmann}}
--- a/src/Doc/Codegen/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-# ad-hoc patching of temporary path from sources
-perl -i -pe 's/\{\\isachardollar\}ISABELLE\{\\isacharunderscore\}TMP\{\\isacharslash\}examples/examples/g' *.tex
-
-isabelle logo Isar
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
-# clean up afterwards
-rm -rf "${ISABELLE_TMP}/examples"
--- a/src/Doc/Codegen/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Codegen/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -13,7 +13,7 @@
\hyphenation{Isar}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_isar}
+\title{\includegraphics[scale=0.5]{isabelle_logo}
\\[4ex] Code generation from Isabelle/HOL theories}
\author{\emph{Florian Haftmann with contributions from Lukas Bulwahn}}
--- a/src/Doc/Corec/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Corec/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Corec/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -48,7 +48,7 @@
\isadroptag{theory}
-\title{%\includegraphics[scale=0.5]{isabelle_hol} \\[4ex]
+\title{%\includegraphics[scale=0.5]{isabelle_logo} \\[4ex]
Defining Nonprimitively (Co)recursive Functions in Isabelle/HOL}
\author{Jasmin Christian Blanchette, Aymeric Bouzy, \\
Andreas Lochbihler, Andrei Popescu, and \\
--- a/src/Doc/Datatypes/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Datatypes/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Datatypes/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -48,7 +48,7 @@
\isadroptag{theory}
-\title{%\includegraphics[scale=0.5]{isabelle_hol} \\[4ex]
+\title{%\includegraphics[scale=0.5]{isabelle_logo} \\[4ex]
Defining (Co)datatypes and Primitively (Co)recursive Functions in Isabelle/HOL}
\author{Julian Biendarra, Jasmin Blanchette, \\
Martin Desharnais, Lorenz Panny, \\
--- a/src/Doc/Eisbach/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo Eisbach
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Eisbach/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Eisbach/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -16,7 +16,7 @@
\hyphenation{Eisbach}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_eisbach}
+\title{\includegraphics[scale=0.5]{isabelle_logo}
\\[4ex] The Eisbach User Manual}
\author{Daniel Matichuk \\
Makarius Wenzel \\
--- a/src/Doc/Functions/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/How_to_Prove_it/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/How_to_Prove_it/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -5,7 +5,7 @@
\begin{document}
\title{How to Prove it in Isabelle/HOL}
-%\subtitle{\includegraphics[scale=.7]{isabelle_hol}}
+%\subtitle{\includegraphics[scale=.7]{isabelle_logo}}
\author{Tobias Nipkow}
\maketitle
--- a/src/Doc/Implementation/Eq.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Eq.thy Sun May 23 20:34:43 2021 +0200
@@ -55,13 +55,13 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Thm.reflexive: "cterm -> thm"} \\
- @{index_ML Thm.symmetric: "thm -> thm"} \\
- @{index_ML Thm.transitive: "thm -> thm -> thm"} \\
- @{index_ML Thm.abstract_rule: "string -> cterm -> thm -> thm"} \\
- @{index_ML Thm.combination: "thm -> thm -> thm"} \\[0.5ex]
- @{index_ML Thm.equal_intr: "thm -> thm -> thm"} \\
- @{index_ML Thm.equal_elim: "thm -> thm -> thm"} \\
+ @{define_ML Thm.reflexive: "cterm -> thm"} \\
+ @{define_ML Thm.symmetric: "thm -> thm"} \\
+ @{define_ML Thm.transitive: "thm -> thm -> thm"} \\
+ @{define_ML Thm.abstract_rule: "string -> cterm -> thm -> thm"} \\
+ @{define_ML Thm.combination: "thm -> thm -> thm"} \\[0.5ex]
+ @{define_ML Thm.equal_intr: "thm -> thm -> thm"} \\
+ @{define_ML Thm.equal_elim: "thm -> thm -> thm"} \\
\end{mldecls}
See also \<^file>\<open>~~/src/Pure/thm.ML\<close> for further description of these inference
@@ -82,9 +82,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_structure Conv} \\
- @{index_ML_type conv} \\
- @{index_ML Simplifier.asm_full_rewrite : "Proof.context -> conv"} \\
+ @{define_ML_structure Conv} \\
+ @{define_ML_type conv} \\
+ @{define_ML Simplifier.asm_full_rewrite : "Proof.context -> conv"} \\
\end{mldecls}
\<^descr> \<^ML_structure>\<open>Conv\<close> is a library of combinators to build conversions,
@@ -109,11 +109,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML rewrite_rule: "Proof.context -> thm list -> thm -> thm"} \\
- @{index_ML rewrite_goals_rule: "Proof.context -> thm list -> thm -> thm"} \\
- @{index_ML rewrite_goal_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML rewrite_goals_tac: "Proof.context -> thm list -> tactic"} \\
- @{index_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\
+ @{define_ML rewrite_rule: "Proof.context -> thm list -> thm -> thm"} \\
+ @{define_ML rewrite_goals_rule: "Proof.context -> thm list -> thm -> thm"} \\
+ @{define_ML rewrite_goal_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML rewrite_goals_tac: "Proof.context -> thm list -> tactic"} \\
+ @{define_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>rewrite_rule\<close>~\<open>ctxt rules thm\<close> rewrites the whole theorem by the
--- a/src/Doc/Implementation/Integration.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Integration.thy Sun May 23 20:34:43 2021 +0200
@@ -37,11 +37,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Toplevel.state} \\
- @{index_ML_exception Toplevel.UNDEF} \\
- @{index_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\
- @{index_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\
- @{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
+ @{define_ML_type Toplevel.state} \\
+ @{define_ML_exception Toplevel.UNDEF} \\
+ @{define_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\
+ @{define_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\
+ @{define_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Toplevel.state\<close> represents Isar toplevel states, which are
@@ -95,17 +95,17 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Toplevel.keep: "(Toplevel.state -> unit) ->
+ @{define_ML Toplevel.keep: "(Toplevel.state -> unit) ->
Toplevel.transition -> Toplevel.transition"} \\
- @{index_ML Toplevel.theory: "(theory -> theory) ->
+ @{define_ML Toplevel.theory: "(theory -> theory) ->
Toplevel.transition -> Toplevel.transition"} \\
- @{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) ->
+ @{define_ML Toplevel.theory_to_proof: "(theory -> Proof.state) ->
Toplevel.transition -> Toplevel.transition"} \\
- @{index_ML Toplevel.proof: "(Proof.state -> Proof.state) ->
+ @{define_ML Toplevel.proof: "(Proof.state -> Proof.state) ->
Toplevel.transition -> Toplevel.transition"} \\
- @{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) ->
+ @{define_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) ->
Toplevel.transition -> Toplevel.transition"} \\
- @{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) ->
+ @{define_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) ->
Toplevel.transition -> Toplevel.transition"} \\
\end{mldecls}
@@ -150,10 +150,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML use_thy: "string -> unit"} \\
- @{index_ML Thy_Info.get_theory: "string -> theory"} \\
- @{index_ML Thy_Info.remove_thy: "string -> unit"} \\
- @{index_ML Thy_Info.register_thy: "theory -> unit"} \\
+ @{define_ML use_thy: "string -> unit"} \\
+ @{define_ML Thy_Info.get_theory: "string -> theory"} \\
+ @{define_ML Thy_Info.remove_thy: "string -> unit"} \\
+ @{define_ML Thy_Info.register_thy: "theory -> unit"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>use_thy\<close>~\<open>A\<close> ensures that theory \<open>A\<close> is fully up-to-date wrt.\ the
--- a/src/Doc/Implementation/Isar.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Isar.thy Sun May 23 20:34:43 2021 +0200
@@ -60,16 +60,16 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Proof.state} \\
- @{index_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\
- @{index_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\
- @{index_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\
- @{index_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\
- @{index_ML Proof.goal: "Proof.state ->
+ @{define_ML_type Proof.state} \\
+ @{define_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\
+ @{define_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\
+ @{define_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\
+ @{define_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\
+ @{define_ML Proof.goal: "Proof.state ->
{context: Proof.context, facts: thm list, goal: thm}"} \\
- @{index_ML Proof.raw_goal: "Proof.state ->
+ @{define_ML Proof.raw_goal: "Proof.state ->
{context: Proof.context, facts: thm list, goal: thm}"} \\
- @{index_ML Proof.theorem: "Method.text option ->
+ @{define_ML Proof.theorem: "Method.text option ->
(thm list list -> Proof.context -> Proof.context) ->
(term * term list) list list -> Proof.context -> Proof.state"} \\
\end{mldecls}
@@ -272,13 +272,13 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Proof.method} \\
- @{index_ML CONTEXT_METHOD: "(thm list -> context_tactic) -> Proof.method"} \\
- @{index_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\
- @{index_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\
- @{index_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\
- @{index_ML Method.insert_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser ->
+ @{define_ML_type Proof.method} \\
+ @{define_ML CONTEXT_METHOD: "(thm list -> context_tactic) -> Proof.method"} \\
+ @{define_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\
+ @{define_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\
+ @{define_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\
+ @{define_ML Method.insert_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser ->
string -> theory -> theory"} \\
\end{mldecls}
@@ -286,7 +286,7 @@
\<^descr> \<^ML>\<open>CONTEXT_METHOD\<close>~\<open>(fn facts => context_tactic)\<close> wraps \<open>context_tactic\<close>
depending on goal facts as a general proof method that may change the proof
- context dynamically. A typical operation is \<^ML>\<open>Proof_Context.update_cases\<close>, which is wrapped up as combinator @{index_ML
+ context dynamically. A typical operation is \<^ML>\<open>Proof_Context.update_cases\<close>, which is wrapped up as combinator @{define_ML
CONTEXT_CASES} for convenience.
\<^descr> \<^ML>\<open>METHOD\<close>~\<open>(fn facts => tactic)\<close> wraps \<open>tactic\<close> depending on goal facts
@@ -495,12 +495,12 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type attribute} \\
- @{index_ML Thm.rule_attribute: "thm list ->
+ @{define_ML_type attribute} \\
+ @{define_ML Thm.rule_attribute: "thm list ->
(Context.generic -> thm -> thm) -> attribute"} \\
- @{index_ML Thm.declaration_attribute: "
+ @{define_ML Thm.declaration_attribute: "
(thm -> Context.generic -> Context.generic) -> attribute"} \\
- @{index_ML Attrib.setup: "binding -> attribute context_parser ->
+ @{define_ML Attrib.setup: "binding -> attribute context_parser ->
string -> theory -> theory"} \\
\end{mldecls}
--- a/src/Doc/Implementation/Local_Theory.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Local_Theory.thy Sun May 23 20:34:43 2021 +0200
@@ -90,11 +90,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type local_theory: Proof.context} \\
- @{index_ML Named_Target.init: "string list -> string -> theory -> local_theory"} \\[1ex]
- @{index_ML Local_Theory.define: "(binding * mixfix) * (Attrib.binding * term) ->
+ @{define_ML_type local_theory = Proof.context} \\
+ @{define_ML Named_Target.init: "string list -> string -> theory -> local_theory"} \\[1ex]
+ @{define_ML Local_Theory.define: "(binding * mixfix) * (Attrib.binding * term) ->
local_theory -> (term * (string * thm)) * local_theory"} \\
- @{index_ML Local_Theory.note: "Attrib.binding * thm list ->
+ @{define_ML Local_Theory.note: "Attrib.binding * thm list ->
local_theory -> (string * thm list) * local_theory"} \\
\end{mldecls}
--- a/src/Doc/Implementation/Logic.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Logic.thy Sun May 23 20:34:43 2021 +0200
@@ -102,22 +102,22 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type class: string} \\
- @{index_ML_type sort: "class list"} \\
- @{index_ML_type arity: "string * sort list * sort"} \\
- @{index_ML_type typ} \\
- @{index_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\
- @{index_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
+ @{define_ML_type class = string} \\
+ @{define_ML_type sort = "class list"} \\
+ @{define_ML_type arity = "string * sort list * sort"} \\
+ @{define_ML_type typ} \\
+ @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\
+ @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
- @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
- @{index_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\
- @{index_ML Sign.add_type_abbrev: "Proof.context ->
+ @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
+ @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
+ @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\
+ @{define_ML Sign.add_type_abbrev: "Proof.context ->
binding * string list * typ -> theory -> theory"} \\
- @{index_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
- @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
- @{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
+ @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
+ @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
+ @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>class\<close> represents type classes.
@@ -313,30 +313,30 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type term} \\
- @{index_ML_op "aconv": "term * term -> bool"} \\
- @{index_ML Term.map_types: "(typ -> typ) -> term -> term"} \\
- @{index_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
- @{index_ML Term.map_aterms: "(term -> term) -> term -> term"} \\
- @{index_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
+ @{define_ML_type term} \\
+ @{define_ML_infix "aconv": "term * term -> bool"} \\
+ @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\
+ @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
+ @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\
+ @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML fastype_of: "term -> typ"} \\
- @{index_ML lambda: "term -> term -> term"} \\
- @{index_ML betapply: "term * term -> term"} \\
- @{index_ML incr_boundvars: "int -> term -> term"} \\
- @{index_ML Sign.declare_const: "Proof.context ->
+ @{define_ML fastype_of: "term -> typ"} \\
+ @{define_ML lambda: "term -> term -> term"} \\
+ @{define_ML betapply: "term * term -> term"} \\
+ @{define_ML incr_boundvars: "int -> term -> term"} \\
+ @{define_ML Sign.declare_const: "Proof.context ->
(binding * typ) * mixfix -> theory -> term * theory"} \\
- @{index_ML Sign.add_abbrev: "string -> binding * term ->
+ @{define_ML Sign.add_abbrev: "string -> binding * term ->
theory -> (term * term) * theory"} \\
- @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
- @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
+ @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
+ @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>term\<close> represents de-Bruijn terms, with comments in
abstractions, and explicitly named free variables and constants; this is a
- datatype with constructors @{index_ML Bound}, @{index_ML Free}, @{index_ML
- Var}, @{index_ML Const}, @{index_ML Abs}, @{index_ML_op "$"}.
+ datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML
+ Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}.
\<^descr> \<open>t\<close>~\<^ML_text>\<open>aconv\<close>~\<open>u\<close> checks \<open>\<alpha>\<close>-equivalence of two terms. This is the
basic equality relation on type \<^ML_type>\<open>term\<close>; raw datatype equality
@@ -564,41 +564,41 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Logic.all: "term -> term -> term"} \\
- @{index_ML Logic.mk_implies: "term * term -> term"} \\
+ @{define_ML Logic.all: "term -> term -> term"} \\
+ @{define_ML Logic.mk_implies: "term * term -> term"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type ctyp} \\
- @{index_ML_type cterm} \\
- @{index_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\
- @{index_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\
- @{index_ML Thm.apply: "cterm -> cterm -> cterm"} \\
- @{index_ML Thm.lambda: "cterm -> cterm -> cterm"} \\
- @{index_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\
- @{index_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\
+ @{define_ML_type ctyp} \\
+ @{define_ML_type cterm} \\
+ @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\
+ @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\
+ @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\
+ @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\
+ @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\
+ @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type thm} \\
- @{index_ML Thm.transfer: "theory -> thm -> thm"} \\
- @{index_ML Thm.assume: "cterm -> thm"} \\
- @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
- @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
- @{index_ML Thm.implies_intr: "cterm -> thm -> thm"} \\
- @{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\
- @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\
- @{index_ML Thm.instantiate: "((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list
+ @{define_ML_type thm} \\
+ @{define_ML Thm.transfer: "theory -> thm -> thm"} \\
+ @{define_ML Thm.assume: "cterm -> thm"} \\
+ @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
+ @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
+ @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\
+ @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\
+ @{define_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\
+ @{define_ML Thm.instantiate: "((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list
-> thm -> thm"} \\
- @{index_ML Thm.add_axiom: "Proof.context ->
+ @{define_ML Thm.add_axiom: "Proof.context ->
binding * term -> theory -> (string * thm) * theory"} \\
- @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory ->
+ @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory ->
(string * ('a -> thm)) * theory"} \\
- @{index_ML Thm.add_def: "Defs.context -> bool -> bool ->
+ @{define_ML Thm.add_def: "Defs.context -> bool -> bool ->
binding * term -> theory -> (string * thm) * theory"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Theory.add_deps: "Defs.context -> string ->
+ @{define_ML Theory.add_deps: "Defs.context -> string ->
Defs.entry -> Defs.entry list -> theory -> theory"} \\
- @{index_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\
+ @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Logic.all\<close>~\<open>a B\<close> produces a Pure quantification \<open>\<And>a. B\<close>, where
@@ -624,7 +624,7 @@
\<^descr> \<^ML>\<open>Thm.apply\<close>, \<^ML>\<open>Thm.lambda\<close>, \<^ML>\<open>Thm.all\<close>, \<^ML>\<open>Drule.mk_implies\<close>
etc.\ compose certified terms (or propositions) incrementally. This is
- equivalent to \<^ML>\<open>Thm.cterm_of\<close> after unchecked \<^ML_op>\<open>$\<close>, \<^ML>\<open>lambda\<close>,
+ equivalent to \<^ML>\<open>Thm.cterm_of\<close> after unchecked \<^ML_infix>\<open>$\<close>, \<^ML>\<open>lambda\<close>,
\<^ML>\<open>Logic.all\<close>, \<^ML>\<open>Logic.mk_implies\<close> etc., but there can be a big
difference in performance when large existing entities are composed by a few
extra constructions on top. There are separate operations to decompose
@@ -796,12 +796,12 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Conjunction.intr: "thm -> thm -> thm"} \\
- @{index_ML Conjunction.elim: "thm -> thm * thm"} \\
- @{index_ML Drule.mk_term: "cterm -> thm"} \\
- @{index_ML Drule.dest_term: "thm -> cterm"} \\
- @{index_ML Logic.mk_type: "typ -> term"} \\
- @{index_ML Logic.dest_type: "term -> typ"} \\
+ @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\
+ @{define_ML Conjunction.elim: "thm -> thm * thm"} \\
+ @{define_ML Drule.mk_term: "cterm -> thm"} \\
+ @{define_ML Drule.dest_term: "thm -> cterm"} \\
+ @{define_ML Logic.mk_type: "typ -> term"} \\
+ @{define_ML Logic.dest_type: "term -> typ"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Conjunction.intr\<close> derives \<open>A &&& B\<close> from \<open>A\<close> and \<open>B\<close>.
@@ -846,8 +846,8 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Thm.extra_shyps: "thm -> sort list"} \\
- @{index_ML Thm.strip_shyps: "thm -> thm"} \\
+ @{define_ML Thm.extra_shyps: "thm -> sort list"} \\
+ @{define_ML Thm.strip_shyps: "thm -> thm"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Thm.extra_shyps\<close>~\<open>thm\<close> determines the extraneous sort hypotheses of
@@ -951,7 +951,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\
+ @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Simplifier.norm_hhf\<close>~\<open>ctxt thm\<close> normalizes the given theorem
@@ -1022,14 +1022,14 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_op "RSN": "thm * (int * thm) -> thm"} \\
- @{index_ML_op "RS": "thm * thm -> thm"} \\
+ @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\
+ @{define_ML_infix "RS": "thm * thm -> thm"} \\
- @{index_ML_op "RLN": "thm list * (int * thm list) -> thm list"} \\
- @{index_ML_op "RL": "thm list * thm list -> thm list"} \\
+ @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\
+ @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\
- @{index_ML_op "MRS": "thm list * thm -> thm"} \\
- @{index_ML_op "OF": "thm * thm list -> thm"} \\
+ @{define_ML_infix "MRS": "thm list * thm -> thm"} \\
+ @{define_ML_infix "OF": "thm * thm list -> thm"} \\
\end{mldecls}
\<^descr> \<open>rule\<^sub>1 RSN (i, rule\<^sub>2)\<close> resolves the conclusion of \<open>rule\<^sub>1\<close> with the
@@ -1184,22 +1184,22 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type proof} \\
- @{index_ML_type proof_body} \\
- @{index_ML Proofterm.proofs: "int Unsynchronized.ref"} \\
- @{index_ML Proofterm.reconstruct_proof:
+ @{define_ML_type proof} \\
+ @{define_ML_type proof_body} \\
+ @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\
+ @{define_ML Proofterm.reconstruct_proof:
"theory -> term -> proof -> proof"} \\
- @{index_ML Proofterm.expand_proof: "theory ->
+ @{define_ML Proofterm.expand_proof: "theory ->
(Proofterm.thm_header -> string option) -> proof -> proof"} \\
- @{index_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\
- @{index_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\
- @{index_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\
+ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\
+ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\
+ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>proof\<close> represents proof terms; this is a datatype with
- constructors @{index_ML Abst}, @{index_ML AbsP}, @{index_ML_op "%"},
- @{index_ML_op "%%"}, @{index_ML PBound}, @{index_ML MinProof}, @{index_ML
- Hyp}, @{index_ML PAxm}, @{index_ML Oracle}, @{index_ML PThm} as explained
+ constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"},
+ @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML
+ Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained
above. %FIXME PClass (!?)
\<^descr> Type \<^ML_type>\<open>proof_body\<close> represents the nested proof information of a
--- a/src/Doc/Implementation/ML.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/ML.thy Sun May 23 20:34:43 2021 +0200
@@ -611,10 +611,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Context.the_generic_context: "unit -> Context.generic"} \\
- @{index_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\
- @{index_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\
- @{index_ML ML_Thms.bind_thm: "string * thm -> unit"} \\
+ @{define_ML Context.the_generic_context: "unit -> Context.generic"} \\
+ @{define_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\
+ @{define_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\
+ @{define_ML ML_Thms.bind_thm: "string * thm -> unit"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Context.the_generic_context ()\<close> refers to the theory context of
@@ -818,10 +818,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_op "|> ": "'a * ('a -> 'b) -> 'b"} \\
- @{index_ML_op "|-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
- @{index_ML_op "#> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
- @{index_ML_op "#-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
+ @{define_ML_infix "|>" : "'a * ('a -> 'b) -> 'b"} \\
+ @{define_ML_infix "|->" : "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
+ @{define_ML_infix "#>" : "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
+ @{define_ML_infix "#->" : "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
\end{mldecls}
\<close>
@@ -853,9 +853,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
- @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
- @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
+ @{define_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
+ @{define_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
+ @{define_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>fold\<close>~\<open>f\<close> lifts the parametrized update function \<open>f\<close> to a list of
@@ -969,10 +969,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML writeln: "string -> unit"} \\
- @{index_ML tracing: "string -> unit"} \\
- @{index_ML warning: "string -> unit"} \\
- @{index_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\
+ @{define_ML writeln: "string -> unit"} \\
+ @{define_ML tracing: "string -> unit"} \\
+ @{define_ML warning: "string -> unit"} \\
+ @{define_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\
\end{mldecls}
\<^descr> \<^ML>\<open>writeln\<close>~\<open>text\<close> outputs \<open>text\<close> as regular message. This is the
@@ -1140,13 +1140,13 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
- @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
- @{index_ML_exception ERROR: string} \\
- @{index_ML_exception Fail: string} \\
- @{index_ML Exn.is_interrupt: "exn -> bool"} \\
- @{index_ML Exn.reraise: "exn -> 'a"} \\
- @{index_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
+ @{define_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
+ @{define_ML can: "('a -> 'b) -> 'a -> bool"} \\
+ @{define_ML_exception ERROR of string} \\
+ @{define_ML_exception Fail of string} \\
+ @{define_ML Exn.is_interrupt: "exn -> bool"} \\
+ @{define_ML Exn.reraise: "exn -> 'a"} \\
+ @{define_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
\end{mldecls}
\<^rail>\<open>
@@ -1284,16 +1284,16 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "Symbol.symbol": string} \\
- @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
- @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
- @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
- @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
- @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
+ @{define_ML_type Symbol.symbol = string} \\
+ @{define_ML Symbol.explode: "string -> Symbol.symbol list"} \\
+ @{define_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
+ @{define_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
+ @{define_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
+ @{define_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type "Symbol.sym"} \\
- @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
+ @{define_ML_type "Symbol.sym"} \\
+ @{define_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Symbol.symbol\<close> represents individual Isabelle symbols.
@@ -1347,7 +1347,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type char} \\
+ @{define_ML_type char} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>char\<close> is \<^emph>\<open>not\<close> used. The smallest textual unit in Isabelle
@@ -1359,7 +1359,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type string} \\
+ @{define_ML_type string} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>string\<close> represents immutable vectors of 8-bit characters.
@@ -1407,7 +1407,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type int} \\
+ @{define_ML_type int} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>int\<close> represents regular mathematical integers, which are
@@ -1425,7 +1425,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Rat.rat} \\
+ @{define_ML_type Rat.rat} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Rat.rat\<close> represents rational numbers, based on the
@@ -1444,8 +1444,8 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Time.time} \\
- @{index_ML seconds: "real -> Time.time"} \\
+ @{define_ML_type Time.time} \\
+ @{define_ML seconds: "real -> Time.time"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Time.time\<close> represents time abstractly according to the
@@ -1463,13 +1463,13 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
- @{index_ML is_some: "'a option -> bool"} \\
- @{index_ML is_none: "'a option -> bool"} \\
- @{index_ML the: "'a option -> 'a"} \\
- @{index_ML these: "'a list option -> 'a list"} \\
- @{index_ML the_list: "'a option -> 'a list"} \\
- @{index_ML the_default: "'a -> 'a option -> 'a"} \\
+ @{define_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
+ @{define_ML is_some: "'a option -> bool"} \\
+ @{define_ML is_none: "'a option -> bool"} \\
+ @{define_ML the: "'a option -> 'a"} \\
+ @{define_ML these: "'a list option -> 'a list"} \\
+ @{define_ML the_list: "'a option -> 'a list"} \\
+ @{define_ML the_default: "'a -> 'a option -> 'a"} \\
\end{mldecls}
\<close>
@@ -1490,11 +1490,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML cons: "'a -> 'a list -> 'a list"} \\
- @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
- @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
- @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
- @{index_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
+ @{define_ML cons: "'a -> 'a list -> 'a list"} \\
+ @{define_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
+ @{define_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
+ @{define_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
+ @{define_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>cons\<close>~\<open>x xs\<close> evaluates to \<open>x :: xs\<close>.
@@ -1563,9 +1563,9 @@
text \<open>
\begin{mldecls}
- @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
- @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
- @{index_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\
+ @{define_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
+ @{define_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
+ @{define_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>AList.lookup\<close>, \<^ML>\<open>AList.defined\<close>, \<^ML>\<open>AList.update\<close> implement the
@@ -1593,10 +1593,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "'a Unsynchronized.ref"} \\
- @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
- @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
- @{index_ML_op ":=": "'a Unsynchronized.ref * 'a -> unit"} \\
+ @{define_ML_type 'a "Unsynchronized.ref"} \\
+ @{define_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
+ @{define_ML "!": "'a Unsynchronized.ref -> 'a"} \\
+ @{define_ML_infix ":=" : "'a Unsynchronized.ref * 'a -> unit"} \\
\end{mldecls}
\<close>
@@ -1608,7 +1608,7 @@
The unwieldy name of \<^ML>\<open>Unsynchronized.ref\<close> for the constructor for
references in Isabelle/ML emphasizes the inconveniences caused by
- mutability. Existing operations \<^ML>\<open>!\<close> and \<^ML_op>\<open>:=\<close> are unchanged,
+ mutability. Existing operations \<^ML>\<open>!\<close> and \<^ML_infix>\<open>:=\<close> are unchanged,
but should be used with special precautions, say in a strictly local
situation that is guaranteed to be restricted to sequential evaluation ---
now and in the future.
@@ -1748,8 +1748,8 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML File.tmp_path: "Path.T -> Path.T"} \\
- @{index_ML serial_string: "unit -> string"} \\
+ @{define_ML File.tmp_path: "Path.T -> Path.T"} \\
+ @{define_ML serial_string: "unit -> string"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>File.tmp_path\<close>~\<open>path\<close> relocates the base component of \<open>path\<close> into the
@@ -1790,9 +1790,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "'a Synchronized.var"} \\
- @{index_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
- @{index_ML Synchronized.guarded_access: "'a Synchronized.var ->
+ @{define_ML_type 'a "Synchronized.var"} \\
+ @{define_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
+ @{define_ML Synchronized.guarded_access: "'a Synchronized.var ->
('a -> ('b * 'a) option) -> 'b"} \\
\end{mldecls}
@@ -1890,12 +1890,12 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "'a Exn.result"} \\
- @{index_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
- @{index_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
- @{index_ML Exn.release: "'a Exn.result -> 'a"} \\
- @{index_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\
- @{index_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
+ @{define_ML_type 'a "Exn.result"} \\
+ @{define_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
+ @{define_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
+ @{define_ML Exn.release: "'a Exn.result -> 'a"} \\
+ @{define_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\
+ @{define_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>'a Exn.result\<close> represents the disjoint sum of ML results
@@ -1945,8 +1945,8 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\
- @{index_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
+ @{define_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\
+ @{define_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Par_List.map\<close>~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> is like \<^ML>\<open>map\<close>~\<open>f [x\<^sub>1, \<dots>,
@@ -2009,10 +2009,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "'a lazy"} \\
- @{index_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\
- @{index_ML Lazy.value: "'a -> 'a lazy"} \\
- @{index_ML Lazy.force: "'a lazy -> 'a"} \\
+ @{define_ML_type 'a "lazy"} \\
+ @{define_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\
+ @{define_ML Lazy.value: "'a -> 'a lazy"} \\
+ @{define_ML Lazy.force: "'a lazy -> 'a"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>'a lazy\<close> represents lazy values over type \<^verbatim>\<open>'a\<close>.
@@ -2090,17 +2090,17 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type "'a future"} \\
- @{index_ML Future.fork: "(unit -> 'a) -> 'a future"} \\
- @{index_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\
- @{index_ML Future.join: "'a future -> 'a"} \\
- @{index_ML Future.joins: "'a future list -> 'a list"} \\
- @{index_ML Future.value: "'a -> 'a future"} \\
- @{index_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\
- @{index_ML Future.cancel: "'a future -> unit"} \\
- @{index_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex]
- @{index_ML Future.promise: "(unit -> unit) -> 'a future"} \\
- @{index_ML Future.fulfill: "'a future -> 'a -> unit"} \\
+ @{define_ML_type 'a "future"} \\
+ @{define_ML Future.fork: "(unit -> 'a) -> 'a future"} \\
+ @{define_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\
+ @{define_ML Future.join: "'a future -> 'a"} \\
+ @{define_ML Future.joins: "'a future list -> 'a list"} \\
+ @{define_ML Future.value: "'a -> 'a future"} \\
+ @{define_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\
+ @{define_ML Future.cancel: "'a future -> unit"} \\
+ @{define_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex]
+ @{define_ML Future.promise: "(unit -> unit) -> 'a future"} \\
+ @{define_ML Future.fulfill: "'a future -> 'a -> unit"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>'a future\<close> represents future values over type \<^verbatim>\<open>'a\<close>.
--- a/src/Doc/Implementation/Prelim.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Prelim.thy Sun May 23 20:34:43 2021 +0200
@@ -108,12 +108,12 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type theory} \\
- @{index_ML Context.eq_thy: "theory * theory -> bool"} \\
- @{index_ML Context.subthy: "theory * theory -> bool"} \\
- @{index_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\
- @{index_ML Theory.parents_of: "theory -> theory list"} \\
- @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
+ @{define_ML_type theory} \\
+ @{define_ML Context.eq_thy: "theory * theory -> bool"} \\
+ @{define_ML Context.subthy: "theory * theory -> bool"} \\
+ @{define_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\
+ @{define_ML Theory.parents_of: "theory -> theory list"} \\
+ @{define_ML Theory.ancestors_of: "theory -> theory list"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>theory\<close> represents theory contexts.
@@ -187,10 +187,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Proof.context} \\
- @{index_ML Proof_Context.init_global: "theory -> Proof.context"} \\
- @{index_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
- @{index_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
+ @{define_ML_type Proof.context} \\
+ @{define_ML Proof_Context.init_global: "theory -> Proof.context"} \\
+ @{define_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
+ @{define_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Proof.context\<close> represents proof contexts.
@@ -236,9 +236,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Context.generic} \\
- @{index_ML Context.theory_of: "Context.generic -> theory"} \\
- @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
+ @{define_ML_type Context.generic} \\
+ @{define_ML Context.theory_of: "Context.generic -> theory"} \\
+ @{define_ML Context.proof_of: "Context.generic -> Proof.context"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Context.generic\<close> is the direct sum of \<^ML_type>\<open>theory\<close>
@@ -339,9 +339,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_functor Theory_Data} \\
- @{index_ML_functor Proof_Data} \\
- @{index_ML_functor Generic_Data} \\
+ @{define_ML_functor Theory_Data} \\
+ @{define_ML_functor Proof_Data} \\
+ @{define_ML_functor Generic_Data} \\
\end{mldecls}
\<^descr> \<^ML_functor>\<open>Theory_Data\<close>\<open>(spec)\<close> declares data for type \<^ML_type>\<open>theory\<close>
@@ -482,15 +482,15 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
- @{index_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\
- @{index_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) ->
+ @{define_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
+ @{define_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\
+ @{define_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) ->
bool Config.T"} \\
- @{index_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) ->
+ @{define_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) ->
int Config.T"} \\
- @{index_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) ->
+ @{define_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) ->
real Config.T"} \\
- @{index_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) ->
+ @{define_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) ->
string Config.T"} \\
\end{mldecls}
@@ -613,18 +613,18 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Name.internal: "string -> string"} \\
- @{index_ML Name.skolem: "string -> string"} \\
+ @{define_ML Name.internal: "string -> string"} \\
+ @{define_ML Name.skolem: "string -> string"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type Name.context} \\
- @{index_ML Name.context: Name.context} \\
- @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
- @{index_ML Name.invent: "Name.context -> string -> int -> string list"} \\
- @{index_ML Name.variant: "string -> Name.context -> string * Name.context"} \\
+ @{define_ML_type Name.context} \\
+ @{define_ML Name.context: Name.context} \\
+ @{define_ML Name.declare: "string -> Name.context -> Name.context"} \\
+ @{define_ML Name.invent: "Name.context -> string -> int -> string list"} \\
+ @{define_ML Name.variant: "string -> Name.context -> string * Name.context"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
+ @{define_ML Variable.names_of: "Proof.context -> Name.context"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Name.internal\<close>~\<open>name\<close> produces an internal name by adding one
@@ -720,7 +720,7 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type indexname: "string * int"} \\
+ @{define_ML_type indexname = "string * int"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>indexname\<close> represents indexed names. This is an
@@ -755,11 +755,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Long_Name.base_name: "string -> string"} \\
- @{index_ML Long_Name.qualifier: "string -> string"} \\
- @{index_ML Long_Name.append: "string -> string -> string"} \\
- @{index_ML Long_Name.implode: "string list -> string"} \\
- @{index_ML Long_Name.explode: "string -> string list"} \\
+ @{define_ML Long_Name.base_name: "string -> string"} \\
+ @{define_ML Long_Name.qualifier: "string -> string"} \\
+ @{define_ML Long_Name.append: "string -> string -> string"} \\
+ @{define_ML Long_Name.implode: "string list -> string"} \\
+ @{define_ML Long_Name.explode: "string -> string list"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Long_Name.base_name\<close>~\<open>name\<close> returns the base name of a long name.
@@ -832,29 +832,29 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type binding} \\
- @{index_ML Binding.empty: binding} \\
- @{index_ML Binding.name: "string -> binding"} \\
- @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
- @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
- @{index_ML Binding.concealed: "binding -> binding"} \\
- @{index_ML Binding.print: "binding -> string"} \\
+ @{define_ML_type binding} \\
+ @{define_ML Binding.empty: binding} \\
+ @{define_ML Binding.name: "string -> binding"} \\
+ @{define_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
+ @{define_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
+ @{define_ML Binding.concealed: "binding -> binding"} \\
+ @{define_ML Binding.print: "binding -> string"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type Name_Space.naming} \\
- @{index_ML Name_Space.global_naming: Name_Space.naming} \\
- @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
- @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
+ @{define_ML_type Name_Space.naming} \\
+ @{define_ML Name_Space.global_naming: Name_Space.naming} \\
+ @{define_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
+ @{define_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML_type Name_Space.T} \\
- @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
- @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
- @{index_ML Name_Space.declare: "Context.generic -> bool ->
+ @{define_ML_type Name_Space.T} \\
+ @{define_ML Name_Space.empty: "string -> Name_Space.T"} \\
+ @{define_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
+ @{define_ML Name_Space.declare: "Context.generic -> bool ->
binding -> Name_Space.T -> string * Name_Space.T"} \\
- @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
- @{index_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\
- @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
+ @{define_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
+ @{define_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\
+ @{define_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>binding\<close> represents the abstract concept of name bindings.
--- a/src/Doc/Implementation/Proof.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Proof.thy Sun May 23 20:34:43 2021 +0200
@@ -92,18 +92,18 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Variable.add_fixes: "
+ @{define_ML Variable.add_fixes: "
string list -> Proof.context -> string list * Proof.context"} \\
- @{index_ML Variable.variant_fixes: "
+ @{define_ML Variable.variant_fixes: "
string list -> Proof.context -> string list * Proof.context"} \\
- @{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\
- @{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
- @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
- @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
- @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
+ @{define_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\
+ @{define_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
+ @{define_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
+ @{define_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
+ @{define_ML Variable.import: "bool -> thm list -> Proof.context ->
((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list)
* Proof.context"} \\
- @{index_ML Variable.focus: "binding list option -> term -> Proof.context ->
+ @{define_ML Variable.focus: "binding list option -> term -> Proof.context ->
((string * (string * typ)) list * term) * Proof.context"} \\
\end{mldecls}
@@ -263,14 +263,14 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type Assumption.export} \\
- @{index_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\
- @{index_ML Assumption.add_assms:
+ @{define_ML_type Assumption.export} \\
+ @{define_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\
+ @{define_ML Assumption.add_assms:
"Assumption.export ->
cterm list -> Proof.context -> thm list * Proof.context"} \\
- @{index_ML Assumption.add_assumes: "
+ @{define_ML Assumption.add_assumes: "
cterm list -> Proof.context -> thm list * Proof.context"} \\
- @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
+ @{define_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>Assumption.export\<close> represents arbitrary export rules, which
@@ -359,31 +359,31 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) ->
+ @{define_ML SUBPROOF: "(Subgoal.focus -> tactic) ->
Proof.context -> int -> tactic"} \\
- @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) ->
+ @{define_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) ->
Proof.context -> int -> tactic"} \\
- @{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) ->
+ @{define_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) ->
Proof.context -> int -> tactic"} \\
- @{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) ->
+ @{define_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) ->
Proof.context -> int -> tactic"} \\
- @{index_ML Subgoal.focus: "Proof.context -> int -> binding list option ->
+ @{define_ML Subgoal.focus: "Proof.context -> int -> binding list option ->
thm -> Subgoal.focus * thm"} \\
- @{index_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option ->
+ @{define_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option ->
thm -> Subgoal.focus * thm"} \\
- @{index_ML Subgoal.focus_params: "Proof.context -> int -> binding list option ->
+ @{define_ML Subgoal.focus_params: "Proof.context -> int -> binding list option ->
thm -> Subgoal.focus * thm"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
+ @{define_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\
- @{index_ML Goal.prove_common: "Proof.context -> int option ->
+ @{define_ML Goal.prove_common: "Proof.context -> int option ->
string list -> term list -> term list ->
({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Obtain.result: "(Proof.context -> tactic) -> thm list ->
+ @{define_ML Obtain.result: "(Proof.context -> tactic) -> thm list ->
Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\
\end{mldecls}
--- a/src/Doc/Implementation/Syntax.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Syntax.thy Sun May 23 20:34:43 2021 +0200
@@ -69,16 +69,16 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Syntax.read_typs: "Proof.context -> string list -> typ list"} \\
- @{index_ML Syntax.read_terms: "Proof.context -> string list -> term list"} \\
- @{index_ML Syntax.read_props: "Proof.context -> string list -> term list"} \\[0.5ex]
- @{index_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\
- @{index_ML Syntax.read_term: "Proof.context -> string -> term"} \\
- @{index_ML Syntax.read_prop: "Proof.context -> string -> term"} \\[0.5ex]
- @{index_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\
- @{index_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\
- @{index_ML Syntax.string_of_typ: "Proof.context -> typ -> string"} \\
- @{index_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\
+ @{define_ML Syntax.read_typs: "Proof.context -> string list -> typ list"} \\
+ @{define_ML Syntax.read_terms: "Proof.context -> string list -> term list"} \\
+ @{define_ML Syntax.read_props: "Proof.context -> string list -> term list"} \\[0.5ex]
+ @{define_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\
+ @{define_ML Syntax.read_term: "Proof.context -> string -> term"} \\
+ @{define_ML Syntax.read_prop: "Proof.context -> string -> term"} \\[0.5ex]
+ @{define_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\
+ @{define_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\
+ @{define_ML Syntax.string_of_typ: "Proof.context -> typ -> string"} \\
+ @{define_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Syntax.read_typs\<close>~\<open>ctxt strs\<close> parses and checks a simultaneous list
@@ -158,11 +158,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\
- @{index_ML Syntax.parse_term: "Proof.context -> string -> term"} \\
- @{index_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\[0.5ex]
- @{index_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\
- @{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
+ @{define_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\
+ @{define_ML Syntax.parse_term: "Proof.context -> string -> term"} \\
+ @{define_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\[0.5ex]
+ @{define_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\
+ @{define_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Syntax.parse_typ\<close>~\<open>ctxt str\<close> parses a source string as pre-type that
@@ -219,11 +219,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\
- @{index_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\
- @{index_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\[0.5ex]
- @{index_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\
- @{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
+ @{define_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\
+ @{define_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\
+ @{define_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\[0.5ex]
+ @{define_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\
+ @{define_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Syntax.check_typs\<close>~\<open>ctxt Ts\<close> checks a simultaneous list of pre-types
--- a/src/Doc/Implementation/Tactic.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/Tactic.thy Sun May 23 20:34:43 2021 +0200
@@ -61,10 +61,10 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Goal.init: "cterm -> thm"} \\
- @{index_ML Goal.finish: "Proof.context -> thm -> thm"} \\
- @{index_ML Goal.protect: "int -> thm -> thm"} \\
- @{index_ML Goal.conclude: "thm -> thm"} \\
+ @{define_ML Goal.init: "cterm -> thm"} \\
+ @{define_ML Goal.finish: "Proof.context -> thm -> thm"} \\
+ @{define_ML Goal.protect: "int -> thm -> thm"} \\
+ @{define_ML Goal.conclude: "thm -> thm"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Goal.init\<close>~\<open>C\<close> initializes a tactical goal from the well-formed
@@ -156,15 +156,15 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML_type tactic: "thm -> thm Seq.seq"} \\
- @{index_ML no_tac: tactic} \\
- @{index_ML all_tac: tactic} \\
- @{index_ML print_tac: "Proof.context -> string -> tactic"} \\[1ex]
- @{index_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex]
- @{index_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\
- @{index_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\
- @{index_ML SELECT_GOAL: "tactic -> int -> tactic"} \\
- @{index_ML PREFER_GOAL: "tactic -> int -> tactic"} \\
+ @{define_ML_type tactic = "thm -> thm Seq.seq"} \\
+ @{define_ML no_tac: tactic} \\
+ @{define_ML all_tac: tactic} \\
+ @{define_ML print_tac: "Proof.context -> string -> tactic"} \\[1ex]
+ @{define_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex]
+ @{define_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\
+ @{define_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\
+ @{define_ML SELECT_GOAL: "tactic -> int -> tactic"} \\
+ @{define_ML PREFER_GOAL: "tactic -> int -> tactic"} \\
\end{mldecls}
\<^descr> Type \<^ML_type>\<open>tactic\<close> represents tactics. The well-formedness conditions
@@ -243,17 +243,17 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML resolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML eresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML dresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML forward_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML biresolve_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\[1ex]
- @{index_ML assume_tac: "Proof.context -> int -> tactic"} \\
- @{index_ML eq_assume_tac: "int -> tactic"} \\[1ex]
- @{index_ML match_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML ematch_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML dmatch_tac: "Proof.context -> thm list -> int -> tactic"} \\
- @{index_ML bimatch_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\
+ @{define_ML resolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML eresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML dresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML forward_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML biresolve_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\[1ex]
+ @{define_ML assume_tac: "Proof.context -> int -> tactic"} \\
+ @{define_ML eq_assume_tac: "int -> tactic"} \\[1ex]
+ @{define_ML match_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML ematch_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML dmatch_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{define_ML bimatch_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>resolve_tac\<close>~\<open>ctxt thms i\<close> refines the goal state using the given
@@ -351,23 +351,23 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML Rule_Insts.res_inst_tac: "Proof.context ->
+ @{define_ML Rule_Insts.res_inst_tac: "Proof.context ->
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->
thm -> int -> tactic"} \\
- @{index_ML Rule_Insts.eres_inst_tac: "Proof.context ->
+ @{define_ML Rule_Insts.eres_inst_tac: "Proof.context ->
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->
thm -> int -> tactic"} \\
- @{index_ML Rule_Insts.dres_inst_tac: "Proof.context ->
+ @{define_ML Rule_Insts.dres_inst_tac: "Proof.context ->
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->
thm -> int -> tactic"} \\
- @{index_ML Rule_Insts.forw_inst_tac: "Proof.context ->
+ @{define_ML Rule_Insts.forw_inst_tac: "Proof.context ->
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->
thm -> int -> tactic"} \\
- @{index_ML Rule_Insts.subgoal_tac: "Proof.context -> string ->
+ @{define_ML Rule_Insts.subgoal_tac: "Proof.context -> string ->
(binding * string option * mixfix) list -> int -> tactic"} \\
- @{index_ML Rule_Insts.thin_tac: "Proof.context -> string ->
+ @{define_ML Rule_Insts.thin_tac: "Proof.context -> string ->
(binding * string option * mixfix) list -> int -> tactic"} \\
- @{index_ML rename_tac: "string list -> int -> tactic"} \\
+ @{define_ML rename_tac: "string list -> int -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>Rule_Insts.res_inst_tac\<close>~\<open>ctxt insts thm i\<close> instantiates the rule
@@ -415,9 +415,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML rotate_tac: "int -> int -> tactic"} \\
- @{index_ML distinct_subgoals_tac: tactic} \\
- @{index_ML flexflex_tac: "Proof.context -> tactic"} \\
+ @{define_ML rotate_tac: "int -> int -> tactic"} \\
+ @{define_ML distinct_subgoals_tac: tactic} \\
+ @{define_ML flexflex_tac: "Proof.context -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>rotate_tac\<close>~\<open>n i\<close> rotates the premises of subgoal \<open>i\<close> by \<open>n\<close>
@@ -450,9 +450,9 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML compose_tac: "Proof.context -> (bool * thm * int) -> int -> tactic"} \\
- @{index_ML Drule.compose: "thm * int * thm -> thm"} \\
- @{index_ML_op COMP: "thm * thm -> thm"} \\
+ @{define_ML compose_tac: "Proof.context -> (bool * thm * int) -> int -> tactic"} \\
+ @{define_ML Drule.compose: "thm * int * thm -> thm"} \\
+ @{define_ML_infix COMP: "thm * thm -> thm"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>compose_tac\<close>~\<open>ctxt (flag, rule, m) i\<close> refines subgoal \<open>i\<close> using
@@ -494,53 +494,53 @@
text \<open>
Sequential composition and alternative choices are the most basic ways to
combine tactics, similarly to ``\<^verbatim>\<open>,\<close>'' and ``\<^verbatim>\<open>|\<close>'' in Isar method notation.
- This corresponds to \<^ML_op>\<open>THEN\<close> and \<^ML_op>\<open>ORELSE\<close> in ML, but there
+ This corresponds to \<^ML_infix>\<open>THEN\<close> and \<^ML_infix>\<open>ORELSE\<close> in ML, but there
are further possibilities for fine-tuning alternation of tactics such as
- \<^ML_op>\<open>APPEND\<close>. Further details become visible in ML due to explicit
+ \<^ML_infix>\<open>APPEND\<close>. Further details become visible in ML due to explicit
subgoal addressing.
\<close>
text %mlref \<open>
\begin{mldecls}
- @{index_ML_op "THEN": "tactic * tactic -> tactic"} \\
- @{index_ML_op "ORELSE": "tactic * tactic -> tactic"} \\
- @{index_ML_op "APPEND": "tactic * tactic -> tactic"} \\
- @{index_ML "EVERY": "tactic list -> tactic"} \\
- @{index_ML "FIRST": "tactic list -> tactic"} \\[0.5ex]
+ @{define_ML_infix "THEN": "tactic * tactic -> tactic"} \\
+ @{define_ML_infix "ORELSE": "tactic * tactic -> tactic"} \\
+ @{define_ML_infix "APPEND": "tactic * tactic -> tactic"} \\
+ @{define_ML "EVERY": "tactic list -> tactic"} \\
+ @{define_ML "FIRST": "tactic list -> tactic"} \\[0.5ex]
- @{index_ML_op "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
- @{index_ML_op "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
- @{index_ML_op "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
- @{index_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\
- @{index_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\
+ @{define_ML_infix "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{define_ML_infix "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{define_ML_infix "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{define_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\
+ @{define_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\
\end{mldecls}
- \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_op>\<open>THEN\<close>~\<open>tac\<^sub>2\<close> is the sequential composition of \<open>tac\<^sub>1\<close> and
+ \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>2\<close> is the sequential composition of \<open>tac\<^sub>1\<close> and
\<open>tac\<^sub>2\<close>. Applied to a goal state, it returns all states reachable in two
steps by applying \<open>tac\<^sub>1\<close> followed by \<open>tac\<^sub>2\<close>. First, it applies \<open>tac\<^sub>1\<close> to
the goal state, getting a sequence of possible next states; then, it applies
\<open>tac\<^sub>2\<close> to each of these and concatenates the results to produce again one
flat sequence of states.
- \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>tac\<^sub>2\<close> makes a choice between \<open>tac\<^sub>1\<close> and
+ \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac\<^sub>2\<close> makes a choice between \<open>tac\<^sub>1\<close> and
\<open>tac\<^sub>2\<close>. Applied to a state, it tries \<open>tac\<^sub>1\<close> and returns the result if
non-empty; if \<open>tac\<^sub>1\<close> fails then it uses \<open>tac\<^sub>2\<close>. This is a deterministic
choice: if \<open>tac\<^sub>1\<close> succeeds then \<open>tac\<^sub>2\<close> is excluded from the result.
- \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_op>\<open>APPEND\<close>~\<open>tac\<^sub>2\<close> concatenates the possible results of
- \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>. Unlike \<^ML_op>\<open>ORELSE\<close> there is \<^emph>\<open>no commitment\<close> to
- either tactic, so \<^ML_op>\<open>APPEND\<close> helps to avoid incompleteness during
+ \<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>APPEND\<close>~\<open>tac\<^sub>2\<close> concatenates the possible results of
+ \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>. Unlike \<^ML_infix>\<open>ORELSE\<close> there is \<^emph>\<open>no commitment\<close> to
+ either tactic, so \<^ML_infix>\<open>APPEND\<close> helps to avoid incompleteness during
search, at the cost of potential inefficiencies.
- \<^descr> \<^ML>\<open>EVERY\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_op>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>THEN\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>EVERY []\<close> is the same as
+ \<^descr> \<^ML>\<open>EVERY\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>EVERY []\<close> is the same as
\<^ML>\<open>all_tac\<close>: it always succeeds.
- \<^descr> \<^ML>\<open>FIRST\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>FIRST []\<close> is the
+ \<^descr> \<^ML>\<open>FIRST\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>FIRST []\<close> is the
same as \<^ML>\<open>no_tac\<close>: it always fails.
- \<^descr> \<^ML_op>\<open>THEN'\<close> is the lifted version of \<^ML_op>\<open>THEN\<close>, for tactics
- with explicit subgoal addressing. So \<open>(tac\<^sub>1\<close>~\<^ML_op>\<open>THEN'\<close>~\<open>tac\<^sub>2) i\<close> is
- the same as \<open>(tac\<^sub>1 i\<close>~\<^ML_op>\<open>THEN\<close>~\<open>tac\<^sub>2 i)\<close>.
+ \<^descr> \<^ML_infix>\<open>THEN'\<close> is the lifted version of \<^ML_infix>\<open>THEN\<close>, for tactics
+ with explicit subgoal addressing. So \<open>(tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN'\<close>~\<open>tac\<^sub>2) i\<close> is
+ the same as \<open>(tac\<^sub>1 i\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>2 i)\<close>.
The other primed tacticals work analogously.
\<close>
@@ -555,11 +555,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML "TRY": "tactic -> tactic"} \\
- @{index_ML "REPEAT": "tactic -> tactic"} \\
- @{index_ML "REPEAT1": "tactic -> tactic"} \\
- @{index_ML "REPEAT_DETERM": "tactic -> tactic"} \\
- @{index_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\
+ @{define_ML "TRY": "tactic -> tactic"} \\
+ @{define_ML "REPEAT": "tactic -> tactic"} \\
+ @{define_ML "REPEAT1": "tactic -> tactic"} \\
+ @{define_ML "REPEAT_DETERM": "tactic -> tactic"} \\
+ @{define_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>TRY\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns the resulting
@@ -567,7 +567,7 @@
applies \<open>tac\<close> at most once.
Note that for tactics with subgoal addressing, the combinator can be applied
- via functional composition: \<^ML>\<open>TRY\<close>~\<^ML_op>\<open>o\<close>~\<open>tac\<close>. There is no need
+ via functional composition: \<^ML>\<open>TRY\<close>~\<^ML_infix>\<open>o\<close>~\<open>tac\<close>. There is no need
for \<^verbatim>\<open>TRY'\<close>.
\<^descr> \<^ML>\<open>REPEAT\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and, recursively, to
@@ -592,10 +592,10 @@
text %mlex \<open>
The basic tactics and tacticals considered above follow some algebraic laws:
- \<^item> \<^ML>\<open>all_tac\<close> is the identity element of the tactical \<^ML_op>\<open>THEN\<close>.
+ \<^item> \<^ML>\<open>all_tac\<close> is the identity element of the tactical \<^ML_infix>\<open>THEN\<close>.
- \<^item> \<^ML>\<open>no_tac\<close> is the identity element of \<^ML_op>\<open>ORELSE\<close> and \<^ML_op>\<open>APPEND\<close>. Also, it is a zero element for \<^ML_op>\<open>THEN\<close>, which means that
- \<open>tac\<close>~\<^ML_op>\<open>THEN\<close>~\<^ML>\<open>no_tac\<close> is equivalent to \<^ML>\<open>no_tac\<close>.
+ \<^item> \<^ML>\<open>no_tac\<close> is the identity element of \<^ML_infix>\<open>ORELSE\<close> and \<^ML_infix>\<open>APPEND\<close>. Also, it is a zero element for \<^ML_infix>\<open>THEN\<close>, which means that
+ \<open>tac\<close>~\<^ML_infix>\<open>THEN\<close>~\<^ML>\<open>no_tac\<close> is equivalent to \<^ML>\<open>no_tac\<close>.
\<^item> \<^ML>\<open>TRY\<close> and \<^ML>\<open>REPEAT\<close> can be expressed as (recursive) functions over
more basic combinators (ignoring some internal implementation tricks):
@@ -607,7 +607,7 @@
\<close>
text \<open>
- If \<open>tac\<close> can return multiple outcomes then so can \<^ML>\<open>REPEAT\<close>~\<open>tac\<close>. \<^ML>\<open>REPEAT\<close> uses \<^ML_op>\<open>ORELSE\<close> and not \<^ML_op>\<open>APPEND\<close>, it applies \<open>tac\<close>
+ If \<open>tac\<close> can return multiple outcomes then so can \<^ML>\<open>REPEAT\<close>~\<open>tac\<close>. \<^ML>\<open>REPEAT\<close> uses \<^ML_infix>\<open>ORELSE\<close> and not \<^ML_infix>\<open>APPEND\<close>, it applies \<open>tac\<close>
as many times as possible in each outcome.
\begin{warn}
@@ -641,20 +641,20 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML ALLGOALS: "(int -> tactic) -> tactic"} \\
- @{index_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\
- @{index_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\
- @{index_ML HEADGOAL: "(int -> tactic) -> tactic"} \\
- @{index_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\
- @{index_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\
- @{index_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\
+ @{define_ML ALLGOALS: "(int -> tactic) -> tactic"} \\
+ @{define_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\
+ @{define_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\
+ @{define_ML HEADGOAL: "(int -> tactic) -> tactic"} \\
+ @{define_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\
+ @{define_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\
+ @{define_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\
\end{mldecls}
- \<^descr> \<^ML>\<open>ALLGOALS\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_op>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>THEN\<close>~\<open>tac 1\<close>. It applies the \<open>tac\<close> to all the subgoals, counting downwards.
+ \<^descr> \<^ML>\<open>ALLGOALS\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac 1\<close>. It applies the \<open>tac\<close> to all the subgoals, counting downwards.
- \<^descr> \<^ML>\<open>SOMEGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>tac 1\<close>. It applies \<open>tac\<close> to one subgoal, counting downwards.
+ \<^descr> \<^ML>\<open>SOMEGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac 1\<close>. It applies \<open>tac\<close> to one subgoal, counting downwards.
- \<^descr> \<^ML>\<open>FIRSTGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac 1\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>ORELSE\<close>~\<open>tac n\<close>. It applies \<open>tac\<close> to one subgoal, counting upwards.
+ \<^descr> \<^ML>\<open>FIRSTGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac 1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac n\<close>. It applies \<open>tac\<close> to one subgoal, counting upwards.
\<^descr> \<^ML>\<open>HEADGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac 1\<close>. It applies \<open>tac\<close>
unconditionally to the first subgoal.
@@ -666,7 +666,7 @@
upwards.
\<^descr> \<^ML>\<open>RANGE\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>k] i\<close> is equivalent to \<open>tac\<^sub>k (i + k -
- 1)\<close>~\<^ML_op>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_op>\<open>THEN\<close>~\<open>tac\<^sub>1 i\<close>. It applies the given list of
+ 1)\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>1 i\<close>. It applies the given list of
tactics to the corresponding range of subgoals, counting downwards.
\<close>
@@ -689,8 +689,8 @@
text \<open>
\begin{mldecls}
- @{index_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\
- @{index_ML CHANGED: "tactic -> tactic"} \\
+ @{define_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\
+ @{define_ML CHANGED: "tactic -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>FILTER\<close>~\<open>sat tac\<close> applies \<open>tac\<close> to the goal state and returns a
@@ -706,15 +706,15 @@
text \<open>
\begin{mldecls}
- @{index_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
- @{index_ML DEPTH_SOLVE: "tactic -> tactic"} \\
- @{index_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\
+ @{define_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
+ @{define_ML DEPTH_SOLVE: "tactic -> tactic"} \\
+ @{define_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>DEPTH_FIRST\<close>~\<open>sat tac\<close> returns the goal state if \<open>sat\<close> returns true.
Otherwise it applies \<open>tac\<close>, then recursively searches from each element of
the resulting sequence. The code uses a stack for efficiency, in effect
- applying \<open>tac\<close>~\<^ML_op>\<open>THEN\<close>~\<^ML>\<open>DEPTH_FIRST\<close>~\<open>sat tac\<close> to the state.
+ applying \<open>tac\<close>~\<^ML_infix>\<open>THEN\<close>~\<^ML>\<open>DEPTH_FIRST\<close>~\<open>sat tac\<close> to the state.
\<^descr> \<^ML>\<open>DEPTH_SOLVE\<close>\<open>tac\<close> uses \<^ML>\<open>DEPTH_FIRST\<close> to search for states having
no subgoals.
@@ -729,9 +729,9 @@
text \<open>
\begin{mldecls}
- @{index_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
- @{index_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
- @{index_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
+ @{define_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
+ @{define_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
+ @{define_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
\end{mldecls}
These search strategies will find a solution if one exists. However, they do
@@ -763,10 +763,10 @@
text \<open>
\begin{mldecls}
- @{index_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\
- @{index_ML IF_UNSOLVED: "tactic -> tactic"} \\
- @{index_ML SOLVE: "tactic -> tactic"} \\
- @{index_ML DETERM: "tactic -> tactic"} \\
+ @{define_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\
+ @{define_ML IF_UNSOLVED: "tactic -> tactic"} \\
+ @{define_ML SOLVE: "tactic -> tactic"} \\
+ @{define_ML DETERM: "tactic -> tactic"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>COND\<close>~\<open>sat tac\<^sub>1 tac\<^sub>2\<close> applies \<open>tac\<^sub>1\<close> to the goal state if it
@@ -792,10 +792,10 @@
text \<open>
\begin{mldecls}
- @{index_ML has_fewer_prems: "int -> thm -> bool"} \\
- @{index_ML Thm.eq_thm: "thm * thm -> bool"} \\
- @{index_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\
- @{index_ML size_of_thm: "thm -> int"} \\
+ @{define_ML has_fewer_prems: "int -> thm -> bool"} \\
+ @{define_ML Thm.eq_thm: "thm * thm -> bool"} \\
+ @{define_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\
+ @{define_ML size_of_thm: "thm -> int"} \\
\end{mldecls}
\<^descr> \<^ML>\<open>has_fewer_prems\<close>~\<open>n thm\<close> reports whether \<open>thm\<close> has fewer than \<open>n\<close>
--- a/src/Doc/Implementation/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo Isar
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Implementation/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Implementation/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -16,7 +16,7 @@
\hyphenation{Isar}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_isar}
+\title{\includegraphics[scale=0.5]{isabelle_logo}
\\[4ex] The Isabelle/Isar Implementation}
\author{\emph{Makarius Wenzel} \\[3ex]
With Contributions by
--- a/src/Doc/Intro/document/build Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Intro/document/build Sun May 23 20:34:43 2021 +0200
@@ -2,9 +2,18 @@
set -e
-FORMAT="$1"
-VARIANT="$2"
+$ISABELLE_LUALATEX root
+
+if [ -f manual.bib -o -f root.bib ]
+then
+ $ISABELLE_BIBTEX root
+ $ISABELLE_LUALATEX root
+fi
-isabelle logo
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
+$ISABELLE_LUALATEX root
+if [ -f root.idx ]
+then
+ "$ISABELLE_HOME/src/Doc/sedindex" root
+ $ISABELLE_LUALATEX root
+fi
--- a/src/Doc/Intro/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Intro/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -6,7 +6,7 @@
%prth *(\(.*\)); \1;
%{\\out \(.*\)} {\\out val it = "\1" : thm}
-\title{\includegraphics[scale=0.5]{isabelle} \\[4ex] Old Introduction to Isabelle}
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] Old Introduction to Isabelle}
\author{{\em Lawrence C. Paulson}\\
Computer Laboratory \\ University of Cambridge \\
\texttt{lcp@cl.cam.ac.uk}\\[3ex]
--- a/src/Doc/Isar_Ref/Document_Preparation.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Document_Preparation.thy Sun May 23 20:34:43 2021 +0200
@@ -100,11 +100,22 @@
@{antiquotation_def subgoals} & : & \<open>antiquotation\<close> \\
@{antiquotation_def prf} & : & \<open>antiquotation\<close> \\
@{antiquotation_def full_prf} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_text} & : & \<open>antiquotation\<close> \\
@{antiquotation_def ML} & : & \<open>antiquotation\<close> \\
- @{antiquotation_def ML_op} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_def} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_ref} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_infix} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_infix_def} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_infix_ref} & : & \<open>antiquotation\<close> \\
@{antiquotation_def ML_type} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_type_def} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_type_ref} & : & \<open>antiquotation\<close> \\
@{antiquotation_def ML_structure} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_structure_def} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_structure_ref} & : & \<open>antiquotation\<close> \\
@{antiquotation_def ML_functor} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_functor_def} & : & \<open>antiquotation\<close> \\
+ @{antiquotation_def ML_functor_ref} & : & \<open>antiquotation\<close> \\
@{antiquotation_def emph} & : & \<open>antiquotation\<close> \\
@{antiquotation_def bold} & : & \<open>antiquotation\<close> \\
@{antiquotation_def verbatim} & : & \<open>antiquotation\<close> \\
@@ -192,9 +203,10 @@
@@{antiquotation subgoals} options |
@@{antiquotation prf} options @{syntax thms} |
@@{antiquotation full_prf} options @{syntax thms} |
+ @@{antiquotation ML_text} options @{syntax text} |
@@{antiquotation ML} options @{syntax text} |
- @@{antiquotation ML_op} options @{syntax text} |
- @@{antiquotation ML_type} options @{syntax text} |
+ @@{antiquotation ML_infix} options @{syntax text} |
+ @@{antiquotation ML_type} options @{syntax typeargs} @{syntax text} |
@@{antiquotation ML_structure} options @{syntax text} |
@@{antiquotation ML_functor} options @{syntax text} |
@@{antiquotation emph} options @{syntax text} |
@@ -280,10 +292,22 @@
\<^descr> \<open>@{full_prf a\<^sub>1 \<dots> a\<^sub>n}\<close> is like \<open>@{prf a\<^sub>1 \<dots> a\<^sub>n}\<close>, but prints the full
proof terms, i.e.\ also displays information omitted in the compact proof
term, which is denoted by ``\<open>_\<close>'' placeholders there.
-
- \<^descr> \<open>@{ML s}\<close>, \<open>@{ML_op s}\<close>, \<open>@{ML_type s}\<close>, \<open>@{ML_structure s}\<close>, and
+
+ \<^descr> \<open>@{ML_text s}\<close> prints ML text verbatim: only the token language is
+ checked.
+
+ \<^descr> \<open>@{ML s}\<close>, \<open>@{ML_infix s}\<close>, \<open>@{ML_type s}\<close>, \<open>@{ML_structure s}\<close>, and
\<open>@{ML_functor s}\<close> check text \<open>s\<close> as ML value, infix operator, type,
- structure, and functor respectively. The source is printed verbatim.
+ exception, structure, and functor respectively. The source is printed
+ verbatim. The variants \<open>@{ML_def s}\<close> and \<open>@{ML_ref s}\<close> etc. maintain the
+ document index: ``def'' means to make a bold entry, ``ref'' means to make a
+ regular entry.
+
+ There are two forms for type constructors, with or without separate type
+ arguments: this impacts only the index entry. For example, \<open>@{ML_type_ref
+ \<open>'a list\<close>}\<close> makes an entry literally for ``\<open>'a list\<close>'' (sorted under the
+ letter ``a''), but \<open>@{ML_type_ref 'a \<open>list\<close>}\<close> makes an entry for the
+ constructor name ``\<open>list\<close>''.
\<^descr> \<open>@{emph s}\<close> prints document source recursively, with {\LaTeX} markup
\<^verbatim>\<open>\emph{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>.
--- a/src/Doc/Isar_Ref/Generic.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Generic.thy Sun May 23 20:34:43 2021 +0200
@@ -153,7 +153,7 @@
\<^descr> @{attribute THEN}~\<open>a\<close> composes rules by resolution; it resolves with the
first premise of \<open>a\<close> (an alternative position may be also specified). See
- also \<^ML_op>\<open>RS\<close> in @{cite "isabelle-implementation"}.
+ also \<^ML_infix>\<open>RS\<close> in @{cite "isabelle-implementation"}.
\<^descr> @{attribute unfolded}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> and @{attribute folded}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close>
expand and fold back again the given definitions throughout a rule.
@@ -638,7 +638,7 @@
of the term ordering.
The default is lexicographic ordering of term structure, but this could be
- also changed locally for special applications via @{index_ML
+ also changed locally for special applications via @{define_ML
Simplifier.set_term_ord} in Isabelle/ML.
\<^medskip>
@@ -866,9 +866,9 @@
text \<open>
\begin{mldecls}
- @{index_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) ->
+ @{define_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) ->
Proof.context -> Proof.context"} \\
- @{index_ML Simplifier.prems_of: "Proof.context -> thm list"} \\
+ @{define_ML Simplifier.prems_of: "Proof.context -> thm list"} \\
\end{mldecls}
The subgoaler is the tactic used to solve subgoals arising out of
@@ -906,13 +906,13 @@
text \<open>
\begin{mldecls}
- @{index_ML_type solver} \\
- @{index_ML Simplifier.mk_solver: "string ->
+ @{define_ML_type solver} \\
+ @{define_ML Simplifier.mk_solver: "string ->
(Proof.context -> int -> tactic) -> solver"} \\
- @{index_ML_op setSolver: "Proof.context * solver -> Proof.context"} \\
- @{index_ML_op addSolver: "Proof.context * solver -> Proof.context"} \\
- @{index_ML_op setSSolver: "Proof.context * solver -> Proof.context"} \\
- @{index_ML_op addSSolver: "Proof.context * solver -> Proof.context"} \\
+ @{define_ML_infix setSolver: "Proof.context * solver -> Proof.context"} \\
+ @{define_ML_infix addSolver: "Proof.context * solver -> Proof.context"} \\
+ @{define_ML_infix setSSolver: "Proof.context * solver -> Proof.context"} \\
+ @{define_ML_infix addSSolver: "Proof.context * solver -> Proof.context"} \\
\end{mldecls}
A solver is a tactic that attempts to solve a subgoal after simplification.
@@ -992,17 +992,17 @@
text \<open>
\begin{mldecls}
- @{index_ML_op setloop: "Proof.context *
+ @{define_ML_infix setloop: "Proof.context *
(Proof.context -> int -> tactic) -> Proof.context"} \\
- @{index_ML_op addloop: "Proof.context *
+ @{define_ML_infix addloop: "Proof.context *
(string * (Proof.context -> int -> tactic))
-> Proof.context"} \\
- @{index_ML_op delloop: "Proof.context * string -> Proof.context"} \\
- @{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
- @{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
- @{index_ML Splitter.add_split_bang: "
+ @{define_ML_infix delloop: "Proof.context * string -> Proof.context"} \\
+ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
+ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
+ @{define_ML Splitter.add_split_bang: "
thm -> Proof.context -> Proof.context"} \\
- @{index_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\
+ @{define_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\
\end{mldecls}
The looper is a list of tactics that are applied after simplification, in
@@ -1624,23 +1624,23 @@
text \<open>
\begin{mldecls}
- @{index_ML_type wrapper: "(int -> tactic) -> (int -> tactic)"} \\[0.5ex]
- @{index_ML_op addSWrapper: "Proof.context *
+ @{define_ML_type wrapper = "(int -> tactic) -> (int -> tactic)"} \\[0.5ex]
+ @{define_ML_infix addSWrapper: "Proof.context *
(string * (Proof.context -> wrapper)) -> Proof.context"} \\
- @{index_ML_op addSbefore: "Proof.context *
+ @{define_ML_infix addSbefore: "Proof.context *
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
- @{index_ML_op addSafter: "Proof.context *
+ @{define_ML_infix addSafter: "Proof.context *
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
- @{index_ML_op delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
- @{index_ML_op addWrapper: "Proof.context *
+ @{define_ML_infix delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
+ @{define_ML_infix addWrapper: "Proof.context *
(string * (Proof.context -> wrapper)) -> Proof.context"} \\
- @{index_ML_op addbefore: "Proof.context *
+ @{define_ML_infix addbefore: "Proof.context *
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
- @{index_ML_op addafter: "Proof.context *
+ @{define_ML_infix addafter: "Proof.context *
(string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
- @{index_ML_op delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
- @{index_ML addSss: "Proof.context -> Proof.context"} \\
- @{index_ML addss: "Proof.context -> Proof.context"} \\
+ @{define_ML_infix delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
+ @{define_ML addSss: "Proof.context -> Proof.context"} \\
+ @{define_ML addss: "Proof.context -> Proof.context"} \\
\end{mldecls}
The proof strategy of the Classical Reasoner is simple. Perform as
--- a/src/Doc/Isar_Ref/Inner_Syntax.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Inner_Syntax.thy Sun May 23 20:34:43 2021 +0200
@@ -220,8 +220,8 @@
text \<open>
\begin{mldecls}
- @{index_ML print_mode_value: "unit -> string list"} \\
- @{index_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\
+ @{define_ML print_mode_value: "unit -> string list"} \\
+ @{define_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\
\end{mldecls}
The \<^emph>\<open>print mode\<close> facility allows to modify various operations for printing.
--- a/src/Doc/Isar_Ref/Outer_Syntax.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Outer_Syntax.thy Sun May 23 20:34:43 2021 +0200
@@ -343,12 +343,16 @@
no infixes.
\<^rail>\<open>
- @{syntax_def typespec}:
- (() | @{syntax type_ident} | '(' ( @{syntax type_ident} + ',' ) ')') @{syntax name}
+ @{syntax_def typeargs}:
+ (() | @{syntax type_ident} | '(' ( @{syntax type_ident} + ',' ) ')')
;
- @{syntax_def typespec_sorts}:
+ @{syntax_def typeargs_sorts}:
(() | (@{syntax type_ident} ('::' @{syntax sort})?) |
- '(' ( (@{syntax type_ident} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name}
+ '(' ( (@{syntax type_ident} ('::' @{syntax sort})?) + ',' ) ')')
+ ;
+ @{syntax_def typespec}: @{syntax typeargs} @{syntax name}
+ ;
+ @{syntax_def typespec_sorts}: @{syntax typeargs_sorts} @{syntax name}
\<close>
\<close>
--- a/src/Doc/Isar_Ref/Proof.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Proof.thy Sun May 23 20:34:43 2021 +0200
@@ -626,7 +626,7 @@
Structural composition ``\<open>m\<^sub>1\<close>\<^verbatim>\<open>;\<close>~\<open>m\<^sub>2\<close>'' means that method \<open>m\<^sub>1\<close> is
applied with restriction to the first subgoal, then \<open>m\<^sub>2\<close> is applied
consecutively with restriction to each subgoal that has newly emerged due to
- \<open>m\<^sub>1\<close>. This is analogous to the tactic combinator \<^ML_op>\<open>THEN_ALL_NEW\<close> in
+ \<open>m\<^sub>1\<close>. This is analogous to the tactic combinator \<^ML_infix>\<open>THEN_ALL_NEW\<close> in
Isabelle/ML, see also @{cite "isabelle-implementation"}. For example, \<open>(rule
r; blast)\<close> applies rule \<open>r\<close> and then solves all new subgoals by \<open>blast\<close>.
--- a/src/Doc/Isar_Ref/Spec.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Spec.thy Sun May 23 20:34:43 2021 +0200
@@ -520,8 +520,11 @@
@{command_def "locale_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
\end{tabular}
- \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes}
- \indexisarelem{defines}\indexisarelem{notes}
+ @{index_ref \<open>\<^theory_text>\<open>fixes\<close> (element)\<close>}
+ @{index_ref \<open>\<^theory_text>\<open>constrains\<close> (element)\<close>}
+ @{index_ref \<open>\<^theory_text>\<open>assumes\<close> (element)\<close>}
+ @{index_ref \<open>\<^theory_text>\<open>defines\<close> (element)\<close>}
+ @{index_ref \<open>\<^theory_text>\<open>notes\<close> (element)\<close>}
\<^rail>\<open>
@@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'?
;
--- a/src/Doc/Isar_Ref/Symbols.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/Symbols.thy Sun May 23 20:34:43 2021 +0200
@@ -32,7 +32,7 @@
\begin{center}
\begin{isabellebody}
- \input{syms}
+ @{show_symbols}
\end{isabellebody}
\end{center}
\<close>
--- a/src/Doc/Isar_Ref/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo Isar
-./showsymbols "$ISABELLE_HOME/lib/texinputs/isabellesym.sty" > syms.tex
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Isar_Ref/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Isar_Ref/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -23,7 +23,7 @@
\hyphenation{Isar}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_isar} \\[4ex] The Isabelle/Isar Reference Manual}
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] The Isabelle/Isar Reference Manual}
\author{\emph{Makarius Wenzel} \\[3ex]
With Contributions by
Clemens Ballarin,
--- a/src/Doc/Isar_Ref/document/showsymbols Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-#!/usr/bin/env perl
-
-print "\\begin{supertabular}{ll\@{\\qquad}ll}\n";
-
-$eol = "&";
-
-while (<ARGV>) {
- if (m/^\\newcommand\{\\isasym([A-Za-z]+)\}/) {
- print "\\verb,\\<$1>, & {\\isasym$1} $eol\n";
- if ("$eol" eq "&") {
- $eol = "\\\\";
- } else {
- $eol = "&";
- }
- }
-}
-
-if ("$eol" eq "\\\\") {
- print "$eol\n";
-}
-
-print "\\end{supertabular}\n";
-
--- a/src/Doc/JEdit/JEdit.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/JEdit/JEdit.thy Sun May 23 20:34:43 2021 +0200
@@ -1911,7 +1911,7 @@
text \<open>
Citations are managed by {\LaTeX} and Bib{\TeX} in \<^verbatim>\<open>.bib\<close> files. The
- Isabelle session build process and the @{tool latex} tool @{cite
+ Isabelle session build process and the @{tool document} tool @{cite
"isabelle-system"} are smart enough to assemble the result, based on the
session directory layout.
--- a/src/Doc/JEdit/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo jEdit
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/JEdit/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/JEdit/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -20,7 +20,7 @@
\isabellestyle{literal}
\def\isastylett{\footnotesize\tt}
-\title{\includegraphics[scale=0.5]{isabelle_jedit} \\[4ex] Isabelle/jEdit}
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] Isabelle/jEdit}
\author{\emph{Makarius Wenzel}}
--- a/src/Doc/Locales/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Logics/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Logics/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Logics/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -8,7 +8,7 @@
%%% to index constants: \\tt \([a-zA-Z0-9][a-zA-Z0-9_]*\) \\cdx{\1}
%%% to deverbify: \\verb|\([^|]*\)| \\ttindex{\1}
%% run ../sedindex logics to prepare index file
-\title{\includegraphics[scale=0.5]{isabelle} \\[4ex] Isabelle's Logics}
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] Isabelle's Logics}
\author{{\em Lawrence C. Paulson}\\
Computer Laboratory \\ University of Cambridge \\
--- a/src/Doc/Logics_ZF/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo ZF
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Logics_ZF/document/logics.sty Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Logics_ZF/document/logics.sty Sun May 23 20:34:43 2021 +0200
@@ -1,1 +1,179 @@
-% logics.sty : Logics Manuals Page Layout
%
\typeout{Document Style logics. Released 18 August 2003}
\hyphenation{Isa-belle man-u-script man-u-scripts ap-pen-dix mut-u-al-ly}
\hyphenation{data-type data-types co-data-type co-data-types }
%usage: \iflabelundefined{LABEL}{if not defined}{if defined}
\newcommand{\iflabelundefined}[1]{\@ifundefined{r@#1}}
%%%INDEXING use isa-index to process the index
\newcommand\seealso[2]{\emph{see also} #1}
\usepackage{makeidx}
%index, putting page numbers of definitions in boldface
\def\bold#1{\textbf{#1}}
\newcommand\fnote[1]{#1n}
\newcommand\indexbold[1]{\index{#1|bold}}
% The alternative to \protect\isa in the indexing macros is
% \noexpand\noexpand \noexpand\isa
% need TWO levels of \noexpand to delay the expansion of \isa:
% the \noexpand\noexpand will leave one \noexpand, to be given to the
% (still unexpanded) \isa token. See TeX by Topic, page 122.
%%%% for indexing constants, symbols, theorems, ...
\newcommand\cdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (constant)}}
\newcommand\sdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (symbol)}}
\newcommand\tdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (theorem)}}
\newcommand\tdxbold[1]{\isa{#1}\index{#1@\protect\isa{#1} (theorem)|bold}}
\newcommand\cldx[1]{\isa{#1}\index{#1@\protect\isa{#1} (class)}}
\newcommand\tydx[1]{\isa{#1}\index{#1@\protect\isa{#1} (type)}}
\newcommand\thydx[1]{\isa{#1}\index{#1@\protect\isa{#1} (theory)}}
\newcommand\attrdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (attribute)}}
\newcommand\cmmdx[1]{\index{#1@\protect\isacommand{#1} (command)}}
\newcommand\commdx[1]{\isacommand{#1}\index{#1@\protect\isacommand{#1} (command)}}
\newcommand\methdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (method)}}
\newcommand\tooldx[1]{\isa{#1}\index{#1@\protect\isa{#1} (tool)}}
\newcommand\settdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (setting)}}
%set argument in \bf font and index in ROMAN font (for definitions in text!)
\newcommand\bfindex[1]{{\bf#1}\index{#1|bold}\@}
\newcommand\rmindex[1]{{#1}\index{#1}\@}
\newcommand\ttindex[1]{\texttt{#1}\index{#1@\texttt{#1}}\@}
\newcommand\ttindexbold[1]{\texttt{#1}\index{#1@\texttt{#1}|bold}\@}
\newcommand{\indexboldpos}[2]{#1\@}
\newcommand{\ttindexboldpos}[2]{\isa{#1}\@}
%\newtheorem{theorem}{Theorem}[section]
\newtheorem{Exercise}{Exercise}[section]
\newenvironment{exercise}{\begin{Exercise}\rm}{\end{Exercise}}
\newcommand{\ttlbr}{\texttt{[|}}
\newcommand{\ttrbr}{\texttt{|]}}
\newcommand{\ttor}{\texttt{|}}
\newcommand{\ttall}{\texttt{!}}
\newcommand{\ttuniquex}{\texttt{?!}}
\newcommand{\ttEXU}{\texttt{EX!}}
\newcommand{\ttAnd}{\texttt{!!}}
\newcommand{\isasymignore}{}
\newcommand{\isasymimp}{\isasymlongrightarrow}
\newcommand{\isasymImp}{\isasymLongrightarrow}
\newcommand{\isasymFun}{\isasymRightarrow}
\newcommand{\isasymuniqex}{\isamath{\exists!\,}}
\renewcommand{\S}{Sect.\ts}
\renewenvironment{isamarkuptxt}{\begin{isamarkuptext}}{\end{isamarkuptext}}
\newif\ifremarks
\newcommand{\REMARK}[1]{\ifremarks\marginpar{\raggedright\footnotesize#1}\fi}
%names of Isabelle rules
\newcommand{\rulename}[1]{\hfill(#1)}
\newcommand{\rulenamedx}[1]{\hfill(#1\index{#1@\protect\isa{#1} (theorem)|bold})}
%%%% meta-logical connectives
\let\Forall=\bigwedge
\let\Imp=\Longrightarrow
\let\To=\Rightarrow
\newcommand{\Var}[1]{{?\!#1}}
%%% underscores as ordinary characters, not for subscripting
%% use @ or \sb for subscripting; use \at for @
%% only works in \tt font
%% must not make _ an active char; would make \ttindex fail!
\gdef\underscoreoff{\catcode`\@=8\catcode`\_=\other}
\gdef\underscoreon{\catcode`\_=8\makeatother}
\chardef\other=12
\chardef\at=`\@
% alternative underscore
\def\_{\leavevmode\kern.06em\vbox{\hrule height.2ex width.3em}\hskip0.1em}
%%%% ``WARNING'' environment
\def\dbend{\vtop to 0pt{\vss\hbox{\Huge\bf!}\vss}}
\newenvironment{warn}{\medskip\medbreak\begingroup \clubpenalty=10000
\small %%WAS\baselineskip=0.9\baselineskip
\noindent \hangindent\parindent \hangafter=-2
\hbox to0pt{\hskip-\hangindent\dbend\hfill}\ignorespaces}%
{\par\endgroup\medbreak}
%%%% Standard logical symbols
\let\turn=\vdash
\let\conj=\wedge
\let\disj=\vee
\let\imp=\rightarrow
\let\bimp=\leftrightarrow
\newcommand\all[1]{\forall#1.} %quantification
\newcommand\ex[1]{\exists#1.}
\newcommand{\pair}[1]{\langle#1\rangle}
\newcommand{\lparr}{\mathopen{(\!|}}
\newcommand{\rparr}{\mathclose{|\!)}}
\newcommand{\fs}{\mathpunct{,\,}}
\newcommand{\ty}{\mathrel{::}}
\newcommand{\asn}{\mathrel{:=}}
\newcommand{\more}{\ldots}
\newcommand{\record}[1]{\lparr #1 \rparr}
\newcommand{\dtt}{\mathord.}
\newcommand\lbrakk{\mathopen{[\![}}
\newcommand\rbrakk{\mathclose{]\!]}}
\newcommand\List[1]{\lbrakk#1\rbrakk} %was \obj
\newcommand\vpile[1]{\begin{array}{c}#1\end{array}}
\newenvironment{matharray}[1]{\[\begin{array}{#1}}{\end{array}\]}
\newcommand{\Text}[1]{\mbox{#1}}
\DeclareMathSymbol{\dshsym}{\mathalpha}{letters}{"2D}
\newcommand{\dsh}{\mathit{\dshsym}}
\let\int=\cap
\let\un=\cup
\let\inter=\bigcap
\let\union=\bigcup
\def\ML{{\sc ml}}
\def\AST{{\sc ast}}
%macros to change the treatment of symbols
\def\relsemicolon{\mathcode`\;="303B} %treat ; like a relation
\def\binperiod{\mathcode`\.="213A} %treat . like a binary operator
\def\binvert{\mathcode`\|="226A} %treat | like a binary operator
%redefinition of \sloppy and \fussy to use \emergencystretch
\def\sloppy{\tolerance2000 \hfuzz.5pt \vfuzz.5pt \emergencystretch=15pt}
\def\fussy{\tolerance200 \hfuzz.1pt \vfuzz.1pt \emergencystretch=0pt}
%non-bf version of description
\def\descrlabel#1{\hspace\labelsep #1}
\def\descr{\list{}{\labelwidth\z@ \itemindent-\leftmargin\let\makelabel\descrlabel}}
\let\enddescr\endlist
% The mathcodes for the letters A, ..., Z, a, ..., z are changed to
% generate text italic rather than math italic by default. This makes
% multi-letter identifiers look better. The mathcode for character c
% is set to |"7000| (variable family) + |"400| (text italic) + |c|.
%
\DeclareSymbolFont{italics}{\encodingdefault}{\rmdefault}{m}{it}%
\def\@setmcodes#1#2#3{{\count0=#1 \count1=#3
\loop \global\mathcode\count0=\count1 \ifnum \count0<#2
\advance\count0 by1 \advance\count1 by1 \repeat}}
\@setmcodes{`A}{`Z}{"7\hexnumber@\symitalics41}
\@setmcodes{`a}{`z}{"7\hexnumber@\symitalics61}
%%% \dquotes permits usage of "..." for \hbox{...}
%%% also taken from under.sty
{\catcode`\"=\active
\gdef\dquotes{\catcode`\"=\active \let"=\@mathText}%
\gdef\@mathText#1"{\hbox{\mathTextFont #1\/}}}
\def\mathTextFont{\frenchspacing\tt}
\def\dquotesoff{\catcode`\"=\other}
\ No newline at end of file
+% logics.sty : Logics Manuals Page Layout
+%
+\typeout{Document Style logics. Released 18 August 2003}
+
+\hyphenation{Isa-belle man-u-script man-u-scripts ap-pen-dix mut-u-al-ly}
+\hyphenation{data-type data-types co-data-type co-data-types }
+
+%usage: \iflabelundefined{LABEL}{if not defined}{if defined}
+\newcommand{\iflabelundefined}[1]{\@ifundefined{r@#1}}
+
+
+%%%INDEXING use isa-index to process the index
+
+\newcommand\seealso[2]{\emph{see also} #1}
+\usepackage{makeidx}
+
+%index, putting page numbers of definitions in boldface
+\def\bold#1{\textbf{#1}}
+\newcommand\fnote[1]{#1n}
+\newcommand\indexbold[1]{\index{#1|bold}}
+
+% The alternative to \protect\isa in the indexing macros is
+% \noexpand\noexpand \noexpand\isa
+% need TWO levels of \noexpand to delay the expansion of \isa:
+% the \noexpand\noexpand will leave one \noexpand, to be given to the
+% (still unexpanded) \isa token. See TeX by Topic, page 122.
+
+%%%% for indexing constants, symbols, theorems, ...
+\newcommand\cdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (constant)}}
+\newcommand\sdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (symbol)}}
+
+\newcommand\tdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (theorem)}}
+\newcommand\tdxbold[1]{\isa{#1}\index{#1@\protect\isa{#1} (theorem)|bold}}
+
+\newcommand\cldx[1]{\isa{#1}\index{#1@\protect\isa{#1} (class)}}
+\newcommand\tydx[1]{\isa{#1}\index{#1@\protect\isa{#1} (type)}}
+\newcommand\thydx[1]{\isa{#1}\index{#1@\protect\isa{#1} (theory)}}
+
+\newcommand\attrdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (attribute)}}
+\newcommand\cmmdx[1]{\index{#1@\protect\isacommand{#1} (command)}}
+\newcommand\commdx[1]{\isacommand{#1}\index{#1@\protect\isacommand{#1} (command)}}
+\newcommand\methdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (method)}}
+
+%set argument in \bf font and index in ROMAN font (for definitions in text!)
+\newcommand\bfindex[1]{{\bf#1}\index{#1|bold}\@}
+
+\newcommand\rmindex[1]{{#1}\index{#1}\@}
+\newcommand\ttindex[1]{\texttt{#1}\index{#1@\texttt{#1}}\@}
+\newcommand\ttindexbold[1]{\texttt{#1}\index{#1@\texttt{#1}|bold}\@}
+
+\newcommand{\indexboldpos}[2]{#1\@}
+\newcommand{\ttindexboldpos}[2]{\isa{#1}\@}
+
+%\newtheorem{theorem}{Theorem}[section]
+\newtheorem{Exercise}{Exercise}[section]
+\newenvironment{exercise}{\begin{Exercise}\rm}{\end{Exercise}}
+\newcommand{\ttlbr}{\texttt{[|}}
+\newcommand{\ttrbr}{\texttt{|]}}
+\newcommand{\ttor}{\texttt{|}}
+\newcommand{\ttall}{\texttt{!}}
+\newcommand{\ttuniquex}{\texttt{?!}}
+\newcommand{\ttEXU}{\texttt{EX!}}
+\newcommand{\ttAnd}{\texttt{!!}}
+
+\newcommand{\isasymignore}{}
+\newcommand{\isasymimp}{\isasymlongrightarrow}
+\newcommand{\isasymImp}{\isasymLongrightarrow}
+\newcommand{\isasymFun}{\isasymRightarrow}
+\newcommand{\isasymuniqex}{\isamath{\exists!\,}}
+\renewcommand{\S}{Sect.\ts}
+
+\renewenvironment{isamarkuptxt}{\begin{isamarkuptext}}{\end{isamarkuptext}}
+
+\newif\ifremarks
+\newcommand{\REMARK}[1]{\ifremarks\marginpar{\raggedright\footnotesize#1}\fi}
+
+%names of Isabelle rules
+\newcommand{\rulename}[1]{\hfill(#1)}
+\newcommand{\rulenamedx}[1]{\hfill(#1\index{#1@\protect\isa{#1} (theorem)|bold})}
+
+%%%% meta-logical connectives
+
+\let\Forall=\bigwedge
+\let\Imp=\Longrightarrow
+\let\To=\Rightarrow
+\newcommand{\Var}[1]{{?\!#1}}
+
+%%% underscores as ordinary characters, not for subscripting
+%% use @ or \sb for subscripting; use \at for @
+%% only works in \tt font
+%% must not make _ an active char; would make \ttindex fail!
+\gdef\underscoreoff{\catcode`\@=8\catcode`\_=\other}
+\gdef\underscoreon{\catcode`\_=8\makeatother}
+\chardef\other=12
+\chardef\at=`\@
+
+% alternative underscore
+\def\_{\leavevmode\kern.06em\vbox{\hrule height.2ex width.3em}\hskip0.1em}
+
+
+%%%% ``WARNING'' environment
+\def\dbend{\vtop to 0pt{\vss\hbox{\Huge\bf!}\vss}}
+\newenvironment{warn}{\medskip\medbreak\begingroup \clubpenalty=10000
+ \small %%WAS\baselineskip=0.9\baselineskip
+ \noindent \hangindent\parindent \hangafter=-2
+ \hbox to0pt{\hskip-\hangindent\dbend\hfill}\ignorespaces}%
+ {\par\endgroup\medbreak}
+
+
+%%%% Standard logical symbols
+\let\turn=\vdash
+\let\conj=\wedge
+\let\disj=\vee
+\let\imp=\rightarrow
+\let\bimp=\leftrightarrow
+\newcommand\all[1]{\forall#1.} %quantification
+\newcommand\ex[1]{\exists#1.}
+\newcommand{\pair}[1]{\langle#1\rangle}
+
+\newcommand{\lparr}{\mathopen{(\!|}}
+\newcommand{\rparr}{\mathclose{|\!)}}
+\newcommand{\fs}{\mathpunct{,\,}}
+\newcommand{\ty}{\mathrel{::}}
+\newcommand{\asn}{\mathrel{:=}}
+\newcommand{\more}{\ldots}
+\newcommand{\record}[1]{\lparr #1 \rparr}
+\newcommand{\dtt}{\mathord.}
+
+\newcommand\lbrakk{\mathopen{[\![}}
+\newcommand\rbrakk{\mathclose{]\!]}}
+\newcommand\List[1]{\lbrakk#1\rbrakk} %was \obj
+\newcommand\vpile[1]{\begin{array}{c}#1\end{array}}
+\newenvironment{matharray}[1]{\[\begin{array}{#1}}{\end{array}\]}
+\newcommand{\Text}[1]{\mbox{#1}}
+
+\DeclareMathSymbol{\dshsym}{\mathalpha}{letters}{"2D}
+\newcommand{\dsh}{\mathit{\dshsym}}
+
+\let\int=\cap
+\let\un=\cup
+\let\inter=\bigcap
+\let\union=\bigcup
+
+\def\ML{{\sc ml}}
+\def\AST{{\sc ast}}
+
+%macros to change the treatment of symbols
+\def\relsemicolon{\mathcode`\;="303B} %treat ; like a relation
+\def\binperiod{\mathcode`\.="213A} %treat . like a binary operator
+\def\binvert{\mathcode`\|="226A} %treat | like a binary operator
+
+%redefinition of \sloppy and \fussy to use \emergencystretch
+\def\sloppy{\tolerance2000 \hfuzz.5pt \vfuzz.5pt \emergencystretch=15pt}
+\def\fussy{\tolerance200 \hfuzz.1pt \vfuzz.1pt \emergencystretch=0pt}
+
+%non-bf version of description
+\def\descrlabel#1{\hspace\labelsep #1}
+\def\descr{\list{}{\labelwidth\z@ \itemindent-\leftmargin\let\makelabel\descrlabel}}
+\let\enddescr\endlist
+
+% The mathcodes for the letters A, ..., Z, a, ..., z are changed to
+% generate text italic rather than math italic by default. This makes
+% multi-letter identifiers look better. The mathcode for character c
+% is set to |"7000| (variable family) + |"400| (text italic) + |c|.
+%
+\DeclareSymbolFont{italics}{\encodingdefault}{\rmdefault}{m}{it}%
+\def\@setmcodes#1#2#3{{\count0=#1 \count1=#3
+ \loop \global\mathcode\count0=\count1 \ifnum \count0<#2
+ \advance\count0 by1 \advance\count1 by1 \repeat}}
+\@setmcodes{`A}{`Z}{"7\hexnumber@\symitalics41}
+\@setmcodes{`a}{`z}{"7\hexnumber@\symitalics61}
+
+%%% \dquotes permits usage of "..." for \hbox{...}
+%%% also taken from under.sty
+{\catcode`\"=\active
+\gdef\dquotes{\catcode`\"=\active \let"=\@mathText}%
+\gdef\@mathText#1"{\hbox{\mathTextFont #1\/}}}
+\def\mathTextFont{\frenchspacing\tt}
+\def\dquotesoff{\catcode`\"=\other}
--- a/src/Doc/Logics_ZF/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Logics_ZF/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -13,7 +13,7 @@
%%% to index constants: \\tt \([a-zA-Z0-9][a-zA-Z0-9_]*\) \\cdx{\1}
%%% to deverbify: \\verb|\([^|]*\)| \\ttindex{\1}
-\title{\includegraphics[scale=0.5]{isabelle_zf} \\[4ex]
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex]
Isabelle's Logics: FOL and ZF}
\author{{\em Lawrence C. Paulson}\\
--- a/src/Doc/Main/Main_Doc.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Main/Main_Doc.thy Sun May 23 20:34:43 2021 +0200
@@ -4,14 +4,14 @@
begin
setup \<open>
- Thy_Output.antiquotation_pretty_source \<^binding>\<open>term_type_only\<close> (Args.term -- Args.typ_abbrev)
+ Document_Output.antiquotation_pretty_source \<^binding>\<open>term_type_only\<close> (Args.term -- Args.typ_abbrev)
(fn ctxt => fn (t, T) =>
(if fastype_of t = Sign.certify_typ (Proof_Context.theory_of ctxt) T then ()
else error "term_type_only: type mismatch";
Syntax.pretty_typ ctxt T))
\<close>
setup \<open>
- Thy_Output.antiquotation_pretty_source \<^binding>\<open>expanded_typ\<close> Args.typ
+ Document_Output.antiquotation_pretty_source \<^binding>\<open>expanded_typ\<close> Args.typ
Syntax.pretty_typ
\<close>
(*>*)
--- a/src/Doc/Main/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle latex -o "$FORMAT"
-isabelle latex -o "$FORMAT"
-
--- a/src/Doc/Nitpick/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo Nitpick
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Nitpick/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Nitpick/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -48,7 +48,7 @@
%\renewcommand\labelitemi{$\bullet$}
\renewcommand\labelitemi{\raise.065ex\hbox{\small\textbullet}}
-\title{\includegraphics[scale=0.5]{isabelle_nitpick} \\[4ex]
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex]
Picking Nits \\[\smallskipamount]
\Large A User's Guide to Nitpick for Isabelle/HOL}
\author{\hbox{} \\
--- a/src/Doc/Prog_Prove/LaTeXsugar.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Prog_Prove/LaTeXsugar.thy Sun May 23 20:34:43 2021 +0200
@@ -44,10 +44,11 @@
"_asm" :: "prop \<Rightarrow> asms" ("_")
setup \<open>
- Thy_Output.antiquotation_pretty_source \<^binding>\<open>const_typ\<close> (Scan.lift Args.embedded_inner_syntax)
+ Document_Output.antiquotation_pretty_source \<^binding>\<open>const_typ\<close>
+ (Scan.lift Args.embedded_inner_syntax)
(fn ctxt => fn c =>
let val tc = Proof_Context.read_const {proper = false, strict = false} ctxt c in
- Pretty.block [Thy_Output.pretty_term ctxt tc, Pretty.str " ::",
+ Pretty.block [Document_Output.pretty_term ctxt tc, Pretty.str " ::",
Pretty.brk 1, Syntax.pretty_typ ctxt (fastype_of tc)]
end)
\<close>
--- a/src/Doc/Prog_Prove/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo HOL
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Prog_Prove/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Prog_Prove/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -7,7 +7,7 @@
\begin{document}
\title{Programming and Proving in Isabelle/HOL}
-\subtitle{\includegraphics[scale=.7]{isabelle_hol}}
+\subtitle{\includegraphics[scale=.7]{isabelle_logo}}
\author{Tobias Nipkow}
\maketitle
--- a/src/Doc/ROOT Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/ROOT Sun May 23 20:34:43 2021 +0200
@@ -1,23 +1,23 @@
chapter Doc
session Classes (doc) in "Classes" = HOL +
- options [document_variants = "classes", quick_and_dirty]
+ options [document_logo = "Isar", document_bibliography,
+ document_variants = "classes", quick_and_dirty]
theories [document = false] Setup
theories Classes
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Codegen (doc) in "Codegen" = HOL +
- options [document_variants = "codegen", print_mode = "no_brackets,iff"]
+ options [document_logo = "Isar", document_bibliography, document_variants = "codegen",
+ print_mode = "no_brackets,iff"]
sessions
"HOL-Library"
theories [document = false]
@@ -32,35 +32,31 @@
Adaptation
Further
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Corec (doc) in "Corec" = Datatypes +
- options [document_variants = "corec"]
+ options [document_bibliography, document_variants = "corec"]
theories
Corec
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Datatypes (doc) in "Datatypes" = HOL +
- options [document_variants = "datatypes"]
+ options [document_bibliography, document_variants = "datatypes"]
sessions
"HOL-Library"
theories [document = false]
@@ -68,20 +64,18 @@
theories
Datatypes
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Eisbach (doc) in "Eisbach" = HOL +
- options [document_variants = "eisbach", quick_and_dirty,
- print_mode = "no_brackets,iff", show_question_marks = false]
+ options [document_logo = "Eisbach", document_bibliography, document_variants = "eisbach",
+ quick_and_dirty, print_mode = "no_brackets,iff", show_question_marks = false]
sessions
"HOL-Eisbach"
theories [document = false]
@@ -90,7 +84,6 @@
Preface
Manual
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
@@ -99,22 +92,20 @@
"underscore.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Functions (doc) in "Functions" = HOL +
- options [document_variants = "functions", skip_proofs = false, quick_and_dirty]
+ options [document_bibliography, document_variants = "functions",
+ skip_proofs = false, quick_and_dirty]
theories Functions
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"conclusion.tex"
"intro.tex"
"root.tex"
@@ -130,9 +121,9 @@
"prelude.tex"
session Intro (doc) in "Intro" = Pure +
- options [document_variants = "intro"]
+ options [document_logo = "_", document_bibliography, document_build = "build",
+ document_variants = "intro"]
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
@@ -146,7 +137,8 @@
"root.tex"
session Implementation (doc) in "Implementation" = HOL +
- options [document_variants = "implementation", quick_and_dirty]
+ options [document_logo = "Isar", document_bibliography,
+ document_variants = "implementation", quick_and_dirty]
theories
Eq
Integration
@@ -160,7 +152,6 @@
theories [parallel_proofs = 0]
Logic
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
@@ -169,12 +160,12 @@
"underscore.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
session Isar_Ref (doc) in "Isar_Ref" = HOL +
- options [document_variants = "isar-ref", quick_and_dirty, thy_output_source]
+ options [document_logo = "Isar", document_bibliography, document_variants = "isar-ref",
+ quick_and_dirty, thy_output_source]
sessions
"HOL-Library"
theories
@@ -193,7 +184,6 @@
Quick_Reference
Symbols
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
@@ -202,15 +192,14 @@
"underscore.sty"
"manual.bib"
document_files
- "build"
"isar-vm.pdf"
"isar-vm.svg"
"root.tex"
- "showsymbols"
"style.sty"
session JEdit (doc) in "JEdit" = HOL +
- options [document_variants = "jedit", thy_output_source]
+ options [document_logo = "jEdit", document_bibliography, document_variants = "jedit",
+ thy_output_source]
theories
JEdit
document_files (in "..")
@@ -219,7 +208,6 @@
"isar.sty"
"manual.bib"
"pdfsetup.sty"
- "prepare_document"
"ttbox.sty"
"underscore.sty"
document_files (in "../Isar_Ref/document")
@@ -227,7 +215,6 @@
document_files
"auto-tools.png"
"bibtex-mode.png"
- "build"
"cite-completion.png"
"isabelle-jedit.png"
"markdown-document.png"
@@ -247,54 +234,52 @@
"theories.png"
session Sugar (doc) in "Sugar" = HOL +
- options [document_variants = "sugar"]
+ options [document_bibliography, document_variants = "sugar"]
sessions
"HOL-Library"
theories Sugar
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
document_files
- "build"
"root.bib"
"root.tex"
session Locales (doc) in "Locales" = HOL +
- options [document_variants = "locales", thy_output_margin = 65, skip_proofs = false]
+ options [document_bibliography, document_variants = "locales",
+ thy_output_margin = 65, skip_proofs = false]
theories
Examples1
Examples2
Examples3
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
document_files
- "build"
"root.bib"
"root.tex"
session Logics (doc) in "Logics" = Pure +
- options [document_variants = "logics"]
+ options [document_logo = "_", document_bibliography, document_build = "build",
+ document_variants = "logics"]
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"ttbox.sty"
"manual.bib"
+ document_files (in "../Intro/document")
+ "build"
document_files
"CTT.tex"
"HOL.tex"
"LK.tex"
"Sequents.tex"
- "build"
"preface.tex"
"root.tex"
"syntax.tex"
session Logics_ZF (doc) in "Logics_ZF" = ZF +
- options [document_variants = "logics-ZF", print_mode = "brackets",
- thy_output_source]
+ options [document_logo = "ZF", document_bibliography, document_build = "build",
+ document_variants = "logics-ZF", print_mode = "brackets", thy_output_source]
sessions
FOL
theories
@@ -304,17 +289,17 @@
If
ZF_Isar
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"isar.sty"
"ttbox.sty"
"manual.bib"
+ document_files (in "../Intro/document")
+ "build"
document_files (in "../Logics/document")
"syntax.tex"
document_files
"FOL.tex"
"ZF.tex"
- "build"
"logics.sty"
"root.tex"
@@ -322,25 +307,22 @@
options [document_variants = "main"]
theories Main_Doc
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
document_files
- "build"
"root.tex"
session Nitpick (doc) in "Nitpick" = Pure +
- options [document_variants = "nitpick"]
+ options [document_logo = "Nitpick", document_bibliography, document_variants = "nitpick"]
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"manual.bib"
document_files
- "build"
"root.tex"
session Prog_Prove (doc) in "Prog_Prove" = HOL +
- options [document_variants = "prog-prove", show_question_marks = false]
+ options [document_logo = "HOL", document_bibliography, document_variants = "prog-prove",
+ show_question_marks = false]
theories
Basics
Bool_nat_list
@@ -351,11 +333,9 @@
document_files (in ".")
"MyList.thy"
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
document_files
"bang.pdf"
- "build"
"intro-isabelle.tex"
"prelude.tex"
"root.bib"
@@ -363,18 +343,17 @@
"svmono.cls"
session Sledgehammer (doc) in "Sledgehammer" = Pure +
- options [document_variants = "sledgehammer"]
+ options [document_logo = "S/H", document_bibliography, document_variants = "sledgehammer"]
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"manual.bib"
document_files
- "build"
"root.tex"
session System (doc) in "System" = Pure +
- options [document_variants = "system", thy_output_source]
+ options [document_logo = "_", document_bibliography, document_variants = "system",
+ thy_output_source]
sessions
"HOL-Library"
theories
@@ -386,7 +365,6 @@
Phabricator
Misc
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
@@ -397,11 +375,11 @@
document_files (in "../Isar_Ref/document")
"style.sty"
document_files
- "build"
"root.tex"
session Tutorial (doc) in "Tutorial" = HOL +
- options [document_variants = "tutorial", print_mode = "brackets", skip_proofs = false]
+ options [document_logo = "HOL", document_bibliography, document_build = "build",
+ document_variants = "tutorial", print_mode = "brackets", skip_proofs = false]
directories "Advanced" "CTL" "CodeGen" "Datatype" "Documents" "Fun" "Ifexpr"
"Inductive" "Misc" "Protocol" "Rules" "Sets" "ToyList" "Trie" "Types"
theories [document = false]
@@ -494,7 +472,7 @@
"types0.tex"
session Typeclass_Hierarchy (doc) in "Typeclass_Hierarchy" = HOL +
- options [document_variants = "typeclass_hierarchy"]
+ options [document_logo = "Isar", document_bibliography, document_variants = "typeclass_hierarchy"]
sessions
"HOL-Library"
theories [document = false]
@@ -502,13 +480,11 @@
theories
Typeclass_Hierarchy
document_files (in "..")
- "prepare_document"
"pdfsetup.sty"
"iman.sty"
"extra.sty"
"isar.sty"
"manual.bib"
document_files
- "build"
"root.tex"
"style.sty"
--- a/src/Doc/Sledgehammer/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo -o isabelle_sledgehammer.pdf "S/H"
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
--- a/src/Doc/Sledgehammer/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Sledgehammer/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -54,7 +54,7 @@
%\renewcommand\labelitemi{$\bullet$}
\renewcommand\labelitemi{\raise.065ex\hbox{\small\textbullet}}
-\title{\includegraphics[scale=0.5]{isabelle_sledgehammer} \\[4ex]
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex]
Hammering Away \\[\smallskipamount]
\Large A User's Guide to Sledgehammer for Isabelle/HOL}
\author{\hbox{} \\
--- a/src/Doc/Sugar/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/System/Environment.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/System/Environment.thy Sun May 23 20:34:43 2021 +0200
@@ -181,9 +181,10 @@
\<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the line editor for the
@{tool_ref console} interface.
- \<^descr>[@{setting_def ISABELLE_PDFLATEX}, @{setting_def ISABELLE_BIBTEX}] refer to
- {\LaTeX} related tools for Isabelle document preparation (see also
- \secref{sec:tool-latex}).
+ \<^descr>[@{setting_def ISABELLE_PDFLATEX}, @{setting_def ISABELLE_LUALATEX},
+ @{setting_def ISABELLE_BIBTEX}, @{setting_def ISABELLE_MAKEINDEX}] refer to
+ {\LaTeX}-related tools for Isabelle document preparation (see also
+ \secref{sec:tool-document}).
\<^descr>[@{setting_def ISABELLE_TOOLS}] is a colon separated list of directories
that are scanned by @{executable isabelle} for external utility programs
--- a/src/Doc/System/Presentation.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/System/Presentation.thy Sun May 23 20:34:43 2021 +0200
@@ -22,8 +22,7 @@
with the PDF into the given directory (relative to the session directory).
Alternatively, @{tool_ref document} may be used to turn the generated
- {\LaTeX} sources of a session (exports from its build database) into PDF,
- using suitable invocations of @{tool_ref latex}.
+ {\LaTeX} sources of a session (exports from its build database) into PDF.
\<close>
@@ -161,12 +160,20 @@
\<^medskip> Isabelle is usually smart enough to create the PDF from the given
\<^verbatim>\<open>root.tex\<close> and optional \<^verbatim>\<open>root.bib\<close> (bibliography) and \<^verbatim>\<open>root.idx\<close> (index)
- using standard {\LaTeX} tools. Alternatively, \isakeyword{document\_files}
- in the session \<^verbatim>\<open>ROOT\<close> may include an executable \<^verbatim>\<open>build\<close> script to take
- care of that. It is invoked with command-line arguments for the document
- format (\<^verbatim>\<open>pdf\<close>) and the document variant name. The script needs to produce
- corresponding output files, e.g.\ \<^verbatim>\<open>root.pdf\<close> for default document variants
- (the main work can be delegated to @{tool latex}). \<close>
+ using standard {\LaTeX} tools. Actual command-lines are given by settings
+ @{setting_ref ISABELLE_PDFLATEX}, @{setting_ref ISABELLE_LUALATEX},
+ @{setting_ref ISABELLE_BIBTEX}, @{setting_ref ISABELLE_MAKEINDEX}: these
+ variables are used without quoting in shell scripts, and thus may contain
+ additional options.
+
+ Alternatively, the session \<^verbatim>\<open>ROOT\<close> may include an option
+ \<^verbatim>\<open>document_build=build\<close> together with an executable \<^verbatim>\<open>build\<close> script in
+ \isakeyword{document\_files}: it is invoked with command-line arguments for
+ the document format (\<^verbatim>\<open>pdf\<close>) and the document variant name. The script needs
+ to produce corresponding output files, e.g.\ \<^verbatim>\<open>root.pdf\<close> for default
+ document variants.
+\<close>
+
subsubsection \<open>Examples\<close>
@@ -177,47 +184,4 @@
@{verbatim [display] \<open>isabelle document -v -V -O. FOL\<close>}
\<close>
-
-section \<open>Running {\LaTeX} within the Isabelle environment
- \label{sec:tool-latex}\<close>
-
-text \<open>
- The @{tool_def latex} tool provides the basic interface for Isabelle
- document preparation. Its usage is:
- @{verbatim [display]
-\<open>Usage: isabelle latex [OPTIONS] [FILE]
-
- Options are:
- -o FORMAT specify output format: pdf (default), bbl, idx, sty
-
- Run LaTeX (and related tools) on FILE (default root.tex),
- producing the specified output format.\<close>}
-
- Appropriate {\LaTeX}-related programs are run on the input file, according
- to the given output format: @{executable pdflatex}, @{executable bibtex}
- (for \<^verbatim>\<open>bbl\<close>), and @{executable makeindex} (for \<^verbatim>\<open>idx\<close>). The actual commands
- are determined from the settings environment (@{setting ISABELLE_PDFLATEX}
- etc.).
-
- The \<^verbatim>\<open>sty\<close> output format causes the Isabelle style files to be updated from
- the distribution. This is useful in special situations where the document
- sources are to be processed another time by separate tools.
-\<close>
-
-
-subsubsection \<open>Examples\<close>
-
-text \<open>
- Invoking @{tool latex} by hand may be occasionally useful when debugging
- failed attempts of the automatic document preparation stage of batch-mode
- Isabelle. The abortive process leaves the sources at a certain place within
- @{setting ISABELLE_BROWSER_INFO}, see the runtime error message for details.
- This enables users to inspect {\LaTeX} runs in further detail, e.g.\ like
- this:
-
- @{verbatim [display]
-\<open>cd "$(isabelle getenv -b ISABELLE_BROWSER_INFO)/Unsorted/Test/document"
-isabelle latex -o pdf\<close>}
-\<close>
-
end
--- a/src/Doc/System/Sessions.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/System/Sessions.thy Sun May 23 20:34:43 2021 +0200
@@ -238,6 +238,16 @@
is occasionally useful to control the global visibility of commands via
session options (e.g.\ in \<^verbatim>\<open>ROOT\<close>).
+ \<^item> @{system_option_def "document_bibliography"} explicitly enables the use
+ of \<^verbatim>\<open>bibtex\<close>; the default is to check the presence of \<^verbatim>\<open>root.bib\<close>, but it
+ could have a different name.
+
+ \<^item> @{system_option_def "document_preprocessor"} specifies the name of an
+ executable that is run within the document output directory, after
+ preparing the document sources and before the actual build process. This
+ allows to apply adhoc patches, without requiring a separate \<^verbatim>\<open>build\<close>
+ script.
+
\<^item> @{system_option_def "threads"} determines the number of worker threads
for parallel checking of theories and proofs. The default \<open>0\<close> means that a
sensible maximum value is determined by the underlying hardware. For
--- a/src/Doc/System/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/System/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/System/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -17,7 +17,7 @@
\isabellestyle{literal}
\def\isastylett{\footnotesize\tt}
-\title{\includegraphics[scale=0.5]{isabelle} \\[4ex] The Isabelle System Manual}
+\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] The Isabelle System Manual}
\author{\emph{Makarius Wenzel}}
--- a/src/Doc/Tutorial/document/build Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Tutorial/document/build Sun May 23 20:34:43 2021 +0200
@@ -2,13 +2,9 @@
set -e
-FORMAT="$1"
-VARIANT="$2"
-
-isabelle logo HOL
-isabelle latex -o "$FORMAT"
-isabelle latex -o bbl
+$ISABELLE_LUALATEX root
+$ISABELLE_BIBTEX root
+$ISABELLE_LUALATEX root
+$ISABELLE_LUALATEX root
./isa-index root
-isabelle latex -o "$FORMAT"
-[ -f root.out ] && "$ISABELLE_HOME/src/Doc/fixbookmarks" root.out
-isabelle latex -o "$FORMAT"
+$ISABELLE_LUALATEX root
--- a/src/Doc/Tutorial/document/isa-index Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Tutorial/document/isa-index Sun May 23 20:34:43 2021 +0200
@@ -20,4 +20,4 @@
s~\*\(\".\".\".\)~\1@\\\\isa {\1}~g
s~\*\(\".\".\)~\1@\\\\isa {\1}~g
s~\*\(\".\)~\1@\\\\isa {\1}~g
-s~\*\([^ |!@{}][^ |!@{}]*\)~\1@\\\\isa {\1}~g" $1.idx | makeindex -c -q -o $1.ind
+s~\*\([^ |!@{}][^ |!@{}]*\)~\1@\\\\isa {\1}~g" $1.idx | $ISABELLE_MAKEINDEX -o $1.ind
--- a/src/Doc/Tutorial/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Tutorial/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -29,7 +29,7 @@
\begin{document}
\title{
\begin{center}
-\includegraphics[scale=.8]{isabelle_hol}
+\includegraphics[scale=.8]{isabelle_logo}
\\ \vspace{0.5cm} A Proof Assistant for Higher-Order Logic
\end{center}}
\author{Tobias Nipkow \quad Lawrence C. Paulson \quad Markus Wenzel%\\[1ex]
--- a/src/Doc/Tutorial/document/tutorial.sty Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Tutorial/document/tutorial.sty Sun May 23 20:34:43 2021 +0200
@@ -42,8 +42,6 @@
\newcommand\cmmdx[1]{\index{#1@\protect\isacommand{#1} (command)}}
\newcommand\commdx[1]{\isacommand{#1}\index{#1@\protect\isacommand{#1} (command)}}
\newcommand\methdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (method)}}
-\newcommand\tooldx[1]{\isa{#1}\index{#1@\protect\isa{#1} (tool)}}
-\newcommand\settdx[1]{\isa{#1}\index{#1@\protect\isa{#1} (setting)}}
\newcommand\pgdx[1]{\pgmenu{#1}\index{#1@\protect\pgmenu{#1} (Proof General)}}
%set argument in \bf font and index in ROMAN font (for definitions in text!)
--- a/src/Doc/Typeclass_Hierarchy/document/build Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-VARIANT="$2"
-
-"$ISABELLE_TOOL" logo Isar
-"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
-
--- a/src/Doc/Typeclass_Hierarchy/document/root.tex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/Typeclass_Hierarchy/document/root.tex Sun May 23 20:34:43 2021 +0200
@@ -9,7 +9,7 @@
\hyphenation{Isar}
\isadroptag{theory}
-\title{\includegraphics[scale=0.5]{isabelle_isar}
+\title{\includegraphics[scale=0.5]{isabelle_logo}
\\[4ex] The {Isabelle/HOL} type-class hierarchy}
\author{\emph{Florian Haftmann}}
--- a/src/Doc/antiquote_setup.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/antiquote_setup.ML Sun May 23 20:34:43 2021 +0200
@@ -33,104 +33,20 @@
| clean_name s = s |> translate (fn "_" => "-" | "\<hyphen>" => "-" | c => c);
-(* ML text *)
-
-local
-
-fun ml_val (toks1, []) = ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read ");"
- | ml_val (toks1, toks2) =
- ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read ");";
-
-fun ml_op (toks1, []) = ML_Lex.read "fn _ => (op " @ toks1 @ ML_Lex.read ");"
- | ml_op (toks1, toks2) =
- ML_Lex.read "fn _ => (op " @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read ");";
-
-fun ml_type (toks1, []) = ML_Lex.read "val _ = NONE : (" @ toks1 @ ML_Lex.read ") option;"
- | ml_type (toks1, toks2) =
- ML_Lex.read "val _ = [NONE : (" @ toks1 @ ML_Lex.read ") option, NONE : (" @
- toks2 @ ML_Lex.read ") option];";
-
-fun ml_exception (toks1, []) = ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : exn);"
- | ml_exception (toks1, toks2) =
- ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read " -> exn);";
-
-fun ml_structure (toks, _) =
- ML_Lex.read "functor XXX() = struct structure XX = " @ toks @ ML_Lex.read " end;";
-
-fun ml_functor (Antiquote.Text tok :: _, _) =
- ML_Lex.read "ML_Env.check_functor " @
- ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok))
- | ml_functor _ = raise Fail "Bad ML functor specification";
-
-val is_name =
- ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.Long_Ident);
-
-fun ml_name txt =
- (case filter is_name (ML_Lex.tokenize txt) of
- toks as [_] => ML_Lex.flatten toks
- | _ => error ("Single ML name expected in input: " ^ quote txt));
-
-fun prep_ml source =
- (#1 (Input.source_content source), ML_Lex.read_source source);
-
-fun index_ml name kind ml = Thy_Output.antiquotation_raw name
- (Scan.lift (Args.text_input -- Scan.option (Args.colon |-- Args.text_input)))
- (fn ctxt => fn (source1, opt_source2) =>
- let
- val (txt1, toks1) = prep_ml source1;
- val (txt2, toks2) =
- (case opt_source2 of
- SOME source => prep_ml source
- | NONE => ("", []));
-
- val txt =
- if txt2 = "" then txt1
- else if kind = "type" then txt1 ^ " = " ^ txt2
- else if kind = "exception" then txt1 ^ " of " ^ txt2
- else if Symbol_Pos.is_identifier (Long_Name.base_name (ml_name txt1))
- then txt1 ^ ": " ^ txt2
- else txt1 ^ " : " ^ txt2;
- val txt' = if kind = "" then txt else kind ^ " " ^ txt;
-
- val pos = Input.pos_of source1;
- val _ =
- ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos (ml (toks1, toks2))
- handle ERROR msg => error (msg ^ Position.here pos);
- val kind' = if kind = "" then "ML" else "ML " ^ kind;
- in
- Latex.block
- [Latex.string ("\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string (ml_name txt1) ^ "}"),
- Thy_Output.verbatim ctxt txt']
- end);
-
-in
-
-val _ =
- Theory.setup
- (index_ml \<^binding>\<open>index_ML\<close> "" ml_val #>
- index_ml \<^binding>\<open>index_ML_op\<close> "infix" ml_op #>
- index_ml \<^binding>\<open>index_ML_type\<close> "type" ml_type #>
- index_ml \<^binding>\<open>index_ML_exception\<close> "exception" ml_exception #>
- index_ml \<^binding>\<open>index_ML_structure\<close> "structure" ml_structure #>
- index_ml \<^binding>\<open>index_ML_functor\<close> "functor" ml_functor);
-
-end;
-
-
(* named theorems *)
val _ =
- Theory.setup (Thy_Output.antiquotation_raw \<^binding>\<open>named_thms\<close>
+ Theory.setup (Document_Output.antiquotation_raw \<^binding>\<open>named_thms\<close>
(Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name)))
(fn ctxt =>
map (fn (thm, name) =>
Output.output
(Document_Antiquotation.format ctxt
- (Document_Antiquotation.delimit ctxt (Thy_Output.pretty_thm ctxt thm))) ^
+ (Document_Antiquotation.delimit ctxt (Document_Output.pretty_thm ctxt thm))) ^
enclose "\\rulename{" "}" (Output.output name))
#> space_implode "\\par\\smallskip%\n"
#> Latex.string #> single
- #> Thy_Output.isabelle ctxt));
+ #> Document_Output.isabelle ctxt));
(* Isabelle/Isar entities (with index) *)
@@ -150,7 +66,7 @@
val arg = enclose "{" "}" o clean_string;
fun entity check markup binding index =
- Thy_Output.antiquotation_raw
+ Document_Output.antiquotation_raw
(binding |> Binding.map_name (fn name => name ^
(case index of NONE => "" | SOME true => "_def" | SOME false => "_ref")))
(Scan.lift (Scan.optional (Args.parens Args.name) "" -- Args.name_position))
@@ -207,4 +123,36 @@
end;
+
+(* show symbols *)
+
+val _ =
+ Theory.setup (Document_Output.antiquotation_raw \<^binding>\<open>show_symbols\<close> (Scan.succeed ())
+ (fn _ => fn _ =>
+ let
+ val symbol_name =
+ unprefix "\\newcommand{\\isasym"
+ #> raw_explode
+ #> take_prefix Symbol.is_ascii_letter
+ #> implode;
+
+ val symbols =
+ File.read \<^file>\<open>~~/lib/texinputs/isabellesym.sty\<close>
+ |> split_lines
+ |> map_filter (fn line =>
+ (case try symbol_name line of
+ NONE => NONE
+ | SOME "" => NONE
+ | SOME name => SOME ("\\verb,\\" ^ "<" ^ name ^ ">, & {\\isasym" ^ name ^ "}")));
+
+ val eol = "\\\\\n";
+ fun table (a :: b :: rest) = a ^ " & " ^ b ^ eol :: table rest
+ | table [a] = [a ^ eol]
+ | table [] = [];
+ in
+ Latex.string
+ ("\\begin{supertabular}{ll@{\\qquad}ll}\n" ^ implode (table symbols) ^
+ "\\end{supertabular}\n")
+ end))
+
end;
--- a/src/Doc/fixbookmarks Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-#!/usr/bin/env bash
-
-perl -pi -e 's/\\([a-zA-Z]+)\s*/$1/g; s/\$//g; s/^BOOKMARK/\\BOOKMARK/g;' "$@"
--- a/src/Doc/iman.sty Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/iman.sty Sun May 23 20:34:43 2021 +0200
@@ -31,16 +31,10 @@
\newcommand\mltydx[1]{{\tt#1}\index{#1@{\tt#1} ML type}}
\newcommand\xdx[1]{{\tt#1}\index{#1@{\tt#1} exception}}
-\newcommand\ndx[1]{{\tt#1}\index{#1@{\tt#1} nonterminal}}
-\newcommand\ndxbold[1]{{\tt#1}\index{#1@{\tt#1} nonterminal|bold}}
-
\newcommand\cldx[1]{{\tt#1}\index{#1@{\tt#1} class}}
\newcommand\tydx[1]{\textit{#1}\index{#1@{\textit{#1}} type}}
\newcommand\thydx[1]{{\tt#1}\index{#1@{\tt#1} theory}}
-\newcommand\tooldx[1]{{\tt#1}\index{#1@{\tt#1} tool}}
-\newcommand\settdx[1]{{\tt#1}\index{#1@{\tt#1} setting}}
-
%set argument in \tt font; at the same time, index using * prefix
\newcommand\rmindex[1]{{#1}\index{#1}\@}
\newcommand\ttindex[1]{{\tt#1}\index{*#1}\@}
--- a/src/Doc/isar.sty Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/isar.sty Sun May 23 20:34:43 2021 +0200
@@ -9,10 +9,6 @@
\newcommand{\isasystem}[1]{{\def\isacharminus{-}\def\isacharunderscore{\_}\isadigitreset\tt #1}}
\newcommand{\isatool}[1]{{\def\isacharminus{-}\def\isacharunderscore{\_}\isadigitreset\tt isabelle #1}}
-\newcommand{\indexoutertoken}[1]{\indexdef{}{syntax}{#1}}
-\newcommand{\indexouternonterm}[1]{\indexdef{}{syntax}{#1}}
-\newcommand{\indexisarelem}[1]{\indexdef{}{element}{#1}}
-
\newcommand{\isasymIF}{\isakeyword{if}}
\newcommand{\isasymFOR}{\isakeyword{for}}
\newcommand{\isasymAND}{\isakeyword{and}}
--- a/src/Doc/more_antiquote.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/more_antiquote.ML Sun May 23 20:34:43 2021 +0200
@@ -10,7 +10,7 @@
(* class specifications *)
val _ =
- Theory.setup (Thy_Output.antiquotation_pretty \<^binding>\<open>class_spec\<close> (Scan.lift Args.name)
+ Theory.setup (Document_Output.antiquotation_pretty \<^binding>\<open>class_spec\<close> (Scan.lift Args.name)
(fn ctxt => fn s =>
let
val thy = Proof_Context.theory_of ctxt;
@@ -21,7 +21,7 @@
(* code theorem antiquotation *)
val _ =
- Theory.setup (Thy_Output.antiquotation_pretty \<^binding>\<open>code_thms\<close> Args.term
+ Theory.setup (Document_Output.antiquotation_pretty \<^binding>\<open>code_thms\<close> Args.term
(fn ctxt => fn raw_const =>
let
val thy = Proof_Context.theory_of ctxt;
@@ -33,6 +33,6 @@
|> these
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
|> map (HOLogic.mk_obj_eq o Variable.import_vars ctxt o Axclass.overload ctxt);
- in Pretty.chunks (map (Thy_Output.pretty_thm ctxt) thms) end));
+ in Pretty.chunks (map (Document_Output.pretty_thm ctxt) thms) end));
end;
--- a/src/Doc/prepare_document Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-#!/usr/bin/env bash
-
-set -e
-
-FORMAT="$1"
-
-isabelle latex -o sty
-cp "$ISABELLE_HOME/src/Doc/pdfsetup.sty" .
-
-isabelle latex -o "$FORMAT"
-isabelle latex -o bbl
-[ -f root.idx ] && "$ISABELLE_HOME/src/Doc/sedindex" root
-isabelle latex -o "$FORMAT"
-[ -f root.out ] && "$ISABELLE_HOME/src/Doc/fixbookmarks" root.out
-isabelle latex -o "$FORMAT"
-
--- a/src/Doc/sedindex Wed May 19 14:17:40 2021 +0100
+++ b/src/Doc/sedindex Sun May 23 20:34:43 2021 +0200
@@ -18,4 +18,4 @@
s~\*\(\".\".\".\)~\1@{\\\\tt \1}~g
s~\*\(\".\".\)~\1@{\\\\tt \1}~g
s~\*\(\".\)~\1@{\\\\tt \1}~g
-s~\*\([^ |!@{}][^ |!@{}]*\)~\1@{\\\\tt \1}~g" $1.idx | makeindex -c -q -o $1.ind
+s~\*\([^ |!@{}][^ |!@{}]*\)~\1@{\\\\tt \1}~g" $1.idx | $ISABELLE_MAKEINDEX -o $1.ind
--- a/src/HOL/Library/Code_Lazy.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/HOL/Library/Code_Lazy.thy Sun May 23 20:34:43 2021 +0200
@@ -182,7 +182,7 @@
\<open>object Lazy {
final class Lazy[A] (f: Unit => A) {
var evaluated = false;
- lazy val x: A = f ()
+ lazy val x: A = f(())
def get() : A = {
evaluated = true;
@@ -210,7 +210,7 @@
x: Lazy[A],
dummy: Term) : Term = {
if (x.evaluated)
- app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(abs("_")(unitT)(term_of(x.get)))
+ app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(abs("_")(unitT)(term_of(x.get())))
else
app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(const("Pure.dummy_pattern")(funT(unitT)(ty)))
}
--- a/src/HOL/Library/LaTeXsugar.thy Wed May 19 14:17:40 2021 +0100
+++ b/src/HOL/Library/LaTeXsugar.thy Sun May 23 20:34:43 2021 +0200
@@ -97,11 +97,11 @@
"_asm" :: "prop \<Rightarrow> asms" ("_")
setup \<open>
- Thy_Output.antiquotation_pretty_source_embedded \<^binding>\<open>const_typ\<close>
+ Document_Output.antiquotation_pretty_source_embedded \<^binding>\<open>const_typ\<close>
(Scan.lift Args.embedded_inner_syntax)
(fn ctxt => fn c =>
let val tc = Proof_Context.read_const {proper = false, strict = false} ctxt c in
- Pretty.block [Thy_Output.pretty_term ctxt tc, Pretty.str " ::",
+ Pretty.block [Document_Output.pretty_term ctxt tc, Pretty.str " ::",
Pretty.brk 1, Syntax.pretty_typ ctxt (fastype_of tc)]
end)
\<close>
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Sun May 23 20:34:43 2021 +0200
@@ -1200,7 +1200,7 @@
local
fun antiquote_setup binding co =
- Thy_Output.antiquotation_pretty_source_embedded binding
+ Document_Output.antiquotation_pretty_source_embedded binding
((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) --
Args.type_name {proper = true, strict = true})
(fn ctxt => fn (pos, type_name) =>
--- a/src/HOL/Tools/value_command.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/HOL/Tools/value_command.ML Sun May 23 20:34:43 2021 +0200
@@ -73,10 +73,10 @@
>> (fn ((name, modes), t) => Toplevel.keep (value_cmd name modes t)));
val _ = Theory.setup
- (Thy_Output.antiquotation_pretty_source_embedded \<^binding>\<open>value\<close>
+ (Document_Output.antiquotation_pretty_source_embedded \<^binding>\<open>value\<close>
(Scan.lift opt_evaluator -- Term_Style.parse -- Args.term)
(fn ctxt => fn ((name, style), t) =>
- Thy_Output.pretty_term ctxt (style (value_select name ctxt t)))
+ Document_Output.pretty_term ctxt (style (value_select name ctxt t)))
#> add_evaluator (\<^binding>\<open>simp\<close>, Code_Simp.dynamic_value) #> snd
#> add_evaluator (\<^binding>\<open>nbe\<close>, Nbe.dynamic_value) #> snd
#> add_evaluator (\<^binding>\<open>code\<close>, Code_Evaluation.dynamic_value_strict) #> snd);
--- a/src/Pure/Admin/build_csdp.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Admin/build_csdp.scala Sun May 23 20:34:43 2021 +0200
@@ -23,13 +23,13 @@
def print: Option[String] =
if (changed.isEmpty) None
else
- Some(" * " + platform + ":\n" + changed.map(p => " " + p._1 + "=" + p._2)
+ Some(" * " + platform + ":\n" + changed.map(p => " " + Properties.Eq(p))
.mkString("\n"))
def change(path: Path): Unit =
{
- def change_line(line: String, entry: (String, String)): String =
- line.replaceAll(entry._1 + "=.*", entry._1 + "=" + entry._2)
+ def change_line(line: String, p: (String, String)): String =
+ line.replaceAll(p._1 + "=.*", Properties.Eq(p))
File.change(path, s =>
split_lines(s).map(line => changed.foldLeft(line)(change_line)).mkString("\n"))
}
--- a/src/Pure/Admin/build_doc.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Admin/build_doc.scala Sun May 23 20:34:43 2021 +0200
@@ -55,7 +55,7 @@
progress.echo("Documentation " + quote(doc) + " ...")
using(store.open_database_context())(db_context =>
- Presentation.build_documents(session, deps, db_context,
+ Document_Build.build_documents(Document_Build.context(session, deps, db_context),
output_pdf = Some(Path.explode("~~/doc"))))
None
}
--- a/src/Pure/Admin/build_log.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Admin/build_log.scala Sun May 23 20:34:43 2021 +0200
@@ -60,15 +60,10 @@
object Entry
{
def unapply(s: String): Option[Entry] =
- s.indexOf('=') match {
- case -1 => None
- case i =>
- val a = s.substring(0, i)
- val b = Library.perhaps_unquote(s.substring(i + 1))
- Some((a, b))
- }
- def apply(a: String, b: String): String = a + "=" + quote(b)
- def getenv(a: String): String = apply(a, Isabelle_System.getenv(a))
+ for { (a, b) <- Properties.Eq.unapply(s) }
+ yield (a, Library.perhaps_unquote(b))
+ def getenv(a: String): String =
+ Properties.Eq(a, quote(Isabelle_System.getenv(a)))
}
def show(): String =
@@ -231,11 +226,8 @@
/* settings */
- def get_setting(a: String): Option[Settings.Entry] =
- lines.find(_.startsWith(a + "=")) match {
- case Some(line) => Settings.Entry.unapply(line)
- case None => None
- }
+ def get_setting(name: String): Option[Settings.Entry] =
+ lines.collectFirst({ case Settings.Entry(a, b) if a == name => a -> b })
def get_all_settings: Settings.T =
for { c <- Settings.all_settings; entry <- get_setting(c.name) }
--- a/src/Pure/Admin/components.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Admin/components.scala Sun May 23 20:34:43 2021 +0200
@@ -339,7 +339,7 @@
Build and publish Isabelle components as .tar.gz archives on SSH server,
depending on system options:
-""" + Library.prefix_lines(" ", show_options) + "\n",
+""" + Library.indent_lines(2, show_options) + "\n",
"P" -> (_ => publish = true),
"f" -> (_ => force = true),
"o:" -> (arg => options = options + arg),
--- a/src/Pure/General/exn.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/General/exn.scala Sun May 23 20:34:43 2021 +0200
@@ -119,8 +119,7 @@
/* message */
def user_message(exn: Throwable): Option[String] =
- if (exn.getClass == classOf[RuntimeException] ||
- exn.getClass == classOf[User_Error])
+ if (exn.isInstanceOf[User_Error] || exn.getClass == classOf[RuntimeException])
{
Some(proper_string(exn.getMessage) getOrElse "Error")
}
--- a/src/Pure/General/http.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/General/http.scala Sun May 23 20:34:43 2021 +0200
@@ -200,10 +200,10 @@
sealed case class Arg(method: String, uri: URI, request: Bytes)
{
def decode_properties: Properties.T =
- space_explode('&', request.text).map(s =>
- space_explode('=', s) match {
- case List(a, b) => Url.decode(a) -> Url.decode(b)
- case _ => error("Malformed key-value pair in HTTP/POST: " + quote(s))
+ space_explode('&', request.text).map(
+ {
+ case Properties.Eq(a, b) => Url.decode(a) -> Url.decode(b)
+ case s => error("Malformed key-value pair in HTTP/POST: " + quote(s))
})
}
--- a/src/Pure/General/path.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/General/path.scala Sun May 23 20:34:43 2021 +0200
@@ -263,7 +263,7 @@
case Path.Variable(s) =>
val path = Path.explode(Isabelle_System.getenv_strict(s, env))
if (path.elems.exists(_.isInstanceOf[Path.Variable]))
- error("Illegal path variable nesting: " + s + "=" + path.toString)
+ error("Illegal path variable nesting: " + Properties.Eq(s, path.toString))
else path.elems
case x => List(x)
}
--- a/src/Pure/General/properties.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/General/properties.scala Sun May 23 20:34:43 2021 +0200
@@ -14,6 +14,18 @@
type Entry = (java.lang.String, java.lang.String)
type T = List[Entry]
+ object Eq
+ {
+ def apply(a: java.lang.String, b: java.lang.String): java.lang.String = a + "=" + b
+ def apply(entry: Entry): java.lang.String = apply(entry._1, entry._2)
+
+ def unapply(str: java.lang.String): Option[Entry] =
+ {
+ val i = str.indexOf('=')
+ if (i <= 0) None else Some((str.substring(0, i), str.substring(i + 1)))
+ }
+ }
+
def defined(props: T, name: java.lang.String): java.lang.Boolean =
props.exists({ case (x, _) => x == name })
--- a/src/Pure/ML/ml_file.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/ML/ml_file.ML Sun May 23 20:34:43 2021 +0200
@@ -21,7 +21,7 @@
val provide = Resources.provide_file file;
val source = Token.file_source file;
- val _ = Thy_Output.check_comments (Context.proof_of gthy) (Input.source_explode source);
+ val _ = Document_Output.check_comments (Context.proof_of gthy) (Input.source_explode source);
val flags: ML_Compiler.flags =
{environment = environment, redirect = true, verbose = true,
--- a/src/Pure/ML/ml_statistics.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/ML/ml_statistics.scala Sun May 23 20:34:43 2021 +0200
@@ -58,12 +58,7 @@
{
def progress_stdout(line: String): Unit =
{
- val props =
- Library.space_explode(',', line).flatMap((entry: String) =>
- Library.space_explode('=', entry) match {
- case List(a, b) => Some((a, b))
- case _ => None
- })
+ val props = Library.space_explode(',', line).flatMap(Properties.Eq.unapply)
if (props.nonEmpty) consume(props)
}
--- a/src/Pure/PIDE/command.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/PIDE/command.ML Sun May 23 20:34:43 2021 +0200
@@ -226,7 +226,7 @@
else Toplevel.command_errors int tr st;
fun check_token_comments ctxt tok =
- (Thy_Output.check_comments ctxt (Input.source_explode (Token.input_of tok)); [])
+ (Document_Output.check_comments ctxt (Input.source_explode (Token.input_of tok)); [])
handle exn =>
if Exn.is_interrupt exn then Exn.reraise exn
else Runtime.exn_messages exn;
--- a/src/Pure/PIDE/prover.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/PIDE/prover.scala Sun May 23 20:34:43 2021 +0200
@@ -50,7 +50,7 @@
kind + " [[" + res + "]]"
else
kind + " " +
- (for ((x, y) <- properties) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]"
+ (properties.map(Properties.Eq.apply)).mkString("{", ",", "}") + " [[" + res + "]]"
}
}
--- a/src/Pure/PIDE/resources.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/PIDE/resources.ML Sun May 23 20:34:43 2021 +0200
@@ -198,7 +198,7 @@
in (name, multi) end;
val _ = Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>scala_function\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>scala_function\<close>
(Scan.lift Parse.embedded_position) (#1 oo check_scala_function) #>
ML_Antiquotation.inline_embedded \<^binding>\<open>scala_function\<close>
(Args.context -- Scan.lift Parse.embedded_position
@@ -424,11 +424,11 @@
in
val _ = Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>session\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>session\<close>
(Scan.lift Parse.embedded_position) check_session #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>path\<close> (document_antiq check_path) (K I) #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>file\<close> (document_antiq check_file) (K I) #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>dir\<close> (document_antiq check_dir) (K I) #>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>path\<close> (document_antiq check_path) (K I) #>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>file\<close> (document_antiq check_file) (K I) #>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>dir\<close> (document_antiq check_dir) (K I) #>
ML_Antiquotation.value_embedded \<^binding>\<open>path\<close> (ML_antiq check_path) #>
ML_Antiquotation.value_embedded \<^binding>\<open>file\<close> (ML_antiq check_file) #>
ML_Antiquotation.value_embedded \<^binding>\<open>dir\<close> (ML_antiq check_dir) #>
--- a/src/Pure/PIDE/yxml.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/PIDE/yxml.scala Sun May 23 20:34:43 2021 +0200
@@ -71,12 +71,7 @@
else err("unbalanced element " + quote(name))
private def parse_attrib(source: CharSequence): (String, String) =
- {
- val s = source.toString
- val i = s.indexOf('=')
- if (i <= 0) err_attribute()
- (s.substring(0, i), s.substring(i + 1))
- }
+ Properties.Eq.unapply(source.toString) getOrElse err_attribute()
def parse_body(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body =
--- a/src/Pure/ROOT.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/ROOT.ML Sun May 23 20:34:43 2021 +0200
@@ -304,7 +304,7 @@
ML_file "ML/ml_antiquotations2.ML";
ML_file "ML/ml_pid.ML";
ML_file "Thy/document_antiquotation.ML";
-ML_file "Thy/thy_output.ML";
+ML_file "Thy/document_output.ML";
ML_file "Thy/document_antiquotations.ML";
ML_file "General/graph_display.ML";
ML_file "pure_syn.ML";
--- a/src/Pure/ROOT.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/ROOT.scala Sun May 23 20:34:43 2021 +0200
@@ -21,4 +21,3 @@
val proper_string = Library.proper_string _
def proper_list[A](list: List[A]): Option[List[A]] = Library.proper_list(list)
}
-
--- a/src/Pure/System/isabelle_platform.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/isabelle_platform.scala Sun May 23 20:34:43 2021 +0200
@@ -29,10 +29,7 @@
val result = ssh.execute("bash -c " + Bash.string(script)).check
new Isabelle_Platform(
result.out_lines.map(line =>
- space_explode('=', line) match {
- case List(a, b) => (a, b)
- case _ => error("Bad output: " + quote(result.out))
- }))
+ Properties.Eq.unapply(line) getOrElse error("Bad output: " + quote(result.out))))
}
}
--- a/src/Pure/System/isabelle_system.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/isabelle_system.scala Sun May 23 20:34:43 2021 +0200
@@ -128,11 +128,11 @@
if (rc != 0) error(output)
val entries =
- (for (entry <- space_explode('\u0000', File.read(dump)) if entry != "") yield {
- val i = entry.indexOf('=')
- if (i <= 0) entry -> ""
- else entry.substring(0, i) -> entry.substring(i + 1)
- }).toMap
+ space_explode('\u0000', File.read(dump)).flatMap(
+ {
+ case Properties.Eq(a, b) => Some(a -> b)
+ case s => if (s.isEmpty || s.startsWith("=")) None else Some(s -> "")
+ }).toMap
entries + ("PATH" -> entries("PATH_JVM")) - "PATH_JVM"
}
finally { dump.delete }
--- a/src/Pure/System/isabelle_tool.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/isabelle_tool.ML Sun May 23 20:34:43 2021 +0200
@@ -38,7 +38,7 @@
val _ =
Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>tool\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>tool\<close>
(Scan.lift Parse.embedded_position) check);
end;
--- a/src/Pure/System/isabelle_tool.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/isabelle_tool.scala Sun May 23 20:34:43 2021 +0200
@@ -187,6 +187,7 @@
Build_Docker.isabelle_tool,
Build_Job.isabelle_tool,
Doc.isabelle_tool,
+ Document_Build.isabelle_tool,
Dump.isabelle_tool,
Export.isabelle_tool,
ML_Process.isabelle_tool,
@@ -198,7 +199,6 @@
Phabricator.isabelle_tool2,
Phabricator.isabelle_tool3,
Phabricator.isabelle_tool4,
- Presentation.isabelle_tool,
Profiling_Report.isabelle_tool,
Server.isabelle_tool,
Sessions.isabelle_tool,
--- a/src/Pure/System/options.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/options.scala Sun May 23 20:34:43 2021 +0200
@@ -355,12 +355,10 @@
}
def + (str: String): Options =
- {
- str.indexOf('=') match {
- case -1 => this + (str, None)
- case i => this + (str.substring(0, i), str.substring(i + 1))
+ str match {
+ case Properties.Eq(a, b) => this + (a, b)
+ case _ => this + (str, None)
}
- }
def ++ (specs: List[Options.Spec]): Options =
specs.foldLeft(this) { case (x, (y, z)) => x + (y, z) }
--- a/src/Pure/System/scala_compiler.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/System/scala_compiler.ML Sun May 23 20:34:43 2021 +0200
@@ -65,21 +65,21 @@
val _ =
Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>scala\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>scala\<close>
(Scan.lift Args.embedded_position)
(fn _ => fn (s, pos) => (static_check (s, pos); s)) #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>scala_type\<close>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>scala_type\<close>
(Scan.lift (Args.embedded_position -- (types >> print_types)))
(fn _ => fn ((t, pos), type_args) =>
(static_check ("type _Test_" ^ type_args ^ " = " ^ t ^ type_args, pos);
scala_name (t ^ type_args))) #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>scala_object\<close>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>scala_object\<close>
(Scan.lift Args.embedded_position)
(fn _ => fn (x, pos) => (static_check ("val _test_ = " ^ x, pos); scala_name x)) #>
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>scala_method\<close>
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>scala_method\<close>
(Scan.lift (class -- Args.embedded_position -- types -- args))
(fn _ => fn (((class_context, (method, pos)), method_types), method_args) =>
let
--- a/src/Pure/Thy/bibtex.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/bibtex.ML Sun May 23 20:34:43 2021 +0200
@@ -41,7 +41,7 @@
val _ =
Theory.setup
(Document_Antiquotation.setup_option \<^binding>\<open>cite_macro\<close> (Config.put cite_macro) #>
- Thy_Output.antiquotation_raw \<^binding>\<open>cite\<close>
+ Document_Output.antiquotation_raw \<^binding>\<open>cite\<close>
(Scan.lift
(Scan.option (Parse.verbatim || Parse.cartouche) -- Parse.and_list1 Args.name_position))
(fn ctxt => fn (opt, citations) =>
--- a/src/Pure/Thy/bibtex.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/bibtex.scala Sun May 23 20:34:43 2021 +0200
@@ -55,11 +55,13 @@
val log_path = dir + Path.explode(root_name).ext("blg")
if (log_path.is_file) {
val Error1 = """^(I couldn't open database file .+)$""".r
- val Error2 = """^(.+)---line (\d+) of file (.+)""".r
+ val Error2 = """^(I found no .+)$""".r
+ val Error3 = """^(.+)---line (\d+) of file (.+)""".r
Line.logical_lines(File.read(log_path)).flatMap(line =>
line match {
case Error1(msg) => Some("Bibtex error: " + msg)
- case Error2(msg, Value.Int(l), file) =>
+ case Error2(msg) => Some("Bibtex error: " + msg)
+ case Error3(msg, Value.Int(l), file) =>
val path = File.standard_path(file)
if (Path.is_wellformed(path)) {
val pos = Position.Line_File(l, (dir + Path.explode(path)).canonical.implode)
--- a/src/Pure/Thy/document_antiquotation.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/document_antiquotation.ML Sun May 23 20:34:43 2021 +0200
@@ -35,6 +35,7 @@
val integer: string -> int
val evaluate: (Symbol_Pos.T list -> Latex.text list) -> Proof.context ->
Antiquote.text_antiquote -> Latex.text list
+ val approx_content: Proof.context -> string -> string
end;
structure Document_Antiquotation: DOCUMENT_ANTIQUOTATION =
@@ -187,15 +188,52 @@
in
+fun read_antiq ctxt ({range = (pos, _), body, ...}: Antiquote.antiq) =
+ let val keywords = Thy_Header.get_keywords' ctxt;
+ in Token.read_antiq keywords parse_antiq (body, pos) end;
+
fun evaluate eval_text ctxt antiq =
(case antiq of
Antiquote.Text ss => eval_text ss
| Antiquote.Control {name, body, ...} =>
eval ctxt Position.none
([], Token.make_src name (if null body then [] else [Token.read_cartouche body]))
- | Antiquote.Antiq {range = (pos, _), body, ...} =>
- let val keywords = Thy_Header.get_keywords' ctxt;
- in eval ctxt pos (Token.read_antiq keywords parse_antiq (body, pos)) end);
+ | Antiquote.Antiq antiq => eval ctxt (#1 (#range antiq)) (read_antiq ctxt antiq));
+
+end;
+
+
+(* approximative content, e.g. for index entries *)
+
+local
+
+val strip_symbols = [Symbol.open_, Symbol.close, "'", "\"", "`"];
+
+fun strip_symbol sym =
+ if member (op =) strip_symbols sym then ""
+ else
+ (case Symbol.decode sym of
+ Symbol.Char s => s
+ | Symbol.UTF8 s => s
+ | Symbol.Sym s => s
+ | Symbol.Control s => s
+ | _ => "");
+
+fun strip_antiq _ (Antiquote.Text body) = map Symbol_Pos.symbol body
+ | strip_antiq _ (Antiquote.Control {body, ...}) = map Symbol_Pos.symbol body
+ | strip_antiq ctxt (Antiquote.Antiq antiq) =
+ read_antiq ctxt antiq |> #2 |> tl
+ |> maps (Symbol.explode o Token.content_of);
+
+in
+
+fun approx_content ctxt =
+ Symbol_Pos.explode0
+ #> trim (Symbol.is_blank o Symbol_Pos.symbol)
+ #> Antiquote.parse_comments Position.none
+ #> maps (strip_antiq ctxt)
+ #> map strip_symbol
+ #> implode;
end;
--- a/src/Pure/Thy/document_antiquotations.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/document_antiquotations.ML Sun May 23 20:34:43 2021 +0200
@@ -14,14 +14,14 @@
type style = term -> term;
fun pretty_term_style ctxt (style: style, t) =
- Thy_Output.pretty_term ctxt (style t);
+ Document_Output.pretty_term ctxt (style t);
fun pretty_thms_style ctxt (style: style, ths) =
- map (fn th => Thy_Output.pretty_term ctxt (style (Thm.full_prop_of th))) ths;
+ map (fn th => Document_Output.pretty_term ctxt (style (Thm.full_prop_of th))) ths;
fun pretty_term_typ ctxt (style: style, t) =
let val t' = style t
- in Thy_Output.pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end;
+ in Document_Output.pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end;
fun pretty_term_typeof ctxt (style: style, t) =
Syntax.pretty_typ ctxt (Term.fastype_of (style t));
@@ -31,7 +31,7 @@
val t = Const (c, Consts.type_scheme (Proof_Context.consts_of ctxt) c)
handle TYPE (msg, _, _) => error msg;
val (t', _) = yield_singleton (Variable.import_terms true) t ctxt;
- in Thy_Output.pretty_term ctxt t' end;
+ in Document_Output.pretty_term ctxt t' end;
fun pretty_abbrev ctxt s =
let
@@ -62,14 +62,12 @@
fun pretty_theory ctxt (name, pos) =
(Theory.check {long = true} ctxt (name, pos); Pretty.str name);
-val basic_entity = Thy_Output.antiquotation_pretty_source_embedded;
+val basic_entity = Document_Output.antiquotation_pretty_source_embedded;
fun basic_entities name scan pretty =
Document_Antiquotation.setup name scan
(fn {context = ctxt, source = src, argument = xs} =>
- Thy_Output.pretty_items_source ctxt {embedded = false} src (map (pretty ctxt) xs));
-
-in
+ Document_Output.pretty_items_source ctxt {embedded = false} src (map (pretty ctxt) xs));
val _ = Theory.setup
(basic_entity \<^binding>\<open>prop\<close> (Term_Style.parse -- Args.prop) pretty_term_style #>
@@ -87,9 +85,9 @@
basic_entities \<^binding>\<open>full_prf\<close> Attrib.thms (pretty_prf true) #>
Document_Antiquotation.setup \<^binding>\<open>thm\<close> (Term_Style.parse -- Attrib.thms)
(fn {context = ctxt, source = src, argument = arg} =>
- Thy_Output.pretty_items_source ctxt {embedded = false} src (pretty_thms_style ctxt arg)));
+ Document_Output.pretty_items_source ctxt {embedded = false} src (pretty_thms_style ctxt arg)));
-end;
+in end;
(* Markdown errors *)
@@ -102,48 +100,80 @@
error ("Bad Markdown structure: illegal " ^ quote (Binding.name_of binding) ^
Position.here (Position.no_range_position (#1 (Token.range_of src)))))
-in
-
val _ =
Theory.setup
(markdown_error \<^binding>\<open>item\<close> #>
markdown_error \<^binding>\<open>enum\<close> #>
markdown_error \<^binding>\<open>descr\<close>);
-end;
+in end;
(* control spacing *)
val _ =
Theory.setup
- (Thy_Output.antiquotation_raw \<^binding>\<open>noindent\<close> (Scan.succeed ())
+ (Document_Output.antiquotation_raw \<^binding>\<open>noindent\<close> (Scan.succeed ())
(fn _ => fn () => Latex.string "\\noindent") #>
- Thy_Output.antiquotation_raw \<^binding>\<open>smallskip\<close> (Scan.succeed ())
+ Document_Output.antiquotation_raw \<^binding>\<open>smallskip\<close> (Scan.succeed ())
(fn _ => fn () => Latex.string "\\smallskip") #>
- Thy_Output.antiquotation_raw \<^binding>\<open>medskip\<close> (Scan.succeed ())
+ Document_Output.antiquotation_raw \<^binding>\<open>medskip\<close> (Scan.succeed ())
(fn _ => fn () => Latex.string "\\medskip") #>
- Thy_Output.antiquotation_raw \<^binding>\<open>bigskip\<close> (Scan.succeed ())
+ Document_Output.antiquotation_raw \<^binding>\<open>bigskip\<close> (Scan.succeed ())
(fn _ => fn () => Latex.string "\\bigskip"));
-(* control style *)
+(* nested document text *)
local
-fun control_antiquotation name s1 s2 =
- Thy_Output.antiquotation_raw_embedded name (Scan.lift Args.cartouche_input)
- (fn ctxt => Latex.enclose_block s1 s2 o Thy_Output.output_document ctxt {markdown = false});
-
-in
+fun nested_antiquotation name s1 s2 =
+ Document_Output.antiquotation_raw_embedded name (Scan.lift Args.cartouche_input)
+ (fn ctxt => fn txt =>
+ (Context_Position.reports ctxt (Document_Output.document_reports txt);
+ Latex.enclose_block s1 s2 (Document_Output.output_document ctxt {markdown = false} txt)));
val _ =
Theory.setup
- (control_antiquotation \<^binding>\<open>footnote\<close> "\\footnote{" "}" #>
- control_antiquotation \<^binding>\<open>emph\<close> "\\emph{" "}" #>
- control_antiquotation \<^binding>\<open>bold\<close> "\\textbf{" "}");
+ (nested_antiquotation \<^binding>\<open>footnote\<close> "\\footnote{" "}" #>
+ nested_antiquotation \<^binding>\<open>emph\<close> "\\emph{" "}" #>
+ nested_antiquotation \<^binding>\<open>bold\<close> "\\textbf{" "}");
+
+in end;
+
+
+(* index entries *)
+
+local
+
+val index_like = Parse.$$$ "(" |-- Parse.!!! (Parse.$$$ "is" |-- Args.name --| Parse.$$$ ")");
+val index_args = Parse.enum1 "!" (Args.embedded_input -- Scan.option index_like);
+
+fun output_text ctxt = Latex.block o Document_Output.output_document ctxt {markdown = false};
-end;
+fun index binding def =
+ Document_Output.antiquotation_raw binding (Scan.lift index_args)
+ (fn ctxt => fn args =>
+ let
+ val _ = Context_Position.reports ctxt (maps (Document_Output.document_reports o #1) args);
+ fun make_item (txt, opt_like) =
+ let
+ val text = output_text ctxt txt;
+ val like =
+ (case opt_like of
+ SOME s => s
+ | NONE => Document_Antiquotation.approx_content ctxt (Input.string_of txt));
+ val _ =
+ if is_none opt_like andalso Context_Position.is_visible ctxt then
+ writeln ("(" ^ Markup.markup Markup.keyword2 "is" ^ " " ^ quote like ^ ")" ^
+ Position.here (Input.pos_of txt))
+ else ();
+ in {text = text, like = like} end;
+ in Latex.index_entry {items = map make_item args, def = def} end);
+
+val _ = Theory.setup (index \<^binding>\<open>index_ref\<close> false #> index \<^binding>\<open>index_def\<close> true);
+
+in end;
(* quasi-formal text (unchecked) *)
@@ -161,18 +191,18 @@
Input.source_content #> #1 #> Document_Antiquotation.prepare_lines ctxt;
fun text_antiquotation name =
- Thy_Output.antiquotation_raw_embedded name (Scan.lift Args.text_input)
+ Document_Output.antiquotation_raw_embedded name (Scan.lift Args.text_input)
(fn ctxt => fn text =>
let
val _ = report_text ctxt text;
in
prepare_text ctxt text
- |> Thy_Output.output_source ctxt
- |> Thy_Output.isabelle ctxt
+ |> Document_Output.output_source ctxt
+ |> Document_Output.isabelle ctxt
end);
val theory_text_antiquotation =
- Thy_Output.antiquotation_raw_embedded \<^binding>\<open>theory_text\<close> (Scan.lift Args.text_input)
+ Document_Output.antiquotation_raw_embedded \<^binding>\<open>theory_text\<close> (Scan.lift Args.text_input)
(fn ctxt => fn text =>
let
val keywords = Thy_Header.get_keywords' ctxt;
@@ -186,21 +216,17 @@
in
prepare_text ctxt text
|> Token.explode0 keywords
- |> maps (Thy_Output.output_token ctxt)
- |> Thy_Output.isabelle ctxt
+ |> maps (Document_Output.output_token ctxt)
+ |> Document_Output.isabelle ctxt
end);
-in
-
val _ =
Theory.setup
(text_antiquotation \<^binding>\<open>text\<close> #>
text_antiquotation \<^binding>\<open>cartouche\<close> #>
theory_text_antiquotation);
-end;
-
-
+in end;
(* goal state *)
@@ -208,19 +234,17 @@
local
fun goal_state name main =
- Thy_Output.antiquotation_pretty name (Scan.succeed ())
+ Document_Output.antiquotation_pretty name (Scan.succeed ())
(fn ctxt => fn () =>
Goal_Display.pretty_goal
(Config.put Goal_Display.show_main_goal main ctxt)
(#goal (Proof.goal (Toplevel.proof_of (Toplevel.presentation_state ctxt)))));
-in
-
val _ = Theory.setup
(goal_state \<^binding>\<open>goals\<close> true #>
goal_state \<^binding>\<open>subgoals\<close> false);
-end;
+in end;
(* embedded lemma *)
@@ -241,15 +265,15 @@
|> Proof.theorem NONE (K I) [[(prop, [])]]
|> Proof.global_terminal_proof (m1, m2);
in
- Thy_Output.pretty_source ctxt {embedded = false}
- [hd src, prop_tok] (Thy_Output.pretty_term ctxt prop)
+ Document_Output.pretty_source ctxt {embedded = false}
+ [hd src, prop_tok] (Document_Output.pretty_term ctxt prop)
end));
(* verbatim text *)
val _ = Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>verbatim\<close> (Scan.lift Args.text_input)
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>verbatim\<close> (Scan.lift Args.text_input)
(fn ctxt => fn text =>
let
val pos = Input.pos_of text;
@@ -263,14 +287,14 @@
(* bash functions *)
val _ = Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>bash_function\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>bash_function\<close>
(Scan.lift Args.embedded_position) Isabelle_System.check_bash_function);
(* system options *)
val _ = Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>system_option\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>system_option\<close>
(Scan.lift Args.embedded_position)
(fn ctxt => fn (name, pos) =>
let val _ = Completion.check_option (Options.default ()) ctxt (name, pos);
@@ -281,30 +305,104 @@
local
-fun ml_text name ml =
- Thy_Output.antiquotation_verbatim_embedded name (Scan.lift Args.text_input)
- (fn ctxt => fn text =>
- let val _ = ML_Context.eval_in (SOME ctxt) ML_Compiler.flags (Input.pos_of text) (ml text)
- in #1 (Input.source_content text) end);
+fun test_val (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read ");"
+ | test_val (ml1, ml2) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read ");";
+
+fun test_op (ml1, ml2) = test_val (ML_Lex.read "op " @ ml1, ml2);
+
+fun test_type (ml1, []) = ML_Lex.read "val _ = NONE : (" @ ml1 @ ML_Lex.read ") option;"
+ | test_type (ml1, ml2) =
+ ML_Lex.read "val _ = [NONE : (" @ ml1 @ ML_Lex.read ") option, NONE : (" @
+ ml2 @ ML_Lex.read ") option];";
+
+fun test_exn (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : exn);"
+ | test_exn (ml1, ml2) =
+ ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read " -> exn);";
+
+fun test_struct (ml, _) =
+ ML_Lex.read "functor XXX() = struct structure XX = " @ ml @ ML_Lex.read " end;";
+
+fun test_functor (Antiquote.Text tok :: _, _) =
+ ML_Lex.read "ML_Env.check_functor " @
+ ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok))
+ | test_functor _ = raise Fail "Bad ML functor specification";
+
+val parse_ml0 = Args.text_input >> (fn source => ("", (source, Input.empty)));
+
+val parse_ml =
+ Args.text_input -- Scan.optional (Args.colon |-- Args.text_input) Input.empty >> pair "";
+
+val parse_exn =
+ Args.text_input -- Scan.optional (Args.$$$ "of" |-- Args.text_input) Input.empty >> pair "";
+
+val parse_type =
+ (Parse.type_args >> (fn [] => "" | [a] => a ^ " " | bs => enclose "(" ") " (commas bs))) --
+ (Args.text_input -- Scan.optional (Args.$$$ "=" |-- Args.text_input) Input.empty);
+
+fun eval ctxt pos ml =
+ ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos ml
+ handle ERROR msg => error (msg ^ Position.here pos);
-fun ml_enclose bg en source =
- ML_Lex.read bg @ ML_Lex.read_source source @ ML_Lex.read en;
+fun make_text sep sources =
+ let
+ val (txt1, txt2) = apply2 (#1 o Input.source_content) sources;
+ val is_ident =
+ (case try List.last (Symbol.explode txt1) of
+ NONE => false
+ | SOME s => Symbol.is_ascii_letdig s);
+ val txt =
+ if txt2 = "" then txt1
+ else if sep = ":" andalso is_ident then txt1 ^ ": " ^ txt2
+ else txt1 ^ " " ^ sep ^ " " ^ txt2
+ in (txt, txt1) end;
+
+fun antiquotation_ml parse test kind show_kind binding index =
+ Document_Output.antiquotation_raw binding (Scan.lift parse)
+ (fn ctxt => fn (txt0, sources) =>
+ let
+ val (ml1, ml2) = apply2 ML_Lex.read_source sources;
+ val ml0 = ML_Lex.read_source (Input.string txt0);
+ val _ = test (ml0 @ ml1, ml2) |> eval ctxt (Input.pos_of (#1 sources));
+
+ val sep = if kind = "type" then "=" else if kind = "exception" then "of" else ":";
+ val (txt, idx) = make_text sep sources;
+
+ val main_text =
+ Document_Output.verbatim ctxt
+ ((if kind = "" orelse not show_kind then "" else kind ^ " ") ^ txt0 ^ txt);
+
+ val index_text =
+ index |> Option.map (fn def =>
+ let
+ val ctxt' = Config.put Document_Antiquotation.thy_output_display false ctxt;
+ val kind' = if kind = "" then " (ML)" else " (ML " ^ kind ^ ")";
+ val txt' = Latex.block [Document_Output.verbatim ctxt' idx, Latex.string kind'];
+ val like = Document_Antiquotation.approx_content ctxt' idx;
+ in Latex.index_entry {items = [{text = txt', like = like}], def = def} end);
+ in Latex.block (the_list index_text @ [main_text]) end);
+
+fun antiquotation_ml0 test kind =
+ antiquotation_ml parse_ml0 test kind false;
+
+fun antiquotation_ml1 parse test kind binding =
+ antiquotation_ml parse test kind true binding (SOME true);
in
-val _ = Theory.setup
- (ml_text \<^binding>\<open>ML\<close> (ml_enclose "fn _ => (" ");") #>
- ml_text \<^binding>\<open>ML_op\<close> (ml_enclose "fn _ => (op " ");") #>
- ml_text \<^binding>\<open>ML_type\<close> (ml_enclose "val _ = NONE : (" ") option;") #>
- ml_text \<^binding>\<open>ML_structure\<close>
- (ml_enclose "functor XXX() = struct structure XX = " " end;") #>
-
- ml_text \<^binding>\<open>ML_functor\<close> (* FIXME formal treatment of functor name (!?) *)
- (fn source =>
- ML_Lex.read ("ML_Env.check_functor " ^
- ML_Syntax.print_string (#1 (Input.source_content source)))) #>
-
- ml_text \<^binding>\<open>ML_text\<close> (K []));
+val _ =
+ Theory.setup
+ (Latex.index_variants (antiquotation_ml0 test_val "") \<^binding>\<open>ML\<close> #>
+ Latex.index_variants (antiquotation_ml0 test_op "infix") \<^binding>\<open>ML_infix\<close> #>
+ Latex.index_variants (antiquotation_ml0 test_type "type") \<^binding>\<open>ML_type\<close> #>
+ Latex.index_variants (antiquotation_ml0 test_struct "structure") \<^binding>\<open>ML_structure\<close> #>
+ Latex.index_variants (antiquotation_ml0 test_functor "functor") \<^binding>\<open>ML_functor\<close> #>
+ antiquotation_ml0 (K []) "text" \<^binding>\<open>ML_text\<close> NONE #>
+ antiquotation_ml1 parse_ml test_val "" \<^binding>\<open>define_ML\<close> #>
+ antiquotation_ml1 parse_ml test_op "infix" \<^binding>\<open>define_ML_infix\<close> #>
+ antiquotation_ml1 parse_type test_type "type" \<^binding>\<open>define_ML_type\<close> #>
+ antiquotation_ml1 parse_exn test_exn "exception" \<^binding>\<open>define_ML_exception\<close> #>
+ antiquotation_ml1 parse_ml0 test_struct "structure" \<^binding>\<open>define_ML_structure\<close> #>
+ antiquotation_ml1 parse_ml0 test_functor "functor" \<^binding>\<open>define_ML_functor\<close>);
end;
@@ -315,7 +413,7 @@
translate_string (fn c => if c = "%" orelse c = "#" orelse c = "^" then "\\" ^ c else c);
val _ = Theory.setup
- (Thy_Output.antiquotation_raw_embedded \<^binding>\<open>url\<close> (Scan.lift Args.embedded_input)
+ (Document_Output.antiquotation_raw_embedded \<^binding>\<open>url\<close> (Scan.lift Args.embedded_input)
(fn ctxt => fn source =>
let
val url = Input.string_of source;
@@ -332,19 +430,17 @@
local
fun entity_antiquotation name check bg en =
- Thy_Output.antiquotation_raw name (Scan.lift Args.name_position)
+ Document_Output.antiquotation_raw name (Scan.lift Args.name_position)
(fn ctxt => fn (name, pos) =>
let val _ = check ctxt (name, pos)
in Latex.enclose_block bg en [Latex.string (Output.output name)] end);
-in
-
val _ =
Theory.setup
(entity_antiquotation \<^binding>\<open>command\<close> Outer_Syntax.check_command "\\isacommand{" "}" #>
entity_antiquotation \<^binding>\<open>method\<close> Method.check_name "\\isa{" "}" #>
entity_antiquotation \<^binding>\<open>attribute\<close> Attrib.check_name "\\isa{" "}");
-end;
+in end;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/document_build.scala Sun May 23 20:34:43 2021 +0200
@@ -0,0 +1,535 @@
+/* Title: Pure/Thy/document_build.scala
+ Author: Makarius
+
+Build theory document (PDF) from session database.
+*/
+
+package isabelle
+
+
+object Document_Build
+{
+ /* document variants */
+
+ sealed case class Content(path: Path, bytes: Bytes)
+ {
+ def write(dir: Path): Unit = Bytes.write(dir + path, bytes)
+ }
+
+ abstract class Document_Name
+ {
+ def name: String
+ def path: Path = Path.basic(name)
+
+ override def toString: String = name
+ }
+
+ object Document_Variant
+ {
+ def parse(name: String, tags: String): Document_Variant =
+ Document_Variant(name, Library.space_explode(',', tags))
+
+ def parse(opt: String): Document_Variant =
+ Library.space_explode('=', opt) match {
+ case List(name) => Document_Variant(name, Nil)
+ case List(name, tags) => parse(name, tags)
+ case _ => error("Malformed document variant: " + quote(opt))
+ }
+ }
+
+ sealed case class Document_Variant(name: String, tags: List[String]) extends Document_Name
+ {
+ def print_tags: String = tags.mkString(",")
+ def print: String = if (tags.isEmpty) name else name + "=" + print_tags
+
+ def isabelletags: Content =
+ {
+ val path = Path.explode("isabelletags.sty")
+ val text =
+ Library.terminate_lines(
+ tags.map(tag =>
+ tag.toList match {
+ case '/' :: cs => "\\isafoldtag{" + cs.mkString + "}"
+ case '-' :: cs => "\\isadroptag{" + cs.mkString + "}"
+ case '+' :: cs => "\\isakeeptag{" + cs.mkString + "}"
+ case cs => "\\isakeeptag{" + cs.mkString + "}"
+ }))
+ Content(path, Bytes(text))
+ }
+ }
+
+ sealed case class Document_Input(name: String, sources: SHA1.Digest)
+ extends Document_Name
+
+ sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes)
+ extends Document_Name
+ {
+ def log: String = log_xz.uncompress().text
+ def log_lines: List[String] = split_lines(log)
+
+ def write(db: SQL.Database, session_name: String): Unit =
+ write_document(db, session_name, this)
+
+ def write(dir: Path): Path =
+ {
+ val path = dir + Path.basic(name).pdf
+ Isabelle_System.make_directory(path.expand.dir)
+ Bytes.write(path, pdf)
+ path
+ }
+ }
+
+
+ /* SQL data model */
+
+ object Data
+ {
+ val session_name = SQL.Column.string("session_name").make_primary_key
+ val name = SQL.Column.string("name").make_primary_key
+ val sources = SQL.Column.string("sources")
+ val log_xz = SQL.Column.bytes("log_xz")
+ val pdf = SQL.Column.bytes("pdf")
+
+ val table = SQL.Table("isabelle_documents", List(session_name, name, sources, log_xz, pdf))
+
+ def where_equal(session_name: String, name: String = ""): SQL.Source =
+ "WHERE " + Data.session_name.equal(session_name) +
+ (if (name == "") "" else " AND " + Data.name.equal(name))
+ }
+
+ def read_documents(db: SQL.Database, session_name: String): List[Document_Input] =
+ {
+ val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name))
+ db.using_statement(select)(stmt =>
+ stmt.execute_query().iterator(res =>
+ {
+ val name = res.string(Data.name)
+ val sources = res.string(Data.sources)
+ Document_Input(name, SHA1.fake(sources))
+ }).toList)
+ }
+
+ def read_document(db: SQL.Database, session_name: String, name: String): Option[Document_Output] =
+ {
+ val select = Data.table.select(sql = Data.where_equal(session_name, name))
+ db.using_statement(select)(stmt =>
+ {
+ val res = stmt.execute_query()
+ if (res.next()) {
+ val name = res.string(Data.name)
+ val sources = res.string(Data.sources)
+ val log_xz = res.bytes(Data.log_xz)
+ val pdf = res.bytes(Data.pdf)
+ Some(Document_Output(name, SHA1.fake(sources), log_xz, pdf))
+ }
+ else None
+ })
+ }
+
+ def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit =
+ {
+ db.using_statement(Data.table.insert())(stmt =>
+ {
+ stmt.string(1) = session_name
+ stmt.string(2) = doc.name
+ stmt.string(3) = doc.sources.toString
+ stmt.bytes(4) = doc.log_xz
+ stmt.bytes(5) = doc.pdf
+ stmt.execute()
+ })
+ }
+
+
+ /* context */
+
+ val isabelle_styles: List[Path] =
+ List("comment.sty", "isabelle.sty", "isabellesym.sty", "pdfsetup.sty", "railsetup.sty").
+ map(name => Path.explode("~~/lib/texinputs") + Path.basic(name))
+
+ def context(
+ session: String,
+ deps: Sessions.Deps,
+ db_context: Sessions.Database_Context,
+ progress: Progress = new Progress): Context =
+ {
+ val info = deps.sessions_structure(session)
+ val base = deps(session)
+ val hierarchy = deps.sessions_structure.hierarchy(session)
+ new Context(info, base, hierarchy, db_context, progress)
+ }
+
+ final class Context private[Document_Build](
+ info: Sessions.Info,
+ base: Sessions.Base,
+ hierarchy: List[String],
+ db_context: Sessions.Database_Context,
+ val progress: Progress = new Progress)
+ {
+ /* session info */
+
+ def session: String = info.name
+ def options: Options = info.options
+
+ def document_bibliography: Boolean = options.bool("document_bibliography")
+
+ def document_preprocessor: Option[String] =
+ proper_string(options.string("document_preprocessor"))
+
+ def document_logo: Option[String] =
+ options.string("document_logo") match {
+ case "" => None
+ case "_" => Some("")
+ case name => Some(name)
+ }
+
+ def document_build: String = options.string("document_build")
+
+ def get_engine(): Engine =
+ {
+ val name = document_build
+ engines.find(_.name == name).getOrElse(error("Bad document_build engine " + quote(name)))
+ }
+
+ def get_export(theory: String, name: String): Export.Entry =
+ db_context.get_export(hierarchy, theory, name)
+
+
+ /* document content */
+
+ def documents: List[Document_Variant] = info.documents
+
+ def session_theories: List[Document.Node.Name] = base.session_theories
+ def document_theories: List[Document.Node.Name] = session_theories ::: base.document_theories
+
+ lazy val tex_files: List[Content] =
+ for (name <- document_theories)
+ yield {
+ val path = Path.basic(tex_name(name))
+ val bytes = get_export(name.theory, document_tex_name(name)).uncompressed
+ Content(path, bytes)
+ }
+
+ lazy val session_graph: Content =
+ {
+ val path = Presentation.session_graph_path
+ val bytes = graphview.Graph_File.make_pdf(options, base.session_graph_display)
+ Content(path, bytes)
+ }
+
+ lazy val session_tex: Content =
+ {
+ val path = Path.basic("session.tex")
+ val text =
+ Library.terminate_lines(
+ base.session_theories.map(name => "\\input{" + tex_name(name) + "}"))
+ Content(path, Bytes(text))
+ }
+
+ lazy val isabelle_logo: Option[Content] =
+ {
+ document_logo.map(logo_name =>
+ Isabelle_System.with_tmp_file("logo", ext = "pdf")(tmp_path =>
+ {
+ Logo.create_logo(logo_name, output_file = tmp_path, quiet = true)
+ val path = Path.basic("isabelle_logo.pdf")
+ val bytes = Bytes.read(tmp_path)
+ Content(path, bytes)
+ }))
+ }
+
+
+ /* document directory */
+
+ def prepare_directory(dir: Path, doc: Document_Variant): Directory =
+ {
+ val doc_dir = dir + Path.basic(doc.name)
+ Isabelle_System.make_directory(doc_dir)
+
+
+ /* sources */
+
+ isabelle_styles.foreach(Isabelle_System.copy_file(_, doc_dir))
+ doc.isabelletags.write(doc_dir)
+
+ for ((base_dir, src) <- info.document_files) {
+ Isabelle_System.copy_file_base(info.dir + base_dir, src, doc_dir)
+ }
+
+ session_tex.write(doc_dir)
+ tex_files.foreach(_.write(doc_dir))
+
+ val root_name1 = "root_" + doc.name
+ val root_name = if ((doc_dir + Path.explode(root_name1).tex).is_file) root_name1 else "root"
+
+ for (name <- document_preprocessor) {
+ def message(s: String): String = s + " for document_preprocessor=" + quote(name)
+ val path = doc_dir + Path.explode(name)
+ if (path.is_file) {
+ try { Isabelle_System.bash(File.bash_path(path), cwd = doc_dir.file).check }
+ catch { case ERROR(msg) => cat_error(msg, message("The error(s) above occurred")) }
+ }
+ else error(message("Missing executable"))
+ }
+
+ val digests1 = List(doc.print, document_logo.toString, document_build).map(SHA1.digest)
+ val digests2 = File.find_files(doc_dir.file, follow_links = true).map(SHA1.digest)
+ val sources = SHA1.digest_set(digests1 ::: digests2)
+
+
+ /* derived material (without SHA1 digest) */
+
+ isabelle_logo.foreach(_.write(doc_dir))
+
+ session_graph.write(doc_dir)
+
+ Directory(doc_dir, doc, root_name, sources)
+ }
+
+ def old_document(directory: Directory): Option[Document_Output] =
+ for {
+ old_doc <- db_context.input_database(session)(read_document(_, _, directory.doc.name))
+ if old_doc.sources == directory.sources
+ }
+ yield old_doc
+ }
+
+ sealed case class Directory(
+ doc_dir: Path,
+ doc: Document_Variant,
+ root_name: String,
+ sources: SHA1.Digest)
+ {
+ def root_name_script(ext: String = ""): String =
+ Bash.string(if (ext.isEmpty) root_name else root_name + "." + ext)
+
+ def conditional_script(ext: String, exe: String, after: String = ""): String =
+ "if [ -f " + root_name_script(ext) + " ]\n" +
+ "then\n" +
+ " " + exe + " " + root_name_script() + "\n" +
+ (if (after.isEmpty) "" else " " + after) +
+ "fi\n"
+
+ def log_errors(): List[String] =
+ Latex.latex_errors(doc_dir, root_name) :::
+ Bibtex.bibtex_errors(doc_dir, root_name)
+
+ def make_document(log: List[String], errors: List[String]): Document_Output =
+ {
+ val root_pdf = Path.basic(root_name).pdf
+ val result_pdf = doc_dir + root_pdf
+
+ if (errors.nonEmpty) {
+ val errors1 = errors ::: List("Failed to build document " + quote(doc.name))
+ throw new Build_Error(log, Exn.cat_message(errors1: _*))
+ }
+ else if (!result_pdf.is_file) {
+ val message = "Bad document result: expected to find " + root_pdf
+ throw new Build_Error(log, message)
+ }
+ else {
+ val log_xz = Bytes(cat_lines(log)).compress()
+ val pdf = Bytes.read(result_pdf)
+ Document_Output(doc.name, sources, log_xz, pdf)
+ }
+ }
+ }
+
+
+ /* build engines */
+
+ lazy val engines: List[Engine] = Isabelle_System.make_services(classOf[Engine])
+
+ abstract class Engine(val name: String) extends Isabelle_System.Service
+ {
+ override def toString: String = name
+
+ def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory
+ def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output
+ }
+
+ abstract class Bash_Engine(name: String) extends Engine(name)
+ {
+ def prepare_directory(context: Context, dir: Path, doc: Document_Variant): Directory =
+ context.prepare_directory(dir, doc)
+
+ def use_pdflatex: Boolean = false
+ def latex_script(context: Context, directory: Directory): String =
+ (if (use_pdflatex) "$ISABELLE_PDFLATEX" else "$ISABELLE_LUALATEX") +
+ " " + directory.root_name_script() + "\n"
+
+ def bibtex_script(context: Context, directory: Directory, latex: Boolean = false): String =
+ {
+ val ext = if (context.document_bibliography) "aux" else "bib"
+ directory.conditional_script(ext, "$ISABELLE_BIBTEX",
+ after = if (latex) latex_script(context, directory) else "")
+ }
+
+ def makeindex_script(context: Context, directory: Directory, latex: Boolean = false): String =
+ directory.conditional_script("idx", "$ISABELLE_MAKEINDEX",
+ after = if (latex) latex_script(context, directory) else "")
+
+ def use_build_script: Boolean = false
+ def build_script(context: Context, directory: Directory): String =
+ {
+ val has_build_script = (directory.doc_dir + Path.explode("build")).is_file
+
+ if (!use_build_script && has_build_script) {
+ error("Unexpected document build script for option document_build=" +
+ quote(context.document_build))
+ }
+ else if (use_build_script && !has_build_script) error("Missing document build script")
+ else if (has_build_script) "./build pdf " + Bash.string(directory.doc.name)
+ else {
+ "set -e\n" +
+ latex_script(context, directory) +
+ bibtex_script(context, directory, latex = true) +
+ makeindex_script(context, directory) +
+ latex_script(context, directory) +
+ makeindex_script(context, directory, latex = true)
+ }
+ }
+
+ def build_document(context: Context, directory: Directory, verbose: Boolean): Document_Output =
+ {
+ val result =
+ context.progress.bash(
+ build_script(context, directory),
+ cwd = directory.doc_dir.file,
+ echo = verbose,
+ watchdog = Time.seconds(0.5))
+
+ val log = result.out_lines ::: result.err_lines
+ val errors = (if (result.ok) Nil else List(result.err)) ::: directory.log_errors()
+ directory.make_document(log, errors)
+ }
+ }
+
+ class LuaLaTeX_Engine extends Bash_Engine("lualatex")
+ class PDFLaTeX_Engine extends Bash_Engine("pdflatex") { override def use_pdflatex: Boolean = true }
+ class Build_Engine extends Bash_Engine("build") { override def use_build_script: Boolean = true }
+
+
+ /* build documents */
+
+ def tex_name(name: Document.Node.Name): String = name.theory_base_name + ".tex"
+ def document_tex_name(name: Document.Node.Name): String = "document/" + tex_name(name)
+
+ class Build_Error(val log_lines: List[String], val message: String)
+ extends Exn.User_Error(message)
+
+ def build_documents(
+ context: Context,
+ output_sources: Option[Path] = None,
+ output_pdf: Option[Path] = None,
+ verbose: Boolean = false): List[Document_Output] =
+ {
+ val progress = context.progress
+ val engine = context.get_engine()
+
+ val documents =
+ for (doc <- context.documents)
+ yield {
+ Isabelle_System.with_tmp_dir("document")(tmp_dir =>
+ {
+ progress.echo("Preparing " + context.session + "/" + doc.name + " ...")
+ val start = Time.now()
+
+ output_sources.foreach(engine.prepare_directory(context, _, doc))
+ val directory = engine.prepare_directory(context, tmp_dir, doc)
+
+ val document =
+ context.old_document(directory) getOrElse
+ engine.build_document(context, directory, verbose)
+
+ val stop = Time.now()
+ val timing = stop - start
+
+ progress.echo("Finished " + context.session + "/" + doc.name +
+ " (" + timing.message_hms + " elapsed time)")
+
+ document
+ })
+ }
+
+ for (dir <- output_pdf; doc <- documents) {
+ val path = doc.write(dir)
+ progress.echo("Document at " + path.absolute)
+ }
+
+ documents
+ }
+
+
+ /* Isabelle tool wrapper */
+
+ val isabelle_tool =
+ Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args =>
+ {
+ var output_sources: Option[Path] = None
+ var output_pdf: Option[Path] = None
+ var verbose_latex = false
+ var dirs: List[Path] = Nil
+ var options = Options.init()
+ var verbose_build = false
+
+ val getopts = Getopts(
+ """
+Usage: isabelle document [OPTIONS] SESSION
+
+ Options are:
+ -O DIR output directory for LaTeX sources and resulting PDF
+ -P DIR output directory for resulting PDF
+ -S DIR output directory for LaTeX sources
+ -V verbose latex
+ -d DIR include session directory
+ -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)
+ -v verbose build
+
+ Prepare the theory document of a session.
+""",
+ "O:" -> (arg =>
+ {
+ val dir = Path.explode(arg)
+ output_sources = Some(dir)
+ output_pdf = Some(dir)
+ }),
+ "P:" -> (arg => { output_pdf = Some(Path.explode(arg)) }),
+ "S:" -> (arg => { output_sources = Some(Path.explode(arg)) }),
+ "V" -> (_ => verbose_latex = true),
+ "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))),
+ "o:" -> (arg => options = options + arg),
+ "v" -> (_ => verbose_build = true))
+
+ val more_args = getopts(args)
+ val session =
+ more_args match {
+ case List(a) => a
+ case _ => getopts.usage()
+ }
+
+ val progress = new Console_Progress(verbose = verbose_build)
+ val store = Sessions.store(options)
+
+ progress.interrupt_handler {
+ val res =
+ Build.build(options, selection = Sessions.Selection.session(session),
+ dirs = dirs, progress = progress, verbose = verbose_build)
+ if (!res.ok) error("Failed to build session " + quote(session))
+
+ val deps =
+ Sessions.load_structure(options + "document=pdf", dirs = dirs).
+ selection_deps(Sessions.Selection.session(session))
+
+ if (output_sources.isEmpty && output_pdf.isEmpty) {
+ progress.echo_warning("No output directory")
+ }
+
+ using(store.open_database_context())(db_context =>
+ {
+ build_documents(context(session, deps, db_context, progress = progress),
+ output_sources = output_sources, output_pdf = output_pdf,
+ verbose = verbose_latex)
+ })
+ }
+ })
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/document_output.ML Sun May 23 20:34:43 2021 +0200
@@ -0,0 +1,577 @@
+(* Title: Pure/Thy/document_output.ML
+ Author: Makarius
+
+Theory document output.
+*)
+
+signature DOCUMENT_OUTPUT =
+sig
+ val document_reports: Input.source -> Position.report list
+ val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text list
+ val check_comments: Proof.context -> Symbol_Pos.T list -> unit
+ val output_token: Proof.context -> Token.T -> Latex.text list
+ val output_source: Proof.context -> string -> Latex.text list
+ type segment =
+ {span: Command_Span.span, command: Toplevel.transition,
+ prev_state: Toplevel.state, state: Toplevel.state}
+ val present_thy: Options.T -> theory -> segment list -> Latex.text list
+ val pretty_term: Proof.context -> term -> Pretty.T
+ val pretty_thm: Proof.context -> thm -> Pretty.T
+ val lines: Latex.text list -> Latex.text list
+ val items: Latex.text list -> Latex.text list
+ val isabelle: Proof.context -> Latex.text list -> Latex.text
+ val isabelle_typewriter: Proof.context -> Latex.text list -> Latex.text
+ val typewriter: Proof.context -> string -> Latex.text
+ val verbatim: Proof.context -> string -> Latex.text
+ val source: Proof.context -> {embedded: bool} -> Token.src -> Latex.text
+ val pretty: Proof.context -> Pretty.T -> Latex.text
+ val pretty_source: Proof.context -> {embedded: bool} -> Token.src -> Pretty.T -> Latex.text
+ val pretty_items: Proof.context -> Pretty.T list -> Latex.text
+ val pretty_items_source: Proof.context -> {embedded: bool} -> Token.src ->
+ Pretty.T list -> Latex.text
+ val antiquotation_pretty:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
+ val antiquotation_pretty_embedded:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
+ val antiquotation_pretty_source:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
+ val antiquotation_pretty_source_embedded:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
+ val antiquotation_raw:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
+ val antiquotation_raw_embedded:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
+ val antiquotation_verbatim:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
+ val antiquotation_verbatim_embedded:
+ binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
+end;
+
+structure Document_Output: DOCUMENT_OUTPUT =
+struct
+
+(* output document source *)
+
+fun document_reports txt =
+ let val pos = Input.pos_of txt in
+ [(pos, Markup.language_document (Input.is_delimited txt)),
+ (pos, Markup.plain_text)]
+ end;
+
+val output_symbols = single o Latex.symbols_output;
+
+fun output_comment ctxt (kind, syms) =
+ (case kind of
+ Comment.Comment =>
+ Input.cartouche_content syms
+ |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false)
+ {markdown = false}
+ |> Latex.enclose_body "%\n\\isamarkupcmt{" "%\n}"
+ | Comment.Cancel =>
+ Symbol_Pos.cartouche_content syms
+ |> output_symbols
+ |> Latex.enclose_body "%\n\\isamarkupcancel{" "}"
+ | Comment.Latex => [Latex.symbols (Symbol_Pos.cartouche_content syms)]
+ | Comment.Marker => [])
+and output_comment_document ctxt (comment, syms) =
+ (case comment of
+ SOME kind => output_comment ctxt (kind, syms)
+ | NONE => [Latex.symbols syms])
+and output_document_text ctxt syms =
+ Comment.read_body syms |> maps (output_comment_document ctxt)
+and output_document ctxt {markdown} source =
+ let
+ val pos = Input.pos_of source;
+ val syms = Input.source_explode source;
+
+ val output_antiquotes =
+ maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt);
+
+ fun output_line line =
+ (if Markdown.line_is_item line then [Latex.string "\\item "] else []) @
+ output_antiquotes (Markdown.line_content line);
+
+ fun output_block (Markdown.Par lines) =
+ Latex.block (separate (Latex.string "\n") (map (Latex.block o output_line) lines))
+ | output_block (Markdown.List {kind, body, ...}) =
+ Latex.environment_block (Markdown.print_kind kind) (output_blocks body)
+ and output_blocks blocks = separate (Latex.string "\n\n") (map output_block blocks);
+ in
+ if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then []
+ else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms
+ then
+ let
+ val ants = Antiquote.parse_comments pos syms;
+ val reports = Antiquote.antiq_reports ants;
+ val blocks = Markdown.read_antiquotes ants;
+ val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks);
+ in output_blocks blocks end
+ else
+ let
+ val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms);
+ val reports = Antiquote.antiq_reports ants;
+ val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants);
+ in output_antiquotes ants end
+ end;
+
+
+(* output tokens with formal comments *)
+
+local
+
+val output_symbols_antiq =
+ (fn Antiquote.Text syms => output_symbols syms
+ | Antiquote.Control {name = (name, _), body, ...} =>
+ Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) ::
+ output_symbols body
+ | Antiquote.Antiq {body, ...} =>
+ Latex.enclose_body "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (output_symbols body));
+
+fun output_comment_symbols ctxt {antiq} (comment, syms) =
+ (case (comment, antiq) of
+ (NONE, false) => output_symbols syms
+ | (NONE, true) =>
+ Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms
+ |> maps output_symbols_antiq
+ | (SOME comment, _) => output_comment ctxt (comment, syms));
+
+fun output_body ctxt antiq bg en syms =
+ Comment.read_body syms
+ |> maps (output_comment_symbols ctxt {antiq = antiq})
+ |> Latex.enclose_body bg en;
+
+in
+
+fun output_token ctxt tok =
+ let
+ fun output antiq bg en =
+ output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok));
+ in
+ (case Token.kind_of tok of
+ Token.Comment NONE => []
+ | Token.Comment (SOME Comment.Marker) => []
+ | Token.Command => output false "\\isacommand{" "}"
+ | Token.Keyword =>
+ if Symbol.is_ascii_identifier (Token.content_of tok)
+ then output false "\\isakeyword{" "}"
+ else output false "" ""
+ | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}"
+ | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}"
+ | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}"
+ | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}"
+ | _ => output false "" "")
+ end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok));
+
+fun output_source ctxt s =
+ output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none));
+
+fun check_comments ctxt =
+ Comment.read_body #> List.app (fn (comment, syms) =>
+ let
+ val pos = #1 (Symbol_Pos.range syms);
+ val _ =
+ comment |> Option.app (fn kind =>
+ Context_Position.reports ctxt (map (pair pos) (Comment.kind_markups kind)));
+ val _ = output_comment_symbols ctxt {antiq = false} (comment, syms);
+ in if comment = SOME Comment.Comment then check_comments ctxt syms else () end);
+
+end;
+
+
+
+(** present theory source **)
+
+(* presentation tokens *)
+
+datatype token =
+ Ignore
+ | Token of Token.T
+ | Heading of string * Input.source
+ | Body of string * Input.source
+ | Raw of Input.source;
+
+fun token_with pred (Token tok) = pred tok
+ | token_with _ _ = false;
+
+val white_token = token_with Document_Source.is_white;
+val white_comment_token = token_with Document_Source.is_white_comment;
+val blank_token = token_with Token.is_blank;
+val newline_token = token_with Token.is_newline;
+
+fun present_token ctxt tok =
+ (case tok of
+ Ignore => []
+ | Token tok => output_token ctxt tok
+ | Heading (cmd, source) =>
+ Latex.enclose_body ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n"
+ (output_document ctxt {markdown = false} source)
+ | Body (cmd, source) =>
+ [Latex.environment_block ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)]
+ | Raw source =>
+ Latex.string "%\n" :: output_document ctxt {markdown = true} source @ [Latex.string "\n"]);
+
+
+(* command spans *)
+
+type command = string * Position.T; (*name, position*)
+type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*)
+
+datatype span = Span of command * (source * source * source * source) * bool;
+
+fun make_span cmd src =
+ let
+ fun chop_newline (tok :: toks) =
+ if newline_token (fst tok) then ([tok], toks, true)
+ else ([], tok :: toks, false)
+ | chop_newline [] = ([], [], false);
+ val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =
+ src
+ |> chop_prefix (white_token o fst)
+ ||>> chop_suffix (white_token o fst)
+ ||>> chop_prefix (white_comment_token o fst)
+ ||> chop_newline;
+ in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;
+
+
+(* present spans *)
+
+local
+
+fun err_bad_nesting here =
+ error ("Bad nesting of commands in presentation" ^ here);
+
+fun edge which f (x: string option, y) =
+ if x = y then I
+ else (case which (x, y) of NONE => I | SOME txt => cons (Latex.string (f txt)));
+
+val begin_tag = edge #2 Latex.begin_tag;
+val end_tag = edge #1 Latex.end_tag;
+fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;
+fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;
+
+fun document_tag cmd_pos state state' tagging_stack =
+ let
+ val ctxt' = Toplevel.presentation_context state';
+ val nesting = Toplevel.level state' - Toplevel.level state;
+
+ val (tagging, taggings) = tagging_stack;
+ val (tag', tagging') = Document_Source.update_tagging ctxt' tagging;
+
+ val tagging_stack' =
+ if nesting = 0 andalso not (Toplevel.is_proof state) then tagging_stack
+ else if nesting >= 0 then (tagging', replicate nesting tagging @ taggings)
+ else
+ (case drop (~ nesting - 1) taggings of
+ tg :: tgs => (tg, tgs)
+ | [] => err_bad_nesting (Position.here cmd_pos));
+ in (tag', tagging_stack') end;
+
+fun read_tag s =
+ (case space_explode "%" s of
+ ["", b] => (SOME b, NONE)
+ | [a, b] => (NONE, SOME (a, b))
+ | _ => error ("Bad document_tags specification: " ^ quote s));
+
+in
+
+fun make_command_tag options keywords =
+ let
+ val document_tags =
+ map read_tag (space_explode "," (Options.string options \<^system_option>\<open>document_tags\<close>));
+ val document_tags_default = map_filter #1 document_tags;
+ val document_tags_command = map_filter #2 document_tags;
+ in
+ fn cmd_name => fn state => fn state' => fn active_tag =>
+ let
+ val keyword_tags =
+ if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"]
+ else Keyword.command_tags keywords cmd_name;
+ val command_tags =
+ the_list (AList.lookup (op =) document_tags_command cmd_name) @
+ keyword_tags @ document_tags_default;
+ val active_tag' =
+ (case command_tags of
+ default_tag :: _ => SOME default_tag
+ | [] =>
+ if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state
+ then active_tag
+ else NONE);
+ in active_tag' end
+ end;
+
+fun present_span command_tag span state state'
+ (tagging_stack, active_tag, newline, latex, present_cont) =
+ let
+ val ctxt' = Toplevel.presentation_context state';
+ val present = fold (fn (tok, (flag, 0)) =>
+ fold cons (present_token ctxt' tok)
+ #> cons (Latex.string flag)
+ | _ => I);
+
+ val Span ((cmd_name, cmd_pos), srcs, span_newline) = span;
+
+ val (tag', tagging_stack') = document_tag cmd_pos state state' tagging_stack;
+ val active_tag' =
+ if is_some tag' then Option.map #1 tag'
+ else command_tag cmd_name state state' active_tag;
+ val edge = (active_tag, active_tag');
+
+ val newline' =
+ if is_none active_tag' then span_newline else newline;
+
+ val latex' =
+ latex
+ |> end_tag edge
+ |> close_delim (fst present_cont) edge
+ |> snd present_cont
+ |> open_delim (present (#1 srcs)) edge
+ |> begin_tag edge
+ |> present (#2 srcs);
+ val present_cont' =
+ if newline then (present (#3 srcs), present (#4 srcs))
+ else (I, present (#3 srcs) #> present (#4 srcs));
+ in (tagging_stack', active_tag', newline', latex', present_cont') end;
+
+fun present_trailer ((_, tags), active_tag, _, latex, present_cont) =
+ if not (null tags) then err_bad_nesting " at end of theory"
+ else
+ latex
+ |> end_tag (active_tag, NONE)
+ |> close_delim (fst present_cont) (active_tag, NONE)
+ |> snd present_cont;
+
+end;
+
+
+(* present_thy *)
+
+local
+
+val markup_true = "\\isamarkuptrue%\n";
+val markup_false = "\\isamarkupfalse%\n";
+
+val opt_newline = Scan.option (Scan.one Token.is_newline);
+
+val ignore =
+ Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
+ >> pair (d + 1)) ||
+ Scan.depend (fn d => Scan.one Token.is_end_ignore --|
+ (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)
+ >> pair (d - 1));
+
+val locale =
+ Scan.option ((Parse.$$$ "(" -- Document_Source.improper -- Parse.$$$ "in") |--
+ Parse.!!!
+ (Document_Source.improper |-- Parse.name --| (Document_Source.improper -- Parse.$$$ ")")));
+
+in
+
+type segment =
+ {span: Command_Span.span, command: Toplevel.transition,
+ prev_state: Toplevel.state, state: Toplevel.state};
+
+fun present_thy options thy (segments: segment list) =
+ let
+ val keywords = Thy_Header.get_keywords thy;
+
+
+ (* tokens *)
+
+ val ignored = Scan.state --| ignore
+ >> (fn d => (NONE, (Ignore, ("", d))));
+
+ fun markup pred mk flag = Scan.peek (fn d =>
+ Document_Source.improper |--
+ Parse.position (Scan.one (fn tok =>
+ Token.is_command tok andalso pred keywords (Token.content_of tok))) --
+ (Document_Source.annotation |--
+ Parse.!!!! ((Document_Source.improper -- locale -- Document_Source.improper) |--
+ Parse.document_source --| Document_Source.improper_end))
+ >> (fn ((tok, pos'), source) =>
+ let val name = Token.content_of tok
+ in (SOME (name, pos'), (mk (name, source), (flag, d))) end));
+
+ val command = Scan.peek (fn d =>
+ Scan.optional (Scan.one Token.is_command_modifier ::: Document_Source.improper) [] --
+ Scan.one Token.is_command --| Document_Source.annotation
+ >> (fn (cmd_mod, cmd) =>
+ map (fn tok => (NONE, (Token tok, ("", d)))) cmd_mod @
+ [(SOME (Token.content_of cmd, Token.pos_of cmd),
+ (Token cmd, (markup_false, d)))]));
+
+ val cmt = Scan.peek (fn d =>
+ Scan.one Document_Source.is_black_comment >> (fn tok => (NONE, (Token tok, ("", d)))));
+
+ val other = Scan.peek (fn d =>
+ Parse.not_eof >> (fn tok => (NONE, (Token tok, ("", d)))));
+
+ val tokens =
+ (ignored ||
+ markup Keyword.is_document_heading Heading markup_true ||
+ markup Keyword.is_document_body Body markup_true ||
+ markup Keyword.is_document_raw (Raw o #2) "") >> single ||
+ command ||
+ (cmt || other) >> single;
+
+
+ (* spans *)
+
+ val is_eof = fn (_, (Token x, _)) => Token.is_eof x | _ => false;
+ val stopper = Scan.stopper (K (NONE, (Token Token.eof, ("", 0)))) is_eof;
+
+ val cmd = Scan.one (is_some o fst);
+ val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
+
+ val white_comments = Scan.many (white_comment_token o fst o snd);
+ val blank = Scan.one (blank_token o fst o snd);
+ val newline = Scan.one (newline_token o fst o snd);
+ val before_cmd =
+ Scan.option (newline -- white_comments) --
+ Scan.option (newline -- white_comments) --
+ Scan.option (blank -- white_comments) -- cmd;
+
+ val span =
+ Scan.repeat non_cmd -- cmd --
+ Scan.repeat (Scan.unless before_cmd non_cmd) --
+ Scan.option (newline >> (single o snd))
+ >> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>
+ make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));
+
+ val spans = segments
+ |> maps (Command_Span.content o #span)
+ |> drop_suffix Token.is_space
+ |> Source.of_list
+ |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat))
+ |> Source.source stopper (Scan.error (Scan.bulk span))
+ |> Source.exhaust;
+
+ val command_results =
+ segments |> map_filter (fn {command, state, ...} =>
+ if Toplevel.is_ignored command then NONE else SOME (command, state));
+
+
+ (* present commands *)
+
+ val command_tag = make_command_tag options keywords;
+
+ fun present_command tr span st st' =
+ Toplevel.setmp_thread_position tr (present_span command_tag span st st');
+
+ fun present _ [] = I
+ | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest;
+ in
+ if length command_results = length spans then
+ (([], []), NONE, true, [], (I, I))
+ |> present (Toplevel.init_toplevel ()) (spans ~~ command_results)
+ |> present_trailer
+ |> rev
+ else error "Messed-up outer syntax for presentation"
+ end;
+
+end;
+
+
+
+(** standard output operations **)
+
+(* pretty printing *)
+
+fun pretty_term ctxt t =
+ Syntax.pretty_term (Proof_Context.augment t ctxt) t;
+
+fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
+
+
+(* default output *)
+
+val lines = separate (Latex.string "\\isanewline%\n");
+val items = separate (Latex.string "\\isasep\\isanewline%\n");
+
+fun isabelle ctxt body =
+ if Config.get ctxt Document_Antiquotation.thy_output_display
+ then Latex.environment_block "isabelle" body
+ else Latex.block (Latex.enclose_body "\\isa{" "}" body);
+
+fun isabelle_typewriter ctxt body =
+ if Config.get ctxt Document_Antiquotation.thy_output_display
+ then Latex.environment_block "isabellett" body
+ else Latex.block (Latex.enclose_body "\\isatt{" "}" body);
+
+fun typewriter ctxt s =
+ isabelle_typewriter ctxt [Latex.string (Latex.output_ascii s)];
+
+fun verbatim ctxt =
+ if Config.get ctxt Document_Antiquotation.thy_output_display
+ then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt
+ else split_lines #> map (typewriter ctxt) #> lines #> Latex.block;
+
+fun token_source ctxt {embedded} tok =
+ if Token.is_kind Token.Cartouche tok andalso embedded andalso
+ not (Config.get ctxt Document_Antiquotation.thy_output_source_cartouche)
+ then Token.content_of tok
+ else Token.unparse tok;
+
+fun is_source ctxt =
+ Config.get ctxt Document_Antiquotation.thy_output_source orelse
+ Config.get ctxt Document_Antiquotation.thy_output_source_cartouche;
+
+fun source ctxt embedded =
+ Token.args_of_src
+ #> map (token_source ctxt embedded #> Document_Antiquotation.prepare_lines ctxt)
+ #> space_implode " "
+ #> output_source ctxt
+ #> isabelle ctxt;
+
+fun pretty ctxt =
+ Document_Antiquotation.output ctxt #> Latex.string #> single #> isabelle ctxt;
+
+fun pretty_source ctxt embedded src prt =
+ if is_source ctxt then source ctxt embedded src else pretty ctxt prt;
+
+fun pretty_items ctxt =
+ map (Document_Antiquotation.output ctxt #> Latex.string) #> items #> isabelle ctxt;
+
+fun pretty_items_source ctxt embedded src prts =
+ if is_source ctxt then source ctxt embedded src else pretty_items ctxt prts;
+
+
+(* antiquotation variants *)
+
+local
+
+fun gen_setup name embedded =
+ if embedded
+ then Document_Antiquotation.setup_embedded name
+ else Document_Antiquotation.setup name;
+
+fun gen_antiquotation_pretty name embedded scan f =
+ gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x));
+
+fun gen_antiquotation_pretty_source name embedded scan f =
+ gen_setup name embedded scan
+ (fn {context = ctxt, source = src, argument = x} =>
+ pretty_source ctxt {embedded = embedded} src (f ctxt x));
+
+fun gen_antiquotation_raw name embedded scan f =
+ gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => f ctxt x);
+
+fun gen_antiquotation_verbatim name embedded scan f =
+ gen_antiquotation_raw name embedded scan (fn ctxt => verbatim ctxt o f ctxt);
+
+in
+
+fun antiquotation_pretty name = gen_antiquotation_pretty name false;
+fun antiquotation_pretty_embedded name = gen_antiquotation_pretty name true;
+
+fun antiquotation_pretty_source name = gen_antiquotation_pretty_source name false;
+fun antiquotation_pretty_source_embedded name = gen_antiquotation_pretty_source name true;
+
+fun antiquotation_raw name = gen_antiquotation_raw name false;
+fun antiquotation_raw_embedded name = gen_antiquotation_raw name true;
+
+fun antiquotation_verbatim name = gen_antiquotation_verbatim name false;
+fun antiquotation_verbatim_embedded name = gen_antiquotation_verbatim name true;
+
+end;
+
+end;
--- a/src/Pure/Thy/latex.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/latex.ML Sun May 23 20:34:43 2021 +0200
@@ -1,7 +1,7 @@
(* Title: Pure/Thy/latex.ML
- Author: Markus Wenzel, TU Muenchen
+ Author: Makarius
-LaTeX presentation elements -- based on outer lexical syntax.
+LaTeX output of theory sources.
*)
signature LATEX =
@@ -29,6 +29,12 @@
val environment: string -> string -> string
val isabelle_body: string -> text list -> text list
val theory_entry: string -> string
+ val index_escape: string -> string
+ type index_item = {text: text, like: string}
+ val index_item: index_item -> text
+ type index_entry = {items: index_item list, def: bool}
+ val index_entry: index_entry -> text
+ val index_variants: (binding -> bool option -> 'a -> 'a) -> binding -> 'a -> 'a
val latexN: string
val latex_output: string -> string * int
val latex_markup: string * Properties.T -> Markup.output
@@ -47,6 +53,9 @@
val text = Text;
val block = Block;
+fun map_text f (Text (s, pos)) = Text (f s, pos)
+ | map_text f (Block texts) = Block ((map o map_text) f texts);
+
fun output_text texts =
let
fun output (Text (s, _)) = Buffer.add s
@@ -243,6 +252,36 @@
fun theory_entry name = "\\input{" ^ name ^ ".tex}\n\n";
+(* index entries *)
+
+type index_item = {text: text, like: string};
+type index_entry = {items: index_item list, def: bool};
+
+val index_escape =
+ translate_string (fn s =>
+ if member_string "!\"@|" s then "\\char" ^ string_of_int (ord s)
+ else if member_string "\\{}#" s then "\"" ^ s else s);
+
+fun index_item (item: index_item) =
+ let
+ val like_text =
+ if #like item = "" then []
+ else [string (index_escape (#like item) ^ "@")];
+ in block (like_text @ [map_text index_escape (#text item)]) end;
+
+fun index_entry (entry: index_entry) =
+ (separate (string "!") (map index_item (#items entry)) @
+ [string ("|" ^ index_escape (if #def entry then "isaindexdef" else "isaindexref"))])
+ |> enclose_block "\\index{" "}";
+
+
+fun index_binding NONE = I
+ | index_binding (SOME def) = Binding.map_name (suffix (if def then "_def" else "_ref"));
+
+fun index_variants setup binding =
+ fold (fn index => setup (index_binding index binding) index) [NONE, SOME true, SOME false];
+
+
(* print mode *)
val latexN = "latex";
--- a/src/Pure/Thy/latex.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/latex.scala Sun May 23 20:34:43 2021 +0200
@@ -22,7 +22,7 @@
val root_log_path = dir + Path.explode(root_name).ext("log")
if (root_log_path.is_file) {
for { (msg, pos) <- filter_errors(dir, File.read(root_log_path)) }
- yield "Latex error" + Position.here(pos) + ":\n" + Library.prefix_lines(" ", msg)
+ yield "Latex error" + Position.here(pos) + ":\n" + Library.indent_lines(2, msg)
}
else Nil
}
--- a/src/Pure/Thy/presentation.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/presentation.scala Sun May 23 20:34:43 2021 +0200
@@ -1,7 +1,7 @@
-/* Title: Pure/Thy/present.scala
+/* Title: Pure/Thy/presentation.scala
Author: Makarius
-HTML/PDF presentation of theory documents.
+HTML presentation of PIDE document content.
*/
package isabelle
@@ -14,6 +14,8 @@
{
/** HTML documents **/
+ /* HTML context */
+
val fonts_path = Path.explode("fonts")
sealed case class HTML_Document(title: String, content: String)
@@ -74,7 +76,7 @@
language = Markup.Elements(Markup.Language.DOCUMENT))
- /* HTML */
+ /* HTML output */
private val div_elements =
Set(HTML.div.name, HTML.pre.name, HTML.par.name, HTML.list.name, HTML.enum.name,
@@ -169,120 +171,7 @@
- /** PDF LaTeX documents **/
-
- /* document info */
-
- abstract class Document_Name
- {
- def name: String
- def path: Path = Path.basic(name)
-
- override def toString: String = name
- }
-
- object Document_Variant
- {
- def parse(name: String, tags: String): Document_Variant =
- Document_Variant(name, Library.space_explode(',', tags))
-
- def parse(opt: String): Document_Variant =
- Library.space_explode('=', opt) match {
- case List(name) => Document_Variant(name, Nil)
- case List(name, tags) => parse(name, tags)
- case _ => error("Malformed document variant: " + quote(opt))
- }
- }
-
- sealed case class Document_Variant(name: String, tags: List[String]) extends Document_Name
- {
- def print_tags: String = tags.mkString(",")
- def print: String = if (tags.isEmpty) name else name + "=" + print_tags
-
- def latex_sty: String =
- Library.terminate_lines(
- tags.map(tag =>
- tag.toList match {
- case '/' :: cs => "\\isafoldtag{" + cs.mkString + "}"
- case '-' :: cs => "\\isadroptag{" + cs.mkString + "}"
- case '+' :: cs => "\\isakeeptag{" + cs.mkString + "}"
- case cs => "\\isakeeptag{" + cs.mkString + "}"
- }))
- }
-
- sealed case class Document_Input(name: String, sources: SHA1.Digest)
- extends Document_Name
-
- sealed case class Document_Output(name: String, sources: SHA1.Digest, log_xz: Bytes, pdf: Bytes)
- extends Document_Name
- {
- def log: String = log_xz.uncompress().text
- def log_lines: List[String] = split_lines(log)
-
- def write(db: SQL.Database, session_name: String): Unit =
- write_document(db, session_name, this)
- }
-
-
- /* SQL data model */
-
- object Data
- {
- val session_name = SQL.Column.string("session_name").make_primary_key
- val name = SQL.Column.string("name").make_primary_key
- val sources = SQL.Column.string("sources")
- val log_xz = SQL.Column.bytes("log_xz")
- val pdf = SQL.Column.bytes("pdf")
-
- val table = SQL.Table("isabelle_documents", List(session_name, name, sources, log_xz, pdf))
-
- def where_equal(session_name: String, name: String = ""): SQL.Source =
- "WHERE " + Data.session_name.equal(session_name) +
- (if (name == "") "" else " AND " + Data.name.equal(name))
- }
-
- def read_documents(db: SQL.Database, session_name: String): List[Document_Input] =
- {
- val select = Data.table.select(List(Data.name, Data.sources), Data.where_equal(session_name))
- db.using_statement(select)(stmt =>
- stmt.execute_query().iterator(res =>
- {
- val name = res.string(Data.name)
- val sources = res.string(Data.sources)
- Document_Input(name, SHA1.fake(sources))
- }).toList)
- }
-
- def read_document(db: SQL.Database, session_name: String, name: String): Option[Document_Output] =
- {
- val select = Data.table.select(sql = Data.where_equal(session_name, name))
- db.using_statement(select)(stmt =>
- {
- val res = stmt.execute_query()
- if (res.next()) {
- val name = res.string(Data.name)
- val sources = res.string(Data.sources)
- val log_xz = res.bytes(Data.log_xz)
- val pdf = res.bytes(Data.pdf)
- Some(Document_Output(name, SHA1.fake(sources), log_xz, pdf))
- }
- else None
- })
- }
-
- def write_document(db: SQL.Database, session_name: String, doc: Document_Output): Unit =
- {
- db.using_statement(Data.table.insert())(stmt =>
- {
- stmt.string(1) = session_name
- stmt.string(2) = doc.name
- stmt.string(3) = doc.sources.toString
- stmt.bytes(4) = doc.log_xz
- stmt.bytes(5) = doc.pdf
- stmt.execute()
- })
- }
-
+ /** HTML presentation **/
/* presentation context */
@@ -311,9 +200,6 @@
}
-
- /** HTML presentation **/
-
/* maintain chapter index */
private val sessions_path = Path.basic(".sessions")
@@ -498,7 +384,7 @@
val documents =
for {
doc <- info.document_variants
- document <- db_context.input_database(session)(read_document(_, _, doc.name))
+ document <- db_context.input_database(session)(Document_Build.read_document(_, _, doc.name))
} yield {
if (verbose) progress.echo("Presenting document " + session + "/" + doc.name)
Bytes.write(session_dir + doc.path.pdf, document.pdf)
@@ -613,259 +499,4 @@
html_context.head(title, List(HTML.par(view_links))) ::
html_context.contents("Theories", theories))
}
-
-
-
- /** build documents **/
-
- val session_tex_path = Path.explode("session.tex")
-
- def tex_name(name: Document.Node.Name): String = name.theory_base_name + ".tex"
- def document_tex_name(name: Document.Node.Name): String = "document/" + tex_name(name)
-
- class Build_Error(val log_lines: List[String], val message: String)
- extends Exn.User_Error(message)
-
- def build_documents(
- session: String,
- deps: Sessions.Deps,
- db_context: Sessions.Database_Context,
- progress: Progress = new Progress,
- output_sources: Option[Path] = None,
- output_pdf: Option[Path] = None,
- verbose: Boolean = false,
- verbose_latex: Boolean = false): List[Document_Output] =
- {
- /* session info */
-
- val info = deps.sessions_structure(session)
- val hierarchy = deps.sessions_structure.hierarchy(session)
- val options = info.options
- val base = deps(session)
- val graph_pdf = graphview.Graph_File.make_pdf(options, base.session_graph_display)
-
-
- /* prepare document directory */
-
- lazy val tex_files =
- for (name <- base.session_theories ::: base.document_theories)
- yield {
- val entry = db_context.get_export(hierarchy, name.theory, document_tex_name(name))
- Path.basic(tex_name(name)) -> entry.uncompressed
- }
-
- def prepare_dir1(dir: Path, doc: Document_Variant): (Path, String) =
- {
- val doc_dir = dir + Path.basic(doc.name)
- Isabelle_System.make_directory(doc_dir)
-
- Isabelle_System.bash("isabelle latex -o sty", cwd = doc_dir.file).check
- File.write(doc_dir + Path.explode("isabelletags.sty"), doc.latex_sty)
- for ((base_dir, src) <- info.document_files) {
- Isabelle_System.copy_file_base(info.dir + base_dir, src, doc_dir)
- }
-
- File.write(doc_dir + session_tex_path,
- Library.terminate_lines(
- base.session_theories.map(name => "\\input{" + tex_name(name) + "}")))
-
- for ((path, tex) <- tex_files) Bytes.write(doc_dir + path, tex)
-
- val root1 = "root_" + doc.name
- val root = if ((doc_dir + Path.explode(root1).tex).is_file) root1 else "root"
-
- (doc_dir, root)
- }
-
- def prepare_dir2(dir: Path, doc: Document_Variant): Unit =
- {
- val doc_dir = dir + Path.basic(doc.name)
-
- // non-deterministic, but irrelevant
- Bytes.write(doc_dir + session_graph_path, graph_pdf)
- }
-
-
- /* produce documents */
-
- val documents =
- for (doc <- info.documents)
- yield {
- Isabelle_System.with_tmp_dir("document")(tmp_dir =>
- {
- progress.echo("Preparing " + session + "/" + doc.name + " ...")
- val start = Time.now()
-
-
- // prepare sources
-
- val (doc_dir, root) = prepare_dir1(tmp_dir, doc)
- val digests = File.find_files(doc_dir.file, follow_links = true).map(SHA1.digest)
- val sources = SHA1.digest_set(digests)
- prepare_dir2(tmp_dir, doc)
-
- for (dir <- output_sources) {
- prepare_dir1(dir, doc)
- prepare_dir2(dir, doc)
- }
-
-
- // old document from database
-
- val old_document =
- for {
- old_doc <- db_context.input_database(session)(read_document(_, _, doc.name))
- if old_doc.sources == sources
- }
- yield {
- Bytes.write(doc_dir + doc.path.pdf, old_doc.pdf)
- old_doc
- }
-
- old_document getOrElse {
- // bash scripts
-
- def root_bash(ext: String): String = Bash.string(root + "." + ext)
-
- def latex_bash(fmt: String = "pdf", ext: String = "tex"): String =
- "isabelle latex -o " + Bash.string(fmt) + " " + Bash.string(root + "." + ext)
-
- def bash(items: String*): Process_Result =
- progress.bash(items.mkString(" && "), cwd = doc_dir.file,
- echo = verbose_latex, watchdog = Time.seconds(0.5))
-
-
- // prepare document
-
- val result =
- if ((doc_dir + Path.explode("build")).is_file) {
- bash("./build pdf " + Bash.string(doc.name))
- }
- else {
- bash(
- latex_bash(),
- "{ [ ! -f " + root_bash("bib") + " ] || " + latex_bash("bbl") + "; }",
- "{ [ ! -f " + root_bash("idx") + " ] || " + latex_bash("idx") + "; }",
- latex_bash(),
- latex_bash())
- }
-
-
- // result
-
- val root_pdf = Path.basic(root).pdf
- val result_path = doc_dir + root_pdf
-
- val log_lines = result.out_lines ::: result.err_lines
-
- val errors =
- (if (result.ok) Nil else List(result.err)) :::
- Latex.latex_errors(doc_dir, root) :::
- Bibtex.bibtex_errors(doc_dir, root)
-
- if (errors.nonEmpty) {
- val message =
- Exn.cat_message(
- (errors ::: List("Failed to build document " + quote(doc.name))):_*)
- throw new Build_Error(log_lines, message)
- }
- else if (!result_path.is_file) {
- val message = "Bad document result: expected to find " + root_pdf
- throw new Build_Error(log_lines, message)
- }
- else {
- val stop = Time.now()
- val timing = stop - start
- progress.echo("Finished " + session + "/" + doc.name +
- " (" + timing.message_hms + " elapsed time)")
-
- val log_xz = Bytes(cat_lines(log_lines)).compress()
- val pdf = Bytes.read(result_path)
- Document_Output(doc.name, sources, log_xz, pdf)
- }
- }
- })
- }
-
- for (dir <- output_pdf; doc <- documents) {
- Isabelle_System.make_directory(dir)
- val path = dir + doc.path.pdf
- Bytes.write(path, doc.pdf)
- progress.echo("Document at " + path.absolute)
- }
-
- documents
- }
-
-
- /* Isabelle tool wrapper */
-
- val isabelle_tool =
- Isabelle_Tool("document", "prepare session theory document", Scala_Project.here, args =>
- {
- var output_sources: Option[Path] = None
- var output_pdf: Option[Path] = None
- var verbose_latex = false
- var dirs: List[Path] = Nil
- var options = Options.init()
- var verbose_build = false
-
- val getopts = Getopts(
- """
-Usage: isabelle document [OPTIONS] SESSION
-
- Options are:
- -O DIR output directory for LaTeX sources and resulting PDF
- -P DIR output directory for resulting PDF
- -S DIR output directory for LaTeX sources
- -V verbose latex
- -d DIR include session directory
- -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)
- -v verbose build
-
- Prepare the theory document of a session.
-""",
- "O:" -> (arg =>
- {
- val dir = Path.explode(arg)
- output_sources = Some(dir)
- output_pdf = Some(dir)
- }),
- "P:" -> (arg => { output_pdf = Some(Path.explode(arg)) }),
- "S:" -> (arg => { output_sources = Some(Path.explode(arg)) }),
- "V" -> (_ => verbose_latex = true),
- "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))),
- "o:" -> (arg => options = options + arg),
- "v" -> (_ => verbose_build = true))
-
- val more_args = getopts(args)
- val session =
- more_args match {
- case List(a) => a
- case _ => getopts.usage()
- }
-
- val progress = new Console_Progress(verbose = verbose_build)
- val store = Sessions.store(options)
-
- progress.interrupt_handler {
- val res =
- Build.build(options, selection = Sessions.Selection.session(session),
- dirs = dirs, progress = progress, verbose = verbose_build)
- if (!res.ok) error("Failed to build session " + quote(session))
-
- val deps =
- Sessions.load_structure(options + "document=pdf", dirs = dirs).
- selection_deps(Sessions.Selection.session(session))
-
- if (output_sources.isEmpty && output_pdf.isEmpty) {
- progress.echo_warning("No output directory")
- }
-
- using(store.open_database_context())(db_context =>
- build_documents(session, deps, db_context, progress = progress,
- output_sources = output_sources, output_pdf = output_pdf,
- verbose = true, verbose_latex = verbose_latex))
- }
- })
}
--- a/src/Pure/Thy/sessions.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/sessions.scala Sun May 23 20:34:43 2021 +0200
@@ -514,11 +514,11 @@
case doc => error("Bad document specification " + quote(doc))
}
- def document_variants: List[Presentation.Document_Variant] =
+ def document_variants: List[Document_Build.Document_Variant] =
{
val variants =
Library.space_explode(':', options.string("document_variants")).
- map(Presentation.Document_Variant.parse)
+ map(Document_Build.Document_Variant.parse)
val dups = Library.duplicates(variants.map(_.name))
if (dups.nonEmpty) error("Duplicate document variants: " + commas_quote(dups))
@@ -526,7 +526,7 @@
variants
}
- def documents: List[Presentation.Document_Variant] =
+ def documents: List[Document_Build.Document_Variant] =
{
val variants = document_variants
if (!document_enabled || document_files.isEmpty) Nil else variants
@@ -1429,10 +1429,10 @@
db.using_statement(
Export.Data.table.delete(Export.Data.session_name.where_equal(name)))(_.execute())
- db.create_table(Presentation.Data.table)
+ db.create_table(Document_Build.Data.table)
db.using_statement(
- Presentation.Data.table.delete(
- Presentation.Data.session_name.where_equal(name)))(_.execute())
+ Document_Build.Data.table.delete(
+ Document_Build.Data.session_name.where_equal(name)))(_.execute())
}
}
--- a/src/Pure/Thy/thy_info.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Thy/thy_info.ML Sun May 23 20:34:43 2021 +0200
@@ -9,7 +9,7 @@
sig
type presentation_context =
{options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T,
- segments: Thy_Output.segment list}
+ segments: Document_Output.segment list}
val adjust_pos_properties: presentation_context -> Position.T -> Properties.T
val apply_presentation: presentation_context -> theory -> unit
val add_presentation: (presentation_context -> theory -> unit) -> theory -> theory
@@ -42,7 +42,7 @@
type presentation_context =
{options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T,
- segments: Thy_Output.segment list};
+ segments: Document_Output.segment list};
fun adjust_pos_properties (context: presentation_context) pos =
Position.offset_properties_of (#adjust_pos context pos) @ Position.id_properties_of pos;
@@ -71,7 +71,7 @@
if exists (Toplevel.is_skipped_proof o #state) segments then ()
else
let
- val body = Thy_Output.present_thy options thy segments;
+ val body = Document_Output.present_thy options thy segments;
in
if Options.string options "document" = "false" then ()
else
--- a/src/Pure/Thy/thy_output.ML Wed May 19 14:17:40 2021 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,570 +0,0 @@
-(* Title: Pure/Thy/thy_output.ML
- Author: Makarius
-
-Theory document output.
-*)
-
-signature THY_OUTPUT =
-sig
- val output_document: Proof.context -> {markdown: bool} -> Input.source -> Latex.text list
- val check_comments: Proof.context -> Symbol_Pos.T list -> unit
- val output_token: Proof.context -> Token.T -> Latex.text list
- val output_source: Proof.context -> string -> Latex.text list
- type segment =
- {span: Command_Span.span, command: Toplevel.transition,
- prev_state: Toplevel.state, state: Toplevel.state}
- val present_thy: Options.T -> theory -> segment list -> Latex.text list
- val pretty_term: Proof.context -> term -> Pretty.T
- val pretty_thm: Proof.context -> thm -> Pretty.T
- val lines: Latex.text list -> Latex.text list
- val items: Latex.text list -> Latex.text list
- val isabelle: Proof.context -> Latex.text list -> Latex.text
- val isabelle_typewriter: Proof.context -> Latex.text list -> Latex.text
- val typewriter: Proof.context -> string -> Latex.text
- val verbatim: Proof.context -> string -> Latex.text
- val source: Proof.context -> {embedded: bool} -> Token.src -> Latex.text
- val pretty: Proof.context -> Pretty.T -> Latex.text
- val pretty_source: Proof.context -> {embedded: bool} -> Token.src -> Pretty.T -> Latex.text
- val pretty_items: Proof.context -> Pretty.T list -> Latex.text
- val pretty_items_source: Proof.context -> {embedded: bool} -> Token.src ->
- Pretty.T list -> Latex.text
- val antiquotation_pretty:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
- val antiquotation_pretty_embedded:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
- val antiquotation_pretty_source:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
- val antiquotation_pretty_source_embedded:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Pretty.T) -> theory -> theory
- val antiquotation_raw:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
- val antiquotation_raw_embedded:
- binding -> 'a context_parser -> (Proof.context -> 'a -> Latex.text) -> theory -> theory
- val antiquotation_verbatim:
- binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
- val antiquotation_verbatim_embedded:
- binding -> 'a context_parser -> (Proof.context -> 'a -> string) -> theory -> theory
-end;
-
-structure Thy_Output: THY_OUTPUT =
-struct
-
-(* output document source *)
-
-val output_symbols = single o Latex.symbols_output;
-
-fun output_comment ctxt (kind, syms) =
- (case kind of
- Comment.Comment =>
- Input.cartouche_content syms
- |> output_document (ctxt |> Config.put Document_Antiquotation.thy_output_display false)
- {markdown = false}
- |> Latex.enclose_body "%\n\\isamarkupcmt{" "%\n}"
- | Comment.Cancel =>
- Symbol_Pos.cartouche_content syms
- |> output_symbols
- |> Latex.enclose_body "%\n\\isamarkupcancel{" "}"
- | Comment.Latex => [Latex.symbols (Symbol_Pos.cartouche_content syms)]
- | Comment.Marker => [])
-and output_comment_document ctxt (comment, syms) =
- (case comment of
- SOME kind => output_comment ctxt (kind, syms)
- | NONE => [Latex.symbols syms])
-and output_document_text ctxt syms =
- Comment.read_body syms |> maps (output_comment_document ctxt)
-and output_document ctxt {markdown} source =
- let
- val pos = Input.pos_of source;
- val syms = Input.source_explode source;
-
- val output_antiquotes =
- maps (Document_Antiquotation.evaluate (output_document_text ctxt) ctxt);
-
- fun output_line line =
- (if Markdown.line_is_item line then [Latex.string "\\item "] else []) @
- output_antiquotes (Markdown.line_content line);
-
- fun output_block (Markdown.Par lines) =
- Latex.block (separate (Latex.string "\n") (map (Latex.block o output_line) lines))
- | output_block (Markdown.List {kind, body, ...}) =
- Latex.environment_block (Markdown.print_kind kind) (output_blocks body)
- and output_blocks blocks = separate (Latex.string "\n\n") (map output_block blocks);
- in
- if Toplevel.is_skipped_proof (Toplevel.presentation_state ctxt) then []
- else if markdown andalso exists (Markdown.is_control o Symbol_Pos.symbol) syms
- then
- let
- val ants = Antiquote.parse_comments pos syms;
- val reports = Antiquote.antiq_reports ants;
- val blocks = Markdown.read_antiquotes ants;
- val _ = Context_Position.reports ctxt (reports @ Markdown.reports blocks);
- in output_blocks blocks end
- else
- let
- val ants = Antiquote.parse_comments pos (trim (Symbol.is_blank o Symbol_Pos.symbol) syms);
- val reports = Antiquote.antiq_reports ants;
- val _ = Context_Position.reports ctxt (reports @ Markdown.text_reports ants);
- in output_antiquotes ants end
- end;
-
-
-(* output tokens with formal comments *)
-
-local
-
-val output_symbols_antiq =
- (fn Antiquote.Text syms => output_symbols syms
- | Antiquote.Control {name = (name, _), body, ...} =>
- Latex.string (Latex.output_symbols [Symbol.encode (Symbol.Control name)]) ::
- output_symbols body
- | Antiquote.Antiq {body, ...} =>
- Latex.enclose_body "%\n\\isaantiq\n" "{}%\n\\endisaantiq\n" (output_symbols body));
-
-fun output_comment_symbols ctxt {antiq} (comment, syms) =
- (case (comment, antiq) of
- (NONE, false) => output_symbols syms
- | (NONE, true) =>
- Antiquote.parse_comments (#1 (Symbol_Pos.range syms)) syms
- |> maps output_symbols_antiq
- | (SOME comment, _) => output_comment ctxt (comment, syms));
-
-fun output_body ctxt antiq bg en syms =
- Comment.read_body syms
- |> maps (output_comment_symbols ctxt {antiq = antiq})
- |> Latex.enclose_body bg en;
-
-in
-
-fun output_token ctxt tok =
- let
- fun output antiq bg en =
- output_body ctxt antiq bg en (Input.source_explode (Token.input_of tok));
- in
- (case Token.kind_of tok of
- Token.Comment NONE => []
- | Token.Comment (SOME Comment.Marker) => []
- | Token.Command => output false "\\isacommand{" "}"
- | Token.Keyword =>
- if Symbol.is_ascii_identifier (Token.content_of tok)
- then output false "\\isakeyword{" "}"
- else output false "" ""
- | Token.String => output false "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}"
- | Token.Alt_String => output false "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}"
- | Token.Verbatim => output true "{\\isacharverbatimopen}" "{\\isacharverbatimclose}"
- | Token.Cartouche => output false "{\\isacartoucheopen}" "{\\isacartoucheclose}"
- | _ => output false "" "")
- end handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok));
-
-fun output_source ctxt s =
- output_body ctxt false "" "" (Symbol_Pos.explode (s, Position.none));
-
-fun check_comments ctxt =
- Comment.read_body #> List.app (fn (comment, syms) =>
- let
- val pos = #1 (Symbol_Pos.range syms);
- val _ =
- comment |> Option.app (fn kind =>
- Context_Position.reports ctxt (map (pair pos) (Comment.kind_markups kind)));
- val _ = output_comment_symbols ctxt {antiq = false} (comment, syms);
- in if comment = SOME Comment.Comment then check_comments ctxt syms else () end);
-
-end;
-
-
-
-(** present theory source **)
-
-(* presentation tokens *)
-
-datatype token =
- Ignore
- | Token of Token.T
- | Heading of string * Input.source
- | Body of string * Input.source
- | Raw of Input.source;
-
-fun token_with pred (Token tok) = pred tok
- | token_with _ _ = false;
-
-val white_token = token_with Document_Source.is_white;
-val white_comment_token = token_with Document_Source.is_white_comment;
-val blank_token = token_with Token.is_blank;
-val newline_token = token_with Token.is_newline;
-
-fun present_token ctxt tok =
- (case tok of
- Ignore => []
- | Token tok => output_token ctxt tok
- | Heading (cmd, source) =>
- Latex.enclose_body ("%\n\\isamarkup" ^ cmd ^ "{") "%\n}\n"
- (output_document ctxt {markdown = false} source)
- | Body (cmd, source) =>
- [Latex.environment_block ("isamarkup" ^ cmd) (output_document ctxt {markdown = true} source)]
- | Raw source =>
- Latex.string "%\n" :: output_document ctxt {markdown = true} source @ [Latex.string "\n"]);
-
-
-(* command spans *)
-
-type command = string * Position.T; (*name, position*)
-type source = (token * (string * int)) list; (*token, markup flag, meta-comment depth*)
-
-datatype span = Span of command * (source * source * source * source) * bool;
-
-fun make_span cmd src =
- let
- fun chop_newline (tok :: toks) =
- if newline_token (fst tok) then ([tok], toks, true)
- else ([], tok :: toks, false)
- | chop_newline [] = ([], [], false);
- val (((src_prefix, src_main), src_suffix1), (src_suffix2, src_appendix, newline)) =
- src
- |> chop_prefix (white_token o fst)
- ||>> chop_suffix (white_token o fst)
- ||>> chop_prefix (white_comment_token o fst)
- ||> chop_newline;
- in Span (cmd, (src_prefix, src_main, src_suffix1 @ src_suffix2, src_appendix), newline) end;
-
-
-(* present spans *)
-
-local
-
-fun err_bad_nesting here =
- error ("Bad nesting of commands in presentation" ^ here);
-
-fun edge which f (x: string option, y) =
- if x = y then I
- else (case which (x, y) of NONE => I | SOME txt => cons (Latex.string (f txt)));
-
-val begin_tag = edge #2 Latex.begin_tag;
-val end_tag = edge #1 Latex.end_tag;
-fun open_delim delim e = edge #2 Latex.begin_delim e #> delim #> edge #2 Latex.end_delim e;
-fun close_delim delim e = edge #1 Latex.begin_delim e #> delim #> edge #1 Latex.end_delim e;
-
-fun document_tag cmd_pos state state' tagging_stack =
- let
- val ctxt' = Toplevel.presentation_context state';
- val nesting = Toplevel.level state' - Toplevel.level state;
-
- val (tagging, taggings) = tagging_stack;
- val (tag', tagging') = Document_Source.update_tagging ctxt' tagging;
-
- val tagging_stack' =
- if nesting = 0 andalso not (Toplevel.is_proof state) then tagging_stack
- else if nesting >= 0 then (tagging', replicate nesting tagging @ taggings)
- else
- (case drop (~ nesting - 1) taggings of
- tg :: tgs => (tg, tgs)
- | [] => err_bad_nesting (Position.here cmd_pos));
- in (tag', tagging_stack') end;
-
-fun read_tag s =
- (case space_explode "%" s of
- ["", b] => (SOME b, NONE)
- | [a, b] => (NONE, SOME (a, b))
- | _ => error ("Bad document_tags specification: " ^ quote s));
-
-in
-
-fun make_command_tag options keywords =
- let
- val document_tags =
- map read_tag (space_explode "," (Options.string options \<^system_option>\<open>document_tags\<close>));
- val document_tags_default = map_filter #1 document_tags;
- val document_tags_command = map_filter #2 document_tags;
- in
- fn cmd_name => fn state => fn state' => fn active_tag =>
- let
- val keyword_tags =
- if cmd_name = "end" andalso Toplevel.is_end_theory state' then ["theory"]
- else Keyword.command_tags keywords cmd_name;
- val command_tags =
- the_list (AList.lookup (op =) document_tags_command cmd_name) @
- keyword_tags @ document_tags_default;
- val active_tag' =
- (case command_tags of
- default_tag :: _ => SOME default_tag
- | [] =>
- if Keyword.is_vacuous keywords cmd_name andalso Toplevel.is_proof state
- then active_tag
- else NONE);
- in active_tag' end
- end;
-
-fun present_span command_tag span state state'
- (tagging_stack, active_tag, newline, latex, present_cont) =
- let
- val ctxt' = Toplevel.presentation_context state';
- val present = fold (fn (tok, (flag, 0)) =>
- fold cons (present_token ctxt' tok)
- #> cons (Latex.string flag)
- | _ => I);
-
- val Span ((cmd_name, cmd_pos), srcs, span_newline) = span;
-
- val (tag', tagging_stack') = document_tag cmd_pos state state' tagging_stack;
- val active_tag' =
- if is_some tag' then Option.map #1 tag'
- else command_tag cmd_name state state' active_tag;
- val edge = (active_tag, active_tag');
-
- val newline' =
- if is_none active_tag' then span_newline else newline;
-
- val latex' =
- latex
- |> end_tag edge
- |> close_delim (fst present_cont) edge
- |> snd present_cont
- |> open_delim (present (#1 srcs)) edge
- |> begin_tag edge
- |> present (#2 srcs);
- val present_cont' =
- if newline then (present (#3 srcs), present (#4 srcs))
- else (I, present (#3 srcs) #> present (#4 srcs));
- in (tagging_stack', active_tag', newline', latex', present_cont') end;
-
-fun present_trailer ((_, tags), active_tag, _, latex, present_cont) =
- if not (null tags) then err_bad_nesting " at end of theory"
- else
- latex
- |> end_tag (active_tag, NONE)
- |> close_delim (fst present_cont) (active_tag, NONE)
- |> snd present_cont;
-
-end;
-
-
-(* present_thy *)
-
-local
-
-val markup_true = "\\isamarkuptrue%\n";
-val markup_false = "\\isamarkupfalse%\n";
-
-val opt_newline = Scan.option (Scan.one Token.is_newline);
-
-val ignore =
- Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
- >> pair (d + 1)) ||
- Scan.depend (fn d => Scan.one Token.is_end_ignore --|
- (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)
- >> pair (d - 1));
-
-val locale =
- Scan.option ((Parse.$$$ "(" -- Document_Source.improper -- Parse.$$$ "in") |--
- Parse.!!!
- (Document_Source.improper |-- Parse.name --| (Document_Source.improper -- Parse.$$$ ")")));
-
-in
-
-type segment =
- {span: Command_Span.span, command: Toplevel.transition,
- prev_state: Toplevel.state, state: Toplevel.state};
-
-fun present_thy options thy (segments: segment list) =
- let
- val keywords = Thy_Header.get_keywords thy;
-
-
- (* tokens *)
-
- val ignored = Scan.state --| ignore
- >> (fn d => (NONE, (Ignore, ("", d))));
-
- fun markup pred mk flag = Scan.peek (fn d =>
- Document_Source.improper |--
- Parse.position (Scan.one (fn tok =>
- Token.is_command tok andalso pred keywords (Token.content_of tok))) --
- (Document_Source.annotation |--
- Parse.!!!! ((Document_Source.improper -- locale -- Document_Source.improper) |--
- Parse.document_source --| Document_Source.improper_end))
- >> (fn ((tok, pos'), source) =>
- let val name = Token.content_of tok
- in (SOME (name, pos'), (mk (name, source), (flag, d))) end));
-
- val command = Scan.peek (fn d =>
- Scan.optional (Scan.one Token.is_command_modifier ::: Document_Source.improper) [] --
- Scan.one Token.is_command --| Document_Source.annotation
- >> (fn (cmd_mod, cmd) =>
- map (fn tok => (NONE, (Token tok, ("", d)))) cmd_mod @
- [(SOME (Token.content_of cmd, Token.pos_of cmd),
- (Token cmd, (markup_false, d)))]));
-
- val cmt = Scan.peek (fn d =>
- Scan.one Document_Source.is_black_comment >> (fn tok => (NONE, (Token tok, ("", d)))));
-
- val other = Scan.peek (fn d =>
- Parse.not_eof >> (fn tok => (NONE, (Token tok, ("", d)))));
-
- val tokens =
- (ignored ||
- markup Keyword.is_document_heading Heading markup_true ||
- markup Keyword.is_document_body Body markup_true ||
- markup Keyword.is_document_raw (Raw o #2) "") >> single ||
- command ||
- (cmt || other) >> single;
-
-
- (* spans *)
-
- val is_eof = fn (_, (Token x, _)) => Token.is_eof x | _ => false;
- val stopper = Scan.stopper (K (NONE, (Token Token.eof, ("", 0)))) is_eof;
-
- val cmd = Scan.one (is_some o fst);
- val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
-
- val white_comments = Scan.many (white_comment_token o fst o snd);
- val blank = Scan.one (blank_token o fst o snd);
- val newline = Scan.one (newline_token o fst o snd);
- val before_cmd =
- Scan.option (newline -- white_comments) --
- Scan.option (newline -- white_comments) --
- Scan.option (blank -- white_comments) -- cmd;
-
- val span =
- Scan.repeat non_cmd -- cmd --
- Scan.repeat (Scan.unless before_cmd non_cmd) --
- Scan.option (newline >> (single o snd))
- >> (fn (((toks1, (cmd, tok2)), toks3), tok4) =>
- make_span (the cmd) (toks1 @ (tok2 :: (toks3 @ the_default [] tok4))));
-
- val spans = segments
- |> maps (Command_Span.content o #span)
- |> drop_suffix Token.is_space
- |> Source.of_list
- |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk tokens >> flat))
- |> Source.source stopper (Scan.error (Scan.bulk span))
- |> Source.exhaust;
-
- val command_results =
- segments |> map_filter (fn {command, state, ...} =>
- if Toplevel.is_ignored command then NONE else SOME (command, state));
-
-
- (* present commands *)
-
- val command_tag = make_command_tag options keywords;
-
- fun present_command tr span st st' =
- Toplevel.setmp_thread_position tr (present_span command_tag span st st');
-
- fun present _ [] = I
- | present st ((span, (tr, st')) :: rest) = present_command tr span st st' #> present st' rest;
- in
- if length command_results = length spans then
- (([], []), NONE, true, [], (I, I))
- |> present (Toplevel.init_toplevel ()) (spans ~~ command_results)
- |> present_trailer
- |> rev
- else error "Messed-up outer syntax for presentation"
- end;
-
-end;
-
-
-
-(** standard output operations **)
-
-(* pretty printing *)
-
-fun pretty_term ctxt t =
- Syntax.pretty_term (Proof_Context.augment t ctxt) t;
-
-fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
-
-
-(* default output *)
-
-val lines = separate (Latex.string "\\isanewline%\n");
-val items = separate (Latex.string "\\isasep\\isanewline%\n");
-
-fun isabelle ctxt body =
- if Config.get ctxt Document_Antiquotation.thy_output_display
- then Latex.environment_block "isabelle" body
- else Latex.block (Latex.enclose_body "\\isa{" "}" body);
-
-fun isabelle_typewriter ctxt body =
- if Config.get ctxt Document_Antiquotation.thy_output_display
- then Latex.environment_block "isabellett" body
- else Latex.block (Latex.enclose_body "\\isatt{" "}" body);
-
-fun typewriter ctxt s =
- isabelle_typewriter ctxt [Latex.string (Latex.output_ascii s)];
-
-fun verbatim ctxt =
- if Config.get ctxt Document_Antiquotation.thy_output_display
- then Document_Antiquotation.indent_lines ctxt #> typewriter ctxt
- else split_lines #> map (typewriter ctxt) #> lines #> Latex.block;
-
-fun token_source ctxt {embedded} tok =
- if Token.is_kind Token.Cartouche tok andalso embedded andalso
- not (Config.get ctxt Document_Antiquotation.thy_output_source_cartouche)
- then Token.content_of tok
- else Token.unparse tok;
-
-fun is_source ctxt =
- Config.get ctxt Document_Antiquotation.thy_output_source orelse
- Config.get ctxt Document_Antiquotation.thy_output_source_cartouche;
-
-fun source ctxt embedded =
- Token.args_of_src
- #> map (token_source ctxt embedded #> Document_Antiquotation.prepare_lines ctxt)
- #> space_implode " "
- #> output_source ctxt
- #> isabelle ctxt;
-
-fun pretty ctxt =
- Document_Antiquotation.output ctxt #> Latex.string #> single #> isabelle ctxt;
-
-fun pretty_source ctxt embedded src prt =
- if is_source ctxt then source ctxt embedded src else pretty ctxt prt;
-
-fun pretty_items ctxt =
- map (Document_Antiquotation.output ctxt #> Latex.string) #> items #> isabelle ctxt;
-
-fun pretty_items_source ctxt embedded src prts =
- if is_source ctxt then source ctxt embedded src else pretty_items ctxt prts;
-
-
-(* antiquotation variants *)
-
-local
-
-fun gen_setup name embedded =
- if embedded
- then Document_Antiquotation.setup_embedded name
- else Document_Antiquotation.setup name;
-
-fun gen_antiquotation_pretty name embedded scan f =
- gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => pretty ctxt (f ctxt x));
-
-fun gen_antiquotation_pretty_source name embedded scan f =
- gen_setup name embedded scan
- (fn {context = ctxt, source = src, argument = x} =>
- pretty_source ctxt {embedded = embedded} src (f ctxt x));
-
-fun gen_antiquotation_raw name embedded scan f =
- gen_setup name embedded scan (fn {context = ctxt, argument = x, ...} => f ctxt x);
-
-fun gen_antiquotation_verbatim name embedded scan f =
- gen_antiquotation_raw name embedded scan (fn ctxt => verbatim ctxt o f ctxt);
-
-in
-
-fun antiquotation_pretty name = gen_antiquotation_pretty name false;
-fun antiquotation_pretty_embedded name = gen_antiquotation_pretty name true;
-
-fun antiquotation_pretty_source name = gen_antiquotation_pretty_source name false;
-fun antiquotation_pretty_source_embedded name = gen_antiquotation_pretty_source name true;
-
-fun antiquotation_raw name = gen_antiquotation_raw name false;
-fun antiquotation_raw_embedded name = gen_antiquotation_raw name true;
-
-fun antiquotation_verbatim name = gen_antiquotation_verbatim name false;
-fun antiquotation_verbatim_embedded name = gen_antiquotation_verbatim name true;
-
-end;
-
-end;
--- a/src/Pure/Tools/build.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/build.scala Sun May 23 20:34:43 2021 +0200
@@ -576,7 +576,7 @@
Build and manage Isabelle sessions, depending on implicit settings:
-""" + Library.prefix_lines(" ", Build_Log.Settings.show()) + "\n",
+""" + Library.indent_lines(2, Build_Log.Settings.show()) + "\n",
"B:" -> (arg => base_sessions = base_sessions ::: List(arg)),
"D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))),
"N" -> (_ => numa_shuffling = true),
--- a/src/Pure/Tools/build_job.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/build_job.scala Sun May 23 20:34:43 2021 +0200
@@ -447,10 +447,10 @@
using(store.open_database_context())(db_context =>
{
val documents =
- Presentation.build_documents(session_name, deps, db_context,
+ Document_Build.build_documents(
+ Document_Build.context(session_name, deps, db_context, progress = progress),
output_sources = info.document_output,
output_pdf = info.document_output,
- progress = progress,
verbose = verbose)
db_context.output_database(session_name)(db =>
documents.foreach(_.write(db, session_name)))
@@ -460,7 +460,7 @@
else (Nil, Nil)
}
catch {
- case exn: Presentation.Build_Error => (exn.log_lines, List(exn.message))
+ case exn: Document_Build.Build_Error => (exn.log_lines, List(exn.message))
case Exn.Interrupt.ERROR(msg) => (Nil, List(msg))
}
--- a/src/Pure/Tools/doc.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/doc.ML Sun May 23 20:34:43 2021 +0200
@@ -18,7 +18,7 @@
val _ =
Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>doc\<close>
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>doc\<close>
(Scan.lift Parse.embedded_position) check);
end;
--- a/src/Pure/Tools/dump.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/dump.scala Sun May 23 20:34:43 2021 +0200
@@ -456,7 +456,7 @@
Dump cumulative PIDE session database, with the following aspects:
-""" + Library.prefix_lines(" ", show_aspects) + "\n",
+""" + Library.indent_lines(4, show_aspects) + "\n",
"A:" -> (arg => aspects = Library.distinct(space_explode(',', arg)).map(the_aspect)),
"B:" -> (arg => base_sessions = base_sessions ::: List(arg)),
"D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))),
--- a/src/Pure/Tools/jedit.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/jedit.ML Sun May 23 20:34:43 2021 +0200
@@ -78,7 +78,7 @@
val _ =
Theory.setup
- (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>action\<close> (Scan.lift Args.embedded_position)
+ (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>action\<close> (Scan.lift Args.embedded_position)
(fn ctxt => fn (name, pos) =>
let
val _ =
--- a/src/Pure/Tools/logo.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/logo.scala Sun May 23 20:34:43 2021 +0200
@@ -11,12 +11,19 @@
{
/* create logo */
+ def make_output_file(logo_name: String): Path =
+ {
+ val name = if (logo_name.isEmpty) "isabelle" else "isabelle_" + Word.lowercase(logo_name)
+ Path.explode(name).pdf
+ }
+
def create_logo(logo_name: String, output_file: Path, quiet: Boolean = false): Unit =
{
Isabelle_System.with_tmp_file("logo", ext = "eps")(tmp_file =>
{
val template = File.read(Path.explode("$ISABELLE_HOME/lib/logo/isabelle_any.eps"))
File.write(tmp_file, template.replace("<any>", logo_name))
+
Isabelle_System.bash(
"\"$ISABELLE_EPSTOPDF\" --filter < " + File.bash_path(tmp_file) +
" > " + File.bash_path(output_file)).check
@@ -46,13 +53,12 @@
"q" -> (_ => quiet = true))
val more_args = getopts(args)
- val (logo_name, output_file) =
+ val logo_name =
more_args match {
- case Nil => ("", Path.explode("isabelle").pdf)
- case List(a) => (a, output.getOrElse(Path.explode("isabelle_" + Word.lowercase(a)).pdf))
+ case Nil => ""
+ case List(name) => name
case _ => getopts.usage()
}
-
- create_logo(logo_name, output_file, quiet = quiet)
+ create_logo(logo_name, output getOrElse make_output_file(logo_name), quiet = quiet)
})
}
--- a/src/Pure/Tools/phabricator.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/phabricator.scala Sun May 23 20:34:43 2021 +0200
@@ -91,7 +91,7 @@
NAME="$(echo "$REPLY" | cut -d: -f1)"
ROOT="$(echo "$REPLY" | cut -d: -f2)"
{
-""" + Library.prefix_lines(" ", body) + """
+""" + Library.indent_lines(6, body) + """
} < /dev/null
done
} < """ + File.bash_path(global_config) + "\n" +
--- a/src/Pure/Tools/rail.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/Tools/rail.ML Sun May 23 20:34:43 2021 +0200
@@ -384,7 +384,7 @@
in Latex.string (Latex.environment "railoutput" (implode (map output_rule rules))) end;
val _ = Theory.setup
- (Thy_Output.antiquotation_raw_embedded \<^binding>\<open>rail\<close> (Scan.lift Args.text_input)
+ (Document_Output.antiquotation_raw_embedded \<^binding>\<open>rail\<close> (Scan.lift Args.text_input)
(fn ctxt => output_rules ctxt o read ctxt));
end;
--- a/src/Pure/build-jars Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/build-jars Sun May 23 20:34:43 2021 +0200
@@ -154,6 +154,7 @@
src/Pure/System/system_channel.scala
src/Pure/System/tty_loop.scala
src/Pure/Thy/bibtex.scala
+ src/Pure/Thy/document_build.scala
src/Pure/Thy/export.scala
src/Pure/Thy/export_theory.scala
src/Pure/Thy/file_format.scala
--- a/src/Pure/library.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/library.ML Sun May 23 20:34:43 2021 +0200
@@ -132,6 +132,7 @@
val fold_string: (string -> 'a -> 'a) -> string -> 'a -> 'a
val exists_string: (string -> bool) -> string -> bool
val forall_string: (string -> bool) -> string -> bool
+ val member_string: string -> string -> bool
val first_field: string -> string -> (string * string) option
val enclose: string -> string -> string -> string
val unenclose: string -> string
@@ -703,6 +704,8 @@
fun forall_string pred = not o exists_string (not o pred);
+fun member_string str s = exists_string (fn s' => s = s') str;
+
fun first_field sep str =
let
val n = size sep;
--- a/src/Pure/library.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/library.scala Sun May 23 20:34:43 2021 +0200
@@ -107,6 +107,9 @@
def prefix_lines(prfx: String, str: String): String =
if (str == "") str else cat_lines(split_lines(str).map(prfx + _))
+ def indent_lines(n: Int, str: String): String =
+ prefix_lines(Symbol.spaces(n), str)
+
def first_line(source: CharSequence): String =
{
val lines = separated_chunks(_ == '\n', source)
--- a/src/Pure/pure_syn.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Pure/pure_syn.ML Sun May 23 20:34:43 2021 +0200
@@ -20,12 +20,8 @@
fun output_document state markdown txt =
let
val ctxt = Toplevel.presentation_context state;
- val pos = Input.pos_of txt;
- val _ =
- Context_Position.reports ctxt
- [(pos, Markup.language_document (Input.is_delimited txt)),
- (pos, Markup.plain_text)];
- in Thy_Output.output_document ctxt markdown txt end;
+ val _ = Context_Position.reports ctxt (Document_Output.document_reports txt);
+ in Document_Output.output_document ctxt markdown txt end;
fun document_command markdown (loc, txt) =
Toplevel.keep (fn state =>
--- a/src/Tools/Code/code_target.ML Wed May 19 14:17:40 2021 +0100
+++ b/src/Tools/Code/code_target.ML Sun May 23 20:34:43 2021 +0200
@@ -782,7 +782,7 @@
in
val _ = Theory.setup
- (Thy_Output.antiquotation_raw \<^binding>\<open>code_stmts\<close>
+ (Document_Output.antiquotation_raw \<^binding>\<open>code_stmts\<close>
(parse_const_terms --
Scan.repeats (parse_consts || parse_types || parse_classes || parse_instances)
-- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))
@@ -790,7 +790,7 @@
present_code ctxt consts symbols
target_name "Example" some_width []
|> trim_line
- |> Thy_Output.verbatim (Config.put Document_Antiquotation.thy_output_display true ctxt)));
+ |> Document_Output.verbatim (Config.put Document_Antiquotation.thy_output_display true ctxt)));
end;
--- a/src/Tools/jEdit/src/keymap_merge.scala Wed May 19 14:17:40 2021 +0100
+++ b/src/Tools/jEdit/src/keymap_merge.scala Sun May 23 20:34:43 2021 +0200
@@ -30,7 +30,7 @@
class Shortcut(val property: String, val binding: String)
{
- override def toString: String = property + "=" + binding
+ override def toString: String = Properties.Eq(property, binding)
def primary: Boolean = property.endsWith(".shortcut")