--- a/Admin/components/components.sha1 Thu Nov 05 08:27:14 2015 +0100
+++ b/Admin/components/components.sha1 Thu Nov 05 08:27:22 2015 +0100
@@ -31,6 +31,7 @@
eccff31931fb128c1dd522cfc85495c9b66e67af Haskabelle-2015.tar.gz
683acd94761ef460cca1a628f650355370de5afb hol-light-bundle-0.5-126.tar.gz
20b53cfc3ffc5b15c1eabc91846915b49b4c0367 isabelle_fonts-20151021.tar.gz
+736844204b2ef83974cd9f0a215738b767958c41 isabelle_fonts-20151104.tar.gz
8d83e433c1419e0c0cc5fd1762903d11b4a5752c jdk-6u31.tar.gz
38d2d2a91c66714c18430e136e7e5191af3996e6 jdk-7u11.tar.gz
d765bc4ad2f34d494429b2a8c1563c49db224944 jdk-7u13.tar.gz
--- a/Admin/components/main Thu Nov 05 08:27:14 2015 +0100
+++ b/Admin/components/main Thu Nov 05 08:27:22 2015 +0100
@@ -4,7 +4,7 @@
e-1.8
exec_process-1.0.3
Haskabelle-2015
-isabelle_fonts-20151021
+isabelle_fonts-20151104
jdk-8u66
jedit_build-20151023
jfreechart-1.0.14-1
--- a/NEWS Thu Nov 05 08:27:14 2015 +0100
+++ b/NEWS Thu Nov 05 08:27:22 2015 +0100
@@ -22,6 +22,10 @@
* Toplevel theorem statement 'proposition' is another alias for
'theorem'.
+* Syntax for formal comments "-- text" now also supports the symbolic
+form "\<comment> text". Command-line tool "isabelle update_cartouches -c" helps
+to update old sources.
+
*** Prover IDE -- Isabelle/Scala/jEdit ***
@@ -124,6 +128,9 @@
recursively, adding appropriate text style markup. These are typically
used in the short form \<^emph>\<open>...\<close> and \<^bold>\<open>...\<close>.
+* Document antiquotation @{footnote} outputs LaTeX source recursively,
+marked as \footnote{}. This is typically used in the short form \<^footnote>\<open>...\<close>.
+
*** Isar ***
--- a/etc/symbols Thu Nov 05 08:27:14 2015 +0100
+++ b/etc/symbols Thu Nov 05 08:27:22 2015 +0100
@@ -326,7 +326,6 @@
\<ordfeminine> code: 0x0000aa
\<ordmasculine> code: 0x0000ba
\<section> code: 0x0000a7
-\<paragraph> code: 0x0000b6
\<exclamdown> code: 0x0000a1
\<questiondown> code: 0x0000bf
\<euro> code: 0x0020ac
@@ -349,6 +348,7 @@
\<some> code: 0x0003f5
\<hole> code: 0x002311
\<newline> code: 0x0023ce
+\<comment> code: 0x002015
\<open> code: 0x002039 group: punctuation font: IsabelleText abbrev: <<
\<close> code: 0x00203a group: punctuation font: IsabelleText abbrev: >>
\<here> code: 0x002302 font: IsabelleText
@@ -359,6 +359,7 @@
\<^item> code: 0x0025aa group: control font: IsabelleText
\<^enum> code: 0x0025b8 group: control font: IsabelleText
\<^descr> code: 0x0027a7 group: control font: IsabelleText
+\<^footnote> code: 0x0000b6 group: control
\<^verbatim> code: 0x0025a9 group: control font: IsabelleText
\<^emph> code: 0x002217 group: control font: IsabelleText
\<^bold> code: 0x002759 group: control font: IsabelleText
--- a/lib/Tools/update_cartouches Thu Nov 05 08:27:14 2015 +0100
+++ b/lib/Tools/update_cartouches Thu Nov 05 08:27:22 2015 +0100
@@ -15,6 +15,7 @@
echo "Usage: isabelle $PRG [FILES|DIRS...]"
echo
echo " Options are:"
+ echo " -c replace comment marker \"--\" by symbol \"\\<comment>\""
echo " -t replace @{text} antiquotations within text tokens"
echo
echo " Recursively find .thy files and update theory syntax to use cartouches"
@@ -30,11 +31,15 @@
# options
+COMMENT="false"
TEXT="false"
-while getopts "t" OPT
+while getopts "ct" OPT
do
case "$OPT" in
+ c)
+ COMMENT="true"
+ ;;
t)
TEXT="true"
;;
@@ -57,4 +62,4 @@
## main
find $SPECS -name \*.thy -print0 | \
- xargs -0 "$ISABELLE_TOOL" java isabelle.Update_Cartouches "$TEXT"
+ xargs -0 "$ISABELLE_TOOL" java isabelle.Update_Cartouches "$COMMENT" "$TEXT"
--- a/lib/fonts/IsabelleText.sfd Thu Nov 05 08:27:14 2015 +0100
+++ b/lib/fonts/IsabelleText.sfd Thu Nov 05 08:27:22 2015 +0100
@@ -19,7 +19,7 @@
OS2_WeightWidthSlopeOnly: 0
OS2_UseTypoMetrics: 1
CreationTime: 1050361371
-ModificationTime: 1445439176
+ModificationTime: 1446669919
PfmFamily: 17
TTFWeight: 400
TTFWidth: 5
@@ -2241,11 +2241,11 @@
DisplaySize: -96
AntiAlias: 1
FitToEm: 1
-WinInfo: 9558 18 16
+WinInfo: 8154 18 16
BeginPrivate: 0
EndPrivate
TeXData: 1 0 0 631296 315648 210432 572416 -1048576 210432 783286 444596 497025 792723 393216 433062 380633 303038 157286 324010 404750 52429 2506097 1059062 262144
-BeginChars: 1114189 1393
+BeginChars: 1114189 1394
StartChar: u10000
Encoding: 65536 65536 0
@@ -16604,69 +16604,32 @@
StartChar: endash
Encoding: 8211 8211 178
-Width: 1233
-Flags: W
-TtInstrs:
-PUSHB_7
- 2
- 182
- 0
- 253
- 4
- 1
- 0
-MDAP[rnd]
-MDRP[rnd,white]
-IUP[x]
-SVTCA[y-axis]
-SRP0
-MIRP[rp0,min,rnd,grey]
-MIRP[min,rnd,grey]
-IUP[y]
-EndTTInstrs
-LayerCount: 2
-Fore
-SplineSet
-0 633 m 1,0,-1
- 1233 633 l 1,1,-1
- 1233 492 l 1,2,-1
- 0 492 l 1,3,-1
- 0 633 l 1,0,-1
-EndSplineSet
-Validated: 1
+Width: 1024
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+100 633 m 1,0,-1
+ 924 633 l 1,1,-1
+ 924 489 l 1,2,-1
+ 100 489 l 1,3,-1
+ 100 633 l 1,0,-1
+EndSplineSet
EndChar
StartChar: emdash
Encoding: 8212 8212 179
-Width: 1233
-Flags: W
-TtInstrs:
-PUSHB_6
- 2
- 182
- 0
- 4
- 1
- 0
-MDAP[rnd]
-MDRP[rnd,grey]
-IUP[x]
-SVTCA[y-axis]
-SRP0
-MDRP[rp0,rnd,grey]
-MIRP[min,rnd,grey]
-IUP[y]
-EndTTInstrs
-LayerCount: 2
-Fore
-SplineSet
-0 633 m 1,0,-1
- 1233 633 l 1,1,-1
- 1233 492 l 1,2,-1
- 0 492 l 1,3,-1
- 0 633 l 1,0,-1
-EndSplineSet
-Validated: 1
+Width: 2048
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+100 633 m 1,0,-1
+ 1948 633 l 1,1,-1
+ 1948 489 l 1,2,-1
+ 100 489 l 1,3,-1
+ 100 633 l 1,0,-1
+EndSplineSet
EndChar
StartChar: quotedblleft
@@ -61936,5 +61899,20 @@
6 -78 l 1,112,-1
EndSplineSet
EndChar
+
+StartChar: afii00208
+Encoding: 8213 8213 1393
+Width: 2048
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+0 633 m 1,0,-1
+ 2048 633 l 1,1,-1
+ 2048 489 l 1,2,-1
+ 0 489 l 1,3,-1
+ 0 633 l 1,0,-1
+EndSplineSet
+EndChar
EndChars
EndSplineFont
--- a/lib/fonts/IsabelleTextBold.sfd Thu Nov 05 08:27:14 2015 +0100
+++ b/lib/fonts/IsabelleTextBold.sfd Thu Nov 05 08:27:22 2015 +0100
@@ -20,7 +20,7 @@
OS2_WeightWidthSlopeOnly: 0
OS2_UseTypoMetrics: 1
CreationTime: 1050374980
-ModificationTime: 1445439141
+ModificationTime: 1446669811
PfmFamily: 17
TTFWeight: 700
TTFWidth: 5
@@ -1678,10 +1678,10 @@
DisplaySize: -96
AntiAlias: 1
FitToEm: 1
-WinInfo: 9534 21 15
+WinInfo: 8064 21 15
BeginPrivate: 0
EndPrivate
-BeginChars: 1114115 1385
+BeginChars: 1114115 1386
StartChar: .notdef
Encoding: 1114112 -1 0
@@ -42258,63 +42258,31 @@
StartChar: endash
Encoding: 8211 8211 999
-Width: 1233
-Flags: W
-TtInstrs:
-PUSHB_5
- 2
- 0
- 4
- 1
- 0
-MDAP[rnd]
-MDRP[min,rnd,white]
-IUP[x]
-SVTCA[y-axis]
-SRP0
-MDRP[rp0,rnd,grey]
-MDRP[min,rnd,grey]
-IUP[y]
-EndTTInstrs
-LayerCount: 2
-Fore
-SplineSet
-0 690 m 1,0,-1
- 1233 690 l 1,1,-1
- 1233 444 l 1,2,-1
- 0 444 l 1,3,-1
- 0 690 l 1,0,-1
+Width: 1024
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+110 690 m 1,0,-1
+ 914 690 l 1,1,-1
+ 914 432 l 1,2,-1
+ 110 432 l 1,3,-1
+ 110 690 l 1,0,-1
EndSplineSet
EndChar
StartChar: emdash
Encoding: 8212 8212 1000
-Width: 1233
-Flags: W
-TtInstrs:
-PUSHB_5
- 2
- 0
- 4
- 1
- 0
-MDAP[rnd]
-MDRP[min,rnd,grey]
-IUP[x]
-SVTCA[y-axis]
-SRP0
-MDRP[rp0,rnd,grey]
-MDRP[min,rnd,grey]
-IUP[y]
-EndTTInstrs
-LayerCount: 2
-Fore
-SplineSet
-0 690 m 1,0,-1
- 1233 690 l 1,1,-1
- 1233 444 l 1,2,-1
- 0 444 l 1,3,-1
- 0 690 l 1,0,-1
+Width: 2048
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+110 690 m 1,0,-1
+ 1938 690 l 1,1,-1
+ 1938 432 l 1,2,-1
+ 110 432 l 1,3,-1
+ 110 690 l 1,0,-1
EndSplineSet
EndChar
@@ -68173,5 +68141,20 @@
6 -78 l 1,112,-1
EndSplineSet
EndChar
+
+StartChar: afii00208
+Encoding: 8213 8213 1385
+Width: 2048
+Flags: W
+LayerCount: 2
+Fore
+SplineSet
+0 690 m 1,0,-1
+ 2048 690 l 1,1,-1
+ 2048 432 l 1,2,-1
+ 0 432 l 1,3,-1
+ 0 690 l 1,0,-1
+EndSplineSet
+EndChar
EndChars
EndSplineFont
--- a/lib/texinputs/isabellesym.sty Thu Nov 05 08:27:14 2015 +0100
+++ b/lib/texinputs/isabellesym.sty Thu Nov 05 08:27:22 2015 +0100
@@ -358,3 +358,4 @@
\newcommand{\isasymclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}
\newcommand{\isasymhole}{\isatext{\rm\wasylozenge}} %requires wasysym
\newcommand{\isasymnewline}{\isatext{\fbox{$\hookleftarrow$}}}
+\newcommand{\isasymcomment}{\isatext{---}}
--- a/src/Doc/Eisbach/Manual.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Eisbach/Manual.thy Thu Nov 05 08:27:22 2015 +0100
@@ -14,8 +14,8 @@
terms, facts, or other methods.
\<^medskip>
- The syntax diagram below refers to some syntactic categories that
- are further defined in @{cite "isabelle-isar-ref"}.
+ The syntax diagram below refers to some syntactic categories that are
+ further defined in @{cite "isabelle-isar-ref"}.
@{rail \<open>
@@{command method} name args @'=' method
@@ -56,11 +56,10 @@
by prop_solver\<^sub>1
text \<open>
- In this example, the facts \<open>impI\<close> and \<open>conjE\<close> are static. They
- are evaluated once when the method is defined and cannot be changed later.
- This makes the method stable in the sense of \<^emph>\<open>static scoping\<close>: naming
- another fact \<open>impI\<close> in a later context won't affect the behaviour of
- \<open>prop_solver\<^sub>1\<close>.
+ In this example, the facts \<open>impI\<close> and \<open>conjE\<close> are static. They are evaluated
+ once when the method is defined and cannot be changed later. This makes the
+ method stable in the sense of \<^emph>\<open>static scoping\<close>: naming another fact \<open>impI\<close>
+ in a later context won't affect the behaviour of \<open>prop_solver\<^sub>1\<close>.
\<close>
@@ -100,11 +99,11 @@
subsection \<open>Named theorems\<close>
text \<open>
- A \<open>named theorem\<close> is a fact whose contents are produced dynamically
- within the current proof context. The Isar command @{command_ref
- "named_theorems"} provides simple access to this concept: it declares a
- dynamic fact with corresponding \<^emph>\<open>attribute\<close> for managing
- this particular data slot in the context.
+ A \<^emph>\<open>named theorem\<close> is a fact whose contents are produced dynamically within
+ the current proof context. The Isar command @{command_ref "named_theorems"}
+ provides simple access to this concept: it declares a dynamic fact with
+ corresponding \<^emph>\<open>attribute\<close> for managing this particular data slot in the
+ context.
\<close>
named_theorems intros
@@ -112,7 +111,8 @@
text \<open>
So far \<open>intros\<close> refers to the empty fact. Using the Isar command
@{command_ref "declare"} we may apply declaration attributes to the context.
- Below we declare both \<open>conjI\<close> and \<open>impI\<close> as \<open>intros\<close>, adding them to the named theorem slot.
+ Below we declare both \<open>conjI\<close> and \<open>impI\<close> as \<open>intros\<close>, adding them to the
+ named theorem slot.
\<close>
declare conjI [intros] and impI [intros]
@@ -136,8 +136,8 @@
text \<open>
Often these named theorems need to be augmented on the spot, when a method
is invoked. The @{keyword_def "declares"} keyword in the signature of
- @{command method} adds the common method syntax \<open>method decl: facts\<close>
- for each named theorem \<open>decl\<close>.
+ @{command method} adds the common method syntax \<open>method decl: facts\<close> for
+ each named theorem \<open>decl\<close>.
\<close>
method prop_solver\<^sub>4 declares intros elims =
@@ -170,11 +170,12 @@
section \<open>Higher-order methods\<close>
text \<open>
- The \<^emph>\<open>structured concatenation\<close> combinator ``\<open>method\<^sub>1 ;
- method\<^sub>2\<close>'' was introduced in Isabelle2015, motivated by development of
- Eisbach. It is similar to ``\<open>method\<^sub>1, method\<^sub>2\<close>'', but \<open>method\<^sub>2\<close> is invoked on on \<^emph>\<open>all\<close> subgoals that have newly emerged from
- \<open>method\<^sub>1\<close>. This is useful to handle cases where the number of
- subgoals produced by a method is determined dynamically at run-time.
+ The \<^emph>\<open>structured concatenation\<close> combinator ``\<open>method\<^sub>1 ; method\<^sub>2\<close>'' was
+ introduced in Isabelle2015, motivated by development of Eisbach. It is
+ similar to ``\<open>method\<^sub>1, method\<^sub>2\<close>'', but \<open>method\<^sub>2\<close> is invoked on on \<^emph>\<open>all\<close>
+ subgoals that have newly emerged from \<open>method\<^sub>1\<close>. This is useful to handle
+ cases where the number of subgoals produced by a method is determined
+ dynamically at run-time.
\<close>
method conj_with uses rule =
@@ -190,18 +191,17 @@
method combinators with prefix syntax. For example, to more usefully exploit
Isabelle's backtracking, the explicit requirement that a method solve all
produced subgoals is frequently useful. This can easily be written as a
- \<^emph>\<open>higher-order method\<close> using ``\<open>;\<close>''. The @{keyword "methods"}
- keyword denotes method parameters that are other proof methods to be invoked
- by the method being defined.
+ \<^emph>\<open>higher-order method\<close> using ``\<open>;\<close>''. The @{keyword "methods"} keyword
+ denotes method parameters that are other proof methods to be invoked by the
+ method being defined.
\<close>
method solve methods m = (m ; fail)
text \<open>
- Given some method-argument \<open>m\<close>, \<open>solve \<open>m\<close>\<close> applies the
- method \<open>m\<close> and then fails whenever \<open>m\<close> produces any new unsolved
- subgoals --- i.e. when \<open>m\<close> fails to completely discharge the goal it
- was applied to.
+ Given some method-argument \<open>m\<close>, \<open>solve \<open>m\<close>\<close> applies the method \<open>m\<close> and then
+ fails whenever \<open>m\<close> produces any new unsolved subgoals --- i.e. when \<open>m\<close>
+ fails to completely discharge the goal it was applied to.
\<close>
@@ -222,41 +222,43 @@
(erule notE ; solve \<open>prop_solver\<close>))+
text \<open>
- The only non-trivial part above is the final alternative \<open>(erule notE
- ; solve \<open>prop_solver\<close>)\<close>. Here, in the case that all other alternatives
- fail, the method takes one of the assumptions @{term "\<not> P"} of the current
- goal and eliminates it with the rule \<open>notE\<close>, causing the goal to be
- proved to become @{term P}. The method then recursively invokes itself on
- the remaining goals. The job of the recursive call is to demonstrate that
- there is a contradiction in the original assumptions (i.e.\ that @{term P}
- can be derived from them). Note this recursive invocation is applied with
- the @{method solve} method combinator to ensure that a contradiction will
- indeed be shown. In the case where a contradiction cannot be found,
- backtracking will occur and a different assumption @{term "\<not> Q"} will be
- chosen for elimination.
+ The only non-trivial part above is the final alternative \<open>(erule notE ;
+ solve \<open>prop_solver\<close>)\<close>. Here, in the case that all other alternatives fail,
+ the method takes one of the assumptions @{term "\<not> P"} of the current goal
+ and eliminates it with the rule \<open>notE\<close>, causing the goal to be proved to
+ become @{term P}. The method then recursively invokes itself on the
+ remaining goals. The job of the recursive call is to demonstrate that there
+ is a contradiction in the original assumptions (i.e.\ that @{term P} can be
+ derived from them). Note this recursive invocation is applied with the
+ @{method solve} method combinator to ensure that a contradiction will indeed
+ be shown. In the case where a contradiction cannot be found, backtracking
+ will occur and a different assumption @{term "\<not> Q"} will be chosen for
+ elimination.
Note that the recursive call to @{method prop_solver} does not have any
- parameters passed to it. Recall that fact parameters, e.g.\ \<open>intros\<close>, \<open>elims\<close>, and \<open>subst\<close>, are managed by declarations
- in the current proof context. They will therefore be passed to any recursive
- call to @{method prop_solver} and, more generally, any invocation of a
- method which declares these named theorems.
+ parameters passed to it. Recall that fact parameters, e.g.\ \<open>intros\<close>,
+ \<open>elims\<close>, and \<open>subst\<close>, are managed by declarations in the current proof
+ context. They will therefore be passed to any recursive call to @{method
+ prop_solver} and, more generally, any invocation of a method which declares
+ these named theorems.
\<^medskip>
After declaring some standard rules to the context, the @{method
prop_solver} becomes capable of solving non-trivial propositional
- tautologies.\<close>
+ tautologies.
+\<close>
lemmas [intros] =
- conjI -- \<open>@{thm conjI}\<close>
- impI -- \<open>@{thm impI}\<close>
- disjCI -- \<open>@{thm disjCI}\<close>
- iffI -- \<open>@{thm iffI}\<close>
- notI -- \<open>@{thm notI}\<close>
+ conjI \<comment> \<open>@{thm conjI}\<close>
+ impI \<comment> \<open>@{thm impI}\<close>
+ disjCI \<comment> \<open>@{thm disjCI}\<close>
+ iffI \<comment> \<open>@{thm iffI}\<close>
+ notI \<comment> \<open>@{thm notI}\<close>
lemmas [elims] =
- impCE -- \<open>@{thm impCE}\<close>
- conjE -- \<open>@{thm conjE}\<close>
- disjE -- \<open>@{thm disjE}\<close>
+ impCE \<comment> \<open>@{thm impCE}\<close>
+ conjE \<comment> \<open>@{thm conjE}\<close>
+ disjE \<comment> \<open>@{thm disjE}\<close>
lemma "(A \<or> B) \<and> (A \<longrightarrow> C) \<and> (B \<longrightarrow> C) \<longrightarrow> C"
by prop_solver
@@ -271,15 +273,15 @@
result of backtracking. When designing more sophisticated proof methods this
proves too restrictive and difficult to manage conceptually.
- To address this, we introduce the @{method_def "match"} method, which
- provides more direct access to the higher-order matching facility at the
- core of Isabelle. It is implemented as a separate proof method (in
- Isabelle/ML), and thus can be directly applied to proofs, however it is most
- useful when applied in the context of writing Eisbach method definitions.
+ To address this, we introduce the @{method_def match} method, which provides
+ more direct access to the higher-order matching facility at the core of
+ Isabelle. It is implemented as a separate proof method (in Isabelle/ML), and
+ thus can be directly applied to proofs, however it is most useful when
+ applied in the context of writing Eisbach method definitions.
\<^medskip>
- The syntax diagram below refers to some syntactic categories that
- are further defined in @{cite "isabelle-isar-ref"}.
+ The syntax diagram below refers to some syntactic categories that are
+ further defined in @{cite "isabelle-isar-ref"}.
@{rail \<open>
@@{method match} kind @'in' (pattern '\<Rightarrow>' cartouche + '\<bar>')
@@ -296,10 +298,10 @@
\<close>}
Matching allows methods to introspect the goal state, and to implement more
- explicit control flow. In the basic case, a term or fact \<open>ts\<close> is given
- to match against as a \<^emph>\<open>match target\<close>, along with a collection of
- pattern-method pairs \<open>(p, m)\<close>: roughly speaking, when the pattern
- \<open>p\<close> matches any member of \<open>ts\<close>, the \<^emph>\<open>inner\<close> method \<open>m\<close> will be executed.
+ explicit control flow. In the basic case, a term or fact \<open>ts\<close> is given to
+ match against as a \<^emph>\<open>match target\<close>, along with a collection of
+ pattern-method pairs \<open>(p, m)\<close>: roughly speaking, when the pattern \<open>p\<close>
+ matches any member of \<open>ts\<close>, the \<^emph>\<open>inner\<close> method \<open>m\<close> will be executed.
\<close>
lemma
@@ -310,11 +312,11 @@
by (match X in I: "Q \<longrightarrow> P" and I': "Q" \<Rightarrow> \<open>insert mp [OF I I']\<close>)
text \<open>
- In this example we have a structured Isar proof, with the named
- assumption \<open>X\<close> and a conclusion @{term "P"}. With the match method
- we can find the local facts @{term "Q \<longrightarrow> P"} and @{term "Q"}, binding them to
- separately as \<open>I\<close> and \<open>I'\<close>. We then specialize the
- modus-ponens rule @{thm mp [of Q P]} to these facts to solve the goal.
+ In this example we have a structured Isar proof, with the named assumption
+ \<open>X\<close> and a conclusion @{term "P"}. With the match method we can find the
+ local facts @{term "Q \<longrightarrow> P"} and @{term "Q"}, binding them to separately as
+ \<open>I\<close> and \<open>I'\<close>. We then specialize the modus-ponens rule @{thm mp [of Q P]} to
+ these facts to solve the goal.
\<close>
@@ -324,15 +326,14 @@
In the previous example we were able to match against an assumption out of
the Isar proof state. In general, however, proof subgoals can be
\<^emph>\<open>unstructured\<close>, with goal parameters and premises arising from rule
- application. To address this, @{method match} uses \<^emph>\<open>subgoal focusing\<close>
- to produce structured goals out of
- unstructured ones. In place of fact or term, we may give the
- keyword @{keyword_def "premises"} as the match target. This causes a subgoal
- focus on the first subgoal, lifting local goal parameters to fixed term
- variables and premises into hypothetical theorems. The match is performed
- against these theorems, naming them and binding them as appropriate.
- Similarly giving the keyword @{keyword_def "conclusion"} matches against the
- conclusion of the first subgoal.
+ application. To address this, @{method match} uses \<^emph>\<open>subgoal focusing\<close> to
+ produce structured goals out of unstructured ones. In place of fact or term,
+ we may give the keyword @{keyword_def "premises"} as the match target. This
+ causes a subgoal focus on the first subgoal, lifting local goal parameters
+ to fixed term variables and premises into hypothetical theorems. The match
+ is performed against these theorems, naming them and binding them as
+ appropriate. Similarly giving the keyword @{keyword_def "conclusion"}
+ matches against the conclusion of the first subgoal.
An unstructured version of the previous example can then be similarly solved
through focusing.
@@ -358,16 +359,16 @@
now-bound @{term A} (bound to @{term P}) against the conclusion (also @{term
P}), finally applying the specialized rule to solve the goal.
- Schematic terms like \<open>?P\<close> may also be used to specify match
- variables, but the result of the match is not bound, and thus cannot be used
- in the inner method body.
+ Schematic terms like \<open>?P\<close> may also be used to specify match variables, but
+ the result of the match is not bound, and thus cannot be used in the inner
+ method body.
\<^medskip>
- In the following example we extract the predicate of an
- existentially quantified conclusion in the current subgoal and search the
- current premises for a matching fact. If both matches are successful, we
- then instantiate the existential introduction rule with both the witness and
- predicate, solving with the matched premise.
+ In the following example we extract the predicate of an existentially
+ quantified conclusion in the current subgoal and search the current premises
+ for a matching fact. If both matches are successful, we then instantiate the
+ existential introduction rule with both the witness and predicate, solving
+ with the matched premise.
\<close>
method solve_ex =
@@ -378,15 +379,14 @@
text \<open>
The first @{method match} matches the pattern @{term "\<exists>x. Q x"} against the
current conclusion, binding the term @{term "Q"} in the inner match. Next
- the pattern \<open>Q y\<close> is matched against all premises of the current
- subgoal. In this case @{term "Q"} is fixed and @{term "y"} may be
- instantiated. Once a match is found, the local fact \<open>U\<close> is bound to
- the matching premise and the variable @{term "y"} is bound to the matching
- witness. The existential introduction rule \<open>exI:\<close>~@{thm exI} is then
- instantiated with @{term "y"} as the witness and @{term "Q"} as the
- predicate, with its proof obligation solved by the local fact U (using the
- Isar attribute @{attribute OF}). The following example is a trivial use of
- this method.
+ the pattern \<open>Q y\<close> is matched against all premises of the current subgoal. In
+ this case @{term "Q"} is fixed and @{term "y"} may be instantiated. Once a
+ match is found, the local fact \<open>U\<close> is bound to the matching premise and the
+ variable @{term "y"} is bound to the matching witness. The existential
+ introduction rule \<open>exI:\<close>~@{thm exI} is then instantiated with @{term "y"} as
+ the witness and @{term "Q"} as the predicate, with its proof obligation
+ solved by the local fact U (using the Isar attribute @{attribute OF}). The
+ following example is a trivial use of this method.
\<close>
lemma "halts p \<Longrightarrow> \<exists>x. halts x"
@@ -419,13 +419,12 @@
with a universal quantifier in the premises that matches the type of @{term
y}. Since @{keyword "premises"} causes a focus, however, there are no
subgoal premises to be found and thus @{method my_allE_bad} will always
- fail. If focusing instead left the premises in place, using methods
- like @{method erule} would lead to unintended behaviour, specifically during
+ fail. If focusing instead left the premises in place, using methods like
+ @{method erule} would lead to unintended behaviour, specifically during
backtracking. In our example, @{method erule} could choose an alternate
- premise while backtracking, while leaving \<open>I\<close> bound to the original
- match. In the case of more complex inner methods, where either \<open>I\<close> or
- bound terms are used, this would almost certainly not be the intended
- behaviour.
+ premise while backtracking, while leaving \<open>I\<close> bound to the original match.
+ In the case of more complex inner methods, where either \<open>I\<close> or bound terms
+ are used, this would almost certainly not be the intended behaviour.
An alternative implementation would be to specialize the elimination rule to
the bound term and apply it directly.
@@ -444,13 +443,13 @@
premise, it is not likely the intended behaviour. Repeated application of
this method will produce an infinite stream of duplicate specialized
premises, due to the original premise never being removed. To address this,
- matched premises may be declared with the @{attribute "thin"} attribute.
- This will hide the premise from subsequent inner matches, and remove it from
- the list of premises when the inner method has finished and the subgoal is
+ matched premises may be declared with the @{attribute thin} attribute. This
+ will hide the premise from subsequent inner matches, and remove it from the
+ list of premises when the inner method has finished and the subgoal is
unfocused. It can be considered analogous to the existing \<open>thin_tac\<close>.
- To complete our example, the correct implementation of the method
- will @{attribute "thin"} the premise from the match and then apply it to the
+ To complete our example, the correct implementation of the method will
+ @{attribute thin} the premise from the match and then apply it to the
specialized elimination rule.\<close>
method my_allE for y :: 'a =
@@ -460,13 +459,14 @@
lemma "\<forall>x. P x \<Longrightarrow> \<forall>x. Q x \<Longrightarrow> P y \<and> Q y"
by (my_allE y)+ (rule conjI)
+
subsubsection \<open>Inner focusing\<close>
text \<open>
- Premises are \<^emph>\<open>accumulated\<close> for the purposes of subgoal focusing.
- In contrast to using standard methods like @{method frule} within
- focused match, another @{method match} will have access to all the premises
- of the outer focus.
+ Premises are \<^emph>\<open>accumulated\<close> for the purposes of subgoal focusing. In
+ contrast to using standard methods like @{method frule} within focused
+ match, another @{method match} will have access to all the premises of the
+ outer focus.
\<close>
lemma "A \<Longrightarrow> B \<Longrightarrow> A \<and> B"
@@ -475,25 +475,23 @@
text \<open>
In this example, the inner @{method match} can find the focused premise
- @{term B}. In contrast, the @{method assumption} method would fail here
- due to @{term B} not being logically accessible.
+ @{term B}. In contrast, the @{method assumption} method would fail here due
+ to @{term B} not being logically accessible.
\<close>
- lemma
- "A \<Longrightarrow> A \<and> (B \<longrightarrow> B)"
+ lemma "A \<Longrightarrow> A \<and> (B \<longrightarrow> B)"
by (match premises in H: A \<Rightarrow> \<open>intro conjI, rule H, rule impI,
match premises (local) in A \<Rightarrow> \<open>fail\<close>
\<bar> H': B \<Rightarrow> \<open>rule H'\<close>\<close>)
text \<open>
- In this example, the only premise that exists in the first focus is
- @{term "A"}. Prior to the inner match, the rule \<open>impI\<close> changes
- the goal @{term "B \<longrightarrow> B"} into @{term "B \<Longrightarrow> B"}. A standard premise
- match would also include @{term A} as an original premise of the outer
- match. The \<open>local\<close> argument limits the match to
- newly focused premises.
+ In this example, the only premise that exists in the first focus is @{term
+ "A"}. Prior to the inner match, the rule \<open>impI\<close> changes the goal @{term "B \<longrightarrow>
+ B"} into @{term "B \<Longrightarrow> B"}. A standard premise match would also include @{term
+ A} as an original premise of the outer match. The \<open>local\<close> argument limits
+ the match to newly focused premises.
+\<close>
-\<close>
section \<open>Attributes\<close>
@@ -547,8 +545,7 @@
text \<open>
The @{attribute of} attribute behaves similarly. It is worth noting,
however, that the positional instantiation of @{attribute of} occurs against
- the position of the variables as they are declared \<^emph>\<open>in the match
- pattern\<close>.
+ the position of the variables as they are declared \<^emph>\<open>in the match pattern\<close>.
\<close>
lemma
@@ -559,15 +556,16 @@
\<open>rule I [of x y]\<close>)
text \<open>
- In this example, the order of schematics in \<open>asm\<close> is actually \<open>?y ?x\<close>, but we instantiate our matched rule in the opposite order. This is
- because the effective rule @{term I} was bound from the match, which
- declared the @{typ 'a} slot first and the @{typ 'b} slot second.
+ In this example, the order of schematics in \<open>asm\<close> is actually \<open>?y ?x\<close>, but
+ we instantiate our matched rule in the opposite order. This is because the
+ effective rule @{term I} was bound from the match, which declared the @{typ
+ 'a} slot first and the @{typ 'b} slot second.
To get the dynamic behaviour of @{attribute of} we can choose to invoke it
- \<^emph>\<open>unchecked\<close>. This avoids trying to do any type inference for the
- provided parameters, instead storing them as their most general type and
- doing type matching at run-time. This, like @{attribute OF}, will throw
- errors if the expected slots don't exist or there is a type mismatch.
+ \<^emph>\<open>unchecked\<close>. This avoids trying to do any type inference for the provided
+ parameters, instead storing them as their most general type and doing type
+ matching at run-time. This, like @{attribute OF}, will throw errors if the
+ expected slots don't exist or there is a type mismatch.
\<close>
lemma
@@ -587,11 +585,11 @@
\<open>prop_solver\<close>)
text \<open>
- In this example, the pattern \<open>\<And>x :: 'a. ?P x \<Longrightarrow> ?Q x\<close> matches against
- the only premise, giving an appropriately typed slot for @{term y}. After
- the match, the resulting rule is instantiated to @{term y} and then declared
- as an @{attribute intros} rule. This is then picked up by @{method
- prop_solver} to solve the goal.
+ In this example, the pattern \<open>\<And>x :: 'a. ?P x \<Longrightarrow> ?Q x\<close> matches against the
+ only premise, giving an appropriately typed slot for @{term y}. After the
+ match, the resulting rule is instantiated to @{term y} and then declared as
+ an @{attribute intros} rule. This is then picked up by @{method prop_solver}
+ to solve the goal.
\<close>
@@ -600,8 +598,9 @@
text \<open>
In all previous examples, @{method match} was only ever searching for a
single rule or premise. Each local fact would therefore always have a length
- of exactly one. We may, however, wish to find \<^emph>\<open>all\<close> matching results.
- To achieve this, we can simply mark a given pattern with the \<open>(multi)\<close> argument.
+ of exactly one. We may, however, wish to find \<^emph>\<open>all\<close> matching results. To
+ achieve this, we can simply mark a given pattern with the \<open>(multi)\<close>
+ argument.
\<close>
lemma
@@ -612,21 +611,21 @@
done
text \<open>
- In the first @{method match}, without the \<open>(multi)\<close> argument, @{term
- I} is only ever be bound to one of the members of \<open>asms\<close>. This
- backtracks over both possibilities (see next section), however neither
- assumption in isolation is sufficient to solve to goal. The use of the
- @{method solves} combinator ensures that @{method prop_solver} has no effect
- on the goal when it doesn't solve it, and so the first match leaves the goal
- unchanged. In the second @{method match}, \<open>I\<close> is bound to all of
- \<open>asms\<close>, declaring both results as \<open>intros\<close>. With these rules
- @{method prop_solver} is capable of solving the goal.
+ In the first @{method match}, without the \<open>(multi)\<close> argument, @{term I} is
+ only ever be bound to one of the members of \<open>asms\<close>. This backtracks over
+ both possibilities (see next section), however neither assumption in
+ isolation is sufficient to solve to goal. The use of the @{method solves}
+ combinator ensures that @{method prop_solver} has no effect on the goal when
+ it doesn't solve it, and so the first match leaves the goal unchanged. In
+ the second @{method match}, \<open>I\<close> is bound to all of \<open>asms\<close>, declaring both
+ results as \<open>intros\<close>. With these rules @{method prop_solver} is capable of
+ solving the goal.
Using for-fixed variables in patterns imposes additional constraints on the
- results. In all previous examples, the choice of using \<open>?P\<close> or a
- for-fixed @{term P} only depended on whether or not @{term P} was mentioned
- in another pattern or the inner method. When using a multi-match, however,
- all for-fixed terms must agree in the results.
+ results. In all previous examples, the choice of using \<open>?P\<close> or a for-fixed
+ @{term P} only depended on whether or not @{term P} was mentioned in another
+ pattern or the inner method. When using a multi-match, however, all
+ for-fixed terms must agree in the results.
\<close>
lemma
@@ -641,11 +640,11 @@
text \<open>
Here we have two seemingly-equivalent applications of @{method match},
however only the second one is capable of solving the goal. The first
- @{method match} selects the first and third members of \<open>asms\<close> (those
- that agree on their conclusion), which is not sufficient. The second
- @{method match} selects the first and second members of \<open>asms\<close> (those
- that agree on their assumption), which is enough for @{method prop_solver}
- to solve the goal.
+ @{method match} selects the first and third members of \<open>asms\<close> (those that
+ agree on their conclusion), which is not sufficient. The second @{method
+ match} selects the first and second members of \<open>asms\<close> (those that agree on
+ their assumption), which is enough for @{method prop_solver} to solve the
+ goal.
\<close>
@@ -655,10 +654,10 @@
Dummy patterns may be given as placeholders for unique schematics in
patterns. They implicitly receive all currently bound variables as
arguments, and are coerced into the @{typ prop} type whenever possible. For
- example, the trivial dummy pattern \<open>_\<close> will match any proposition.
- In contrast, by default the pattern \<open>?P\<close> is considered to have type
- @{typ bool}. It will not bind anything with meta-logical connectives (e.g.
- \<open>_ \<Longrightarrow> _\<close> or \<open>_ &&& _\<close>).
+ example, the trivial dummy pattern \<open>_\<close> will match any proposition. In
+ contrast, by default the pattern \<open>?P\<close> is considered to have type @{typ
+ bool}. It will not bind anything with meta-logical connectives (e.g. \<open>_ \<Longrightarrow> _\<close>
+ or \<open>_ &&& _\<close>).
\<close>
lemma
@@ -670,19 +669,19 @@
section \<open>Backtracking\<close>
text \<open>
- Patterns are considered top-down, executing the inner method \<open>m\<close> of
- the first pattern which is satisfied by the current match target. By
- default, matching performs extensive backtracking by attempting all valid
- variable and fact bindings according to the given pattern. In particular,
- all unifiers for a given pattern will be explored, as well as each matching
+ Patterns are considered top-down, executing the inner method \<open>m\<close> of the
+ first pattern which is satisfied by the current match target. By default,
+ matching performs extensive backtracking by attempting all valid variable
+ and fact bindings according to the given pattern. In particular, all
+ unifiers for a given pattern will be explored, as well as each matching
fact. The inner method \<open>m\<close> will be re-executed for each different
variable/fact binding during backtracking. A successful match is considered
a cut-point for backtracking. Specifically, once a match is made no other
pattern-method pairs will be considered.
- The method \<open>foo\<close> below fails for all goals that are conjunctions. Any
- such goal will match the first pattern, causing the second pattern (that
- would otherwise match all goals) to never be considered.
+ The method \<open>foo\<close> below fails for all goals that are conjunctions. Any such
+ goal will match the first pattern, causing the second pattern (that would
+ otherwise match all goals) to never be considered.
\<close>
method foo =
@@ -690,12 +689,12 @@
text \<open>
The failure of an inner method that is executed after a successful match
- will cause the entire match to fail. This distinction is important
- due to the pervasive use of backtracking. When a method is used in a
- combinator chain, its failure
- becomes significant because it signals previously applied methods to move to
- the next result. Therefore, it is necessary for @{method match} to not mask
- such failure. One can always rewrite a match using the combinators ``\<open>?\<close>'' and ``\<open>|\<close>'' to try subsequent patterns in the case of an
+ will cause the entire match to fail. This distinction is important due to
+ the pervasive use of backtracking. When a method is used in a combinator
+ chain, its failure becomes significant because it signals previously applied
+ methods to move to the next result. Therefore, it is necessary for @{method
+ match} to not mask such failure. One can always rewrite a match using the
+ combinators ``\<open>?\<close>'' and ``\<open>|\<close>'' to try subsequent patterns in the case of an
inner-method failure. The following proof method, for example, always
invokes @{method prop_solver} for all goals because its first alternative
either never matches or (if it does match) always fails.
@@ -710,8 +709,8 @@
text \<open>
Backtracking may be controlled more precisely by marking individual patterns
- as \<open>cut\<close>. This causes backtracking to not progress beyond this pattern:
- once a match is found no others will be considered.
+ as \<open>cut\<close>. This causes backtracking to not progress beyond this pattern: once
+ a match is found no others will be considered.
\<close>
method foo\<^sub>2 =
@@ -722,10 +721,10 @@
In this example, once a conjunction is found (@{term "P \<and> Q"}), all possible
implications of @{term "P"} in the premises are considered, evaluating the
inner @{method rule} with each consequent. No other conjunctions will be
- considered, with method failure occurring once all implications of the
- form \<open>P \<longrightarrow> ?U\<close> have been explored. Here the left-right processing of
- individual patterns is important, as all patterns after of the cut will
- maintain their usual backtracking behaviour.
+ considered, with method failure occurring once all implications of the form
+ \<open>P \<longrightarrow> ?U\<close> have been explored. Here the left-right processing of individual
+ patterns is important, as all patterns after of the cut will maintain their
+ usual backtracking behaviour.
\<close>
lemma "A \<and> B \<Longrightarrow> A \<longrightarrow> D \<Longrightarrow> A \<longrightarrow> C \<Longrightarrow> C"
@@ -735,16 +734,16 @@
by (foo\<^sub>2 | prop_solver)
text \<open>
- In this example, the first lemma is solved by \<open>foo\<^sub>2\<close>, by first
- picking @{term "A \<longrightarrow> D"} for \<open>I'\<close>, then backtracking and ultimately
- succeeding after picking @{term "A \<longrightarrow> C"}. In the second lemma, however,
- @{term "C \<and> D"} is matched first, the second pattern in the match cannot be
- found and so the method fails, falling through to @{method prop_solver}.
+ In this example, the first lemma is solved by \<open>foo\<^sub>2\<close>, by first picking
+ @{term "A \<longrightarrow> D"} for \<open>I'\<close>, then backtracking and ultimately succeeding after
+ picking @{term "A \<longrightarrow> C"}. In the second lemma, however, @{term "C \<and> D"} is
+ matched first, the second pattern in the match cannot be found and so the
+ method fails, falling through to @{method prop_solver}.
- More precise control is also possible by giving a positive
- number \<open>n\<close> as an argument to \<open>cut\<close>. This will limit the number
- of backtracking results of that match to be at most \<open>n\<close>.
- The match argument \<open>(cut 1)\<close> is the same as simply \<open>(cut)\<close>.
+ More precise control is also possible by giving a positive number \<open>n\<close> as an
+ argument to \<open>cut\<close>. This will limit the number of backtracking results of
+ that match to be at most \<open>n\<close>. The match argument \<open>(cut 1)\<close> is the same as
+ simply \<open>(cut)\<close>.
\<close>
@@ -769,15 +768,16 @@
text \<open>
Intuitively it seems like this proof should fail to check. The first match
- result, which binds @{term I} to the first two members of \<open>asms\<close>,
- fails the second inner match due to binding @{term P} to @{term A}.
- Backtracking then attempts to bind @{term I} to the third member of \<open>asms\<close>. This passes all inner matches, but fails when @{method rule} cannot
- successfully apply this to the current goal. After this, a valid match that
- is produced by the unifier is one which binds @{term P} to simply \<open>\<lambda>a. A ?x\<close>. The first inner match succeeds because \<open>\<lambda>a. A ?x\<close> does
- not match @{term A}. The next inner match succeeds because @{term I} has
- only been bound to the first member of \<open>asms\<close>. This is due to @{method
- match} considering \<open>\<lambda>a. A ?x\<close> and \<open>\<lambda>a. A ?y\<close> as distinct
- terms.
+ result, which binds @{term I} to the first two members of \<open>asms\<close>, fails the
+ second inner match due to binding @{term P} to @{term A}. Backtracking then
+ attempts to bind @{term I} to the third member of \<open>asms\<close>. This passes all
+ inner matches, but fails when @{method rule} cannot successfully apply this
+ to the current goal. After this, a valid match that is produced by the
+ unifier is one which binds @{term P} to simply \<open>\<lambda>a. A ?x\<close>. The first inner
+ match succeeds because \<open>\<lambda>a. A ?x\<close> does not match @{term A}. The next inner
+ match succeeds because @{term I} has only been bound to the first member of
+ \<open>asms\<close>. This is due to @{method match} considering \<open>\<lambda>a. A ?x\<close> and \<open>\<lambda>a. A ?y\<close>
+ as distinct terms.
The simplest way to address this is to explicitly disallow term bindings
which we would consider invalid.
@@ -797,8 +797,8 @@
text \<open>
The @{method match} method is not aware of the logical content of match
targets. Each pattern is simply matched against the shallow structure of a
- fact or term. Most facts are in \<^emph>\<open>normal form\<close>, which curries premises
- via meta-implication \<open>_ \<Longrightarrow> _\<close>.
+ fact or term. Most facts are in \<^emph>\<open>normal form\<close>, which curries premises via
+ meta-implication \<open>_ \<Longrightarrow> _\<close>.
\<close>
lemma
@@ -821,17 +821,17 @@
text \<open>
This proof will fail to solve the goal. Our match pattern will only match
rules which have a single premise, and conclusion @{term C}, so the first
- member of \<open>asms\<close> is not bound and thus the proof fails. Matching a
- pattern of the form @{term "P \<Longrightarrow> Q"} against this fact will bind @{term "P"}
- to @{term "A"} and @{term Q} to @{term "B \<Longrightarrow> C"}. Our pattern, with a
- concrete @{term "C"} in the conclusion, will fail to match this fact.
+ member of \<open>asms\<close> is not bound and thus the proof fails. Matching a pattern
+ of the form @{term "P \<Longrightarrow> Q"} against this fact will bind @{term "P"} to
+ @{term "A"} and @{term Q} to @{term "B \<Longrightarrow> C"}. Our pattern, with a concrete
+ @{term "C"} in the conclusion, will fail to match this fact.
- To express our desired match, we may \<^emph>\<open>uncurry\<close> our rules before
- matching against them. This forms a meta-conjunction of all premises in a
- fact, so that only one implication remains. For example the uncurried
- version of @{term "A \<Longrightarrow> B \<Longrightarrow> C"} is @{term "A &&& B \<Longrightarrow> C"}. This will now match
- our desired pattern \<open>_ \<Longrightarrow> C\<close>, and can be \<^emph>\<open>curried\<close> after the
- match to put it back into normal form.
+ To express our desired match, we may \<^emph>\<open>uncurry\<close> our rules before matching
+ against them. This forms a meta-conjunction of all premises in a fact, so
+ that only one implication remains. For example the uncurried version of
+ @{term "A \<Longrightarrow> B \<Longrightarrow> C"} is @{term "A &&& B \<Longrightarrow> C"}. This will now match our
+ desired pattern \<open>_ \<Longrightarrow> C\<close>, and can be \<^emph>\<open>curried\<close> after the match to put it
+ back into normal form.
\<close>
lemma
@@ -858,12 +858,12 @@
done
text \<open>
- In the first @{method match} we attempt to find a member of \<open>asms\<close>
- which matches our goal precisely. This fails due to no such member existing.
- The second match reverses the role of the fact in the match, by first giving
- a general pattern @{term P}. This bound pattern is then matched against
- @{term "A y"}. In this case, @{term P} is bound to \<open>A ?x\<close> and so it
- successfully matches.
+ In the first @{method match} we attempt to find a member of \<open>asms\<close> which
+ matches our goal precisely. This fails due to no such member existing. The
+ second match reverses the role of the fact in the match, by first giving a
+ general pattern @{term P}. This bound pattern is then matched against @{term
+ "A y"}. In this case, @{term P} is bound to \<open>A ?x\<close> and so it successfully
+ matches.
\<close>
@@ -883,9 +883,10 @@
\<open>match (y) in "y :: 'b" for y \<Rightarrow> \<open>rule H [where z = y]\<close>\<close>)
text \<open>
- In this example the type \<open>'b\<close> is matched to \<open>'a\<close>, however
- statically they are formally distinct types. The first match binds \<open>'b\<close> while the inner match serves to coerce @{term y} into having the type
- \<open>'b\<close>. This allows the rule instantiation to successfully apply.
+ In this example the type \<open>'b\<close> is matched to \<open>'a\<close>, however statically they
+ are formally distinct types. The first match binds \<open>'b\<close> while the inner
+ match serves to coerce @{term y} into having the type \<open>'b\<close>. This allows the
+ rule instantiation to successfully apply.
\<close>
@@ -922,10 +923,10 @@
text \<open>
A custom rule attribute is a simple way to extend the functionality of
- Eisbach methods. The dummy rule attribute notation (\<open>[[ _ ]]\<close>)
- invokes the given attribute against a dummy fact and evaluates to the result
- of that attribute. When used as a match target, this can serve as an
- effective auxiliary function.
+ Eisbach methods. The dummy rule attribute notation (\<open>[[ _ ]]\<close>) invokes the
+ given attribute against a dummy fact and evaluates to the result of that
+ attribute. When used as a match target, this can serve as an effective
+ auxiliary function.
\<close>
attribute_setup get_split_rule =
--- a/src/Doc/Eisbach/Preface.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Eisbach/Preface.thy Thu Nov 05 08:27:22 2015 +0100
@@ -5,10 +5,10 @@
begin
text \<open>
- \<^emph>\<open>Eisbach\<close> is a collection of tools which form the basis for defining
- new proof methods in Isabelle/Isar~@{cite "Wenzel-PhD"}. It can be thought
- of as a ``proof method language'', but is more precisely an infrastructure
- for defining new proof methods out of existing ones.
+ \<^emph>\<open>Eisbach\<close> is a collection of tools which form the basis for defining new
+ proof methods in Isabelle/Isar~@{cite "Wenzel-PhD"}. It can be thought of as
+ a ``proof method language'', but is more precisely an infrastructure for
+ defining new proof methods out of existing ones.
The core functionality of Eisbach is provided by the Isar @{command method}
command. Here users may define new methods by combining existing ones with
@@ -27,8 +27,8 @@
high barrier-to-entry for many users.
\<^medskip>
- This manual is written for users familiar with Isabelle/Isar, but
- not necessarily Isabelle/ML. It covers the usage of the @{command method} as
+ This manual is written for users familiar with Isabelle/Isar, but not
+ necessarily Isabelle/ML. It covers the usage of the @{command method} as
well as the @{method match} method, as well as discussing their integration
with existing Isar concepts such as @{command named_theorems}.
\<close>
--- a/src/Doc/Implementation/Isar.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Isar.thy Thu Nov 05 08:27:22 2015 +0100
@@ -183,10 +183,10 @@
method space, e.g.\ @{method rule_tac}.
\<^item> A non-trivial method always needs to make progress: an
- identical follow-up goal state has to be avoided.\footnote{This
+ identical follow-up goal state has to be avoided.\<^footnote>\<open>This
enables the user to write method expressions like \<open>meth\<^sup>+\<close>
without looping, while the trivial do-nothing case can be recovered
- via \<open>meth\<^sup>?\<close>.}
+ via \<open>meth\<^sup>?\<close>.\<close>
Exception: trivial stuttering steps, such as ``@{method -}'' or
@{method succeed}.
@@ -275,11 +275,11 @@
When implementing proof methods, it is advisable to study existing
implementations carefully and imitate the typical ``boiler plate''
for context-sensitive parsing and further combinators to wrap-up
- tactic expressions as methods.\footnote{Aliases or abbreviations of
+ tactic expressions as methods.\<^footnote>\<open>Aliases or abbreviations of
the standard method combinators should be avoided. Note that from
Isabelle99 until Isabelle2009 the system did provide various odd
combinations of method syntax wrappers that made applications more
- complicated than necessary.}
+ complicated than necessary.\<close>
\<close>
text %mlref \<open>
--- a/src/Doc/Implementation/Logic.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Logic.thy Thu Nov 05 08:27:22 2015 +0100
@@ -19,11 +19,11 @@
Derivations are relative to a logical theory, which declares type
constructors, constants, and axioms. Theory declarations support
schematic polymorphism, which is strictly speaking outside the
- logic.\footnote{This is the deeper logical reason, why the theory
+ logic.\<^footnote>\<open>This is the deeper logical reason, why the theory
context \<open>\<Theta>\<close> is separate from the proof context \<open>\<Gamma>\<close>
of the core calculus: type constructors, term constants, and facts
(proof constants) may involve arbitrary type schemes, but the type
- of a locally fixed term parameter is also fixed!}
+ of a locally fixed term parameter is also fixed!\<close>
\<close>
@@ -531,9 +531,9 @@
the simple syntactic types of Pure are always inhabitable.
``Assumptions'' \<open>x :: \<tau>\<close> for type-membership are only
present as long as some \<open>x\<^sub>\<tau>\<close> occurs in the statement
- body.\footnote{This is the key difference to ``\<open>\<lambda>HOL\<close>'' in
+ body.\<^footnote>\<open>This is the key difference to ``\<open>\<lambda>HOL\<close>'' in
the PTS framework @{cite "Barendregt-Geuvers:2001"}, where hypotheses
- \<open>x : A\<close> are treated uniformly for propositions and types.}
+ \<open>x : A\<close> are treated uniformly for propositions and types.\<close>
\<^medskip>
The axiomatization of a theory is implicitly closed by
--- a/src/Doc/Implementation/ML.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/ML.thy Thu Nov 05 08:27:22 2015 +0100
@@ -23,11 +23,11 @@
first-hand explanations should help to understand how proper
Isabelle/ML is to be read and written, and to get access to the
wealth of experience that is expressed in the source text and its
- history of changes.\footnote{See
+ history of changes.\<^footnote>\<open>See
@{url "http://isabelle.in.tum.de/repos/isabelle"} for the full
Mercurial history. There are symbolic tags to refer to official
Isabelle releases, as opposed to arbitrary \<^emph>\<open>tip\<close> versions that
- merely reflect snapshots that are never really up-to-date.}\<close>
+ merely reflect snapshots that are never really up-to-date.\<close>\<close>
section \<open>Style and orthography\<close>
@@ -37,10 +37,10 @@
to tell an informed reader what is really going on and how things
really work. This is a non-trivial aim, but it is supported by a
certain style of writing Isabelle/ML that has emerged from long
- years of system development.\footnote{See also the interesting style
+ years of system development.\<^footnote>\<open>See also the interesting style
guide for OCaml
@{url "http://caml.inria.fr/resources/doc/guides/guidelines.en.html"}
- which shares many of our means and ends.}
+ which shares many of our means and ends.\<close>
The main principle behind any coding style is \<^emph>\<open>consistency\<close>.
For a single author of a small program this merely means ``choose
@@ -123,10 +123,10 @@
For historical reasons, many capitalized names omit underscores,
e.g.\ old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}.
Genuine mixed-case names are \<^emph>\<open>not\<close> used, because clear division
- of words is essential for readability.\footnote{Camel-case was
+ of words is essential for readability.\<^footnote>\<open>Camel-case was
invented to workaround the lack of underscore in some early
non-ASCII character sets. Later it became habitual in some language
- communities that are now strong in numbers.}
+ communities that are now strong in numbers.\<close>
A single (capital) character does not count as ``word'' in this
respect: some Isabelle/ML names are suffixed by extra markers like
@@ -279,10 +279,10 @@
paragraph \<open>Line length\<close>
text \<open>is limited to 80 characters according to ancient standards, but we allow
- as much as 100 characters (not more).\footnote{Readability requires to keep
+ as much as 100 characters (not more).\<^footnote>\<open>Readability requires to keep
the beginning of a line in view while watching its end. Modern wide-screen
displays do not change the way how the human brain works. Sources also need
- to be printable on plain paper with reasonable font-size.} The extra 20
+ to be printable on plain paper with reasonable font-size.\<close> The extra 20
characters acknowledge the space requirements due to qualified library
references in Isabelle/ML.\<close>
@@ -327,11 +327,11 @@
\<close>
paragraph \<open>Indentation\<close>
-text \<open>uses plain spaces, never hard tabulators.\footnote{Tabulators were
+text \<open>uses plain spaces, never hard tabulators.\<^footnote>\<open>Tabulators were
invented to move the carriage of a type-writer to certain predefined
positions. In software they could be used as a primitive run-length
compression of consecutive spaces, but the precise result would depend on
- non-standardized text editor configuration.}
+ non-standardized text editor configuration.\<close>
Each level of nesting is indented by 2 spaces, sometimes 1, very
rarely 4, never 8 or any other odd number.
@@ -562,10 +562,10 @@
Removing the above ML declaration from the source text will remove any trace
of this definition, as expected. The Isabelle/ML toplevel environment is
managed in a \<^emph>\<open>stateless\<close> way: in contrast to the raw ML toplevel, there
- are no global side-effects involved here.\footnote{Such a stateless
+ are no global side-effects involved here.\<^footnote>\<open>Such a stateless
compilation environment is also a prerequisite for robust parallel
compilation within independent nodes of the implicit theory development
- graph.}
+ graph.\<close>
\<^medskip>
The next example shows how to embed ML into Isar proofs, using
@@ -578,7 +578,7 @@
ML_prf %"ML" \<open>val a = 1\<close>
{
ML_prf %"ML" \<open>val b = a + 1\<close>
- } -- \<open>Isar block structure ignored by ML environment\<close>
+ } \<comment> \<open>Isar block structure ignored by ML environment\<close>
ML_prf %"ML" \<open>val c = b + 1\<close>
end
@@ -739,9 +739,9 @@
type \<open>\<tau>\<close> is represented by the iterated function space
\<open>\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>. This is isomorphic to the well-known
encoding via tuples \<open>\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>, but the curried
- version fits more smoothly into the basic calculus.\footnote{The
+ version fits more smoothly into the basic calculus.\<^footnote>\<open>The
difference is even more significant in HOL, because the redundant
- tuple structure needs to be accommodated extraneous proof steps.}
+ tuple structure needs to be accommodated extraneous proof steps.\<close>
Currying gives some flexibility due to \<^emph>\<open>partial application\<close>. A
function \<open>f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^sub>2 \<rightarrow> \<tau>\<close> can be applied to \<open>x: \<tau>\<^sub>1\<close>
@@ -1282,9 +1282,9 @@
\<^descr> @{ML "Symbol.explode"}~\<open>str\<close> produces a symbol list
from the packed form. This function supersedes @{ML
"String.explode"} for virtually all purposes of manipulating text in
- Isabelle!\footnote{The runtime overhead for exploded strings is
+ Isabelle!\<^footnote>\<open>The runtime overhead for exploded strings is
mainly that of the list structure: individual symbols that happen to
- be a singleton string do not require extra memory in Poly/ML.}
+ be a singleton string do not require extra memory in Poly/ML.\<close>
\<^descr> @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
"Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
@@ -1396,8 +1396,8 @@
\<^descr> Type @{ML_type int} represents regular mathematical integers, which
are \<^emph>\<open>unbounded\<close>. Overflow is treated properly, but should never happen
- in practice.\footnote{The size limit for integer bit patterns in memory is
- 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.} This works
+ in practice.\<^footnote>\<open>The size limit for integer bit patterns in memory is
+ 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.\<close> This works
uniformly for all supported ML platforms (Poly/ML and SML/NJ).
Literal integers in ML text are forced to be of this one true
@@ -1614,13 +1614,13 @@
sub-components with explicit communication, general asynchronous
interaction etc. Moreover, parallel evaluation is a prerequisite to
make adequate use of the CPU resources that are available on
- multi-core systems.\footnote{Multi-core computing does not mean that
+ multi-core systems.\<^footnote>\<open>Multi-core computing does not mean that
there are ``spare cycles'' to be wasted. It means that the
continued exponential speedup of CPU performance due to ``Moore's
Law'' follows different rules: clock frequency has reached its peak
around 2005, and applications need to be parallelized in order to
avoid a perceived loss of performance. See also
- @{cite "Sutter:2005"}.}
+ @{cite "Sutter:2005"}.\<close>
Isabelle/Isar exploits the inherent structure of theories and proofs to
support \<^emph>\<open>implicit parallelism\<close> to a large extent. LCF-style theorem
@@ -1671,8 +1671,8 @@
\<^item> Global references (or arrays), i.e.\ mutable memory cells that
persist over several invocations of associated
- operations.\footnote{This is independent of the visibility of such
- mutable values in the toplevel scope.}
+ operations.\<^footnote>\<open>This is independent of the visibility of such
+ mutable values in the toplevel scope.\<close>
\<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O
channels, environment variables, current working directory.
--- a/src/Doc/Implementation/Prelim.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Prelim.thy Thu Nov 05 08:27:22 2015 +0100
@@ -471,17 +471,17 @@
begin
declare [[show_types = false]]
- -- \<open>declaration within (local) theory context\<close>
+ \<comment> \<open>declaration within (local) theory context\<close>
notepad
begin
note [[show_types = true]]
- -- \<open>declaration within proof (forward mode)\<close>
+ \<comment> \<open>declaration within proof (forward mode)\<close>
term x
have "x = x"
using [[show_types = false]]
- -- \<open>declaration within proof (backward mode)\<close>
+ \<comment> \<open>declaration within proof (backward mode)\<close>
..
end
--- a/src/Doc/Implementation/Proof.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Proof.thy Thu Nov 05 08:27:22 2015 +0100
@@ -58,14 +58,14 @@
notepad
begin
{
- fix x -- \<open>all potential occurrences of some \<open>x::\<tau>\<close> are fixed\<close>
+ fix x \<comment> \<open>all potential occurrences of some \<open>x::\<tau>\<close> are fixed\<close>
{
- have "x::'a \<equiv> x" -- \<open>implicit type assignment by concrete occurrence\<close>
+ have "x::'a \<equiv> x" \<comment> \<open>implicit type assignment by concrete occurrence\<close>
by (rule reflexive)
}
- thm this -- \<open>result still with fixed type \<open>'a\<close>\<close>
+ thm this \<comment> \<open>result still with fixed type \<open>'a\<close>\<close>
}
- thm this -- \<open>fully general result for arbitrary \<open>?x::?'a\<close>\<close>
+ thm this \<comment> \<open>fully general result for arbitrary \<open>?x::?'a\<close>\<close>
end
text \<open>The Isabelle/Isar proof context manages the details of term
--- a/src/Doc/Implementation/Syntax.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Syntax.thy Thu Nov 05 08:27:22 2015 +0100
@@ -20,9 +20,9 @@
Moreover, type-inference in the style of Hindley-Milner @{cite hindleymilner}
(and extensions) enables users to write \<open>\<forall>x. B x\<close> concisely, when
the type \<open>'a\<close> is already clear from the
- context.\footnote{Type-inference taken to the extreme can easily confuse
+ context.\<^footnote>\<open>Type-inference taken to the extreme can easily confuse
users. Beginners often stumble over unexpectedly general types inferred by
- the system.}
+ the system.\<close>
\<^medskip>
The main inner syntax operations are \<^emph>\<open>read\<close> for
--- a/src/Doc/Implementation/Tactic.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Implementation/Tactic.thy Thu Nov 05 08:27:22 2015 +0100
@@ -18,10 +18,10 @@
Isabelle/Pure represents a goal as a theorem stating that the
subgoals imply the main goal: \<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow>
C\<close>. The outermost goal structure is that of a Horn Clause: i.e.\
- an iterated implication without any quantifiers\footnote{Recall that
+ an iterated implication without any quantifiers\<^footnote>\<open>Recall that
outermost \<open>\<And>x. \<phi>[x]\<close> is always represented via schematic
variables in the body: \<open>\<phi>[?x]\<close>. These variables may get
- instantiated during the course of reasoning.}. For \<open>n = 0\<close>
+ instantiated during the course of reasoning.\<close>. For \<open>n = 0\<close>
a goal is called ``solved''.
The structure of each subgoal \<open>A\<^sub>i\<close> is that of a
@@ -90,11 +90,11 @@
\secref{sec:tactical-goals}) to a lazy sequence of potential
successor states. The underlying sequence implementation is lazy
both in head and tail, and is purely functional in \<^emph>\<open>not\<close>
- supporting memoing.\footnote{The lack of memoing and the strict
+ supporting memoing.\<^footnote>\<open>The lack of memoing and the strict
nature of ML requires some care when working with low-level
sequence operations, to avoid duplicate or premature evaluation of
results. It also means that modified runtime behavior, such as
- timeout, is very hard to achieve for general tactics.}
+ timeout, is very hard to achieve for general tactics.\<close>
An \<^emph>\<open>empty result sequence\<close> means that the tactic has failed: in
a compound tactic expression other tactics might be tried instead,
@@ -319,12 +319,12 @@
\<^descr> @{ML match_tac}, @{ML ematch_tac}, @{ML dmatch_tac}, and @{ML
bimatch_tac} are similar to @{ML resolve_tac}, @{ML eresolve_tac},
@{ML dresolve_tac}, and @{ML biresolve_tac}, respectively, but do
- not instantiate schematic variables in the goal state.%
-\footnote{Strictly speaking, matching means to treat the unknowns in the goal
+ not instantiate schematic variables in the goal state.\<^footnote>\<open>Strictly speaking,
+ matching means to treat the unknowns in the goal
state as constants, but these tactics merely discard unifiers that would
update the goal state. In rare situations (where the conclusion and
goal state have flexible terms at the same position), the tactic
- will fail even though an acceptable unifier exists.}
+ will fail even though an acceptable unifier exists.\<close>
These tactics were written for a specific application within the classical reasoner.
Flexible subgoals are not updated at will, but are left alone.
--- a/src/Doc/Isar_Ref/First_Order_Logic.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/First_Order_Logic.thy Thu Nov 05 08:27:22 2015 +0100
@@ -357,10 +357,10 @@
theorem
assumes "\<exists>x. \<forall>y. R x y"
shows "\<forall>y. \<exists>x. R x y"
-proof -- \<open>\<open>\<forall>\<close> introduction\<close>
- obtain x where "\<forall>y. R x y" using \<open>\<exists>x. \<forall>y. R x y\<close> .. -- \<open>\<open>\<exists>\<close> elimination\<close>
- fix y have "R x y" using \<open>\<forall>y. R x y\<close> .. -- \<open>\<open>\<forall>\<close> destruction\<close>
- then show "\<exists>x. R x y" .. -- \<open>\<open>\<exists>\<close> introduction\<close>
+proof \<comment> \<open>\<open>\<forall>\<close> introduction\<close>
+ obtain x where "\<forall>y. R x y" using \<open>\<exists>x. \<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<exists>\<close> elimination\<close>
+ fix y have "R x y" using \<open>\<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<forall>\<close> destruction\<close>
+ then show "\<exists>x. R x y" .. \<comment> \<open>\<open>\<exists>\<close> introduction\<close>
qed
--- a/src/Doc/Isar_Ref/Generic.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/Generic.thy Thu Nov 05 08:27:22 2015 +0100
@@ -327,9 +327,9 @@
\<^descr> @{method simp_all} is similar to @{method simp}, but acts on
all goals, working backwards from the last to the first one as usual
- in Isabelle.\footnote{The order is irrelevant for goals without
+ in Isabelle.\<^footnote>\<open>The order is irrelevant for goals without
schematic variables, so simplification might actually be performed
- in parallel here.}
+ in parallel here.\<close>
Chained facts are inserted into all subgoals, before the
simplification process starts. Further rule declarations are the
@@ -347,8 +347,8 @@
normalization process, or simplifying assumptions themselves.
Further options allow to fine-tune the behavior of the Simplifier
in this respect, corresponding to a variety of ML tactics as
- follows.\footnote{Unlike the corresponding Isar proof methods, the
- ML tactics do not insist in changing the goal state.}
+ follows.\<^footnote>\<open>Unlike the corresponding Isar proof methods, the
+ ML tactics do not insist in changing the goal state.\<close>
\begin{center}
\small
@@ -1179,9 +1179,9 @@
is easier to automate.
A \<^bold>\<open>sequent\<close> has the form \<open>\<Gamma> \<turnstile> \<Delta>\<close>, where \<open>\<Gamma>\<close>
- and \<open>\<Delta>\<close> are sets of formulae.\footnote{For first-order
+ and \<open>\<Delta>\<close> are sets of formulae.\<^footnote>\<open>For first-order
logic, sequents can equivalently be made from lists or multisets of
- formulae.} The sequent \<open>P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n\<close> is
+ formulae.\<close> The sequent \<open>P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n\<close> is
\<^bold>\<open>valid\<close> if \<open>P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m\<close> implies \<open>Q\<^sub>1 \<or> \<dots> \<or>
Q\<^sub>n\<close>. Thus \<open>P\<^sub>1, \<dots>, P\<^sub>m\<close> represent assumptions, each of which
is true, while \<open>Q\<^sub>1, \<dots>, Q\<^sub>n\<close> represent alternative goals. A
--- a/src/Doc/Isar_Ref/Inner_Syntax.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/Inner_Syntax.thy Thu Nov 05 08:27:22 2015 +0100
@@ -277,8 +277,8 @@
used internally in Isabelle/Pure.
\<^item> \<^verbatim>\<open>xsymbols\<close>: enable proper mathematical symbols
- instead of ASCII art.\footnote{This traditional mode name stems from
- the ``X-Symbol'' package for classic Proof~General with XEmacs.}
+ instead of ASCII art.\<^footnote>\<open>This traditional mode name stems from
+ the ``X-Symbol'' package for classic Proof~General with XEmacs.\<close>
\<^item> \<^verbatim>\<open>latex\<close>: additional mode that is active in {\LaTeX}
document preparation of Isabelle theory sources; allows to provide
@@ -338,9 +338,9 @@
grammar, where for each argument \<open>i\<close> the syntactic category
is determined by \<open>\<tau>\<^sub>i\<close> (with priority \<open>p\<^sub>i\<close>), and the
result category is determined from \<open>\<tau>\<close> (with priority \<open>p\<close>). Priority specifications are optional, with default 0 for
- arguments and 1000 for the result.\footnote{Omitting priorities is
+ arguments and 1000 for the result.\<^footnote>\<open>Omitting priorities is
prone to syntactic ambiguities unless the delimiter tokens determine
- fully bracketed notation, as in \<open>if _ then _ else _ fi\<close>.}
+ fully bracketed notation, as in \<open>if _ then _ else _ fi\<close>.\<close>
Since \<open>\<tau>\<close> may be again a function type, the constant
type scheme may have more argument positions than the mixfix
@@ -1213,10 +1213,10 @@
side-conditions:
\<^item> Rules must be left linear: \<open>lhs\<close> must not contain
- repeated variables.\footnote{The deeper reason for this is that AST
+ repeated variables.\<^footnote>\<open>The deeper reason for this is that AST
equality is not well-defined: different occurrences of the ``same''
AST could be decorated differently by accidental type-constraints or
- source position information, for example.}
+ source position information, for example.\<close>
\<^item> Every variable in \<open>rhs\<close> must also occur in \<open>lhs\<close>.
--- a/src/Doc/Isar_Ref/Outer_Syntax.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/Outer_Syntax.thy Thu Nov 05 08:27:22 2015 +0100
@@ -223,17 +223,18 @@
subsection \<open>Comments \label{sec:comments}\<close>
-text \<open>Large chunks of plain @{syntax text} are usually given @{syntax
- verbatim}, i.e.\ enclosed in \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>,
- or as @{syntax cartouche} \<open>\<open>\<dots>\<close>\<close>. For convenience, any of the
- smaller text units conforming to @{syntax nameref} are admitted as well. A
- marginal @{syntax comment} is of the form \<^verbatim>\<open>--\<close>~@{syntax text}.
- Any number of these may occur within Isabelle/Isar commands.
+text \<open>
+ Large chunks of plain @{syntax text} are usually given @{syntax verbatim},
+ i.e.\ enclosed in \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>, or as @{syntax cartouche} \<open>\<open>\<dots>\<close>\<close>. For
+ convenience, any of the smaller text units conforming to @{syntax nameref}
+ are admitted as well. A marginal @{syntax comment} is of the form
+ \<^verbatim>\<open>--\<close>~@{syntax text} or \<^verbatim>\<open>\<comment>\<close>~@{syntax text}. Any number of these may occur
+ within Isabelle/Isar commands.
@{rail \<open>
@{syntax_def text}: @{syntax verbatim} | @{syntax cartouche} | @{syntax nameref}
;
- @{syntax_def comment}: '--' @{syntax text}
+ @{syntax_def comment}: ('--' | @'\<comment>') @{syntax text}
\<close>}
\<close>
--- a/src/Doc/Isar_Ref/Proof.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/Proof.thy Thu Nov 05 08:27:22 2015 +0100
@@ -530,8 +530,8 @@
\<^medskip>
The Isar calculation proof commands may be defined as
- follows:\footnote{We suppress internal bookkeeping such as proper
- handling of block-structure.}
+ follows:\<^footnote>\<open>We suppress internal bookkeeping such as proper
+ handling of block-structure.\<close>
\begin{matharray}{rcl}
@{command "also"}\<open>\<^sub>0\<close> & \equiv & @{command "note"}~\<open>calculation = this\<close> \\
@@ -718,9 +718,9 @@
If the goal had been \<open>show\<close> (or \<open>thus\<close>), some
pending sub-goal is solved as well by the rule resulting from the
result \<^emph>\<open>exported\<close> into the enclosing goal context. Thus \<open>qed\<close> may fail for two reasons: either \<open>m\<^sub>2\<close> fails, or the
- resulting rule does not fit to any pending goal\footnote{This
+ resulting rule does not fit to any pending goal\<^footnote>\<open>This
includes any additional ``strong'' assumptions as introduced by
- @{command "assume"}.} of the enclosing context. Debugging such a
+ @{command "assume"}.\<close> of the enclosing context. Debugging such a
situation might involve temporarily changing @{command "show"} into
@{command "have"}, or weakening the local context by replacing
occurrences of @{command "assume"} by @{command "presume"}.
--- a/src/Doc/Isar_Ref/Synopsis.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/Isar_Ref/Synopsis.thy Thu Nov 05 08:27:22 2015 +0100
@@ -17,11 +17,11 @@
notepad
begin
txt \<open>Locally fixed entities:\<close>
- fix x -- \<open>local constant, without any type information yet\<close>
- fix x :: 'a -- \<open>variant with explicit type-constraint for subsequent use\<close>
+ fix x \<comment> \<open>local constant, without any type information yet\<close>
+ fix x :: 'a \<comment> \<open>variant with explicit type-constraint for subsequent use\<close>
fix a b
- assume "a = b" -- \<open>type assignment at first occurrence in concrete term\<close>
+ assume "a = b" \<comment> \<open>type assignment at first occurrence in concrete term\<close>
txt \<open>Definitions (non-polymorphic):\<close>
def x \<equiv> "t::'a"
@@ -234,7 +234,7 @@
proof -
term ?thesis
show ?thesis sorry
- term ?thesis -- \<open>static!\<close>
+ term ?thesis \<comment> \<open>static!\<close>
qed
term "\<dots>"
thm this
@@ -345,7 +345,7 @@
moreover
{ assume C have R sorry }
ultimately
- have R by blast -- \<open>``big-bang integration'' of proof blocks (occasionally fragile)\<close>
+ have R by blast \<comment> \<open>``big-bang integration'' of proof blocks (occasionally fragile)\<close>
end
@@ -364,7 +364,7 @@
begin
fix n :: nat
have "P n"
- proof (rule nat.induct) -- \<open>fragile rule application!\<close>
+ proof (rule nat.induct) \<comment> \<open>fragile rule application!\<close>
show "P 0" sorry
next
fix n :: nat
@@ -503,7 +503,7 @@
from \<open>A x 0\<close> show "Q x 0" sorry
next
case (Suc n)
- from \<open>\<And>x. A x n \<Longrightarrow> Q x n\<close> -- \<open>arbitrary instances can be produced here\<close>
+ from \<open>\<And>x. A x n \<Longrightarrow> Q x n\<close> \<comment> \<open>arbitrary instances can be produced here\<close>
and \<open>A x (Suc n)\<close> show "Q x (Suc n)" sorry
qed
end
@@ -675,9 +675,9 @@
begin
assume a: A and b: B
thm conjI
- thm conjI [of A B] -- "instantiation"
- thm conjI [of A B, OF a b] -- "instantiation and composition"
- thm conjI [OF a b] -- "composition via unification (trivial)"
+ thm conjI [of A B] \<comment> "instantiation"
+ thm conjI [of A B, OF a b] \<comment> "instantiation and composition"
+ thm conjI [OF a b] \<comment> "composition via unification (trivial)"
thm conjI [OF \<open>A\<close> \<open>B\<close>]
thm conjI [OF disjI1]
@@ -710,9 +710,9 @@
fix x
assume "A x"
show "B x" sorry
- } -- "implicit block structure made explicit"
+ } \<comment> "implicit block structure made explicit"
note \<open>\<And>x. A x \<Longrightarrow> B x\<close>
- -- "side exit for the resulting rule"
+ \<comment> "side exit for the resulting rule"
qed
end
@@ -726,12 +726,12 @@
notepad
begin
- assume r1: "A \<Longrightarrow> B \<Longrightarrow> C" -- \<open>simple rule (Horn clause)\<close>
+ assume r1: "A \<Longrightarrow> B \<Longrightarrow> C" \<comment> \<open>simple rule (Horn clause)\<close>
- have A sorry -- "prefix of facts via outer sub-proof"
+ have A sorry \<comment> "prefix of facts via outer sub-proof"
then have C
proof (rule r1)
- show B sorry -- "remaining rule premises via inner sub-proof"
+ show B sorry \<comment> "remaining rule premises via inner sub-proof"
qed
have C
@@ -750,7 +750,7 @@
next
- assume r2: "A \<Longrightarrow> (\<And>x. B1 x \<Longrightarrow> B2 x) \<Longrightarrow> C" -- \<open>nested rule\<close>
+ assume r2: "A \<Longrightarrow> (\<And>x. B1 x \<Longrightarrow> B2 x) \<Longrightarrow> C" \<comment> \<open>nested rule\<close>
have A sorry
then have C
@@ -850,31 +850,31 @@
notepad
begin
have "A \<and> B"
- proof -- \<open>two strictly isolated subproofs\<close>
+ proof \<comment> \<open>two strictly isolated subproofs\<close>
show A sorry
next
show B sorry
qed
have "A \<and> B"
- proof -- \<open>one simultaneous sub-proof\<close>
+ proof \<comment> \<open>one simultaneous sub-proof\<close>
show A and B sorry
qed
have "A \<and> B"
- proof -- \<open>two subproofs in the same context\<close>
+ proof \<comment> \<open>two subproofs in the same context\<close>
show A sorry
show B sorry
qed
have "A \<and> B"
- proof -- \<open>swapped order\<close>
+ proof \<comment> \<open>swapped order\<close>
show B sorry
show A sorry
qed
have "A \<and> B"
- proof -- \<open>sequential subproofs\<close>
+ proof \<comment> \<open>sequential subproofs\<close>
show A sorry
show B using \<open>A\<close> sorry
qed
@@ -941,9 +941,9 @@
following typical representatives:
\<close>
-thm exE -- \<open>local parameter\<close>
-thm conjE -- \<open>local premises\<close>
-thm disjE -- \<open>split into cases\<close>
+thm exE \<comment> \<open>local parameter\<close>
+thm conjE \<comment> \<open>local premises\<close>
+thm disjE \<comment> \<open>split into cases\<close>
text \<open>
Combining these characteristics leads to the following general scheme
@@ -1001,7 +1001,7 @@
print_statement disjE
lemma
- assumes A1 and A2 -- \<open>assumptions\<close>
+ assumes A1 and A2 \<comment> \<open>assumptions\<close>
obtains
(case1) x y where "B1 x y" and "C1 x y"
| (case2) x y where "B2 x y" and "C2 x y"
--- a/src/Doc/JEdit/JEdit.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/JEdit/JEdit.thy Thu Nov 05 08:27:22 2015 +0100
@@ -9,52 +9,50 @@
section \<open>Concepts and terminology\<close>
text \<open>
- Isabelle/jEdit is a Prover IDE that integrates \<^emph>\<open>parallel proof
- checking\<close> @{cite "Wenzel:2009" and "Wenzel:2013:ITP"} with
- \<^emph>\<open>asynchronous user interaction\<close> @{cite "Wenzel:2010" and
- "Wenzel:2012:UITP-EPTCS" and "Wenzel:2014:ITP-PIDE" and "Wenzel:2014:UITP"},
- based on a document-oriented approach to \<^emph>\<open>continuous proof processing\<close>
- @{cite "Wenzel:2011:CICM" and "Wenzel:2012"}. Many concepts and system
- components are fit together in order to make this work. The main building
- blocks are as follows.
+ Isabelle/jEdit is a Prover IDE that integrates \<^emph>\<open>parallel proof checking\<close>
+ @{cite "Wenzel:2009" and "Wenzel:2013:ITP"} with \<^emph>\<open>asynchronous user
+ interaction\<close> @{cite "Wenzel:2010" and "Wenzel:2012:UITP-EPTCS" and
+ "Wenzel:2014:ITP-PIDE" and "Wenzel:2014:UITP"}, based on a document-oriented
+ approach to \<^emph>\<open>continuous proof processing\<close> @{cite "Wenzel:2011:CICM" and
+ "Wenzel:2012"}. Many concepts and system components are fit together in
+ order to make this work. The main building blocks are as follows.
- \<^descr>[PIDE] is a general framework for Prover IDEs based on Isabelle/Scala.
- It is built around a concept of parallel and asynchronous document
- processing, which is supported natively by the parallel proof engine that is
- implemented in Isabelle/ML. The traditional prover command loop is given up;
- instead there is direct support for editing of source text, with rich formal
- markup for GUI rendering.
+ \<^descr>[PIDE] is a general framework for Prover IDEs based on Isabelle/Scala. It
+ is built around a concept of parallel and asynchronous document
+ processing, which is supported natively by the parallel proof engine that
+ is implemented in Isabelle/ML. The traditional prover command loop is
+ given up; instead there is direct support for editing of source text, with
+ rich formal markup for GUI rendering.
- \<^descr>[Isabelle/ML] is the implementation and extension language of
- Isabelle, see also @{cite "isabelle-implementation"}. It is integrated
- into the logical context of Isabelle/Isar and allows to manipulate
- logical entities directly. Arbitrary add-on tools may be implemented
- for object-logics such as Isabelle/HOL.
+ \<^descr>[Isabelle/ML] is the implementation and extension language of Isabelle,
+ see also @{cite "isabelle-implementation"}. It is integrated into the
+ logical context of Isabelle/Isar and allows to manipulate logical entities
+ directly. Arbitrary add-on tools may be implemented for object-logics such
+ as Isabelle/HOL.
- \<^descr>[Isabelle/Scala] is the system programming language of
- Isabelle. It extends the pure logical environment of Isabelle/ML
- towards the ``real world'' of graphical user interfaces, text
- editors, IDE frameworks, web services etc. Special infrastructure
- allows to transfer algebraic datatypes and formatted text easily
- between ML and Scala, using asynchronous protocol commands.
+ \<^descr>[Isabelle/Scala] is the system programming language of Isabelle. It
+ extends the pure logical environment of Isabelle/ML towards the ``real
+ world'' of graphical user interfaces, text editors, IDE frameworks, web
+ services etc. Special infrastructure allows to transfer algebraic
+ datatypes and formatted text easily between ML and Scala, using
+ asynchronous protocol commands.
- \<^descr>[jEdit] is a sophisticated text editor implemented in
- Java.\footnote{@{url "http://www.jedit.org"}} It is easily extensible
- by plugins written in languages that work on the JVM, e.g.\
- Scala\footnote{@{url "http://www.scala-lang.org"}}.
+ \<^descr>[jEdit] is a sophisticated text editor implemented in Java.\<^footnote>\<open>@{url
+ "http://www.jedit.org"}\<close> It is easily extensible by plugins written in
+ languages that work on the JVM, e.g.\ Scala\<^footnote>\<open>@{url
+ "http://www.scala-lang.org"}\<close>.
- \<^descr>[Isabelle/jEdit] is the main example application of the PIDE
- framework and the default user-interface for Isabelle. It targets
- both beginners and experts. Technically, Isabelle/jEdit combines a
- slightly modified version of the jEdit code base with a special
- plugin for Isabelle, integrated as standalone application for the
- main operating system platforms: Linux, Windows, Mac OS X.
+ \<^descr>[Isabelle/jEdit] is the main example application of the PIDE framework
+ and the default user-interface for Isabelle. It targets both beginners and
+ experts. Technically, Isabelle/jEdit combines a slightly modified version
+ of the jEdit code base with a special plugin for Isabelle, integrated as
+ standalone application for the main operating system platforms: Linux,
+ Windows, Mac OS X.
-
- The subtle differences of Isabelle/ML versus Standard ML,
- Isabelle/Scala versus Scala, Isabelle/jEdit versus jEdit need to be
- taken into account when discussing any of these PIDE building blocks
- in public forums, mailing lists, or even scientific publications.
+ The subtle differences of Isabelle/ML versus Standard ML, Isabelle/Scala
+ versus Scala, Isabelle/jEdit versus jEdit need to be taken into account when
+ discussing any of these PIDE building blocks in public forums, mailing
+ lists, or even scientific publications.
\<close>
@@ -73,31 +71,31 @@
the jEdit text editor, while preserving its general look-and-feel as far as
possible. The main plugin is called ``Isabelle'' and has its own menu
\<^emph>\<open>Plugins~/ Isabelle\<close> with access to several panels (see also
- \secref{sec:dockables}), as well as \<^emph>\<open>Plugins~/ Plugin Options~/
- Isabelle\<close> (see also \secref{sec:options}).
+ \secref{sec:dockables}), as well as \<^emph>\<open>Plugins~/ Plugin Options~/ Isabelle\<close>
+ (see also \secref{sec:options}).
The options allow to specify a logic session name --- the same selector is
- accessible in the \<^emph>\<open>Theories\<close> panel (\secref{sec:theories}). On
- application startup, the selected logic session image is provided
- automatically by the Isabelle build tool @{cite "isabelle-system"}: if it is
- absent or outdated wrt.\ its sources, the build process updates it before
- entering the Prover IDE. Change of the logic session within Isabelle/jEdit
- requires a restart of the whole application.
+ accessible in the \<^emph>\<open>Theories\<close> panel (\secref{sec:theories}). On application
+ startup, the selected logic session image is provided automatically by the
+ Isabelle build tool @{cite "isabelle-system"}: if it is absent or outdated
+ wrt.\ its sources, the build process updates it before entering the Prover
+ IDE. Change of the logic session within Isabelle/jEdit requires a restart of
+ the whole application.
\<^medskip>
- The main job of the Prover IDE is to manage sources and their
- changes, taking the logical structure as a formal document into account (see
- also \secref{sec:document-model}). The editor and the prover are connected
+ The main job of the Prover IDE is to manage sources and their changes,
+ taking the logical structure as a formal document into account (see also
+ \secref{sec:document-model}). The editor and the prover are connected
asynchronously in a lock-free manner. The prover is free to organize the
checking of the formal text in parallel on multiple cores, and provides
feedback via markup, which is rendered in the editor via colors, boxes,
squiggly underlines, hyperlinks, popup windows, icons, clickable output etc.
- Using the mouse together with the modifier key \<^verbatim>\<open>CONTROL\<close> (Linux,
- Windows) or \<^verbatim>\<open>COMMAND\<close> (Mac OS X) exposes additional formal content
- via tooltips and/or hyperlinks (see also \secref{sec:tooltips-hyperlinks}).
- Output (in popups etc.) may be explored recursively, using the same
- techniques as in the editor source buffer.
+ Using the mouse together with the modifier key \<^verbatim>\<open>CONTROL\<close> (Linux, Windows)
+ or \<^verbatim>\<open>COMMAND\<close> (Mac OS X) exposes additional formal content via tooltips
+ and/or hyperlinks (see also \secref{sec:tooltips-hyperlinks}). Output (in
+ popups etc.) may be explored recursively, using the same techniques as in
+ the editor source buffer.
Thus the Prover IDE gives an impression of direct access to formal content
of the prover within the editor, but in reality only certain aspects are
@@ -109,14 +107,13 @@
subsection \<open>Documentation\<close>
text \<open>
- The \<^emph>\<open>Documentation\<close> panel of Isabelle/jEdit provides access to the
- standard Isabelle documentation: PDF files are opened by regular desktop
- operations of the underlying platform. The section ``Original jEdit
- Documentation'' contains the original \<^emph>\<open>User's Guide\<close> of this
- sophisticated text editor. The same is accessible via the \<^verbatim>\<open>Help\<close>
- menu or \<^verbatim>\<open>F1\<close> keyboard shortcut, using the built-in HTML viewer of
- Java/Swing. The latter also includes \<^emph>\<open>Frequently Asked Questions\<close> and
- documentation of individual plugins.
+ The \<^emph>\<open>Documentation\<close> panel of Isabelle/jEdit provides access to the standard
+ Isabelle documentation: PDF files are opened by regular desktop operations
+ of the underlying platform. The section ``Original jEdit Documentation''
+ contains the original \<^emph>\<open>User's Guide\<close> of this sophisticated text editor. The
+ same is accessible via the \<^verbatim>\<open>Help\<close> menu or \<^verbatim>\<open>F1\<close> keyboard shortcut, using
+ the built-in HTML viewer of Java/Swing. The latter also includes
+ \<^emph>\<open>Frequently Asked Questions\<close> and documentation of individual plugins.
Most of the information about generic jEdit is relevant for Isabelle/jEdit
as well, but one needs to keep in mind that defaults sometimes differ, and
@@ -129,9 +126,9 @@
subsection \<open>Plugins\<close>
text \<open>
- The \<^emph>\<open>Plugin Manager\<close> of jEdit allows to augment editor functionality by
- JVM modules (jars) that are provided by the central plugin repository, which
- is accessible via various mirror sites.
+ The \<^emph>\<open>Plugin Manager\<close> of jEdit allows to augment editor functionality by JVM
+ modules (jars) that are provided by the central plugin repository, which is
+ accessible via various mirror sites.
Connecting to the plugin server-infrastructure of the jEdit project allows
to update bundled plugins or to add further functionality. This needs to be
@@ -142,28 +139,27 @@
at a grand scale.
\<^medskip>
- The main \<^emph>\<open>Isabelle\<close> plugin is an integral part of
- Isabelle/jEdit and needs to remain active at all times! A few additional
- plugins are bundled with Isabelle/jEdit for convenience or out of necessity,
- notably \<^emph>\<open>Console\<close> with its Isabelle/Scala sub-plugin
- (\secref{sec:scala-console}) and \<^emph>\<open>SideKick\<close> with some Isabelle-specific
- parsers for document tree structure (\secref{sec:sidekick}). The
- \<^emph>\<open>Navigator\<close> plugin is particularly important for hyperlinks within the
- formal document-model (\secref{sec:tooltips-hyperlinks}). Further plugins
- (e.g.\ \<^emph>\<open>ErrorList\<close>, \<^emph>\<open>Code2HTML\<close>) are included to saturate the
- dependencies of bundled plugins, but have no particular use in
- Isabelle/jEdit.
+ The main \<^emph>\<open>Isabelle\<close> plugin is an integral part of Isabelle/jEdit and needs
+ to remain active at all times! A few additional plugins are bundled with
+ Isabelle/jEdit for convenience or out of necessity, notably \<^emph>\<open>Console\<close> with
+ its Isabelle/Scala sub-plugin (\secref{sec:scala-console}) and \<^emph>\<open>SideKick\<close>
+ with some Isabelle-specific parsers for document tree structure
+ (\secref{sec:sidekick}). The \<^emph>\<open>Navigator\<close> plugin is particularly important
+ for hyperlinks within the formal document-model
+ (\secref{sec:tooltips-hyperlinks}). Further plugins (e.g.\ \<^emph>\<open>ErrorList\<close>,
+ \<^emph>\<open>Code2HTML\<close>) are included to saturate the dependencies of bundled plugins,
+ but have no particular use in Isabelle/jEdit.
\<close>
subsection \<open>Options \label{sec:options}\<close>
-text \<open>Both jEdit and Isabelle have distinctive management of
- persistent options.
+text \<open>
+ Both jEdit and Isabelle have distinctive management of persistent options.
- Regular jEdit options are accessible via the dialogs \<^emph>\<open>Utilities~/
- Global Options\<close> or \<^emph>\<open>Plugins~/ Plugin Options\<close>, with a second chance to
- flip the two within the central options dialog. Changes are stored in
+ Regular jEdit options are accessible via the dialogs \<^emph>\<open>Utilities~/ Global
+ Options\<close> or \<^emph>\<open>Plugins~/ Plugin Options\<close>, with a second chance to flip the
+ two within the central options dialog. Changes are stored in
@{file_unchecked "$ISABELLE_HOME_USER/jedit/properties"} and
@{file_unchecked "$ISABELLE_HOME_USER/jedit/keymaps"}.
@@ -173,17 +169,17 @@
coverage of sessions and command-line tools like @{tool build} or @{tool
options}.
- Those Isabelle options that are declared as \<^bold>\<open>public\<close> are configurable
- in Isabelle/jEdit via \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>. Moreover,
- there are various options for rendering of document content, which are
- configurable via \<^emph>\<open>Plugin Options~/ Isabelle~/ Rendering\<close>. Thus
- \<^emph>\<open>Plugin Options~/ Isabelle\<close> in jEdit provides a view on a subset of
- Isabelle system options. Note that some of these options affect general
- parameters that are relevant outside Isabelle/jEdit as well, e.g.\
- @{system_option threads} or @{system_option parallel_proofs} for the
- Isabelle build tool @{cite "isabelle-system"}, but it is possible to use the
- settings variable @{setting ISABELLE_BUILD_OPTIONS} to change defaults for
- batch builds without affecting Isabelle/jEdit.
+ Those Isabelle options that are declared as \<^bold>\<open>public\<close> are configurable in
+ Isabelle/jEdit via \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>. Moreover, there
+ are various options for rendering of document content, which are
+ configurable via \<^emph>\<open>Plugin Options~/ Isabelle~/ Rendering\<close>. Thus \<^emph>\<open>Plugin
+ Options~/ Isabelle\<close> in jEdit provides a view on a subset of Isabelle system
+ options. Note that some of these options affect general parameters that are
+ relevant outside Isabelle/jEdit as well, e.g.\ @{system_option threads} or
+ @{system_option parallel_proofs} for the Isabelle build tool @{cite
+ "isabelle-system"}, but it is possible to use the settings variable
+ @{setting ISABELLE_BUILD_OPTIONS} to change defaults for batch builds
+ without affecting Isabelle/jEdit.
The jEdit action @{action_def isabelle.options} opens the options dialog for
the Isabelle plugin; it can be mapped to editor GUI elements as usual.
@@ -193,24 +189,25 @@
Isabelle/jEdit. Editing the machine-generated @{file_unchecked
"$ISABELLE_HOME_USER/jedit/properties"} or @{file_unchecked
"$ISABELLE_HOME_USER/etc/preferences"} manually while the application is
- running is likely to cause surprise due to lost update!\<close>
+ running is likely to cause surprise due to lost update!
+\<close>
subsection \<open>Keymaps\<close>
-text \<open>Keyboard shortcuts used to be managed as jEdit properties in
- the past, but recent versions (2013) have a separate concept of
- \<^emph>\<open>keymap\<close> that is configurable via \<^emph>\<open>Global Options~/
- Shortcuts\<close>. The \<^verbatim>\<open>imported\<close> keymap is derived from the
- initial environment of properties that is available at the first
- start of the editor; afterwards the keymap file takes precedence.
+text \<open>
+ Keyboard shortcuts used to be managed as jEdit properties in the past, but
+ recent versions (2013) have a separate concept of \<^emph>\<open>keymap\<close> that is
+ configurable via \<^emph>\<open>Global Options~/ Shortcuts\<close>. The \<^verbatim>\<open>imported\<close> keymap is
+ derived from the initial environment of properties that is available at the
+ first start of the editor; afterwards the keymap file takes precedence.
This is relevant for Isabelle/jEdit due to various fine-tuning of default
properties, and additional keyboard shortcuts for Isabelle-specific
functionality. Users may change their keymap later, but need to copy some
keyboard shortcuts manually (see also @{file_unchecked
- "$ISABELLE_HOME_USER/jedit/keymaps"} versus \<^verbatim>\<open>shortcut\<close> properties
- in @{file "$ISABELLE_HOME/src/Tools/jEdit/src/jEdit.props"}).
+ "$ISABELLE_HOME_USER/jedit/keymaps"} versus \<^verbatim>\<open>shortcut\<close> properties in @{file
+ "$ISABELLE_HOME/src/Tools/jEdit/src/jEdit.props"}).
\<close>
@@ -239,33 +236,32 @@
Start jEdit with Isabelle plugin setup and open FILES
(default "$USER_HOME/Scratch.thy" or ":" for empty buffer).\<close>}
- The \<^verbatim>\<open>-l\<close> option specifies the session name of the logic
- image to be used for proof processing. Additional session root
- directories may be included via option \<^verbatim>\<open>-d\<close> to augment
- that name space of @{tool build} @{cite "isabelle-system"}.
+ The \<^verbatim>\<open>-l\<close> option specifies the session name of the logic image to be used
+ for proof processing. Additional session root directories may be included
+ via option \<^verbatim>\<open>-d\<close> to augment that name space of @{tool build} @{cite
+ "isabelle-system"}.
- By default, the specified image is checked and built on demand. The
- \<^verbatim>\<open>-s\<close> option determines where to store the result session image
- of @{tool build}. The \<^verbatim>\<open>-n\<close> option bypasses the implicit build
- process for the selected session image.
+ By default, the specified image is checked and built on demand. The \<^verbatim>\<open>-s\<close>
+ option determines where to store the result session image of @{tool build}.
+ The \<^verbatim>\<open>-n\<close> option bypasses the implicit build process for the selected
+ session image.
- The \<^verbatim>\<open>-m\<close> option specifies additional print modes for the prover
- process. Note that the system option @{system_option_ref jedit_print_mode}
- allows to do the same persistently (e.g.\ via the \<^emph>\<open>Plugin Options\<close>
- dialog of Isabelle/jEdit), without requiring command-line invocation.
+ The \<^verbatim>\<open>-m\<close> option specifies additional print modes for the prover process.
+ Note that the system option @{system_option_ref jedit_print_mode} allows to
+ do the same persistently (e.g.\ via the \<^emph>\<open>Plugin Options\<close> dialog of
+ Isabelle/jEdit), without requiring command-line invocation.
- The \<^verbatim>\<open>-J\<close> and \<^verbatim>\<open>-j\<close> options allow to pass additional
- low-level options to the JVM or jEdit, respectively. The defaults are
- provided by the Isabelle settings environment @{cite "isabelle-system"}, but
- note that these only work for the command-line tool described here, and not
- the regular application.
+ The \<^verbatim>\<open>-J\<close> and \<^verbatim>\<open>-j\<close> options allow to pass additional low-level options to
+ the JVM or jEdit, respectively. The defaults are provided by the Isabelle
+ settings environment @{cite "isabelle-system"}, but note that these only
+ work for the command-line tool described here, and not the regular
+ application.
- The \<^verbatim>\<open>-b\<close> and \<^verbatim>\<open>-f\<close> options control the self-build
- mechanism of Isabelle/jEdit. This is only relevant for building from
- sources, which also requires an auxiliary \<^verbatim>\<open>jedit_build\<close> component
- from @{url "http://isabelle.in.tum.de/components"}. The official
- Isabelle release already includes a pre-built version of Isabelle/jEdit.
-\<close>
+ The \<^verbatim>\<open>-b\<close> and \<^verbatim>\<open>-f\<close> options control the self-build mechanism of
+ Isabelle/jEdit. This is only relevant for building from sources, which also
+ requires an auxiliary \<^verbatim>\<open>jedit_build\<close> component from @{url
+ "http://isabelle.in.tum.de/components"}. The official Isabelle release
+ already includes a pre-built version of Isabelle/jEdit. \<close>
chapter \<open>Augmented jEdit functionality\<close>
@@ -283,39 +279,35 @@
Isabelle/jEdit enables platform-specific look-and-feel by default as
follows.
- \<^descr>[Linux:] The platform-independent \<^emph>\<open>Metal\<close> is used by
- default.
+ \<^descr>[Linux:] The platform-independent \<^emph>\<open>Metal\<close> is used by default.
- \<^emph>\<open>GTK+\<close> also works under the side-condition that the overall GTK theme
- is selected in a Swing-friendly way.\footnote{GTK support in Java/Swing was
- once marketed aggressively by Sun, but never quite finished. Today (2015) it
- is lagging behind further development of Swing and GTK. The graphics
- rendering performance can be worse than for other Swing look-and-feels.
- Nonetheless it has its uses for displays with very high resolution (such as
- ``4K'' or ``UHD'' models), because the rendering by the external library is
- subject to global system settings for font scaling.}
+ \<^emph>\<open>GTK+\<close> also works under the side-condition that the overall GTK theme is
+ selected in a Swing-friendly way.\<^footnote>\<open>GTK support in Java/Swing was once
+ marketed aggressively by Sun, but never quite finished. Today (2015) it is
+ lagging behind further development of Swing and GTK. The graphics
+ rendering performance can be worse than for other Swing look-and-feels.
+ Nonetheless it has its uses for displays with very high resolution (such
+ as ``4K'' or ``UHD'' models), because the rendering by the external
+ library is subject to global system settings for font scaling.\<close>
- \<^descr>[Windows:] Regular \<^emph>\<open>Windows\<close> is used by default, but
- \<^emph>\<open>Windows Classic\<close> also works.
-
- \<^descr>[Mac OS X:] Regular \<^emph>\<open>Mac OS X\<close> is used by default.
+ \<^descr>[Windows:] Regular \<^emph>\<open>Windows\<close> is used by default, but \<^emph>\<open>Windows Classic\<close>
+ also works.
- The bundled \<^emph>\<open>MacOSX\<close> plugin provides various functions that are
- expected from applications on that particular platform: quit from menu or
- dock, preferences menu, drag-and-drop of text files on the application,
- full-screen mode for main editor windows. It is advisable to have the
- \<^emph>\<open>MacOSX\<close> plugin enabled all the time on that platform.
+ \<^descr>[Mac OS X:] Regular \<^emph>\<open>Mac OS X\<close> is used by default.
+ The bundled \<^emph>\<open>MacOSX\<close> plugin provides various functions that are expected
+ from applications on that particular platform: quit from menu or dock,
+ preferences menu, drag-and-drop of text files on the application,
+ full-screen mode for main editor windows. It is advisable to have the
+ \<^emph>\<open>MacOSX\<close> plugin enabled all the time on that platform.
- Users may experiment with different look-and-feels, but need to keep
- in mind that this extra variance of GUI functionality is unlikely to
- work in arbitrary combinations. The platform-independent
- \<^emph>\<open>Metal\<close> and \<^emph>\<open>Nimbus\<close> should always work. The historic
- \<^emph>\<open>CDE/Motif\<close> should be ignored.
+ Users may experiment with different look-and-feels, but need to keep in mind
+ that this extra variance of GUI functionality is unlikely to work in
+ arbitrary combinations. The platform-independent \<^emph>\<open>Metal\<close> and \<^emph>\<open>Nimbus\<close>
+ should always work. The historic \<^emph>\<open>CDE/Motif\<close> should be ignored.
- After changing the look-and-feel in \<^emph>\<open>Global Options~/
- Appearance\<close>, it is advisable to restart Isabelle/jEdit in order to
- take full effect.
+ After changing the look-and-feel in \<^emph>\<open>Global Options~/ Appearance\<close>, it is
+ advisable to restart Isabelle/jEdit in order to take full effect.
\<close>
@@ -326,49 +318,47 @@
were considered ``high resolution'' and bitmap fonts with 12 or 14 pixels as
adequate for text rendering. Today (2015), we routinely see ``Full HD''
monitors at $1920 \times 1080$ pixels, and occasionally ``Ultra HD'' at
- $3840 \times 2160$ or more, but GUI rendering did not really progress
- beyond the old standards.
+ $3840 \times 2160$ or more, but GUI rendering did not really progress beyond
+ the old standards.
Isabelle/jEdit defaults are a compromise for reasonable out-of-the box
results on common platforms and medium resolution displays (e.g.\ the ``Full
HD'' category). Subsequently there are further hints to improve on that.
\<^medskip>
- The \<^bold>\<open>operating-system platform\<close> usually provides some
- configuration for global scaling of text fonts, e.g.\ $120\%$--$250\%$ on
- Windows. Changing that only has a partial effect on GUI rendering;
- satisfactory display quality requires further adjustments.
+ The \<^bold>\<open>operating-system platform\<close> usually provides some configuration for
+ global scaling of text fonts, e.g.\ $120\%$--$250\%$ on Windows. Changing
+ that only has a partial effect on GUI rendering; satisfactory display
+ quality requires further adjustments.
\<^medskip>
- The Isabelle/jEdit \<^bold>\<open>application\<close> and its plugins provide
- various font properties that are summarized below.
+ The Isabelle/jEdit \<^bold>\<open>application\<close> and its plugins provide various font
+ properties that are summarized below.
- \<^item> \<^emph>\<open>Global Options / Text Area / Text font\<close>: the main text area
- font, which is also used as reference point for various derived font sizes,
- e.g.\ the Output panel (\secref{sec:output}).
+ \<^item> \<^emph>\<open>Global Options / Text Area / Text font\<close>: the main text area font,
+ which is also used as reference point for various derived font sizes,
+ e.g.\ the Output panel (\secref{sec:output}).
- \<^item> \<^emph>\<open>Global Options / Gutter / Gutter font\<close>: the font for the gutter
- area left of the main text area, e.g.\ relevant for display of line numbers
- (disabled by default).
+ \<^item> \<^emph>\<open>Global Options / Gutter / Gutter font\<close>: the font for the gutter area
+ left of the main text area, e.g.\ relevant for display of line numbers
+ (disabled by default).
- \<^item> \<^emph>\<open>Global Options / Appearance / Button, menu and label font\<close> as
- well as \<^emph>\<open>List and text field font\<close>: this specifies the primary and
- secondary font for the traditional \<^emph>\<open>Metal\<close> look-and-feel
- (\secref{sec:look-and-feel}), which happens to scale better than newer ones
- like \<^emph>\<open>Nimbus\<close>.
+ \<^item> \<^emph>\<open>Global Options / Appearance / Button, menu and label font\<close> as well as
+ \<^emph>\<open>List and text field font\<close>: this specifies the primary and secondary font
+ for the traditional \<^emph>\<open>Metal\<close> look-and-feel (\secref{sec:look-and-feel}),
+ which happens to scale better than newer ones like \<^emph>\<open>Nimbus\<close>.
- \<^item> \<^emph>\<open>Plugin Options / Isabelle / General / Reset Font Size\<close>: the main
- text area font size for action @{action_ref "isabelle.reset-font-size"},
- e.g.\ relevant for quick scaling like in major web browsers.
+ \<^item> \<^emph>\<open>Plugin Options / Isabelle / General / Reset Font Size\<close>: the main text
+ area font size for action @{action_ref "isabelle.reset-font-size"}, e.g.\
+ relevant for quick scaling like in major web browsers.
- \<^item> \<^emph>\<open>Plugin Options / Console / General / Font\<close>: the console window
- font, e.g.\ relevant for Isabelle/Scala command-line.
-
+ \<^item> \<^emph>\<open>Plugin Options / Console / General / Font\<close>: the console window font,
+ e.g.\ relevant for Isabelle/Scala command-line.
- In \figref{fig:isabelle-jedit-hdpi} the \<^emph>\<open>Metal\<close> look-and-feel is
- configured with custom fonts at 30 pixels, and the main text area and
- console at 36 pixels. Despite the old-fashioned appearance of \<^emph>\<open>Metal\<close>,
- this leads to decent rendering quality on all platforms.
+ In \figref{fig:isabelle-jedit-hdpi} the \<^emph>\<open>Metal\<close> look-and-feel is configured
+ with custom fonts at 30 pixels, and the main text area and console at 36
+ pixels. Despite the old-fashioned appearance of \<^emph>\<open>Metal\<close>, this leads to
+ decent rendering quality on all platforms.
\begin{figure}[htb]
\begin{center}
@@ -387,25 +377,25 @@
section \<open>Dockable windows \label{sec:dockables}\<close>
text \<open>
- In jEdit terminology, a \<^emph>\<open>view\<close> is an editor window with one or more
- \<^emph>\<open>text areas\<close> that show the content of one or more \<^emph>\<open>buffers\<close>. A
- regular view may be surrounded by \<^emph>\<open>dockable windows\<close> that show
- additional information in arbitrary format, not just text; a \<^emph>\<open>plain
- view\<close> does not allow dockables. The \<^emph>\<open>dockable window manager\<close> of jEdit
- organizes these dockable windows, either as \<^emph>\<open>floating\<close> windows, or
- \<^emph>\<open>docked\<close> panels within one of the four margins of the view. There may
- be any number of floating instances of some dockable window, but at most one
- docked instance; jEdit actions that address \<^emph>\<open>the\<close> dockable window of a
- particular kind refer to the unique docked instance.
+ In jEdit terminology, a \<^emph>\<open>view\<close> is an editor window with one or more \<^emph>\<open>text
+ areas\<close> that show the content of one or more \<^emph>\<open>buffers\<close>. A regular view may
+ be surrounded by \<^emph>\<open>dockable windows\<close> that show additional information in
+ arbitrary format, not just text; a \<^emph>\<open>plain view\<close> does not allow dockables.
+ The \<^emph>\<open>dockable window manager\<close> of jEdit organizes these dockable windows,
+ either as \<^emph>\<open>floating\<close> windows, or \<^emph>\<open>docked\<close> panels within one of the four
+ margins of the view. There may be any number of floating instances of some
+ dockable window, but at most one docked instance; jEdit actions that address
+ \<^emph>\<open>the\<close> dockable window of a particular kind refer to the unique docked
+ instance.
Dockables are used routinely in jEdit for important functionality like
- \<^emph>\<open>HyperSearch Results\<close> or the \<^emph>\<open>File System Browser\<close>. Plugins often
- provide a central dockable to access their key functionality, which may be
- opened by the user on demand. The Isabelle/jEdit plugin takes this approach
- to the extreme: its plugin menu provides the entry-points to many panels
- that are managed as dockable windows. Some important panels are docked by
- default, e.g.\ \<^emph>\<open>Documentation\<close>, \<^emph>\<open>Output\<close>, \<^emph>\<open>Query\<close>, but the
- user can change this arrangement easily and persistently.
+ \<^emph>\<open>HyperSearch Results\<close> or the \<^emph>\<open>File System Browser\<close>. Plugins often provide
+ a central dockable to access their key functionality, which may be opened by
+ the user on demand. The Isabelle/jEdit plugin takes this approach to the
+ extreme: its plugin menu provides the entry-points to many panels that are
+ managed as dockable windows. Some important panels are docked by default,
+ e.g.\ \<^emph>\<open>Documentation\<close>, \<^emph>\<open>Output\<close>, \<^emph>\<open>Query\<close>, but the user can change this
+ arrangement easily and persistently.
Compared to plain jEdit, dockable window management in Isabelle/jEdit is
slightly augmented according to the the following principles:
@@ -439,19 +429,18 @@
Isabelle sources consist of \<^emph>\<open>symbols\<close> that extend plain ASCII to allow
infinitely many mathematical symbols within the formal sources. This works
without depending on particular encodings and varying Unicode
- standards.\footnote{Raw Unicode characters within formal sources would
- compromise portability and reliability in the face of changing
- interpretation of special features of Unicode, such as Combining Characters
- or Bi-directional Text.} See also @{cite "Wenzel:2011:CICM"}.
+ standards.\<^footnote>\<open>Raw Unicode characters within formal sources would compromise
+ portability and reliability in the face of changing interpretation of
+ special features of Unicode, such as Combining Characters or Bi-directional
+ Text.\<close> See also @{cite "Wenzel:2011:CICM"}.
For the prover back-end, formal text consists of ASCII characters that are
- grouped according to some simple rules, e.g.\ as plain ``\<^verbatim>\<open>a\<close>'' or
- symbolic ``\<^verbatim>\<open>\<alpha>\<close>''. For the editor front-end, a certain subset of
- symbols is rendered physically via Unicode glyphs, in order to show
- ``\<^verbatim>\<open>\<alpha>\<close>'' as ``\<open>\<alpha>\<close>'', for example. This symbol
- interpretation is specified by the Isabelle system distribution in @{file
- "$ISABELLE_HOME/etc/symbols"} and may be augmented by the user in
- @{file_unchecked "$ISABELLE_HOME_USER/etc/symbols"}.
+ grouped according to some simple rules, e.g.\ as plain ``\<^verbatim>\<open>a\<close>'' or symbolic
+ ``\<^verbatim>\<open>\<alpha>\<close>''. For the editor front-end, a certain subset of symbols is rendered
+ physically via Unicode glyphs, in order to show ``\<^verbatim>\<open>\<alpha>\<close>'' as ``\<open>\<alpha>\<close>'', for
+ example. This symbol interpretation is specified by the Isabelle system
+ distribution in @{file "$ISABELLE_HOME/etc/symbols"} and may be augmented by
+ the user in @{file_unchecked "$ISABELLE_HOME_USER/etc/symbols"}.
The appendix of @{cite "isabelle-isar-ref"} gives an overview of the
standard interpretation of finitely many symbols from the infinite
@@ -484,13 +473,13 @@
Note that a Java/AWT/Swing application can load additional fonts only if
they are not installed on the operating system already! Some outdated
- version of \<^verbatim>\<open>IsabelleText\<close> that happens to be provided by the
- operating system would prevent Isabelle/jEdit to use its bundled version.
- This could lead to missing glyphs (black rectangles), when the system
- version of \<^verbatim>\<open>IsabelleText\<close> is older than the application version.
- This problem can be avoided by refraining to ``install'' any version of
- \<^verbatim>\<open>IsabelleText\<close> in the first place, although it is occasionally
- tempting to use the same font in other applications.
+ version of \<^verbatim>\<open>IsabelleText\<close> that happens to be provided by the operating
+ system would prevent Isabelle/jEdit to use its bundled version. This could
+ lead to missing glyphs (black rectangles), when the system version of
+ \<^verbatim>\<open>IsabelleText\<close> is older than the application version. This problem can be
+ avoided by refraining to ``install'' any version of \<^verbatim>\<open>IsabelleText\<close> in the
+ first place, although it is occasionally tempting to use the same font in
+ other applications.
\<close>
paragraph \<open>Input methods.\<close>
@@ -548,17 +537,16 @@
\<open>\<notin>\<close> & \<^verbatim>\<open>\notin\<close> & \<^verbatim>\<open>~:\<close> \\
\end{tabular}
\<^medskip>
-
+
Note that the above abbreviations refer to the input method. The logical
notation provides ASCII alternatives that often coincide, but sometimes
deviate. This occasionally causes user confusion with very old-fashioned
- Isabelle source that use ASCII replacement notation like \<^verbatim>\<open>!\<close> or
- \<^verbatim>\<open>ALL\<close> directly in the text.
+ Isabelle source that use ASCII replacement notation like \<^verbatim>\<open>!\<close> or \<^verbatim>\<open>ALL\<close>
+ directly in the text.
On the other hand, coincidence of symbol abbreviations with ASCII
- replacement syntax syntax helps to update old theory sources via
- explicit completion (see also \<^verbatim>\<open>C+b\<close> explained in
- \secref{sec:completion}).
+ replacement syntax syntax helps to update old theory sources via explicit
+ completion (see also \<^verbatim>\<open>C+b\<close> explained in \secref{sec:completion}).
\<close>
paragraph \<open>Control symbols.\<close>
@@ -596,8 +584,7 @@
Isabelle/jEdit provides SideKick parsers for its main mode for theory files,
as well as some minor modes for the \<^verbatim>\<open>NEWS\<close> file (see
- \figref{fig:sidekick}), session \<^verbatim>\<open>ROOT\<close> files, and system
- \<^verbatim>\<open>options\<close>.
+ \figref{fig:sidekick}), session \<^verbatim>\<open>ROOT\<close> files, and system \<^verbatim>\<open>options\<close>.
\begin{figure}[htb]
\begin{center}
@@ -607,35 +594,33 @@
\label{fig:sidekick}
\end{figure}
- Moreover, the special SideKick parser \<^verbatim>\<open>isabelle-markup\<close>
- provides access to the full (uninterpreted) markup tree of the PIDE
- document model of the current buffer. This is occasionally useful
- for informative purposes, but the amount of displayed information
- might cause problems for large buffers, both for the human and the
- machine.
+ Moreover, the special SideKick parser \<^verbatim>\<open>isabelle-markup\<close> provides access to
+ the full (uninterpreted) markup tree of the PIDE document model of the
+ current buffer. This is occasionally useful for informative purposes, but
+ the amount of displayed information might cause problems for large buffers,
+ both for the human and the machine.
\<close>
section \<open>Scala console \label{sec:scala-console}\<close>
text \<open>
- The \<^emph>\<open>Console\<close> plugin manages various shells (command interpreters),
- e.g.\ \<^emph>\<open>BeanShell\<close>, which is the official jEdit scripting language, and
- the cross-platform \<^emph>\<open>System\<close> shell. Thus the console provides similar
- functionality than the Emacs buffers \<^verbatim>\<open>*scratch*\<close> and
- \<^verbatim>\<open>*shell*\<close>.
+ The \<^emph>\<open>Console\<close> plugin manages various shells (command interpreters), e.g.\
+ \<^emph>\<open>BeanShell\<close>, which is the official jEdit scripting language, and the
+ cross-platform \<^emph>\<open>System\<close> shell. Thus the console provides similar
+ functionality than the Emacs buffers \<^verbatim>\<open>*scratch*\<close> and \<^verbatim>\<open>*shell*\<close>.
- Isabelle/jEdit extends the repertoire of the console by \<^emph>\<open>Scala\<close>, which
- is the regular Scala toplevel loop running inside the same JVM process as
+ Isabelle/jEdit extends the repertoire of the console by \<^emph>\<open>Scala\<close>, which is
+ the regular Scala toplevel loop running inside the same JVM process as
Isabelle/jEdit itself. This means the Scala command interpreter has access
to the JVM name space and state of the running Prover IDE application. The
default environment imports the full content of packages \<^verbatim>\<open>isabelle\<close> and
\<^verbatim>\<open>isabelle.jedit\<close>.
- For example, \<^verbatim>\<open>PIDE\<close> refers to the Isabelle/jEdit plugin object,
- and \<^verbatim>\<open>view\<close> to the current editor view of jEdit. The Scala
- expression \<^verbatim>\<open>PIDE.snapshot(view)\<close> makes a PIDE document snapshot
- of the current buffer within the current editor view.
+ For example, \<^verbatim>\<open>PIDE\<close> refers to the Isabelle/jEdit plugin object, and \<^verbatim>\<open>view\<close>
+ to the current editor view of jEdit. The Scala expression
+ \<^verbatim>\<open>PIDE.snapshot(view)\<close> makes a PIDE document snapshot of the current buffer
+ within the current editor view.
This helps to explore Isabelle/Scala functionality interactively. Some care
is required to avoid interference with the internals of the running
@@ -649,10 +634,10 @@
File specifications in jEdit follow various formats and conventions
according to \<^emph>\<open>Virtual File Systems\<close>, which may be also provided by
additional plugins. This allows to access remote files via the \<^verbatim>\<open>http:\<close>
- protocol prefix, for example. Isabelle/jEdit attempts to work with
- the file-system model of jEdit as far as possible. In particular, theory
- sources are passed directly from the editor to the prover, without
- indirection via physical files.
+ protocol prefix, for example. Isabelle/jEdit attempts to work with the
+ file-system model of jEdit as far as possible. In particular, theory sources
+ are passed directly from the editor to the prover, without indirection via
+ physical files.
Despite the flexibility of URLs in jEdit, local files are particularly
important and are accessible without protocol prefix. Here the path notation
@@ -663,12 +648,11 @@
The Java notation for files needs to be distinguished from the one of
Isabelle, which uses POSIX notation with forward slashes on \<^emph>\<open>all\<close>
- platforms.\footnote{Isabelle/ML on Windows uses Cygwin file-system access
- and Unix-style path notation.} Moreover, environment variables from the
- Isabelle process may be used freely, e.g.\ @{file
- "$ISABELLE_HOME/etc/symbols"} or @{file_unchecked "$POLYML_HOME/README"}.
- There are special shortcuts: @{file "~"} for @{file "$USER_HOME"} and @{file
- "~~"} for @{file "$ISABELLE_HOME"}.
+ platforms.\<^footnote>\<open>Isabelle/ML on Windows uses Cygwin file-system access and
+ Unix-style path notation.\<close> Moreover, environment variables from the Isabelle
+ process may be used freely, e.g.\ @{file "$ISABELLE_HOME/etc/symbols"} or
+ @{file_unchecked "$POLYML_HOME/README"}. There are special shortcuts: @{file
+ "~"} for @{file "$USER_HOME"} and @{file "~~"} for @{file "$ISABELLE_HOME"}.
\<^medskip>
Since jEdit happens to support environment variables within file
@@ -681,20 +665,19 @@
(\secref{sec:command-line}).
Isabelle/jEdit imitates \<^verbatim>\<open>$ISABELLE_HOME\<close> and \<^verbatim>\<open>$ISABELLE_HOME_USER\<close> within
- the Java process environment, in order to
- allow easy access to these important places from the editor. The file
- browser of jEdit also includes \<^emph>\<open>Favorites\<close> for these two important
- locations.
+ the Java process environment, in order to allow easy access to these
+ important places from the editor. The file browser of jEdit also includes
+ \<^emph>\<open>Favorites\<close> for these two important locations.
\<^medskip>
- Path specifications in prover input or output usually include
- formal markup that turns it into a hyperlink (see also
- \secref{sec:tooltips-hyperlinks}). This allows to open the corresponding
- file in the text editor, independently of the path notation.
+ Path specifications in prover input or output usually include formal markup
+ that turns it into a hyperlink (see also \secref{sec:tooltips-hyperlinks}).
+ This allows to open the corresponding file in the text editor, independently
+ of the path notation.
Formally checked paths in prover input are subject to completion
- (\secref{sec:completion}): partial specifications are resolved via
- directory content and possible completions are offered in a popup.
+ (\secref{sec:completion}): partial specifications are resolved via directory
+ content and possible completions are offered in a popup.
\<close>
@@ -720,9 +703,9 @@
text \<open>
As a regular text editor, jEdit maintains a collection of \<^emph>\<open>buffers\<close> to
store text files; each buffer may be associated with any number of visible
- \<^emph>\<open>text areas\<close>. Buffers are subject to an \<^emph>\<open>edit mode\<close> that is
- determined from the file name extension. The following modes are treated
- specifically in Isabelle/jEdit:
+ \<^emph>\<open>text areas\<close>. Buffers are subject to an \<^emph>\<open>edit mode\<close> that is determined
+ from the file name extension. The following modes are treated specifically
+ in Isabelle/jEdit:
\<^medskip>
\begin{tabular}{lll}
@@ -734,17 +717,16 @@
\<^medskip>
All jEdit buffers are automatically added to the PIDE document-model as
- \<^emph>\<open>document nodes\<close>. The overall document structure is defined by the
- theory nodes in two dimensions:
+ \<^emph>\<open>document nodes\<close>. The overall document structure is defined by the theory
+ nodes in two dimensions:
- \<^enum> via \<^bold>\<open>theory imports\<close> that are specified in the \<^emph>\<open>theory
- header\<close> using concrete syntax of the @{command_ref theory} command
- @{cite "isabelle-isar-ref"};
+ \<^enum> via \<^bold>\<open>theory imports\<close> that are specified in the \<^emph>\<open>theory header\<close> using
+ concrete syntax of the @{command_ref theory} command @{cite
+ "isabelle-isar-ref"};
- \<^enum> via \<^bold>\<open>auxiliary files\<close> that are loaded into a theory by special
- \<^emph>\<open>load commands\<close>, notably @{command_ref ML_file} and @{command_ref
- SML_file} @{cite "isabelle-isar-ref"}.
-
+ \<^enum> via \<^bold>\<open>auxiliary files\<close> that are loaded into a theory by special \<^emph>\<open>load
+ commands\<close>, notably @{command_ref ML_file} and @{command_ref SML_file}
+ @{cite "isabelle-isar-ref"}.
In any case, source files are managed by the PIDE infrastructure: the
physical file-system only plays a subordinate role. The relevant version of
@@ -756,12 +738,12 @@
subsection \<open>Theories \label{sec:theories}\<close>
text \<open>
- The \<^emph>\<open>Theories\<close> panel (see also \figref{fig:theories}) provides an
- overview of the status of continuous checking of theory nodes within the
- document model. Unlike batch sessions of @{tool build} @{cite
- "isabelle-system"}, theory nodes are identified by full path names; this allows
- to work with multiple (disjoint) Isabelle sessions simultaneously within the
- same editor session.
+ The \<^emph>\<open>Theories\<close> panel (see also \figref{fig:theories}) provides an overview
+ of the status of continuous checking of theory nodes within the document
+ model. Unlike batch sessions of @{tool build} @{cite "isabelle-system"},
+ theory nodes are identified by full path names; this allows to work with
+ multiple (disjoint) Isabelle sessions simultaneously within the same editor
+ session.
\begin{figure}[htb]
\begin{center}
@@ -780,38 +762,36 @@
@{system_option jedit_auto_load}.
\<^medskip>
- The visible \<^emph>\<open>perspective\<close> of Isabelle/jEdit is defined by the
- collective view on theory buffers via open text areas. The perspective is
- taken as a hint for document processing: the prover ensures that those parts
- of a theory where the user is looking are checked, while other parts that
- are presently not required are ignored. The perspective is changed by
- opening or closing text area windows, or scrolling within a window.
+ The visible \<^emph>\<open>perspective\<close> of Isabelle/jEdit is defined by the collective
+ view on theory buffers via open text areas. The perspective is taken as a
+ hint for document processing: the prover ensures that those parts of a
+ theory where the user is looking are checked, while other parts that are
+ presently not required are ignored. The perspective is changed by opening or
+ closing text area windows, or scrolling within a window.
- The \<^emph>\<open>Theories\<close> panel provides some further options to influence
- the process of continuous checking: it may be switched off globally
- to restrict the prover to superficial processing of command syntax.
- It is also possible to indicate theory nodes as \<^emph>\<open>required\<close> for
- continuous checking: this means such nodes and all their imports are
- always processed independently of the visibility status (if
- continuous checking is enabled). Big theory libraries that are
- marked as required can have significant impact on performance,
+ The \<^emph>\<open>Theories\<close> panel provides some further options to influence the process
+ of continuous checking: it may be switched off globally to restrict the
+ prover to superficial processing of command syntax. It is also possible to
+ indicate theory nodes as \<^emph>\<open>required\<close> for continuous checking: this means
+ such nodes and all their imports are always processed independently of the
+ visibility status (if continuous checking is enabled). Big theory libraries
+ that are marked as required can have significant impact on performance,
though.
\<^medskip>
- Formal markup of checked theory content is turned into GUI
- rendering, based on a standard repertoire known from IDEs for programming
- languages: colors, icons, highlighting, squiggly underlines, tooltips,
- hyperlinks etc. For outer syntax of Isabelle/Isar there is some traditional
- syntax-highlighting via static keywords and tokenization within the editor;
- this buffer syntax is determined from theory imports. In contrast, the
- painting of inner syntax (term language etc.)\ uses semantic information
- that is reported dynamically from the logical context. Thus the prover can
- provide additional markup to help the user to understand the meaning of
- formal text, and to produce more text with some add-on tools (e.g.\
- information messages with \<^emph>\<open>sendback\<close> markup by automated provers or
- disprovers in the background).
+ Formal markup of checked theory content is turned into GUI rendering, based
+ on a standard repertoire known from IDEs for programming languages: colors,
+ icons, highlighting, squiggly underlines, tooltips, hyperlinks etc. For
+ outer syntax of Isabelle/Isar there is some traditional syntax-highlighting
+ via static keywords and tokenization within the editor; this buffer syntax
+ is determined from theory imports. In contrast, the painting of inner syntax
+ (term language etc.)\ uses semantic information that is reported dynamically
+ from the logical context. Thus the prover can provide additional markup to
+ help the user to understand the meaning of formal text, and to produce more
+ text with some add-on tools (e.g.\ information messages with \<^emph>\<open>sendback\<close>
+ markup by automated provers or disprovers in the background).
+\<close>
-\<close>
subsection \<open>Auxiliary files \label{sec:aux-files}\<close>
@@ -826,14 +806,13 @@
treated as changes of the corresponding load command.
\<^medskip>
- As a concession to the massive amount of ML files in Isabelle/HOL
- itself, the content of auxiliary files is only added to the PIDE
- document-model on demand, the first time when opened explicitly in the
- editor. There are further tricks to manage markup of ML files, such that
- Isabelle/HOL may be edited conveniently in the Prover IDE on small machines
- with only 8\,GB of main memory. Using \<^verbatim>\<open>Pure\<close> as logic session
- image, the exploration may start at the top @{file
- "$ISABELLE_HOME/src/HOL/Main.thy"} or the bottom @{file
+ As a concession to the massive amount of ML files in Isabelle/HOL itself,
+ the content of auxiliary files is only added to the PIDE document-model on
+ demand, the first time when opened explicitly in the editor. There are
+ further tricks to manage markup of ML files, such that Isabelle/HOL may be
+ edited conveniently in the Prover IDE on small machines with only 8\,GB of
+ main memory. Using \<^verbatim>\<open>Pure\<close> as logic session image, the exploration may start
+ at the top @{file "$ISABELLE_HOME/src/HOL/Main.thy"} or the bottom @{file
"$ISABELLE_HOME/src/HOL/HOL.thy"}, for example.
Initially, before an auxiliary file is opened in the editor, the prover
@@ -853,30 +832,30 @@
morally unsupported and might lead to confusion.
\<^medskip>
- Output that refers to an auxiliary file is combined with that of
- the corresponding load command, and shown whenever the file or the command
- are active (see also \secref{sec:output}).
+ Output that refers to an auxiliary file is combined with that of the
+ corresponding load command, and shown whenever the file or the command are
+ active (see also \secref{sec:output}).
Warnings, errors, and other useful markup is attached directly to the
positions in the auxiliary file buffer, in the manner of other well-known
IDEs. By using the load command @{command SML_file} as explained in @{file
"$ISABELLE_HOME/src/Tools/SML/Examples.thy"}, Isabelle/jEdit may be used as
fully-featured IDE for Standard ML, independently of theory or proof
- development: the required theory merely serves as some kind of project
- file for a collection of SML source modules.
+ development: the required theory merely serves as some kind of project file
+ for a collection of SML source modules.
\<close>
section \<open>Output \label{sec:output}\<close>
text \<open>
- Prover output consists of \<^emph>\<open>markup\<close> and \<^emph>\<open>messages\<close>. Both are
- directly attached to the corresponding positions in the original source
- text, and visualized in the text area, e.g.\ as text colours for free and
- bound variables, or as squiggly underlines for warnings, errors etc.\ (see
- also \figref{fig:output}). In the latter case, the corresponding messages
- are shown by hovering with the mouse over the highlighted text --- although
- in many situations the user should already get some clue by looking at the
+ Prover output consists of \<^emph>\<open>markup\<close> and \<^emph>\<open>messages\<close>. Both are directly
+ attached to the corresponding positions in the original source text, and
+ visualized in the text area, e.g.\ as text colours for free and bound
+ variables, or as squiggly underlines for warnings, errors etc.\ (see also
+ \figref{fig:output}). In the latter case, the corresponding messages are
+ shown by hovering with the mouse over the highlighted text --- although in
+ many situations the user should already get some clue by looking at the
position of the text highlighting, without the text itself.
\begin{figure}[htb]
@@ -888,11 +867,10 @@
\label{fig:output}
\end{figure}
- The ``gutter area'' on the left-hand-side of the text area uses
- icons to provide a summary of the messages within the adjacent
- line of text. Message priorities are used to prefer errors over
- warnings, warnings over information messages, but plain output is
- ignored.
+ The ``gutter area'' on the left-hand-side of the text area uses icons to
+ provide a summary of the messages within the adjacent line of text. Message
+ priorities are used to prefer errors over warnings, warnings over
+ information messages, but plain output is ignored.
The ``overview area'' on the right-hand-side of the text area uses similar
information to paint small rectangles for the overall status of the whole
@@ -900,16 +878,14 @@
the given window height. Mouse clicks on the overview area position the
cursor approximately to the corresponding line of text in the buffer.
- Another course-grained overview is provided by the \<^emph>\<open>Theories\<close>
- panel, but without direct correspondence to text positions. A
- double-click on one of the theory entries with their status overview
- opens the corresponding text buffer, without changing the cursor
- position.
+ Another course-grained overview is provided by the \<^emph>\<open>Theories\<close> panel, but
+ without direct correspondence to text positions. A double-click on one of
+ the theory entries with their status overview opens the corresponding text
+ buffer, without changing the cursor position.
\<^medskip>
- In addition, the \<^emph>\<open>Output\<close> panel displays prover
- messages that correspond to a given command, within a separate
- window.
+ In addition, the \<^emph>\<open>Output\<close> panel displays prover messages that correspond to
+ a given command, within a separate window.
The cursor position in the presently active text area determines the prover
command whose cumulative message output is appended and shown in that window
@@ -925,36 +901,34 @@
possible to do meaningful proof editing within the primary text area and its
markup, while using secondary output windows only rarely.
- The main purpose of the output window is to ``debug'' unclear
- situations by inspecting internal state of the prover.\footnote{In
- that sense, unstructured tactic scripts depend on continuous
- debugging with internal state inspection.} Consequently, some
- special messages for \<^emph>\<open>tracing\<close> or \<^emph>\<open>proof state\<close> only
+ The main purpose of the output window is to ``debug'' unclear situations by
+ inspecting internal state of the prover.\<^footnote>\<open>In that sense, unstructured tactic
+ scripts depend on continuous debugging with internal state inspection.\<close>
+ Consequently, some special messages for \<^emph>\<open>tracing\<close> or \<^emph>\<open>proof state\<close> only
appear here, and are not attached to the original source.
\<^medskip>
- In any case, prover messages also contain markup that may
- be explored recursively via tooltips or hyperlinks (see
- \secref{sec:tooltips-hyperlinks}), or clicked directly to initiate
- certain actions (see \secref{sec:auto-tools} and
- \secref{sec:sledgehammer}).\<close>
+ In any case, prover messages also contain markup that may be explored
+ recursively via tooltips or hyperlinks (see
+ \secref{sec:tooltips-hyperlinks}), or clicked directly to initiate certain
+ actions (see \secref{sec:auto-tools} and \secref{sec:sledgehammer}).
+\<close>
section \<open>Query \label{sec:query}\<close>
text \<open>
- The \<^emph>\<open>Query\<close> panel provides various GUI forms to request extra
- information from the prover. In old times the user would have issued some
- diagnostic command like @{command find_theorems} and inspected its output,
- but this is now integrated into the Prover IDE.
+ The \<^emph>\<open>Query\<close> panel provides various GUI forms to request extra information
+ from the prover. In old times the user would have issued some diagnostic
+ command like @{command find_theorems} and inspected its output, but this is
+ now integrated into the Prover IDE.
- A \<^emph>\<open>Query\<close> window provides some input fields and buttons for a
- particular query command, with output in a dedicated text area. There are
- various query modes: \<^emph>\<open>Find Theorems\<close>, \<^emph>\<open>Find Constants\<close>,
- \<^emph>\<open>Print Context\<close>, e.g.\ see \figref{fig:query}. As usual in jEdit,
- multiple \<^emph>\<open>Query\<close> windows may be active at the same time: any number of
- floating instances, but at most one docked instance (which is used by
- default).
+ A \<^emph>\<open>Query\<close> window provides some input fields and buttons for a particular
+ query command, with output in a dedicated text area. There are various query
+ modes: \<^emph>\<open>Find Theorems\<close>, \<^emph>\<open>Find Constants\<close>, \<^emph>\<open>Print Context\<close>, e.g.\ see
+ \figref{fig:query}. As usual in jEdit, multiple \<^emph>\<open>Query\<close> windows may be
+ active at the same time: any number of floating instances, but at most one
+ docked instance (which is used by default).
\begin{figure}[htb]
\begin{center}
@@ -967,20 +941,19 @@
\<^medskip>
The following GUI elements are common to all query modes:
- \<^item> The spinning wheel provides feedback about the status of a pending
- query wrt.\ the evaluation of its context and its own operation.
+ \<^item> The spinning wheel provides feedback about the status of a pending query
+ wrt.\ the evaluation of its context and its own operation.
- \<^item> The \<^emph>\<open>Apply\<close> button attaches a fresh query invocation to the
- current context of the command where the cursor is pointing in the text.
+ \<^item> The \<^emph>\<open>Apply\<close> button attaches a fresh query invocation to the current
+ context of the command where the cursor is pointing in the text.
- \<^item> The \<^emph>\<open>Search\<close> field allows to highlight query output according to
- some regular expression, in the notation that is commonly used on the Java
- platform.\footnote{@{url
- "http://docs.oracle.com/javase/7/docs/api/java/util/regex/Pattern.html"}}
- This may serve as an additional visual filter of the result.
+ \<^item> The \<^emph>\<open>Search\<close> field allows to highlight query output according to some
+ regular expression, in the notation that is commonly used on the Java
+ platform.\<^footnote>\<open>@{url
+ "https://docs.oracle.com/javase/8/docs/api/java/util/regex/Pattern.html"}\<close>
+ This may serve as an additional visual filter of the result.
- \<^item> The \<^emph>\<open>Zoom\<close> box controls the font size of the output area.
-
+ \<^item> The \<^emph>\<open>Zoom\<close> box controls the font size of the output area.
All query operations are asynchronous: there is no need to wait for the
evaluation of the document for the query context, nor for the query
@@ -994,26 +967,26 @@
subsection \<open>Find theorems\<close>
text \<open>
- The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Theorems\<close> mode retrieves facts from the
- theory or proof context matching all of given criteria in the \<^emph>\<open>Find\<close>
- text field. A single criterium has the following syntax:
+ The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Theorems\<close> mode retrieves facts from the theory
+ or proof context matching all of given criteria in the \<^emph>\<open>Find\<close> text field. A
+ single criterium has the following syntax:
@{rail \<open>
('-'?) ('name' ':' @{syntax nameref} | 'intro' | 'elim' | 'dest' |
'solves' | 'simp' ':' @{syntax term} | @{syntax term})
\<close>}
- See also the Isar command @{command_ref find_theorems} in
- @{cite "isabelle-isar-ref"}.
+ See also the Isar command @{command_ref find_theorems} in @{cite
+ "isabelle-isar-ref"}.
\<close>
subsection \<open>Find constants\<close>
text \<open>
- The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Constants\<close> mode prints all constants
- whose type meets all of the given criteria in the \<^emph>\<open>Find\<close> text field.
- A single criterium has the following syntax:
+ The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Constants\<close> mode prints all constants whose type
+ meets all of the given criteria in the \<^emph>\<open>Find\<close> text field. A single
+ criterium has the following syntax:
@{rail \<open>
('-'?)
@@ -1028,8 +1001,8 @@
subsection \<open>Print context\<close>
text \<open>
- The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Print Context\<close> mode prints information from
- the theory or proof context, or proof state. See also the Isar commands
+ The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Print Context\<close> mode prints information from the
+ theory or proof context, or proof state. See also the Isar commands
@{command_ref print_context}, @{command_ref print_cases}, @{command_ref
print_term_bindings}, @{command_ref print_theorems}, @{command_ref
print_state} described in @{cite "isabelle-isar-ref"}.
@@ -1040,12 +1013,11 @@
text \<open>
Formally processed text (prover input or output) contains rich markup
- information that can be explored further by using the \<^verbatim>\<open>CONTROL\<close>
- modifier key on Linux and Windows, or \<^verbatim>\<open>COMMAND\<close> on Mac OS X.
- Hovering with the mouse while the modifier is pressed reveals a
- \<^emph>\<open>tooltip\<close> (grey box over the text with a yellow popup) and/or a
- \<^emph>\<open>hyperlink\<close> (black rectangle over the text with change of mouse
- pointer); see also \figref{fig:tooltip}.
+ information that can be explored further by using the \<^verbatim>\<open>CONTROL\<close> modifier
+ key on Linux and Windows, or \<^verbatim>\<open>COMMAND\<close> on Mac OS X. Hovering with the mouse
+ while the modifier is pressed reveals a \<^emph>\<open>tooltip\<close> (grey box over the text
+ with a yellow popup) and/or a \<^emph>\<open>hyperlink\<close> (black rectangle over the text
+ with change of mouse pointer); see also \figref{fig:tooltip}.
\begin{figure}[htb]
\begin{center}
@@ -1055,9 +1027,9 @@
\label{fig:tooltip}
\end{figure}
- Tooltip popups use the same rendering mechanisms as the main text
- area, and further tooltips and/or hyperlinks may be exposed
- recursively by the same mechanism; see \figref{fig:nested-tooltips}.
+ Tooltip popups use the same rendering mechanisms as the main text area, and
+ further tooltips and/or hyperlinks may be exposed recursively by the same
+ mechanism; see \figref{fig:nested-tooltips}.
\begin{figure}[htb]
\begin{center}
@@ -1067,25 +1039,24 @@
\label{fig:nested-tooltips}
\end{figure}
- The tooltip popup window provides some controls to \<^emph>\<open>close\<close> or
- \<^emph>\<open>detach\<close> the window, turning it into a separate \<^emph>\<open>Info\<close>
- window managed by jEdit. The \<^verbatim>\<open>ESCAPE\<close> key closes
- \<^emph>\<open>all\<close> popups, which is particularly relevant when nested
- tooltips are stacking up.
+ The tooltip popup window provides some controls to \<^emph>\<open>close\<close> or \<^emph>\<open>detach\<close> the
+ window, turning it into a separate \<^emph>\<open>Info\<close> window managed by jEdit. The
+ \<^verbatim>\<open>ESCAPE\<close> key closes \<^emph>\<open>all\<close> popups, which is particularly relevant when
+ nested tooltips are stacking up.
\<^medskip>
- A black rectangle in the text indicates a hyperlink that may be
- followed by a mouse click (while the \<^verbatim>\<open>CONTROL\<close> or \<^verbatim>\<open>COMMAND\<close> modifier
- key is still pressed). Such jumps to other text locations
- are recorded by the \<^emph>\<open>Navigator\<close> plugin, which is bundled with
- Isabelle/jEdit and enabled by default, including navigation arrows in the
- main jEdit toolbar.
+ A black rectangle in the text indicates a hyperlink that may be followed by
+ a mouse click (while the \<^verbatim>\<open>CONTROL\<close> or \<^verbatim>\<open>COMMAND\<close> modifier key is still
+ pressed). Such jumps to other text locations are recorded by the
+ \<^emph>\<open>Navigator\<close> plugin, which is bundled with Isabelle/jEdit and enabled by
+ default, including navigation arrows in the main jEdit toolbar.
- Also note that the link target may be a file that is itself not
- subject to formal document processing of the editor session and thus
- prevents further exploration: the chain of hyperlinks may end in
- some source file of the underlying logic image, or within the
- ML bootstrap sources of Isabelle/Pure.\<close>
+ Also note that the link target may be a file that is itself not subject to
+ formal document processing of the editor session and thus prevents further
+ exploration: the chain of hyperlinks may end in some source file of the
+ underlying logic image, or within the ML bootstrap sources of
+ Isabelle/Pure.
+\<close>
section \<open>Completion \label{sec:completion}\<close>
@@ -1100,16 +1071,16 @@
\<^medskip>
\<^emph>\<open>Explicit completion\<close> is triggered by the action @{action_ref
- "isabelle.complete"}, which is bound to the keyboard shortcut \<^verbatim>\<open>C+b\<close>,
- and thus overrides the jEdit default for @{action_ref "complete-word"}.
+ "isabelle.complete"}, which is bound to the keyboard shortcut \<^verbatim>\<open>C+b\<close>, and
+ thus overrides the jEdit default for @{action_ref "complete-word"}.
- \<^emph>\<open>Implicit completion\<close> hooks into the regular keyboard input stream of
- the editor, with some event filtering and optional delays.
+ \<^emph>\<open>Implicit completion\<close> hooks into the regular keyboard input stream of the
+ editor, with some event filtering and optional delays.
\<^medskip>
- Completion options may be configured in \<^emph>\<open>Plugin Options~/
- Isabelle~/ General~/ Completion\<close>. These are explained in further detail
- below, whenever relevant. There is also a summary of options in
+ Completion options may be configured in \<^emph>\<open>Plugin Options~/ Isabelle~/
+ General~/ Completion\<close>. These are explained in further detail below, whenever
+ relevant. There is also a summary of options in
\secref{sec:completion-options}.
The asynchronous nature of PIDE interaction means that information from the
@@ -1132,19 +1103,17 @@
kinds and purposes. The completion mechanism supports this by the following
built-in templates:
- \<^descr> \<^verbatim>\<open>`\<close> (single ASCII back-quote) supports \<^emph>\<open>quotations\<close>
- via text cartouches. There are three selections, which are always presented
- in the same order and do not depend on any context information. The default
- choice produces a template ``\<open>\<open>\<box>\<close>\<close>'', where the box indicates the
- cursor position after insertion; the other choices help to repair the block
- structure of unbalanced text cartouches.
+ \<^descr> \<^verbatim>\<open>`\<close> (single ASCII back-quote) supports \<^emph>\<open>quotations\<close> via text
+ cartouches. There are three selections, which are always presented in the
+ same order and do not depend on any context information. The default
+ choice produces a template ``\<open>\<open>\<box>\<close>\<close>'', where the box indicates the cursor
+ position after insertion; the other choices help to repair the block
+ structure of unbalanced text cartouches.
- \<^descr> \<^verbatim>\<open>@{\<close> is completed to the template ``\<open>@{\<box>}\<close>'',
- where the box indicates the cursor position after insertion. Here it is
- convenient to use the wildcard ``\<^verbatim>\<open>__\<close>'' or a more specific name
- prefix to let semantic completion of name-space entries propose
- antiquotation names.
-
+ \<^descr> \<^verbatim>\<open>@{\<close> is completed to the template ``\<open>@{\<box>}\<close>'', where the box indicates
+ the cursor position after insertion. Here it is convenient to use the
+ wildcard ``\<^verbatim>\<open>__\<close>'' or a more specific name prefix to let semantic
+ completion of name-space entries propose antiquotation names.
With some practice, input of quoted sub-languages and antiquotations of
embedded languages should work fluently. Note that national keyboard layouts
@@ -1193,10 +1162,9 @@
When inserted into the text, the above examples all produce the same Unicode
rendering \<open>\<forall>\<close> of the underlying symbol \<^verbatim>\<open>\<forall>\<close>.
- A symbol abbreviation that is a plain word, like \<^verbatim>\<open>ALL\<close>, is
- treated like a syntax keyword. Non-word abbreviations like \<^verbatim>\<open>-->\<close>
- are inserted more aggressively, except for single-character abbreviations
- like \<^verbatim>\<open>!\<close> above.
+ A symbol abbreviation that is a plain word, like \<^verbatim>\<open>ALL\<close>, is treated like a
+ syntax keyword. Non-word abbreviations like \<^verbatim>\<open>-->\<close> are inserted more
+ aggressively, except for single-character abbreviations like \<^verbatim>\<open>!\<close> above.
\<^medskip>
Symbol completion depends on the semantic language context
@@ -1217,17 +1185,17 @@
@{system_option_ref completion_limit}. The completion mechanism takes this
into account when collecting information on the prover side.
- Already recognized names are \<^emph>\<open>not\<close> completed further, but completion
- may be extended by appending a suffix of underscores. This provokes a failed
+ Already recognized names are \<^emph>\<open>not\<close> completed further, but completion may be
+ extended by appending a suffix of underscores. This provokes a failed
lookup, and another completion attempt while ignoring the underscores. For
- example, in a name space where \<^verbatim>\<open>foo\<close> and \<^verbatim>\<open>foobar\<close>
- are known, the input \<^verbatim>\<open>foo\<close> remains unchanged, but \<^verbatim>\<open>foo_\<close> may be completed
- to \<^verbatim>\<open>foo\<close> or \<^verbatim>\<open>foobar\<close>.
+ example, in a name space where \<^verbatim>\<open>foo\<close> and \<^verbatim>\<open>foobar\<close> are known, the input
+ \<^verbatim>\<open>foo\<close> remains unchanged, but \<^verbatim>\<open>foo_\<close> may be completed to \<^verbatim>\<open>foo\<close> or
+ \<^verbatim>\<open>foobar\<close>.
- The special identifier ``\<^verbatim>\<open>__\<close>'' serves as a wild-card for
- arbitrary completion: it exposes the name-space content to the completion
- mechanism (truncated according to @{system_option completion_limit}). This
- is occasionally useful to explore an unknown name-space, e.g.\ in some
+ The special identifier ``\<^verbatim>\<open>__\<close>'' serves as a wild-card for arbitrary
+ completion: it exposes the name-space content to the completion mechanism
+ (truncated according to @{system_option completion_limit}). This is
+ occasionally useful to explore an unknown name-space, e.g.\ in some
template.
\<close>
@@ -1239,9 +1207,9 @@
source text, e.g.\ for the argument of a load command
(\secref{sec:aux-files}), the completion mechanism explores the directory
content and offers the result as completion popup. Relative path
- specifications are understood wrt.\ the \<^emph>\<open>master directory\<close> of the
- document node (\secref{sec:buffer-node}) of the enclosing editor buffer;
- this requires a proper theory, not an auxiliary file.
+ specifications are understood wrt.\ the \<^emph>\<open>master directory\<close> of the document
+ node (\secref{sec:buffer-node}) of the enclosing editor buffer; this
+ requires a proper theory, not an auxiliary file.
A suffix of slashes may be used to continue the exploration of an already
recognized directory name.
@@ -1274,10 +1242,10 @@
default keyboard shortcut \<^verbatim>\<open>C+b\<close>.
\<^medskip>
- Dictionary lookup uses some educated guesses about lower-case,
- upper-case, and capitalized words. This is oriented on common use in
- English, where this aspect is not decisive for proper spelling, in contrast
- to German, for example.
+ Dictionary lookup uses some educated guesses about lower-case, upper-case,
+ and capitalized words. This is oriented on common use in English, where this
+ aspect is not decisive for proper spelling, in contrast to German, for
+ example.
\<close>
@@ -1296,14 +1264,14 @@
symbol completion for ML source, but within ML strings, comments,
antiquotations.
- The prover may produce \<^emph>\<open>no completion\<close> markup in exceptional
- situations, to tell that some language keywords should be excluded from
- further completion attempts. For example, \<^verbatim>\<open>:\<close> within accepted
- Isar syntax looses its meaning as abbreviation for symbol \<open>\<in>\<close>.
+ The prover may produce \<^emph>\<open>no completion\<close> markup in exceptional situations, to
+ tell that some language keywords should be excluded from further completion
+ attempts. For example, \<^verbatim>\<open>:\<close> within accepted Isar syntax looses its meaning
+ as abbreviation for symbol \<open>\<in>\<close>.
\<^medskip>
- The completion context is \<^emph>\<open>ignored\<close> for built-in templates and
- symbols in their explicit form ``\<^verbatim>\<open>\<foobar>\<close>''; see also
+ The completion context is \<^emph>\<open>ignored\<close> for built-in templates and symbols in
+ their explicit form ``\<^verbatim>\<open>\<foobar>\<close>''; see also
\secref{sec:completion-varieties}. This allows to complete within broken
input that escapes its normal semantic context, e.g.\ antiquotations or
string literals in ML, which do not allow arbitrary backslash sequences.
@@ -1317,56 +1285,55 @@
optional delay after keyboard input according to @{system_option
jedit_completion_delay}.
- \<^descr>[Explicit completion] works via action @{action_ref
- "isabelle.complete"} with keyboard shortcut \<^verbatim>\<open>C+b\<close>. This
- overrides the shortcut for @{action_ref "complete-word"} in jEdit, but it is
- possible to restore the original jEdit keyboard mapping of @{action
- "complete-word"} via \<^emph>\<open>Global Options~/ Shortcuts\<close> and invent a
- different one for @{action "isabelle.complete"}.
+ \<^descr>[Explicit completion] works via action @{action_ref "isabelle.complete"}
+ with keyboard shortcut \<^verbatim>\<open>C+b\<close>. This overrides the shortcut for @{action_ref
+ "complete-word"} in jEdit, but it is possible to restore the original jEdit
+ keyboard mapping of @{action "complete-word"} via \<^emph>\<open>Global Options~/
+ Shortcuts\<close> and invent a different one for @{action "isabelle.complete"}.
\<^descr>[Explicit spell-checker completion] works via @{action_ref
"isabelle.complete-word"}, which is exposed in the jEdit context menu, if
the mouse points to a word that the spell-checker can complete.
- \<^descr>[Implicit completion] works via regular keyboard input of the editor.
- It depends on further side-conditions:
+ \<^descr>[Implicit completion] works via regular keyboard input of the editor. It
+ depends on further side-conditions:
- \<^enum> The system option @{system_option_ref jedit_completion} needs to
- be enabled (default).
+ \<^enum> The system option @{system_option_ref jedit_completion} needs to be
+ enabled (default).
- \<^enum> Completion of syntax keywords requires at least 3 relevant
- characters in the text.
+ \<^enum> Completion of syntax keywords requires at least 3 relevant characters in
+ the text.
- \<^enum> The system option @{system_option_ref jedit_completion_delay}
- determines an additional delay (0.5 by default), before opening a completion
- popup. The delay gives the prover a chance to provide semantic completion
+ \<^enum> The system option @{system_option_ref jedit_completion_delay} determines
+ an additional delay (0.5 by default), before opening a completion popup.
+ The delay gives the prover a chance to provide semantic completion
information, notably the context (\secref{sec:completion-context}).
\<^enum> The system option @{system_option_ref jedit_completion_immediate}
(enabled by default) controls whether replacement text should be inserted
immediately without popup, regardless of @{system_option
- jedit_completion_delay}. This aggressive mode of completion is restricted to
- Isabelle symbols and their abbreviations (\secref{sec:symbols}).
+ jedit_completion_delay}. This aggressive mode of completion is restricted
+ to Isabelle symbols and their abbreviations (\secref{sec:symbols}).
- \<^enum> Completion of symbol abbreviations with only one relevant
- character in the text always enforces an explicit popup,
- regardless of @{system_option_ref jedit_completion_immediate}.
+ \<^enum> Completion of symbol abbreviations with only one relevant character in
+ the text always enforces an explicit popup, regardless of
+ @{system_option_ref jedit_completion_immediate}.
\<close>
subsection \<open>Completion popup \label{sec:completion-popup}\<close>
text \<open>
- A \<^emph>\<open>completion popup\<close> is a minimally invasive GUI component over the
- text area that offers a selection of completion items to be inserted into
- the text, e.g.\ by mouse clicks. Items are sorted dynamically, according to
- the frequency of selection, with persistent history. The popup may interpret
- special keys \<^verbatim>\<open>ENTER\<close>, \<^verbatim>\<open>TAB\<close>, \<^verbatim>\<open>ESCAPE\<close>,
- \<^verbatim>\<open>UP\<close>, \<^verbatim>\<open>DOWN\<close>, \<^verbatim>\<open>PAGE_UP\<close>, \<^verbatim>\<open>PAGE_DOWN\<close>, but all other key events are
- passed to the underlying text area.
- This allows to ignore unwanted completions most of the time and continue
- typing quickly. Thus the popup serves as a mechanism of confirmation of
- proposed items, but the default is to continue without completion.
+ A \<^emph>\<open>completion popup\<close> is a minimally invasive GUI component over the text
+ area that offers a selection of completion items to be inserted into the
+ text, e.g.\ by mouse clicks. Items are sorted dynamically, according to the
+ frequency of selection, with persistent history. The popup may interpret
+ special keys \<^verbatim>\<open>ENTER\<close>, \<^verbatim>\<open>TAB\<close>, \<^verbatim>\<open>ESCAPE\<close>, \<^verbatim>\<open>UP\<close>, \<^verbatim>\<open>DOWN\<close>, \<^verbatim>\<open>PAGE_UP\<close>,
+ \<^verbatim>\<open>PAGE_DOWN\<close>, but all other key events are passed to the underlying text
+ area. This allows to ignore unwanted completions most of the time and
+ continue typing quickly. Thus the popup serves as a mechanism of
+ confirmation of proposed items, but the default is to continue without
+ completion.
The meaning of special keys is as follows:
@@ -1383,9 +1350,9 @@
\end{tabular}
\<^medskip>
- Movement within the popup is only active for multiple items.
- Otherwise the corresponding key event retains its standard meaning
- within the underlying text area.
+ Movement within the popup is only active for multiple items. Otherwise the
+ corresponding key event retains its standard meaning within the underlying
+ text area.
\<close>
@@ -1401,32 +1368,31 @@
all combinations make sense. At least the following important cases are
well-defined:
- \<^descr>[No selection.] The original is removed and the replacement inserted,
- depending on the caret position.
+ \<^descr>[No selection.] The original is removed and the replacement inserted,
+ depending on the caret position.
- \<^descr>[Rectangular selection of zero width.] This special case is treated by
- jEdit as ``tall caret'' and insertion of completion imitates its normal
- behaviour: separate copies of the replacement are inserted for each line of
- the selection.
+ \<^descr>[Rectangular selection of zero width.] This special case is treated by
+ jEdit as ``tall caret'' and insertion of completion imitates its normal
+ behaviour: separate copies of the replacement are inserted for each line
+ of the selection.
- \<^descr>[Other rectangular selection or multiple selections.] Here the original
- is removed and the replacement is inserted for each line (or segment) of the
- selection.
-
+ \<^descr>[Other rectangular selection or multiple selections.] Here the original
+ is removed and the replacement is inserted for each line (or segment) of
+ the selection.
- Support for multiple selections is particularly useful for
- \<^emph>\<open>HyperSearch\<close>: clicking on one of the items in the \<^emph>\<open>HyperSearch
- Results\<close> window makes jEdit select all its occurrences in the corresponding
- line of text. Then explicit completion can be invoked via \<^verbatim>\<open>C+b\<close>,
- e.g.\ to replace occurrences of \<^verbatim>\<open>-->\<close> by \<open>\<longrightarrow>\<close>.
+ Support for multiple selections is particularly useful for \<^emph>\<open>HyperSearch\<close>:
+ clicking on one of the items in the \<^emph>\<open>HyperSearch Results\<close> window makes
+ jEdit select all its occurrences in the corresponding line of text. Then
+ explicit completion can be invoked via \<^verbatim>\<open>C+b\<close>, e.g.\ to replace occurrences
+ of \<^verbatim>\<open>-->\<close> by \<open>\<longrightarrow>\<close>.
\<^medskip>
- Insertion works by removing and inserting pieces of text from the
- buffer. This counts as one atomic operation on the jEdit history. Thus
- unintended completions may be reverted by the regular @{action undo} action
- of jEdit. According to normal jEdit policies, the recovered text after
- @{action undo} is selected: \<^verbatim>\<open>ESCAPE\<close> is required to reset the
- selection and to continue typing more text.
+ Insertion works by removing and inserting pieces of text from the buffer.
+ This counts as one atomic operation on the jEdit history. Thus unintended
+ completions may be reverted by the regular @{action undo} action of jEdit.
+ According to normal jEdit policies, the recovered text after @{action undo}
+ is selected: \<^verbatim>\<open>ESCAPE\<close> is required to reset the selection and to continue
+ typing more text.
\<close>
@@ -1434,8 +1400,8 @@
text \<open>
This is a summary of Isabelle/Scala system options that are relevant for
- completion. They may be configured in \<^emph>\<open>Plugin Options~/ Isabelle~/
- General\<close> as usual.
+ completion. They may be configured in \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>
+ as usual.
\<^item> @{system_option_def completion_limit} specifies the maximum number of
items for various semantic completion operations (name-space entries etc.)
@@ -1444,10 +1410,10 @@
regular jEdit key events (\secref{sec:completion-input}): it allows to
disable implicit completion altogether.
- \<^item> @{system_option_def jedit_completion_select_enter} and
- @{system_option_def jedit_completion_select_tab} enable keys to select a
- completion item from the popup (\secref{sec:completion-popup}). Note that a
- regular mouse click on the list of items is always possible.
+ \<^item> @{system_option_def jedit_completion_select_enter} and @{system_option_def
+ jedit_completion_select_tab} enable keys to select a completion item from
+ the popup (\secref{sec:completion-popup}). Note that a regular mouse click
+ on the list of items is always possible.
\<^item> @{system_option_def jedit_completion_context} specifies whether the
language context provided by the prover should be used at all. Disabling
@@ -1459,17 +1425,17 @@
jedit_completion_immediate} determine the handling of keyboard events for
implicit completion (\secref{sec:completion-input}).
- A @{system_option jedit_completion_delay}~\<^verbatim>\<open>> 0\<close> postpones the
- processing of key events, until after the user has stopped typing for the
- given time span, but @{system_option jedit_completion_immediate}~\<^verbatim>\<open>"= true\<close>
- means that abbreviations of Isabelle symbols are handled nonetheless.
+ A @{system_option jedit_completion_delay}~\<^verbatim>\<open>> 0\<close> postpones the processing of
+ key events, until after the user has stopped typing for the given time span,
+ but @{system_option jedit_completion_immediate}~\<^verbatim>\<open>"= true\<close> means that
+ abbreviations of Isabelle symbols are handled nonetheless.
\<^item> @{system_option_def jedit_completion_path_ignore} specifies ``glob''
patterns to ignore in file-system path completion (separated by colons),
e.g.\ backup files ending with tilde.
- \<^item> @{system_option_def spell_checker} is a global guard for all
- spell-checker operations: it allows to disable that mechanism altogether.
+ \<^item> @{system_option_def spell_checker} is a global guard for all spell-checker
+ operations: it allows to disable that mechanism altogether.
\<^item> @{system_option_def spell_checker_dictionary} determines the current
dictionary, taken from the colon-separated list in the settings variable
@@ -1478,9 +1444,9 @@
permanent dictionary updates is stored in the directory @{file_unchecked
"$ISABELLE_HOME_USER/dictionaries"}, in a separate file for each dictionary.
- \<^item> @{system_option_def spell_checker_elements} specifies a
- comma-separated list of markup elements that delimit words in the source
- that is subject to spell-checking, including various forms of comments.
+ \<^item> @{system_option_def spell_checker_elements} specifies a comma-separated
+ list of markup elements that delimit words in the source that is subject to
+ spell-checking, including various forms of comments.
\<close>
@@ -1489,22 +1455,21 @@
text \<open>
Continuous document processing works asynchronously in the background.
Visible document source that has been evaluated may get augmented by
- additional results of \<^emph>\<open>asynchronous print functions\<close>. The canonical
- example is proof state output, which is always enabled. More heavy-weight
- print functions may be applied, in order to prove or disprove parts of the
- formal text by other means.
+ additional results of \<^emph>\<open>asynchronous print functions\<close>. The canonical example
+ is proof state output, which is always enabled. More heavy-weight print
+ functions may be applied, in order to prove or disprove parts of the formal
+ text by other means.
- Isabelle/HOL provides various automatically tried tools that operate
- on outermost goal statements (e.g.\ @{command lemma}, @{command
- theorem}), independently of the state of the current proof attempt.
- They work implicitly without any arguments. Results are output as
- \<^emph>\<open>information messages\<close>, which are indicated in the text area by
- blue squiggles and a blue information sign in the gutter (see
- \figref{fig:auto-tools}). The message content may be shown as for
- other output (see also \secref{sec:output}). Some tools
- produce output with \<^emph>\<open>sendback\<close> markup, which means that
- clicking on certain parts of the output inserts that text into the
- source in the proper place.
+ Isabelle/HOL provides various automatically tried tools that operate on
+ outermost goal statements (e.g.\ @{command lemma}, @{command theorem}),
+ independently of the state of the current proof attempt. They work
+ implicitly without any arguments. Results are output as \<^emph>\<open>information
+ messages\<close>, which are indicated in the text area by blue squiggles and a blue
+ information sign in the gutter (see \figref{fig:auto-tools}). The message
+ content may be shown as for other output (see also \secref{sec:output}).
+ Some tools produce output with \<^emph>\<open>sendback\<close> markup, which means that clicking
+ on certain parts of the output inserts that text into the source in the
+ proper place.
\begin{figure}[htb]
\begin{center}
@@ -1515,85 +1480,81 @@
\end{figure}
\<^medskip>
- The following Isabelle system options control the behavior
- of automatically tried tools (see also the jEdit dialog window
- \<^emph>\<open>Plugin Options~/ Isabelle~/ General~/ Automatically tried
- tools\<close>):
+ The following Isabelle system options control the behavior of automatically
+ tried tools (see also the jEdit dialog window \<^emph>\<open>Plugin Options~/ Isabelle~/
+ General~/ Automatically tried tools\<close>):
- \<^item> @{system_option_ref auto_methods} controls automatic use of a
- combination of standard proof methods (@{method auto}, @{method
- simp}, @{method blast}, etc.). This corresponds to the Isar command
- @{command_ref "try0"} @{cite "isabelle-isar-ref"}.
+ \<^item> @{system_option_ref auto_methods} controls automatic use of a combination
+ of standard proof methods (@{method auto}, @{method simp}, @{method blast},
+ etc.). This corresponds to the Isar command @{command_ref "try0"} @{cite
+ "isabelle-isar-ref"}.
The tool is disabled by default, since unparameterized invocation of
- standard proof methods often consumes substantial CPU resources
- without leading to success.
+ standard proof methods often consumes substantial CPU resources without
+ leading to success.
- \<^item> @{system_option_ref auto_nitpick} controls a slightly reduced
- version of @{command_ref nitpick}, which tests for counterexamples using
- first-order relational logic. See also the Nitpick manual
- @{cite "isabelle-nitpick"}.
+ \<^item> @{system_option_ref auto_nitpick} controls a slightly reduced version of
+ @{command_ref nitpick}, which tests for counterexamples using first-order
+ relational logic. See also the Nitpick manual @{cite "isabelle-nitpick"}.
- This tool is disabled by default, due to the extra overhead of
- invoking an external Java process for each attempt to disprove a
- subgoal.
+ This tool is disabled by default, due to the extra overhead of invoking an
+ external Java process for each attempt to disprove a subgoal.
\<^item> @{system_option_ref auto_quickcheck} controls automatic use of
- @{command_ref quickcheck}, which tests for counterexamples using a
- series of assignments for free variables of a subgoal.
+ @{command_ref quickcheck}, which tests for counterexamples using a series of
+ assignments for free variables of a subgoal.
- This tool is \<^emph>\<open>enabled\<close> by default. It requires little
- overhead, but is a bit weaker than @{command nitpick}.
+ This tool is \<^emph>\<open>enabled\<close> by default. It requires little overhead, but is a
+ bit weaker than @{command nitpick}.
- \<^item> @{system_option_ref auto_sledgehammer} controls a significantly
- reduced version of @{command_ref sledgehammer}, which attempts to prove
- a subgoal using external automatic provers. See also the
- Sledgehammer manual @{cite "isabelle-sledgehammer"}.
+ \<^item> @{system_option_ref auto_sledgehammer} controls a significantly reduced
+ version of @{command_ref sledgehammer}, which attempts to prove a subgoal
+ using external automatic provers. See also the Sledgehammer manual @{cite
+ "isabelle-sledgehammer"}.
- This tool is disabled by default, due to the relatively heavy nature
- of Sledgehammer.
+ This tool is disabled by default, due to the relatively heavy nature of
+ Sledgehammer.
\<^item> @{system_option_ref auto_solve_direct} controls automatic use of
- @{command_ref solve_direct}, which checks whether the current subgoals
- can be solved directly by an existing theorem. This also helps to
- detect duplicate lemmas.
+ @{command_ref solve_direct}, which checks whether the current subgoals can
+ be solved directly by an existing theorem. This also helps to detect
+ duplicate lemmas.
This tool is \<^emph>\<open>enabled\<close> by default.
- Invocation of automatically tried tools is subject to some global
- policies of parallel execution, which may be configured as follows:
+ Invocation of automatically tried tools is subject to some global policies
+ of parallel execution, which may be configured as follows:
- \<^item> @{system_option_ref auto_time_limit} (default 2.0) determines the
- timeout (in seconds) for each tool execution.
+ \<^item> @{system_option_ref auto_time_limit} (default 2.0) determines the timeout
+ (in seconds) for each tool execution.
- \<^item> @{system_option_ref auto_time_start} (default 1.0) determines the
- start delay (in seconds) for automatically tried tools, after the
- main command evaluation is finished.
+ \<^item> @{system_option_ref auto_time_start} (default 1.0) determines the start
+ delay (in seconds) for automatically tried tools, after the main command
+ evaluation is finished.
- Each tool is submitted independently to the pool of parallel
- execution tasks in Isabelle/ML, using hardwired priorities according
- to its relative ``heaviness''. The main stages of evaluation and
- printing of proof states take precedence, but an already running
- tool is not canceled and may thus reduce reactivity of proof
- document processing.
+ Each tool is submitted independently to the pool of parallel execution tasks
+ in Isabelle/ML, using hardwired priorities according to its relative
+ ``heaviness''. The main stages of evaluation and printing of proof states
+ take precedence, but an already running tool is not canceled and may thus
+ reduce reactivity of proof document processing.
- Users should experiment how the available CPU resources (number of
- cores) are best invested to get additional feedback from prover in
- the background, by using a selection of weaker or stronger tools.
+ Users should experiment how the available CPU resources (number of cores)
+ are best invested to get additional feedback from prover in the background,
+ by using a selection of weaker or stronger tools.
\<close>
section \<open>Sledgehammer \label{sec:sledgehammer}\<close>
-text \<open>The \<^emph>\<open>Sledgehammer\<close> panel (\figref{fig:sledgehammer})
- provides a view on some independent execution of the Isar command
- @{command_ref sledgehammer}, with process indicator (spinning wheel) and
- GUI elements for important Sledgehammer arguments and options. Any
- number of Sledgehammer panels may be active, according to the
- standard policies of Dockable Window Management in jEdit. Closing
- such windows also cancels the corresponding prover tasks.
+text \<open>
+ The \<^emph>\<open>Sledgehammer\<close> panel (\figref{fig:sledgehammer}) provides a view on
+ some independent execution of the Isar command @{command_ref sledgehammer},
+ with process indicator (spinning wheel) and GUI elements for important
+ Sledgehammer arguments and options. Any number of Sledgehammer panels may be
+ active, according to the standard policies of Dockable Window Management in
+ jEdit. Closing such windows also cancels the corresponding prover tasks.
\begin{figure}[htb]
\begin{center}
@@ -1603,34 +1564,37 @@
\label{fig:sledgehammer}
\end{figure}
- The \<^emph>\<open>Apply\<close> button attaches a fresh invocation of @{command
- sledgehammer} to the command where the cursor is pointing in the
- text --- this should be some pending proof problem. Further buttons
- like \<^emph>\<open>Cancel\<close> and \<^emph>\<open>Locate\<close> help to manage the running
- process.
+ The \<^emph>\<open>Apply\<close> button attaches a fresh invocation of @{command sledgehammer}
+ to the command where the cursor is pointing in the text --- this should be
+ some pending proof problem. Further buttons like \<^emph>\<open>Cancel\<close> and \<^emph>\<open>Locate\<close>
+ help to manage the running process.
- Results appear incrementally in the output window of the panel.
- Proposed proof snippets are marked-up as \<^emph>\<open>sendback\<close>, which
- means a single mouse click inserts the text into a suitable place of
- the original source. Some manual editing may be required
- nonetheless, say to remove earlier proof attempts.\<close>
+ Results appear incrementally in the output window of the panel. Proposed
+ proof snippets are marked-up as \<^emph>\<open>sendback\<close>, which means a single mouse
+ click inserts the text into a suitable place of the original source. Some
+ manual editing may be required nonetheless, say to remove earlier proof
+ attempts.
+\<close>
chapter \<open>Isabelle document preparation\<close>
-text \<open>The ultimate purpose of Isabelle is to produce nicely rendered documents
+text \<open>
+ The ultimate purpose of Isabelle is to produce nicely rendered documents
with the Isabelle document preparation system, which is based on {\LaTeX};
see also @{cite "isabelle-system" and "isabelle-isar-ref"}. Isabelle/jEdit
- provides some additional support for document editing.\<close>
+ provides some additional support for document editing.
+\<close>
section \<open>Document outline\<close>
-text \<open>Theory sources may contain document markup commands, such as
- @{command_ref chapter}, @{command_ref section}, @{command subsection}. The
- Isabelle SideKick parser (\secref{sec:sidekick}) represents this document
- outline as structured tree view, with formal statements and proofs nested
- inside; see \figref{fig:sidekick-document}.
+text \<open>
+ Theory sources may contain document markup commands, such as @{command_ref
+ chapter}, @{command_ref section}, @{command subsection}. The Isabelle
+ SideKick parser (\secref{sec:sidekick}) represents this document outline as
+ structured tree view, with formal statements and proofs nested inside; see
+ \figref{fig:sidekick-document}.
\begin{figure}[htb]
\begin{center}
@@ -1641,25 +1605,27 @@
\end{figure}
It is also possible to use text folding according to this structure, by
- adjusting \<^emph>\<open>Utilities / Buffer Options / Folding mode\<close> of jEdit. The
- default mode \<^verbatim>\<open>isabelle\<close> uses the structure of formal definitions,
- statements, and proofs. The alternative mode \<^verbatim>\<open>sidekick\<close> uses the
- document structure of the SideKick parser, as explained above.\<close>
+ adjusting \<^emph>\<open>Utilities / Buffer Options / Folding mode\<close> of jEdit. The default
+ mode \<^verbatim>\<open>isabelle\<close> uses the structure of formal definitions, statements, and
+ proofs. The alternative mode \<^verbatim>\<open>sidekick\<close> uses the document structure of the
+ SideKick parser, as explained above.
+\<close>
section \<open>Citations and Bib{\TeX} entries\<close>
-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
+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-system"} are smart enough to assemble the result, based on the
session directory layout.
The document antiquotation \<open>@{cite}\<close> is described in @{cite
"isabelle-isar-ref"}. Within the Prover IDE it provides semantic markup for
tooltips, hyperlinks, and completion for Bib{\TeX} database entries.
- Isabelle/jEdit does \<^emph>\<open>not\<close> know about the actual Bib{\TeX} environment
- used in {\LaTeX} batch-mode, but it can take citations from those \<^verbatim>\<open>.bib\<close>
- files that happen to be open in the editor; see \figref{fig:cite-completion}.
+ Isabelle/jEdit does \<^emph>\<open>not\<close> know about the actual Bib{\TeX} environment used
+ in {\LaTeX} batch-mode, but it can take citations from those \<^verbatim>\<open>.bib\<close> files
+ that happen to be open in the editor; see \figref{fig:cite-completion}.
\begin{figure}[htb]
\begin{center}
@@ -1669,9 +1635,9 @@
\label{fig:cite-completion}
\end{figure}
- Isabelle/jEdit also provides some support for editing \<^verbatim>\<open>.bib\<close>
- files themselves. There is syntax highlighting based on entry types
- (according to standard Bib{\TeX} styles), a context-menu to compose entries
+ Isabelle/jEdit also provides some support for editing \<^verbatim>\<open>.bib\<close> files
+ themselves. There is syntax highlighting based on entry types (according to
+ standard Bib{\TeX} styles), a context-menu to compose entries
systematically, and a SideKick tree view of the overall content; see
\figref{fig:bibtex-mode}.
@@ -1689,35 +1655,33 @@
section \<open>Timing\<close>
-text \<open>Managed evaluation of commands within PIDE documents includes
- timing information, which consists of elapsed (wall-clock) time, CPU
- time, and GC (garbage collection) time. Note that in a
- multithreaded system it is difficult to measure execution time
- precisely: elapsed time is closer to the real requirements of
- runtime resources than CPU or GC time, which are both subject to
- influences from the parallel environment that are outside the scope
- of the current command transaction.
+text \<open>
+ Managed evaluation of commands within PIDE documents includes timing
+ information, which consists of elapsed (wall-clock) time, CPU time, and GC
+ (garbage collection) time. Note that in a multithreaded system it is
+ difficult to measure execution time precisely: elapsed time is closer to the
+ real requirements of runtime resources than CPU or GC time, which are both
+ subject to influences from the parallel environment that are outside the
+ scope of the current command transaction.
- The \<^emph>\<open>Timing\<close> panel provides an overview of cumulative command
- timings for each document node. Commands with elapsed time below
- the given threshold are ignored in the grand total. Nodes are
- sorted according to their overall timing. For the document node
- that corresponds to the current buffer, individual command timings
- are shown as well. A double-click on a theory node or command moves
- the editor focus to that particular source position.
+ The \<^emph>\<open>Timing\<close> panel provides an overview of cumulative command timings for
+ each document node. Commands with elapsed time below the given threshold are
+ ignored in the grand total. Nodes are sorted according to their overall
+ timing. For the document node that corresponds to the current buffer,
+ individual command timings are shown as well. A double-click on a theory
+ node or command moves the editor focus to that particular source position.
- It is also possible to reveal individual timing information via some
- tooltip for the corresponding command keyword, using the technique
- of mouse hovering with \<^verbatim>\<open>CONTROL\<close>/\<^verbatim>\<open>COMMAND\<close>
- modifier key as explained in \secref{sec:tooltips-hyperlinks}.
- Actual display of timing depends on the global option
- @{system_option_ref jedit_timing_threshold}, which can be configured in
- \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>.
+ It is also possible to reveal individual timing information via some tooltip
+ for the corresponding command keyword, using the technique of mouse hovering
+ with \<^verbatim>\<open>CONTROL\<close>~/ \<^verbatim>\<open>COMMAND\<close> modifier key as explained in
+ \secref{sec:tooltips-hyperlinks}. Actual display of timing depends on the
+ global option @{system_option_ref jedit_timing_threshold}, which can be
+ configured in \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>.
\<^medskip>
- The \<^emph>\<open>Monitor\<close> panel visualizes various data collections about
- recent activity of the Isabelle/ML task farm and the underlying ML runtime
- system. The display is continuously updated according to @{system_option_ref
+ The \<^emph>\<open>Monitor\<close> panel visualizes various data collections about recent
+ activity of the Isabelle/ML task farm and the underlying ML runtime system.
+ The display is continuously updated according to @{system_option_ref
editor_chart_delay}. Note that the painting of the chart takes considerable
runtime itself --- on the Java Virtual Machine that runs Isabelle/Scala, not
Isabelle/ML. Internally, the Isabelle/Scala module \<^verbatim>\<open>isabelle.ML_Statistics\<close>
@@ -1727,32 +1691,30 @@
section \<open>Low-level output\<close>
-text \<open>Prover output is normally shown directly in the main text area
- or secondary \<^emph>\<open>Output\<close> panels, as explained in
- \secref{sec:output}.
+text \<open>
+ Prover output is normally shown directly in the main text area or secondary
+ \<^emph>\<open>Output\<close> panels, as explained in \secref{sec:output}.
- Beyond this, it is occasionally useful to inspect low-level output
- channels via some of the following additional panels:
+ Beyond this, it is occasionally useful to inspect low-level output channels
+ via some of the following additional panels:
- \<^item> \<^emph>\<open>Protocol\<close> shows internal messages between the
- Isabelle/Scala and Isabelle/ML side of the PIDE document editing protocol.
- Recording of messages starts with the first activation of the
- corresponding dockable window; earlier messages are lost.
+ \<^item> \<^emph>\<open>Protocol\<close> shows internal messages between the Isabelle/Scala and
+ Isabelle/ML side of the PIDE document editing protocol. Recording of
+ messages starts with the first activation of the corresponding dockable
+ window; earlier messages are lost.
- Actual display of protocol messages causes considerable slowdown, so
- it is important to undock all \<^emph>\<open>Protocol\<close> panels for production
- work.
+ Actual display of protocol messages causes considerable slowdown, so it is
+ important to undock all \<^emph>\<open>Protocol\<close> panels for production work.
\<^item> \<^emph>\<open>Raw Output\<close> shows chunks of text from the \<^verbatim>\<open>stdout\<close> and \<^verbatim>\<open>stderr\<close>
- channels of the prover process.
- Recording of output starts with the first activation of the
- corresponding dockable window; earlier output is lost.
+ channels of the prover process. Recording of output starts with the first
+ activation of the corresponding dockable window; earlier output is lost.
- The implicit stateful nature of physical I/O channels makes it
- difficult to relate raw output to the actual command from where it
- was originating. Parallel execution may add to the confusion.
- Peeking at physical process I/O is only the last resort to diagnose
- problems with tools that are not PIDE compliant.
+ The implicit stateful nature of physical I/O channels makes it difficult to
+ relate raw output to the actual command from where it was originating.
+ Parallel execution may add to the confusion. Peeking at physical process I/O
+ is only the last resort to diagnose problems with tools that are not PIDE
+ compliant.
Under normal circumstances, prover output always works via managed message
channels (corresponding to @{ML writeln}, @{ML warning}, @{ML
@@ -1762,50 +1724,45 @@
\<^item> \<^emph>\<open>Syslog\<close> shows system messages that might be relevant to diagnose
problems with the startup or shutdown phase of the prover process; this also
- includes raw output on \<^verbatim>\<open>stderr\<close>. Isabelle/ML also provides an
- explicit @{ML Output.system_message} operation, which is occasionally useful
- for diagnostic purposes within the system infrastructure itself.
+ includes raw output on \<^verbatim>\<open>stderr\<close>. Isabelle/ML also provides an explicit @{ML
+ Output.system_message} operation, which is occasionally useful for
+ diagnostic purposes within the system infrastructure itself.
- A limited amount of syslog messages are buffered, independently of
- the docking state of the \<^emph>\<open>Syslog\<close> panel. This allows to
- diagnose serious problems with Isabelle/PIDE process management,
- outside of the actual protocol layer.
+ A limited amount of syslog messages are buffered, independently of the
+ docking state of the \<^emph>\<open>Syslog\<close> panel. This allows to diagnose serious
+ problems with Isabelle/PIDE process management, outside of the actual
+ protocol layer.
- Under normal situations, such low-level system output can be
- ignored.
+ Under normal situations, such low-level system output can be ignored.
\<close>
chapter \<open>Known problems and workarounds \label{sec:problems}\<close>
text \<open>
- \<^item> \<^bold>\<open>Problem:\<close> Odd behavior of some diagnostic commands with
- global side-effects, like writing a physical file.
+ \<^item> \<^bold>\<open>Problem:\<close> Odd behavior of some diagnostic commands with global
+ side-effects, like writing a physical file.
- \<^bold>\<open>Workaround:\<close> Copy/paste complete command text from
- elsewhere, or disable continuous checking temporarily.
-
- \<^item> \<^bold>\<open>Problem:\<close> No direct support to remove document nodes from the
- collection of theories.
+ \<^bold>\<open>Workaround:\<close> Copy/paste complete command text from elsewhere, or disable
+ continuous checking temporarily.
- \<^bold>\<open>Workaround:\<close> Clear the buffer content of unused files and close
- \<^emph>\<open>without\<close> saving changes.
+ \<^item> \<^bold>\<open>Problem:\<close> No direct support to remove document nodes from the collection
+ of theories.
- \<^item> \<^bold>\<open>Problem:\<close> Keyboard shortcuts \<^verbatim>\<open>C+PLUS\<close> and
- \<^verbatim>\<open>C+MINUS\<close> for adjusting the editor font size depend on
- platform details and national keyboards.
+ \<^bold>\<open>Workaround:\<close> Clear the buffer content of unused files and close \<^emph>\<open>without\<close>
+ saving changes.
- \<^bold>\<open>Workaround:\<close> Rebind keys via \<^emph>\<open>Global Options~/
- Shortcuts\<close>.
+ \<^item> \<^bold>\<open>Problem:\<close> Keyboard shortcuts \<^verbatim>\<open>C+PLUS\<close> and \<^verbatim>\<open>C+MINUS\<close> for adjusting the
+ editor font size depend on platform details and national keyboards.
+
+ \<^bold>\<open>Workaround:\<close> Rebind keys via \<^emph>\<open>Global Options~/ Shortcuts\<close>.
\<^item> \<^bold>\<open>Problem:\<close> The Mac OS X key sequence \<^verbatim>\<open>COMMAND+COMMA\<close> for application
- \<^emph>\<open>Preferences\<close> is in conflict with the
- jEdit default keyboard shortcut for \<^emph>\<open>Incremental Search Bar\<close> (action
- @{action_ref "quick-search"}).
+ \<^emph>\<open>Preferences\<close> is in conflict with the jEdit default keyboard shortcut for
+ \<^emph>\<open>Incremental Search Bar\<close> (action @{action_ref "quick-search"}).
- \<^bold>\<open>Workaround:\<close> Rebind key via \<^emph>\<open>Global Options~/
- Shortcuts\<close> according to national keyboard, e.g.\ \<^verbatim>\<open>COMMAND+SLASH\<close>
- on English ones.
+ \<^bold>\<open>Workaround:\<close> Rebind key via \<^emph>\<open>Global Options~/ Shortcuts\<close> according to
+ national keyboard, e.g.\ \<^verbatim>\<open>COMMAND+SLASH\<close> on English ones.
\<^item> \<^bold>\<open>Problem:\<close> On Mac OS X with native Apple look-and-feel, some exotic
national keyboards may cause a conflict of menu accelerator keys with
@@ -1815,44 +1772,42 @@
\<^bold>\<open>Workaround:\<close> Disable the native Apple menu bar via Java runtime option
\<^verbatim>\<open>-Dapple.laf.useScreenMenuBar=false\<close>.
- \<^item> \<^bold>\<open>Problem:\<close> Mac OS X system fonts sometimes lead to
- character drop-outs in the main text area.
+ \<^item> \<^bold>\<open>Problem:\<close> Mac OS X system fonts sometimes lead to character drop-outs in
+ the main text area.
- \<^bold>\<open>Workaround:\<close> Use the default \<^verbatim>\<open>IsabelleText\<close> font.
- (Do not install that font on the system.)
-
- \<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 input methods such as IBus
- tend to disrupt key event handling of Java/AWT/Swing.
+ \<^bold>\<open>Workaround:\<close> Use the default \<^verbatim>\<open>IsabelleText\<close> font. (Do not install that
+ font on the system.)
- \<^bold>\<open>Workaround:\<close> Do not use X11 input methods. Note that environment
- variable \<^verbatim>\<open>XMODIFIERS\<close> is reset by default within Isabelle
- settings.
+ \<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 input methods such as IBus tend to disrupt key
+ event handling of Java/AWT/Swing.
- \<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 window managers that are
- not ``re-parenting'' cause problems with additional windows opened
- by Java. This affects either historic or neo-minimalistic window
- managers like \<^verbatim>\<open>awesome\<close> or \<^verbatim>\<open>xmonad\<close>.
+ \<^bold>\<open>Workaround:\<close> Do not use X11 input methods. Note that environment variable
+ \<^verbatim>\<open>XMODIFIERS\<close> is reset by default within Isabelle settings.
+
+ \<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 window managers that are not ``re-parenting''
+ cause problems with additional windows opened by Java. This affects either
+ historic or neo-minimalistic window managers like \<^verbatim>\<open>awesome\<close> or \<^verbatim>\<open>xmonad\<close>.
\<^bold>\<open>Workaround:\<close> Use a regular re-parenting X11 window manager.
- \<^item> \<^bold>\<open>Problem:\<close> Various forks of Linux/X11 window managers and
- desktop environments (like Gnome) disrupt the handling of menu popups and
- mouse positions of Java/AWT/Swing.
+ \<^item> \<^bold>\<open>Problem:\<close> Various forks of Linux/X11 window managers and desktop
+ environments (like Gnome) disrupt the handling of menu popups and mouse
+ positions of Java/AWT/Swing.
\<^bold>\<open>Workaround:\<close> Use mainstream versions of Linux desktops.
- \<^item> \<^bold>\<open>Problem:\<close> Native Windows look-and-feel with global font
- scaling leads to bad GUI rendering of various tree views.
+ \<^item> \<^bold>\<open>Problem:\<close> Native Windows look-and-feel with global font scaling leads to
+ bad GUI rendering of various tree views.
- \<^bold>\<open>Workaround:\<close> Use \<^emph>\<open>Metal\<close> look-and-feel and re-adjust its
- primary and secondary font as explained in \secref{sec:hdpi}.
+ \<^bold>\<open>Workaround:\<close> Use \<^emph>\<open>Metal\<close> look-and-feel and re-adjust its primary and
+ secondary font as explained in \secref{sec:hdpi}.
\<^item> \<^bold>\<open>Problem:\<close> Full-screen mode via jEdit action @{action_ref
- "toggle-full-screen"} (default keyboard shortcut \<^verbatim>\<open>F11\<close>) works on
- Windows, but not on Mac OS X or various Linux/X11 window managers.
+ "toggle-full-screen"} (default keyboard shortcut \<^verbatim>\<open>F11\<close>) works on Windows,
+ but not on Mac OS X or various Linux/X11 window managers.
- \<^bold>\<open>Workaround:\<close> Use native full-screen control of the window
- manager (notably on Mac OS X).
+ \<^bold>\<open>Workaround:\<close> Use native full-screen control of the window manager (notably
+ on Mac OS X).
\<close>
end
\ No newline at end of file
--- a/src/Doc/System/Basics.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/System/Basics.thy Thu Nov 05 08:27:22 2015 +0100
@@ -1,35 +1,34 @@
+(*:wrap=hard:maxLineLen=78:*)
+
theory Basics
imports Base
begin
chapter \<open>The Isabelle system environment\<close>
-text \<open>This manual describes Isabelle together with related tools and
- user interfaces as seen from a system oriented view. See also the
- \<^emph>\<open>Isabelle/Isar Reference Manual\<close> @{cite "isabelle-isar-ref"} for
- the actual Isabelle input language and related concepts, and
- \<^emph>\<open>The Isabelle/Isar Implementation
+text \<open>
+ This manual describes Isabelle together with related tools and user
+ interfaces as seen from a system oriented view. See also the \<^emph>\<open>Isabelle/Isar
+ Reference Manual\<close> @{cite "isabelle-isar-ref"} for the actual Isabelle input
+ language and related concepts, and \<^emph>\<open>The Isabelle/Isar Implementation
Manual\<close> @{cite "isabelle-implementation"} for the main concepts of the
underlying implementation in Isabelle/ML.
\<^medskip>
- The Isabelle system environment provides the following
- basic infrastructure to integrate tools smoothly.
+ The Isabelle system environment provides the following basic infrastructure
+ to integrate tools smoothly.
- \<^enum> The \<^emph>\<open>Isabelle settings\<close> mechanism provides process
- environment variables to all Isabelle executables (including tools
- and user interfaces).
+ \<^enum> The \<^emph>\<open>Isabelle settings\<close> mechanism provides process environment variables
+ to all Isabelle executables (including tools and user interfaces).
- \<^enum> The raw \<^emph>\<open>Isabelle process\<close> (@{executable_ref
- "isabelle_process"}) runs logic sessions either interactively or in
- batch mode. In particular, this view abstracts over the invocation
- of the actual ML system to be used. Regular users rarely need to
- care about the low-level process.
+ \<^enum> The raw \<^emph>\<open>Isabelle process\<close> (@{executable_ref "isabelle_process"}) runs
+ logic sessions either interactively or in batch mode. In particular, this
+ view abstracts over the invocation of the actual ML system to be used.
+ Regular users rarely need to care about the low-level process.
- \<^enum> The main \<^emph>\<open>Isabelle tool wrapper\<close> (@{executable_ref
- isabelle}) provides a generic startup environment Isabelle related
- utilities, user interfaces etc. Such tools automatically benefit
- from the settings mechanism.
+ \<^enum> The main \<^emph>\<open>Isabelle tool wrapper\<close> (@{executable_ref isabelle}) provides a
+ generic startup environment Isabelle related utilities, user interfaces etc.
+ Such tools automatically benefit from the settings mechanism.
\<close>
@@ -37,299 +36,270 @@
text \<open>
The Isabelle system heavily depends on the \<^emph>\<open>settings
- mechanism\<close>\indexbold{settings}. Essentially, this is a statically
- scoped collection of environment variables, such as @{setting
- ISABELLE_HOME}, @{setting ML_SYSTEM}, @{setting ML_HOME}. These
- variables are \<^emph>\<open>not\<close> intended to be set directly from the shell,
- though. Isabelle employs a somewhat more sophisticated scheme of
- \<^emph>\<open>settings files\<close> --- one for site-wide defaults, another for
- additional user-specific modifications. With all configuration
- variables in clearly defined places, this scheme is more
- maintainable and user-friendly than global shell environment
- variables.
+ mechanism\<close>\indexbold{settings}. Essentially, this is a statically scoped
+ collection of environment variables, such as @{setting ISABELLE_HOME},
+ @{setting ML_SYSTEM}, @{setting ML_HOME}. These variables are \<^emph>\<open>not\<close>
+ intended to be set directly from the shell, though. Isabelle employs a
+ somewhat more sophisticated scheme of \<^emph>\<open>settings files\<close> --- one for
+ site-wide defaults, another for additional user-specific modifications. With
+ all configuration variables in clearly defined places, this scheme is more
+ maintainable and user-friendly than global shell environment variables.
- In particular, we avoid the typical situation where prospective
- users of a software package are told to put several things into
- their shell startup scripts, before being able to actually run the
- program. Isabelle requires none such administrative chores of its
- end-users --- the executables can be invoked straight away.
- Occasionally, users would still want to put the @{file
- "$ISABELLE_HOME/bin"} directory into their shell's search path, but
+ In particular, we avoid the typical situation where prospective users of a
+ software package are told to put several things into their shell startup
+ scripts, before being able to actually run the program. Isabelle requires
+ none such administrative chores of its end-users --- the executables can be
+ invoked straight away. Occasionally, users would still want to put the
+ @{file "$ISABELLE_HOME/bin"} directory into their shell's search path, but
this is not required.
\<close>
subsection \<open>Bootstrapping the environment \label{sec:boot}\<close>
-text \<open>Isabelle executables need to be run within a proper settings
- environment. This is bootstrapped as described below, on the first
- invocation of one of the outer wrapper scripts (such as
- @{executable_ref isabelle}). This happens only once for each
- process tree, i.e.\ the environment is passed to subprocesses
- according to regular Unix conventions.
+text \<open>
+ Isabelle executables need to be run within a proper settings environment.
+ This is bootstrapped as described below, on the first invocation of one of
+ the outer wrapper scripts (such as @{executable_ref isabelle}). This happens
+ only once for each process tree, i.e.\ the environment is passed to
+ subprocesses according to regular Unix conventions.
+
+ \<^enum> The special variable @{setting_def ISABELLE_HOME} is determined
+ automatically from the location of the binary that has been run.
- \<^enum> The special variable @{setting_def ISABELLE_HOME} is
- determined automatically from the location of the binary that has
- been run.
-
- You should not try to set @{setting ISABELLE_HOME} manually. Also
- note that the Isabelle executables either have to be run from their
- original location in the distribution directory, or via the
- executable objects created by the @{tool install} tool. Symbolic
- links are admissible, but a plain copy of the @{file
- "$ISABELLE_HOME/bin"} files will not work!
+ You should not try to set @{setting ISABELLE_HOME} manually. Also note
+ that the Isabelle executables either have to be run from their original
+ location in the distribution directory, or via the executable objects
+ created by the @{tool install} tool. Symbolic links are admissible, but a
+ plain copy of the @{file "$ISABELLE_HOME/bin"} files will not work!
+
+ \<^enum> The file @{file "$ISABELLE_HOME/etc/settings"} is run as a
+ @{executable_ref bash} shell script with the auto-export option for
+ variables enabled.
- \<^enum> The file @{file "$ISABELLE_HOME/etc/settings"} is run as a
- @{executable_ref bash} shell script with the auto-export option for
- variables enabled.
-
- This file holds a rather long list of shell variable assigments,
- thus providing the site-wide default settings. The Isabelle
- distribution already contains a global settings file with sensible
- defaults for most variables. When installing the system, only a few
- of these may have to be adapted (probably @{setting ML_SYSTEM}
- etc.).
-
- \<^enum> The file @{file_unchecked "$ISABELLE_HOME_USER/etc/settings"} (if it
- exists) is run in the same way as the site default settings. Note
- that the variable @{setting ISABELLE_HOME_USER} has already been set
- before --- usually to something like \<^verbatim>\<open>$USER_HOME/.isabelle/IsabelleXXXX\<close>.
-
- Thus individual users may override the site-wide defaults.
- Typically, a user settings file contains only a few lines, with some
- assignments that are actually changed. Never copy the central
- @{file "$ISABELLE_HOME/etc/settings"} file!
+ This file holds a rather long list of shell variable assignments, thus
+ providing the site-wide default settings. The Isabelle distribution
+ already contains a global settings file with sensible defaults for most
+ variables. When installing the system, only a few of these may have to be
+ adapted (probably @{setting ML_SYSTEM} etc.).
+
+ \<^enum> The file @{file_unchecked "$ISABELLE_HOME_USER/etc/settings"} (if it
+ exists) is run in the same way as the site default settings. Note that the
+ variable @{setting ISABELLE_HOME_USER} has already been set before ---
+ usually to something like \<^verbatim>\<open>$USER_HOME/.isabelle/IsabelleXXXX\<close>.
+ Thus individual users may override the site-wide defaults. Typically, a
+ user settings file contains only a few lines, with some assignments that
+ are actually changed. Never copy the central @{file
+ "$ISABELLE_HOME/etc/settings"} file!
- Since settings files are regular GNU @{executable_def bash} scripts,
- one may use complex shell commands, such as \<^verbatim>\<open>if\<close> or
- \<^verbatim>\<open>case\<close> statements to set variables depending on the
- system architecture or other environment variables. Such advanced
- features should be added only with great care, though. In
- particular, external environment references should be kept at a
+ Since settings files are regular GNU @{executable_def bash} scripts, one may
+ use complex shell commands, such as \<^verbatim>\<open>if\<close> or \<^verbatim>\<open>case\<close> statements to set
+ variables depending on the system architecture or other environment
+ variables. Such advanced features should be added only with great care,
+ though. In particular, external environment references should be kept at a
minimum.
\<^medskip>
A few variables are somewhat special:
- \<^item> @{setting_def ISABELLE_PROCESS} and @{setting_def ISABELLE_TOOL} are set
- automatically to the absolute path names of the @{executable
- "isabelle_process"} and @{executable isabelle} executables,
- respectively.
-
- \<^item> @{setting_ref ISABELLE_OUTPUT} will have the identifiers of
- the Isabelle distribution (cf.\ @{setting ISABELLE_IDENTIFIER}) and
- the ML system (cf.\ @{setting ML_IDENTIFIER}) appended automatically
- to its value.
+ \<^item> @{setting_def ISABELLE_PROCESS} and @{setting_def ISABELLE_TOOL} are set
+ automatically to the absolute path names of the @{executable
+ "isabelle_process"} and @{executable isabelle} executables, respectively.
+ \<^item> @{setting_ref ISABELLE_OUTPUT} will have the identifiers of the Isabelle
+ distribution (cf.\ @{setting ISABELLE_IDENTIFIER}) and the ML system (cf.\
+ @{setting ML_IDENTIFIER}) appended automatically to its value.
\<^medskip>
- Note that the settings environment may be inspected with
- the @{tool getenv} tool. This might help to figure out the effect
- of complex settings scripts.\<close>
+ Note that the settings environment may be inspected with the @{tool getenv}
+ tool. This might help to figure out the effect of complex settings scripts.
+\<close>
subsection \<open>Common variables\<close>
text \<open>
- This is a reference of common Isabelle settings variables. Note that
- the list is somewhat open-ended. Third-party utilities or interfaces
- may add their own selection. Variables that are special in some
- sense are marked with \<open>\<^sup>*\<close>.
+ This is a reference of common Isabelle settings variables. Note that the
+ list is somewhat open-ended. Third-party utilities or interfaces may add
+ their own selection. Variables that are special in some sense are marked
+ with \<open>\<^sup>*\<close>.
- \<^descr>[@{setting_def USER_HOME}\<open>\<^sup>*\<close>] Is the cross-platform
- user home directory. On Unix systems this is usually the same as
- @{setting HOME}, but on Windows it is the regular home directory of
- the user, not the one of within the Cygwin root
- file-system.\footnote{Cygwin itself offers another choice whether
- its HOME should point to the @{file_unchecked "/home"} directory tree or the
- Windows user home.}
+ \<^descr>[@{setting_def USER_HOME}\<open>\<^sup>*\<close>] Is the cross-platform user home directory.
+ On Unix systems this is usually the same as @{setting HOME}, but on Windows
+ it is the regular home directory of the user, not the one of within the
+ Cygwin root file-system.\<^footnote>\<open>Cygwin itself offers another choice whether its
+ HOME should point to the @{file_unchecked "/home"} directory tree or the
+ Windows user home.\<close>
- \<^descr>[@{setting_def ISABELLE_HOME}\<open>\<^sup>*\<close>] is the location of the
- top-level Isabelle distribution directory. This is automatically
- determined from the Isabelle executable that has been invoked. Do
- not attempt to set @{setting ISABELLE_HOME} yourself from the shell!
-
- \<^descr>[@{setting_def ISABELLE_HOME_USER}] is the user-specific
- counterpart of @{setting ISABELLE_HOME}. The default value is
- relative to @{file_unchecked "$USER_HOME/.isabelle"}, under rare
- circumstances this may be changed in the global setting file.
- Typically, the @{setting ISABELLE_HOME_USER} directory mimics
- @{setting ISABELLE_HOME} to some extend. In particular, site-wide
- defaults may be overridden by a private
- \<^verbatim>\<open>$ISABELLE_HOME_USER/etc/settings\<close>.
+ \<^descr>[@{setting_def ISABELLE_HOME}\<open>\<^sup>*\<close>] is the location of the top-level
+ Isabelle distribution directory. This is automatically determined from the
+ Isabelle executable that has been invoked. Do not attempt to set @{setting
+ ISABELLE_HOME} yourself from the shell!
- \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is
- automatically set to the general platform family: \<^verbatim>\<open>linux\<close>,
- \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that
+ \<^descr>[@{setting_def ISABELLE_HOME_USER}] is the user-specific counterpart of
+ @{setting ISABELLE_HOME}. The default value is relative to @{file_unchecked
+ "$USER_HOME/.isabelle"}, under rare circumstances this may be changed in the
+ global setting file. Typically, the @{setting ISABELLE_HOME_USER} directory
+ mimics @{setting ISABELLE_HOME} to some extend. In particular, site-wide
+ defaults may be overridden by a private \<^verbatim>\<open>$ISABELLE_HOME_USER/etc/settings\<close>.
+
+ \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is automatically set to the
+ general platform family: \<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that
platform-dependent tools usually need to refer to the more specific
identification according to @{setting ISABELLE_PLATFORM}, @{setting
ISABELLE_PLATFORM32}, @{setting ISABELLE_PLATFORM64}.
- \<^descr>[@{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] is automatically
- set to a symbolic identifier for the underlying hardware and
- operating system. The Isabelle platform identification always
- refers to the 32 bit variant, even this is a 64 bit machine. Note
- that the ML or Java runtime may have a different idea, depending on
- which binaries are actually run.
+ \<^descr>[@{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] is automatically set to a symbolic
+ identifier for the underlying hardware and operating system. The Isabelle
+ platform identification always refers to the 32 bit variant, even this is a
+ 64 bit machine. Note that the ML or Java runtime may have a different idea,
+ depending on which binaries are actually run.
- \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>] is similar to
- @{setting ISABELLE_PLATFORM} but refers to the proper 64 bit variant
- on a platform that supports this; the value is empty for 32 bit.
- Note that the following bash expression (including the quotes)
- prefers the 64 bit platform, if that is available:
+ \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>] is similar to @{setting
+ ISABELLE_PLATFORM} but refers to the proper 64 bit variant on a platform
+ that supports this; the value is empty for 32 bit. Note that the following
+ bash expression (including the quotes) prefers the 64 bit platform, if that
+ is available:
@{verbatim [display] \<open>"${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}"\<close>}
- \<^descr>[@{setting_def ISABELLE_PROCESS}\<open>\<^sup>*\<close>, @{setting
- ISABELLE_TOOL}\<open>\<^sup>*\<close>] are automatically set to the full path
- names of the @{executable "isabelle_process"} and @{executable
- isabelle} executables, respectively. Thus other tools and scripts
- need not assume that the @{file "$ISABELLE_HOME/bin"} directory is
- on the current search path of the shell.
-
- \<^descr>[@{setting_def ISABELLE_IDENTIFIER}\<open>\<^sup>*\<close>] refers
- to the name of this Isabelle distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012\<close>''.
+ \<^descr>[@{setting_def ISABELLE_PROCESS}\<open>\<^sup>*\<close>, @{setting ISABELLE_TOOL}\<open>\<^sup>*\<close>] are
+ automatically set to the full path names of the @{executable
+ "isabelle_process"} and @{executable isabelle} executables, respectively.
+ Thus other tools and scripts need not assume that the @{file
+ "$ISABELLE_HOME/bin"} directory is on the current search path of the shell.
- \<^descr>[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME},
- @{setting_def ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def
- ML_IDENTIFIER}\<open>\<^sup>*\<close>] specify the underlying ML system
- to be used for Isabelle. There is only a fixed set of admissable
- @{setting ML_SYSTEM} names (see the @{file
+ \<^descr>[@{setting_def ISABELLE_IDENTIFIER}\<open>\<^sup>*\<close>] refers to the name of this
+ Isabelle distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012\<close>''.
+
+ \<^descr>[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME}, @{setting_def
+ ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def ML_IDENTIFIER}\<open>\<^sup>*\<close>]
+ specify the underlying ML system to be used for Isabelle. There is only a
+ fixed set of admissable @{setting ML_SYSTEM} names (see the @{file
"$ISABELLE_HOME/etc/settings"} file of the distribution).
-
+
The actual compiler binary will be run from the directory @{setting
- ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the
- command line. The optional @{setting ML_PLATFORM} may specify the
- binary format of ML heap images, which is useful for cross-platform
- installations. The value of @{setting ML_IDENTIFIER} is
- automatically obtained by composing the values of @{setting
- ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version values.
+ ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the command line.
+ The optional @{setting ML_PLATFORM} may specify the binary format of ML heap
+ images, which is useful for cross-platform installations. The value of
+ @{setting ML_IDENTIFIER} is automatically obtained by composing the values
+ of @{setting ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version
+ values.
- \<^descr>[@{setting_def ML_SYSTEM_POLYML}\<open>\<^sup>*\<close>] is \<^verbatim>\<open>true\<close>
- for @{setting ML_SYSTEM} values derived from Poly/ML, as opposed to
- SML/NJ where it is empty. This is particularly useful with the
- build option @{system_option condition}
- (\secref{sec:system-options}) to restrict big sessions to something
- that SML/NJ can still handle.
+ \<^descr>[@{setting_def ML_SYSTEM_POLYML}\<open>\<^sup>*\<close>] is \<^verbatim>\<open>true\<close> for @{setting ML_SYSTEM}
+ values derived from Poly/ML, as opposed to SML/NJ where it is empty. This is
+ particularly useful with the build option @{system_option condition}
+ (\secref{sec:system-options}) to restrict big sessions to something that
+ SML/NJ can still handle.
- \<^descr>[@{setting_def ISABELLE_JDK_HOME}] needs to point to a full JDK
- (Java Development Kit) installation with \<^verbatim>\<open>javac\<close> and
- \<^verbatim>\<open>jar\<close> executables. This is essential for Isabelle/Scala
- and other JVM-based tools to work properly. Note that conventional
- \<^verbatim>\<open>JAVA_HOME\<close> usually points to the JRE (Java Runtime
+ \<^descr>[@{setting_def ISABELLE_JDK_HOME}] needs to point to a full JDK (Java
+ Development Kit) installation with \<^verbatim>\<open>javac\<close> and \<^verbatim>\<open>jar\<close> executables. This is
+ essential for Isabelle/Scala and other JVM-based tools to work properly.
+ Note that conventional \<^verbatim>\<open>JAVA_HOME\<close> usually points to the JRE (Java Runtime
Environment), not JDK.
-
- \<^descr>[@{setting_def ISABELLE_PATH}] is a list of directories
- (separated by colons) where Isabelle logic images may reside. When
- looking up heaps files, the value of @{setting ML_IDENTIFIER} is
- appended to each component internally.
-
- \<^descr>[@{setting_def ISABELLE_OUTPUT}\<open>\<^sup>*\<close>] is a
- directory where output heap files should be stored by default. The
- ML system and Isabelle version identifier is appended here, too.
-
- \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where
- theory browser information (HTML text, graph data, and printable
- documents) is stored (see also \secref{sec:info}). The default
- value is @{file_unchecked "$ISABELLE_HOME_USER/browser_info"}.
-
- \<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to
- load if none is given explicitely by the user. The default value is
- \<^verbatim>\<open>HOL\<close>.
-
- \<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the
- line editor for the @{tool_ref console} interface.
+
+ \<^descr>[@{setting_def ISABELLE_PATH}] is a list of directories (separated by
+ colons) where Isabelle logic images may reside. When looking up heaps files,
+ the value of @{setting ML_IDENTIFIER} is appended to each component
+ internally.
+
+ \<^descr>[@{setting_def ISABELLE_OUTPUT}\<open>\<^sup>*\<close>] is a directory where output heap files
+ should be stored by default. The ML system and Isabelle version identifier
+ is appended here, too.
+
+ \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where theory
+ browser information (HTML text, graph data, and printable documents) is
+ stored (see also \secref{sec:info}). The default value is @{file_unchecked
+ "$ISABELLE_HOME_USER/browser_info"}.
+
+ \<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to load if none
+ is given explicitely by the user. The default value is \<^verbatim>\<open>HOL\<close>.
+
+ \<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the line editor for the
+ @{tool_ref console} interface.
- \<^descr>[@{setting_def ISABELLE_LATEX}, @{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_TOOLS}] is a colon separated list of
- directories that are scanned by @{executable isabelle} for external
- utility programs (see also \secref{sec:isabelle-tool}).
-
- \<^descr>[@{setting_def ISABELLE_DOCS}] is a colon separated list of
- directories with documentation files.
+ \<^descr>[@{setting_def ISABELLE_LATEX}, @{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_TOOLS}] is a colon separated list of directories
+ that are scanned by @{executable isabelle} for external utility programs
+ (see also \secref{sec:isabelle-tool}).
- \<^descr>[@{setting_def PDF_VIEWER}] specifies the program to be used
- for displaying \<^verbatim>\<open>pdf\<close> files.
+ \<^descr>[@{setting_def ISABELLE_DOCS}] is a colon separated list of directories
+ with documentation files.
+
+ \<^descr>[@{setting_def PDF_VIEWER}] specifies the program to be used for displaying
+ \<^verbatim>\<open>pdf\<close> files.
- \<^descr>[@{setting_def DVI_VIEWER}] specifies the program to be used
- for displaying \<^verbatim>\<open>dvi\<close> files.
-
- \<^descr>[@{setting_def ISABELLE_TMP_PREFIX}\<open>\<^sup>*\<close>] is the
- prefix from which any running @{executable "isabelle_process"}
- derives an individual directory for temporary files.
+ \<^descr>[@{setting_def DVI_VIEWER}] specifies the program to be used for displaying
+ \<^verbatim>\<open>dvi\<close> files.
+
+ \<^descr>[@{setting_def ISABELLE_TMP_PREFIX}\<open>\<^sup>*\<close>] is the prefix from which any
+ running @{executable "isabelle_process"} derives an individual directory for
+ temporary files.
\<close>
subsection \<open>Additional components \label{sec:components}\<close>
-text \<open>Any directory may be registered as an explicit \<^emph>\<open>Isabelle
- component\<close>. The general layout conventions are that of the main
- Isabelle distribution itself, and the following two files (both
- optional) have a special meaning:
+text \<open>
+ Any directory may be registered as an explicit \<^emph>\<open>Isabelle component\<close>. The
+ general layout conventions are that of the main Isabelle distribution
+ itself, and the following two files (both optional) have a special meaning:
- \<^item> \<^verbatim>\<open>etc/settings\<close> holds additional settings that are
- initialized when bootstrapping the overall Isabelle environment,
- cf.\ \secref{sec:boot}. As usual, the content is interpreted as a
- \<^verbatim>\<open>bash\<close> script. It may refer to the component's enclosing
- directory via the \<^verbatim>\<open>COMPONENT\<close> shell variable.
+ \<^item> \<^verbatim>\<open>etc/settings\<close> holds additional settings that are initialized when
+ bootstrapping the overall Isabelle environment, cf.\ \secref{sec:boot}. As
+ usual, the content is interpreted as a \<^verbatim>\<open>bash\<close> script. It may refer to the
+ component's enclosing directory via the \<^verbatim>\<open>COMPONENT\<close> shell variable.
- For example, the following setting allows to refer to files within
- the component later on, without having to hardwire absolute paths:
- @{verbatim [display] \<open>MY_COMPONENT_HOME="$COMPONENT"\<close>}
+ For example, the following setting allows to refer to files within the
+ component later on, without having to hardwire absolute paths:
+ @{verbatim [display] \<open>MY_COMPONENT_HOME="$COMPONENT"\<close>}
- Components can also add to existing Isabelle settings such as
- @{setting_def ISABELLE_TOOLS}, in order to provide
- component-specific tools that can be invoked by end-users. For
- example:
- @{verbatim [display] \<open>ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"\<close>}
+ Components can also add to existing Isabelle settings such as
+ @{setting_def ISABELLE_TOOLS}, in order to provide component-specific
+ tools that can be invoked by end-users. For example:
+ @{verbatim [display] \<open>ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"\<close>}
- \<^item> \<^verbatim>\<open>etc/components\<close> holds a list of further
- sub-components of the same structure. The directory specifications
- given here can be either absolute (with leading \<^verbatim>\<open>/\<close>) or
- relative to the component's main directory.
+ \<^item> \<^verbatim>\<open>etc/components\<close> holds a list of further sub-components of the same
+ structure. The directory specifications given here can be either absolute
+ (with leading \<^verbatim>\<open>/\<close>) or relative to the component's main directory.
-
- The root of component initialization is @{setting ISABELLE_HOME}
- itself. After initializing all of its sub-components recursively,
- @{setting ISABELLE_HOME_USER} is included in the same manner (if
- that directory exists). This allows to install private components
- via @{file_unchecked "$ISABELLE_HOME_USER/etc/components"}, although it is
- often more convenient to do that programmatically via the
- \<^verbatim>\<open>init_component\<close> shell function in the \<^verbatim>\<open>etc/settings\<close>
- script of \<^verbatim>\<open>$ISABELLE_HOME_USER\<close> (or any other component
- directory). For example:
+ The root of component initialization is @{setting ISABELLE_HOME} itself.
+ After initializing all of its sub-components recursively, @{setting
+ ISABELLE_HOME_USER} is included in the same manner (if that directory
+ exists). This allows to install private components via @{file_unchecked
+ "$ISABELLE_HOME_USER/etc/components"}, although it is often more convenient
+ to do that programmatically via the \<^verbatim>\<open>init_component\<close> shell function in the
+ \<^verbatim>\<open>etc/settings\<close> script of \<^verbatim>\<open>$ISABELLE_HOME_USER\<close> (or any other component
+ directory). For example:
@{verbatim [display] \<open>init_component "$HOME/screwdriver-2.0"\<close>}
- This is tolerant wrt.\ missing component directories, but might
- produce a warning.
+ This is tolerant wrt.\ missing component directories, but might produce a
+ warning.
\<^medskip>
- More complex situations may be addressed by initializing
- components listed in a given catalog file, relatively to some base
- directory:
+ More complex situations may be addressed by initializing components listed
+ in a given catalog file, relatively to some base directory:
@{verbatim [display] \<open>init_components "$HOME/my_component_store" "some_catalog_file"\<close>}
- The component directories listed in the catalog file are treated as
- relative to the given base directory.
+ The component directories listed in the catalog file are treated as relative
+ to the given base directory.
- See also \secref{sec:tool-components} for some tool-support for
- resolving components that are formally initialized but not installed
- yet.
+ See also \secref{sec:tool-components} for some tool-support for resolving
+ components that are formally initialized but not installed yet.
\<close>
section \<open>The raw Isabelle process \label{sec:isabelle-process}\<close>
text \<open>
- The @{executable_def "isabelle_process"} executable runs bare-bones
- Isabelle logic sessions --- either interactively or in batch mode.
- It provides an abstraction over the underlying ML system, and over
- the actual heap file locations. Its usage is:
+ The @{executable_def "isabelle_process"} executable runs bare-bones Isabelle
+ logic sessions --- either interactively or in batch mode. It provides an
+ abstraction over the underlying ML system, and over the actual heap file
+ locations. Its usage is:
@{verbatim [display]
\<open>Usage: isabelle_process [OPTIONS] [INPUT] [OUTPUT]
@@ -349,123 +319,111 @@
actual file names (containing at least one /).
If INPUT is "RAW_ML_SYSTEM", just start the bare bones ML system.\<close>}
- Input files without path specifications are looked up in the
- @{setting ISABELLE_PATH} setting, which may consist of multiple
- components separated by colons --- these are tried in the given
- order with the value of @{setting ML_IDENTIFIER} appended
- internally. In a similar way, base names are relative to the
- directory specified by @{setting ISABELLE_OUTPUT}. In any case,
- actual file locations may also be given by including at least one
- slash (\<^verbatim>\<open>/\<close>) in the name (hint: use \<^verbatim>\<open>./\<close> to
- refer to the current directory).
+ Input files without path specifications are looked up in the @{setting
+ ISABELLE_PATH} setting, which may consist of multiple components separated
+ by colons --- these are tried in the given order with the value of @{setting
+ ML_IDENTIFIER} appended internally. In a similar way, base names are
+ relative to the directory specified by @{setting ISABELLE_OUTPUT}. In any
+ case, actual file locations may also be given by including at least one
+ slash (\<^verbatim>\<open>/\<close>) in the name (hint: use \<^verbatim>\<open>./\<close> to refer to the current
+ directory).
\<close>
subsubsection \<open>Options\<close>
text \<open>
- If the input heap file does not have write permission bits set, or
- the \<^verbatim>\<open>-r\<close> option is given explicitly, then the session
- started will be read-only. That is, the ML world cannot be
- committed back into the image file. Otherwise, a writable session
- enables commits into either the input file, or into another output
- heap file (if that is given as the second argument on the command
+ If the input heap file does not have write permission bits set, or the \<^verbatim>\<open>-r\<close>
+ option is given explicitly, then the session started will be read-only. That
+ is, the ML world cannot be committed back into the image file. Otherwise, a
+ writable session enables commits into either the input file, or into another
+ output heap file (if that is given as the second argument on the command
line).
- The read-write state of sessions is determined at startup only, it
- cannot be changed intermediately. Also note that heap images may
- require considerable amounts of disk space (hundreds of MB or some GB).
- Users are responsible for themselves to dispose their heap files
- when they are no longer needed.
+ The read-write state of sessions is determined at startup only, it cannot be
+ changed intermediately. Also note that heap images may require considerable
+ amounts of disk space (hundreds of MB or some GB). Users are responsible for
+ themselves to dispose their heap files when they are no longer needed.
\<^medskip>
- The \<^verbatim>\<open>-w\<close> option makes the output heap file
- read-only after terminating. Thus subsequent invocations cause the
- logic image to be read-only automatically.
+ The \<^verbatim>\<open>-w\<close> option makes the output heap file read-only after terminating.
+ Thus subsequent invocations cause the logic image to be read-only
+ automatically.
\<^medskip>
- Using the \<^verbatim>\<open>-e\<close> option, arbitrary ML code may be
- passed to the Isabelle session from the command line. Multiple
- \<^verbatim>\<open>-e\<close> options are evaluated in the given order. Strange things
- may happen when erroneous ML code is provided. Also make sure that
- the ML commands are terminated properly by semicolon.
+ Using the \<^verbatim>\<open>-e\<close> option, arbitrary ML code may be passed to the Isabelle
+ session from the command line. Multiple \<^verbatim>\<open>-e\<close> options are evaluated in the
+ given order. Strange things may happen when erroneous ML code is provided.
+ Also make sure that the ML commands are terminated properly by semicolon.
\<^medskip>
- The \<^verbatim>\<open>-m\<close> option adds identifiers of print modes
- to be made active for this session. Typically, this is used by some
- user interface, e.g.\ to enable output of proper mathematical
- symbols.
+ The \<^verbatim>\<open>-m\<close> option adds identifiers of print modes to be made active for this
+ session. Typically, this is used by some user interface, e.g.\ to enable
+ output of proper mathematical symbols.
\<^medskip>
- Isabelle normally enters an interactive top-level loop
- (after processing the \<^verbatim>\<open>-e\<close> texts). The \<^verbatim>\<open>-q\<close>
- option inhibits interaction, thus providing a pure batch mode
- facility.
+ Isabelle normally enters an interactive top-level loop (after processing the
+ \<^verbatim>\<open>-e\<close> texts). The \<^verbatim>\<open>-q\<close> option inhibits interaction, thus providing a pure
+ batch mode facility.
\<^medskip>
- Option \<^verbatim>\<open>-o\<close> allows to override Isabelle system
- options for this process, see also \secref{sec:system-options}.
- Alternatively, option \<^verbatim>\<open>-O\<close> specifies the full environment of
- system options as a file in YXML format (according to the Isabelle/Scala
- function \<^verbatim>\<open>isabelle.Options.encode\<close>).
+ Option \<^verbatim>\<open>-o\<close> allows to override Isabelle system options for this process,
+ see also \secref{sec:system-options}. Alternatively, option \<^verbatim>\<open>-O\<close> specifies
+ the full environment of system options as a file in YXML format (according
+ to the Isabelle/Scala function \<^verbatim>\<open>isabelle.Options.encode\<close>).
\<^medskip>
- The \<^verbatim>\<open>-P\<close> option starts the Isabelle process wrapper
- for Isabelle/Scala, with a private protocol running over the specified TCP
- socket. Isabelle/ML and Isabelle/Scala provide various programming
- interfaces to invoke protocol functions over untyped strings and XML
- trees.
+ The \<^verbatim>\<open>-P\<close> option starts the Isabelle process wrapper for Isabelle/Scala,
+ with a private protocol running over the specified TCP socket. Isabelle/ML
+ and Isabelle/Scala provide various programming interfaces to invoke protocol
+ functions over untyped strings and XML trees.
\<^medskip>
- The \<^verbatim>\<open>-S\<close> option makes the Isabelle process more
- secure by disabling some critical operations, notably runtime
- compilation and evaluation of ML source code.
+ The \<^verbatim>\<open>-S\<close> option makes the Isabelle process more secure by disabling some
+ critical operations, notably runtime compilation and evaluation of ML source
+ code.
\<close>
subsubsection \<open>Examples\<close>
text \<open>
- Run an interactive session of the default object-logic (as specified
- by the @{setting ISABELLE_LOGIC} setting) like this:
+ Run an interactive session of the default object-logic (as specified by the
+ @{setting ISABELLE_LOGIC} setting) like this:
@{verbatim [display] \<open>isabelle_process\<close>}
- Usually @{setting ISABELLE_LOGIC} refers to one of the standard
- logic images, which are read-only by default. A writable session
- --- based on \<^verbatim>\<open>HOL\<close>, but output to \<^verbatim>\<open>Test\<close> (in the
- directory specified by the @{setting ISABELLE_OUTPUT} setting) ---
- may be invoked as follows:
+ Usually @{setting ISABELLE_LOGIC} refers to one of the standard logic
+ images, which are read-only by default. A writable session --- based on
+ \<^verbatim>\<open>HOL\<close>, but output to \<^verbatim>\<open>Test\<close> (in the directory specified by the @{setting
+ ISABELLE_OUTPUT} setting) --- may be invoked as follows:
@{verbatim [display] \<open>isabelle_process HOL Test\<close>}
- Ending this session normally (e.g.\ by typing control-D) dumps the
- whole ML system state into \<^verbatim>\<open>Test\<close> (be prepared for more
- than 100\,MB):
+ Ending this session normally (e.g.\ by typing control-D) dumps the whole ML
+ system state into \<^verbatim>\<open>Test\<close> (be prepared for more than 100\,MB):
- The \<^verbatim>\<open>Test\<close> session may be continued later (still in
- writable state) by: @{verbatim [display] \<open>isabelle_process Test\<close>}
+ The \<^verbatim>\<open>Test\<close> session may be continued later (still in writable state) by:
+ @{verbatim [display] \<open>isabelle_process Test\<close>}
A read-only \<^verbatim>\<open>Test\<close> session may be started by:
@{verbatim [display] \<open>isabelle_process -r Test\<close>}
\<^bigskip>
- The next example demonstrates batch execution of Isabelle.
- We retrieve the \<^verbatim>\<open>Main\<close> theory value from the theory loader
- within ML (observe the delicate quoting rules for the Bash shell
- vs.\ ML):
+ The next example demonstrates batch execution of Isabelle. We retrieve the
+ \<^verbatim>\<open>Main\<close> theory value from the theory loader within ML (observe the delicate
+ quoting rules for the Bash shell vs.\ ML):
@{verbatim [display] \<open>isabelle_process -e 'Thy_Info.get_theory "Main";' -q -r HOL\<close>}
- Note that the output text will be interspersed with additional junk
- messages by the ML runtime environment. The \<^verbatim>\<open>-W\<close> option
- allows to communicate with the Isabelle process via an external
- program in a more robust fashion.
+ Note that the output text will be interspersed with additional junk messages
+ by the ML runtime environment. The \<^verbatim>\<open>-W\<close> option allows to communicate with
+ the Isabelle process via an external program in a more robust fashion.
\<close>
section \<open>The Isabelle tool wrapper \label{sec:isabelle-tool}\<close>
text \<open>
- All Isabelle related tools and interfaces are called via a common
- wrapper --- @{executable isabelle}:
+ All Isabelle related tools and interfaces are called via a common wrapper
+ --- @{executable isabelle}:
@{verbatim [display]
\<open>Usage: isabelle TOOL [ARGS ...]
@@ -474,20 +432,19 @@
Available tools:
...\<close>}
- In principle, Isabelle tools are ordinary executable scripts that
- are run within the Isabelle settings environment, see
- \secref{sec:settings}. The set of available tools is collected by
- @{executable isabelle} from the directories listed in the @{setting
- ISABELLE_TOOLS} setting. Do not try to call the scripts directly
- from the shell. Neither should you add the tool directories to your
- shell's search path!
+ In principle, Isabelle tools are ordinary executable scripts that are run
+ within the Isabelle settings environment, see \secref{sec:settings}. The set
+ of available tools is collected by @{executable isabelle} from the
+ directories listed in the @{setting ISABELLE_TOOLS} setting. Do not try to
+ call the scripts directly from the shell. Neither should you add the tool
+ directories to your shell's search path!
\<close>
subsubsection \<open>Examples\<close>
-text \<open>Show the list of available documentation of the Isabelle
- distribution:
+text \<open>
+ Show the list of available documentation of the Isabelle distribution:
@{verbatim [display] \<open>isabelle doc\<close>}
View a certain document as follows:
--- a/src/Doc/System/Misc.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/System/Misc.thy Thu Nov 05 08:27:22 2015 +0100
@@ -1,3 +1,5 @@
+(*:wrap=hard:maxLineLen=78:*)
+
theory Misc
imports Base
begin
@@ -5,25 +7,27 @@
chapter \<open>Miscellaneous tools \label{ch:tools}\<close>
text \<open>
- Subsequently we describe various Isabelle related utilities, given
- in alphabetical order.
+ Subsequently we describe various Isabelle related utilities, given in
+ alphabetical order.
\<close>
section \<open>Theory graph browser \label{sec:browse}\<close>
-text \<open>The Isabelle graph browser is a general tool for visualizing
- dependency graphs. Certain nodes of the graph (i.e.\ theories) can
- be grouped together in ``directories'', whose contents may be
- hidden, thus enabling the user to collapse irrelevant portions of
- information. The browser is written in Java, it can be used both as
- a stand-alone application and as an applet.\<close>
+text \<open>
+ The Isabelle graph browser is a general tool for visualizing dependency
+ graphs. Certain nodes of the graph (i.e.\ theories) can be grouped together
+ in ``directories'', whose contents may be hidden, thus enabling the user to
+ collapse irrelevant portions of information. The browser is written in Java,
+ it can be used both as a stand-alone application and as an applet.
+\<close>
subsection \<open>Invoking the graph browser\<close>
-text \<open>The stand-alone version of the graph browser is wrapped up as
- @{tool_def browser}:
+text \<open>
+ The stand-alone version of the graph browser is wrapped up as @{tool_def
+ browser}:
@{verbatim [display]
\<open>Usage: isabelle browser [OPTIONS] [GRAPHFILE]
@@ -32,35 +36,31 @@
-c cleanup -- remove GRAPHFILE after use
-o FILE output to FILE (ps, eps, pdf)\<close>}
- When no file name is specified, the browser automatically changes to
- the directory @{setting ISABELLE_BROWSER_INFO}.
+ When no file name is specified, the browser automatically changes to the
+ directory @{setting ISABELLE_BROWSER_INFO}.
\<^medskip>
- The \<^verbatim>\<open>-b\<close> option indicates that this is for
- administrative build only, i.e.\ no browser popup if no files are
- given.
+ The \<^verbatim>\<open>-b\<close> option indicates that this is for administrative build only, i.e.\
+ no browser popup if no files are given.
- The \<^verbatim>\<open>-c\<close> option causes the input file to be removed
- after use.
+ The \<^verbatim>\<open>-c\<close> option causes the input file to be removed after use.
- The \<^verbatim>\<open>-o\<close> option indicates batch-mode operation, with the
- output written to the indicated file; note that \<^verbatim>\<open>pdf\<close>
- produces an \<^verbatim>\<open>eps\<close> copy as well.
+ The \<^verbatim>\<open>-o\<close> option indicates batch-mode operation, with the output written to
+ the indicated file; note that \<^verbatim>\<open>pdf\<close> produces an \<^verbatim>\<open>eps\<close> copy as well.
\<^medskip>
- The applet version of the browser is part of the standard
- WWW theory presentation, see the link ``theory dependencies'' within
- each session index.
+ The applet version of the browser is part of the standard WWW theory
+ presentation, see the link ``theory dependencies'' within each session
+ index.
\<close>
subsection \<open>Using the graph browser\<close>
text \<open>
- The browser's main window, which is shown in
- \figref{fig:browserwindow}, consists of two sub-windows. In the
- left sub-window, the directory tree is displayed. The graph itself
- is displayed in the right sub-window.
+ The browser's main window, which is shown in \figref{fig:browserwindow},
+ consists of two sub-windows. In the left sub-window, the directory tree is
+ displayed. The graph itself is displayed in the right sub-window.
\begin{figure}[ht]
\includegraphics[width=\textwidth]{browser_screenshot}
@@ -72,63 +72,57 @@
subsubsection \<open>The directory tree window\<close>
text \<open>
- We describe the usage of the directory browser and the meaning of
- the different items in the browser window.
+ We describe the usage of the directory browser and the meaning of the
+ different items in the browser window.
- \<^item> A red arrow before a directory name indicates that the
- directory is currently ``folded'', i.e.~the nodes in this directory
- are collapsed to one single node. In the right sub-window, the names
- of nodes corresponding to folded directories are enclosed in square
- brackets and displayed in red color.
+ \<^item> A red arrow before a directory name indicates that the directory is
+ currently ``folded'', i.e.~the nodes in this directory are collapsed to one
+ single node. In the right sub-window, the names of nodes corresponding to
+ folded directories are enclosed in square brackets and displayed in red
+ color.
- \<^item> A green downward arrow before a directory name indicates that
- the directory is currently ``unfolded''. It can be folded by
- clicking on the directory name. Clicking on the name for a second
- time unfolds the directory again. Alternatively, a directory can
- also be unfolded by clicking on the corresponding node in the right
- sub-window.
+ \<^item> A green downward arrow before a directory name indicates that the
+ directory is currently ``unfolded''. It can be folded by clicking on the
+ directory name. Clicking on the name for a second time unfolds the directory
+ again. Alternatively, a directory can also be unfolded by clicking on the
+ corresponding node in the right sub-window.
- \<^item> Blue arrows stand before ordinary node names. When clicking on
- such a name (i.e.\ that of a theory), the graph display window
- focuses to the corresponding node. Double clicking invokes a text
- viewer window in which the contents of the theory file are
- displayed.
+ \<^item> Blue arrows stand before ordinary node names. When clicking on such a name
+ (i.e.\ that of a theory), the graph display window focuses to the
+ corresponding node. Double clicking invokes a text viewer window in which
+ the contents of the theory file are displayed.
\<close>
subsubsection \<open>The graph display window\<close>
text \<open>
- When pointing on an ordinary node, an upward and a downward arrow is
- shown. Initially, both of these arrows are green. Clicking on the
- upward or downward arrow collapses all predecessor or successor
- nodes, respectively. The arrow's color then changes to red,
- indicating that the predecessor or successor nodes are currently
- collapsed. The node corresponding to the collapsed nodes has the
- name ``\<^verbatim>\<open>[....]\<close>''. To uncollapse the nodes again, simply
- click on the red arrow or on the node with the name ``\<^verbatim>\<open>[....]\<close>''.
- Similar to the directory browser, the contents of
- theory files can be displayed by double clicking on the
- corresponding node.
+ When pointing on an ordinary node, an upward and a downward arrow is shown.
+ Initially, both of these arrows are green. Clicking on the upward or
+ downward arrow collapses all predecessor or successor nodes, respectively.
+ The arrow's color then changes to red, indicating that the predecessor or
+ successor nodes are currently collapsed. The node corresponding to the
+ collapsed nodes has the name ``\<^verbatim>\<open>[....]\<close>''. To uncollapse the nodes again,
+ simply click on the red arrow or on the node with the name ``\<^verbatim>\<open>[....]\<close>''.
+ Similar to the directory browser, the contents of theory files can be
+ displayed by double clicking on the corresponding node.
\<close>
subsubsection \<open>The ``File'' menu\<close>
text \<open>
- Due to Java Applet security restrictions this menu is only available
- in the full application version. The meaning of the menu items is as
- follows:
+ Due to Java Applet security restrictions this menu is only available in the
+ full application version. The meaning of the menu items is as follows:
\<^descr>[Open \dots] Open a new graph file.
- \<^descr>[Export to PostScript] Outputs the current graph in Postscript
- format, appropriately scaled to fit on one single sheet of A4 paper.
- The resulting file can be printed directly.
+ \<^descr>[Export to PostScript] Outputs the current graph in Postscript format,
+ appropriately scaled to fit on one single sheet of A4 paper. The resulting
+ file can be printed directly.
- \<^descr>[Export to EPS] Outputs the current graph in Encapsulated
- Postscript format. The resulting file can be included in other
- documents.
+ \<^descr>[Export to EPS] Outputs the current graph in Encapsulated Postscript
+ format. The resulting file can be included in other documents.
\<^descr>[Quit] Quit the graph browser.
\<close>
@@ -150,22 +144,20 @@
\<^descr>[\<open>vertex_name\<close>] The name of the vertex.
- \<^descr>[\<open>vertex_ID\<close>] The vertex identifier. Note that there may
- be several vertices with equal names, whereas identifiers must be
- unique.
+ \<^descr>[\<open>vertex_ID\<close>] The vertex identifier. Note that there may be several
+ vertices with equal names, whereas identifiers must be unique.
- \<^descr>[\<open>dir_name\<close>] The name of the ``directory'' the vertex
- should be placed in. A ``\<^verbatim>\<open>+\<close>'' sign after \<open>dir_name\<close> indicates that the nodes in the directory are initially
- visible. Directories are initially invisible by default.
+ \<^descr>[\<open>dir_name\<close>] The name of the ``directory'' the vertex should be placed in.
+ A ``\<^verbatim>\<open>+\<close>'' sign after \<open>dir_name\<close> indicates that the nodes in the directory
+ are initially visible. Directories are initially invisible by default.
- \<^descr>[\<open>path\<close>] The path of the corresponding theory file. This
- is specified relatively to the path of the graph definition file.
+ \<^descr>[\<open>path\<close>] The path of the corresponding theory file. This is specified
+ relatively to the path of the graph definition file.
- \<^descr>[List of successor/predecessor nodes] A ``\<^verbatim>\<open><\<close>''
- sign before the list means that successor nodes are listed, a
- ``\<^verbatim>\<open>>\<close>'' sign means that predecessor nodes are listed. If
- neither ``\<^verbatim>\<open><\<close>'' nor ``\<^verbatim>\<open>>\<close>'' is found, the
- browser assumes that successor nodes are listed.
+ \<^descr>[List of successor/predecessor nodes] A ``\<^verbatim>\<open><\<close>'' sign before the list means
+ that successor nodes are listed, a ``\<^verbatim>\<open>>\<close>'' sign means that predecessor
+ nodes are listed. If neither ``\<^verbatim>\<open><\<close>'' nor ``\<^verbatim>\<open>>\<close>'' is found, the browser
+ assumes that successor nodes are listed.
\<close>
@@ -188,33 +180,28 @@
ISABELLE_COMPONENT_REPOSITORY="http://isabelle.in.tum.de/components"\<close>}
- Components are initialized as described in \secref{sec:components}
- in a permissive manner, which can mark components as ``missing''.
- This state is amended by letting @{tool "components"} download and
- unpack components that are published on the default component
- repository @{url "http://isabelle.in.tum.de/components/"} in
- particular.
+ Components are initialized as described in \secref{sec:components} in a
+ permissive manner, which can mark components as ``missing''. This state is
+ amended by letting @{tool "components"} download and unpack components that
+ are published on the default component repository @{url
+ "http://isabelle.in.tum.de/components/"} in particular.
- Option \<^verbatim>\<open>-R\<close> specifies an alternative component
- repository. Note that \<^verbatim>\<open>file:///\<close> URLs can be used for
- local directories.
+ Option \<^verbatim>\<open>-R\<close> specifies an alternative component repository. Note that
+ \<^verbatim>\<open>file:///\<close> URLs can be used for local directories.
- Option \<^verbatim>\<open>-a\<close> selects all missing components to be
- resolved. Explicit components may be named as command
- line-arguments as well. Note that components are uniquely
- identified by their base name, while the installation takes place in
- the location that was specified in the attempt to initialize the
- component before.
+ Option \<^verbatim>\<open>-a\<close> selects all missing components to be resolved. Explicit
+ components may be named as command line-arguments as well. Note that
+ components are uniquely identified by their base name, while the
+ installation takes place in the location that was specified in the attempt
+ to initialize the component before.
- Option \<^verbatim>\<open>-l\<close> lists the current state of available and
- missing components with their location (full name) within the
- file-system.
+ Option \<^verbatim>\<open>-l\<close> lists the current state of available and missing components
+ with their location (full name) within the file-system.
- Option \<^verbatim>\<open>-I\<close> initializes the user settings file to
- subscribe to the standard components specified in the Isabelle
- repository clone --- this does not make any sense for regular
- Isabelle releases. If the file already exists, it needs to be
- edited manually according to the printed explanation.
+ Option \<^verbatim>\<open>-I\<close> initializes the user settings file to subscribe to the standard
+ components specified in the Isabelle repository clone --- this does not make
+ any sense for regular Isabelle releases. If the file already exists, it
+ needs to be edited manually according to the printed explanation.
\<close>
@@ -236,14 +223,14 @@
Run Isabelle process with raw ML console and line editor
(default ISABELLE_LINE_EDITOR).\<close>}
- The \<^verbatim>\<open>-l\<close> option specifies the logic session name. By default,
- its heap image is checked and built on demand, but the option \<^verbatim>\<open>-n\<close> skips that.
+ The \<^verbatim>\<open>-l\<close> option specifies the logic session name. By default, its heap
+ image is checked and built on demand, but the option \<^verbatim>\<open>-n\<close> skips that.
- Options \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-o\<close>, \<^verbatim>\<open>-s\<close> are passed
- directly to @{tool build} (\secref{sec:tool-build}).
+ Options \<^verbatim>\<open>-d\<close>, \<^verbatim>\<open>-o\<close>, \<^verbatim>\<open>-s\<close> are passed directly to @{tool build}
+ (\secref{sec:tool-build}).
- Options \<^verbatim>\<open>-m\<close>, \<^verbatim>\<open>-o\<close> are passed directly to the
- underlying Isabelle process (\secref{sec:isabelle-process}).
+ Options \<^verbatim>\<open>-m\<close>, \<^verbatim>\<open>-o\<close> are passed directly to the underlying Isabelle process
+ (\secref{sec:isabelle-process}).
The Isabelle process is run through the line editor that is specified via
the settings variable @{setting ISABELLE_LINE_EDITOR} (e.g.\
@@ -259,19 +246,18 @@
section \<open>Displaying documents \label{sec:tool-display}\<close>
-text \<open>The @{tool_def display} tool displays documents in DVI or PDF
- format:
+text \<open>
+ The @{tool_def display} tool displays documents in DVI or PDF format:
@{verbatim [display]
\<open>Usage: isabelle display DOCUMENT
Display DOCUMENT (in DVI or PDF format).\<close>}
\<^medskip>
- The settings @{setting DVI_VIEWER} and @{setting
- PDF_VIEWER} determine the programs for viewing the corresponding
- file formats. Normally this opens the document via the desktop
- environment, potentially in an asynchronous manner with re-use of
- previews views.
+ The settings @{setting DVI_VIEWER} and @{setting PDF_VIEWER} determine the
+ programs for viewing the corresponding file formats. Normally this opens the
+ document via the desktop environment, potentially in an asynchronous manner
+ with re-use of previews views.
\<close>
@@ -284,27 +270,27 @@
View Isabelle documentation.\<close>}
- If called without arguments, it lists all available documents. Each
- line starts with an identifier, followed by a short description. Any
- of these identifiers may be specified as arguments, in order to
- display the corresponding document (see also
- \secref{sec:tool-display}).
+ If called without arguments, it lists all available documents. Each line
+ starts with an identifier, followed by a short description. Any of these
+ identifiers may be specified as arguments, in order to display the
+ corresponding document (see also \secref{sec:tool-display}).
\<^medskip>
- The @{setting ISABELLE_DOCS} setting specifies the list of
- directories (separated by colons) to be scanned for documentations.
+ The @{setting ISABELLE_DOCS} setting specifies the list of directories
+ (separated by colons) to be scanned for documentations.
\<close>
section \<open>Shell commands within the settings environment \label{sec:tool-env}\<close>
-text \<open>The @{tool_def env} tool is a direct wrapper for the standard
- \<^verbatim>\<open>/usr/bin/env\<close> command on POSIX systems, running within
- the Isabelle settings environment (\secref{sec:settings}).
+text \<open>
+ The @{tool_def env} tool is a direct wrapper for the standard
+ \<^verbatim>\<open>/usr/bin/env\<close> command on POSIX systems, running within the Isabelle
+ settings environment (\secref{sec:settings}).
- The command-line arguments are that of the underlying version of
- \<^verbatim>\<open>env\<close>. For example, the following invokes an instance of
- the GNU Bash shell within the Isabelle environment:
+ The command-line arguments are that of the underlying version of \<^verbatim>\<open>env\<close>. For
+ example, the following invokes an instance of the GNU Bash shell within the
+ Isabelle environment:
@{verbatim [display] \<open>isabelle env bash\<close>}
\<close>
@@ -325,38 +311,39 @@
Get value of VARNAMES from the Isabelle settings.\<close>}
- With the \<^verbatim>\<open>-a\<close> option, one may inspect the full process
- environment that Isabelle related programs are run in. This usually
- contains much more variables than are actually Isabelle settings.
- Normally, output is a list of lines of the form \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close>. The \<^verbatim>\<open>-b\<close> option
- causes only the values to be printed.
+ With the \<^verbatim>\<open>-a\<close> option, one may inspect the full process environment that
+ Isabelle related programs are run in. This usually contains much more
+ variables than are actually Isabelle settings. Normally, output is a list of
+ lines of the form \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close>. The \<^verbatim>\<open>-b\<close> option causes only the values
+ to be printed.
- Option \<^verbatim>\<open>-d\<close> produces a dump of the complete environment
- to the specified file. Entries are terminated by the ASCII null
- character, i.e.\ the C string terminator.
+ Option \<^verbatim>\<open>-d\<close> produces a dump of the complete environment to the specified
+ file. Entries are terminated by the ASCII null character, i.e.\ the C string
+ terminator.
\<close>
subsubsection \<open>Examples\<close>
-text \<open>Get the location of @{setting ISABELLE_HOME_USER} where
- user-specific information is stored:
+text \<open>
+ Get the location of @{setting ISABELLE_HOME_USER} where user-specific
+ information is stored:
@{verbatim [display] \<open>isabelle getenv ISABELLE_HOME_USER\<close>}
\<^medskip>
- Get the value only of the same settings variable, which is
- particularly useful in shell scripts:
+ Get the value only of the same settings variable, which is particularly
+ useful in shell scripts:
@{verbatim [display] \<open>isabelle getenv -b ISABELLE_OUTPUT\<close>}
\<close>
section \<open>Installing standalone Isabelle executables \label{sec:tool-install}\<close>
-text \<open>By default, the main Isabelle binaries (@{executable
- "isabelle"} etc.) are just run from their location within the
- distribution directory, probably indirectly by the shell through its
- @{setting PATH}. Other schemes of installation are supported by the
- @{tool_def install} tool:
+text \<open>
+ By default, the main Isabelle binaries (@{executable "isabelle"} etc.) are
+ just run from their location within the distribution directory, probably
+ indirectly by the shell through its @{setting PATH}. Other schemes of
+ installation are supported by the @{tool_def install} tool:
@{verbatim [display]
\<open>Usage: isabelle install [OPTIONS] BINDIR
@@ -367,24 +354,26 @@
Install Isabelle executables with absolute references to the
distribution directory.\<close>}
- The \<^verbatim>\<open>-d\<close> option overrides the current Isabelle
- distribution directory as determined by @{setting ISABELLE_HOME}.
+ The \<^verbatim>\<open>-d\<close> option overrides the current Isabelle distribution directory as
+ determined by @{setting ISABELLE_HOME}.
- The \<open>BINDIR\<close> argument tells where executable wrapper scripts
- for @{executable "isabelle_process"} and @{executable isabelle}
- should be placed, which is typically a directory in the shell's
- @{setting PATH}, such as \<^verbatim>\<open>$HOME/bin\<close>.
+ The \<open>BINDIR\<close> argument tells where executable wrapper scripts for
+ @{executable "isabelle_process"} and @{executable isabelle} should be
+ placed, which is typically a directory in the shell's @{setting PATH}, such
+ as \<^verbatim>\<open>$HOME/bin\<close>.
\<^medskip>
- It is also possible to make symbolic links of the main
- Isabelle executables manually, but making separate copies outside
- the Isabelle distribution directory will not work!\<close>
+ It is also possible to make symbolic links of the main Isabelle executables
+ manually, but making separate copies outside the Isabelle distribution
+ directory will not work!
+\<close>
section \<open>Creating instances of the Isabelle logo\<close>
-text \<open>The @{tool_def logo} tool creates instances of the generic
- Isabelle logo as EPS and PDF, for inclusion in {\LaTeX} documents.
+text \<open>
+ The @{tool_def logo} tool creates instances of the generic Isabelle logo as
+ EPS and PDF, for inclusion in {\LaTeX} documents.
@{verbatim [display]
\<open>Usage: isabelle logo [OPTIONS] XYZ
@@ -394,16 +383,15 @@
-n NAME alternative output base name (default "isabelle_xyx")
-q quiet mode\<close>}
- Option \<^verbatim>\<open>-n\<close> specifies an alternative (base) name for the
- generated files. The default is \<^verbatim>\<open>isabelle_\<close>\<open>xyz\<close>
- in lower-case.
+ Option \<^verbatim>\<open>-n\<close> specifies an alternative (base) name for the generated files.
+ The default is \<^verbatim>\<open>isabelle_\<close>\<open>xyz\<close> in lower-case.
Option \<^verbatim>\<open>-q\<close> omits printing of the result file name.
\<^medskip>
- Implementors of Isabelle tools and applications are
- encouraged to make derived Isabelle logos for their own projects
- using this template.\<close>
+ Implementors of Isabelle tools and applications are encouraged to make
+ derived Isabelle logos for their own projects using this template.
+\<close>
section \<open>Output the version identifier of the Isabelle distribution\<close>
@@ -419,53 +407,49 @@
Display Isabelle version information.\<close>}
\<^medskip>
- The default is to output the full version string of the
- Isabelle distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012: May 2012\<close>.
+ The default is to output the full version string of the Isabelle
+ distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2012: May 2012\<close>.
- The \<^verbatim>\<open>-i\<close> option produces a short identification derived
- from the Mercurial id of the @{setting ISABELLE_HOME} directory.
+ The \<^verbatim>\<open>-i\<close> option produces a short identification derived from the Mercurial
+ id of the @{setting ISABELLE_HOME} directory.
\<close>
section \<open>Convert XML to YXML\<close>
text \<open>
- The @{tool_def yxml} tool converts a standard XML document (stdin)
- to the much simpler and more efficient YXML format of Isabelle
- (stdout). The YXML format is defined as follows.
-
- \<^enum> The encoding is always UTF-8.
+ The @{tool_def yxml} tool converts a standard XML document (stdin) to the
+ much simpler and more efficient YXML format of Isabelle (stdout). The YXML
+ format is defined as follows.
- \<^enum> Body text is represented verbatim (no escaping, no special
- treatment of white space, no named entities, no CDATA chunks, no
- comments).
+ \<^enum> The encoding is always UTF-8.
- \<^enum> Markup elements are represented via ASCII control characters
- \<open>\<^bold>X = 5\<close> and \<open>\<^bold>Y = 6\<close> as follows:
+ \<^enum> Body text is represented verbatim (no escaping, no special treatment of
+ white space, no named entities, no CDATA chunks, no comments).
- \begin{tabular}{ll}
- XML & YXML \\\hline
- \<^verbatim>\<open><\<close>\<open>name attribute\<close>\<^verbatim>\<open>=\<close>\<open>value \<dots>\<close>\<^verbatim>\<open>>\<close> &
- \<open>\<^bold>X\<^bold>Yname\<^bold>Yattribute\<close>\<^verbatim>\<open>=\<close>\<open>value\<dots>\<^bold>X\<close> \\
- \<^verbatim>\<open></\<close>\<open>name\<close>\<^verbatim>\<open>>\<close> & \<open>\<^bold>X\<^bold>Y\<^bold>X\<close> \\
- \end{tabular}
+ \<^enum> Markup elements are represented via ASCII control characters \<open>\<^bold>X = 5\<close>
+ and \<open>\<^bold>Y = 6\<close> as follows:
- There is no special case for empty body text, i.e.\ \<^verbatim>\<open><foo/>\<close>
- is treated like \<^verbatim>\<open><foo></foo>\<close>. Also note that
- \<open>\<^bold>X\<close> and \<open>\<^bold>Y\<close> may never occur in
- well-formed XML documents.
+ \begin{tabular}{ll}
+ XML & YXML \\\hline
+ \<^verbatim>\<open><\<close>\<open>name attribute\<close>\<^verbatim>\<open>=\<close>\<open>value \<dots>\<close>\<^verbatim>\<open>>\<close> &
+ \<open>\<^bold>X\<^bold>Yname\<^bold>Yattribute\<close>\<^verbatim>\<open>=\<close>\<open>value\<dots>\<^bold>X\<close> \\
+ \<^verbatim>\<open></\<close>\<open>name\<close>\<^verbatim>\<open>>\<close> & \<open>\<^bold>X\<^bold>Y\<^bold>X\<close> \\
+ \end{tabular}
+ There is no special case for empty body text, i.e.\ \<^verbatim>\<open><foo/>\<close> is treated
+ like \<^verbatim>\<open><foo></foo>\<close>. Also note that \<open>\<^bold>X\<close> and \<open>\<^bold>Y\<close> may never occur in
+ well-formed XML documents.
Parsing YXML is pretty straight-forward: split the text into chunks
- separated by \<open>\<^bold>X\<close>, then split each chunk into
- sub-chunks separated by \<open>\<^bold>Y\<close>. Markup chunks start
- with an empty sub-chunk, and a second empty sub-chunk indicates
- close of an element. Any other non-empty chunk consists of plain
- text. For example, see @{file "~~/src/Pure/PIDE/yxml.ML"} or
- @{file "~~/src/Pure/PIDE/yxml.scala"}.
+ separated by \<open>\<^bold>X\<close>, then split each chunk into sub-chunks separated by \<open>\<^bold>Y\<close>.
+ Markup chunks start with an empty sub-chunk, and a second empty sub-chunk
+ indicates close of an element. Any other non-empty chunk consists of plain
+ text. For example, see @{file "~~/src/Pure/PIDE/yxml.ML"} or @{file
+ "~~/src/Pure/PIDE/yxml.scala"}.
- YXML documents may be detected quickly by checking that the first
- two characters are \<open>\<^bold>X\<^bold>Y\<close>.
+ YXML documents may be detected quickly by checking that the first two
+ characters are \<open>\<^bold>X\<^bold>Y\<close>.
\<close>
end
\ No newline at end of file
--- a/src/Doc/System/Presentation.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/System/Presentation.thy Thu Nov 05 08:27:22 2015 +0100
@@ -1,24 +1,27 @@
+(*:wrap=hard:maxLineLen=78:*)
+
theory Presentation
imports Base
begin
chapter \<open>Presenting theories \label{ch:present}\<close>
-text \<open>Isabelle provides several ways to present the outcome of
- formal developments, including WWW-based browsable libraries or
- actual printable documents. Presentation is centered around the
- concept of \<^emph>\<open>sessions\<close> (\chref{ch:session}). The global session
- structure is that of a tree, with Isabelle Pure at its root, further
- object-logics derived (e.g.\ HOLCF from HOL, and HOL from Pure), and
- application sessions further on in the hierarchy.
+text \<open>
+ Isabelle provides several ways to present the outcome of formal
+ developments, including WWW-based browsable libraries or actual printable
+ documents. Presentation is centered around the concept of \<^emph>\<open>sessions\<close>
+ (\chref{ch:session}). The global session structure is that of a tree, with
+ Isabelle Pure at its root, further object-logics derived (e.g.\ HOLCF from
+ HOL, and HOL from Pure), and application sessions further on in the
+ hierarchy.
- The tools @{tool_ref mkroot} and @{tool_ref build} provide the
- primary means for managing Isabelle sessions, including proper setup
- for presentation; @{tool build} takes care to have @{executable_ref
- "isabelle_process"} run any additional stages required for document
- preparation, notably the @{tool_ref document} and @{tool_ref latex}.
- The complete tool chain for managing batch-mode Isabelle sessions is
- illustrated in \figref{fig:session-tools}.
+ The tools @{tool_ref mkroot} and @{tool_ref build} provide the primary means
+ for managing Isabelle sessions, including proper setup for presentation;
+ @{tool build} takes care to have @{executable_ref "isabelle_process"} run
+ any additional stages required for document preparation, notably the
+ @{tool_ref document} and @{tool_ref latex}. The complete tool chain for
+ managing batch-mode Isabelle sessions is illustrated in
+ \figref{fig:session-tools}.
\begin{figure}[htbp]
\begin{center}
@@ -53,58 +56,55 @@
text \<open>
\index{theory browsing information|bold}
- As a side-effect of building sessions, Isabelle is able to generate
- theory browsing information, including HTML documents that show the
- theory sources and the relationship with its ancestors and
- descendants. Besides the HTML file that is generated for every
- theory, Isabelle stores links to all theories of a session in an
- index file. As a second hierarchy, groups of sessions are organized
- as \<^emph>\<open>chapters\<close>, with a separate index. Note that the implicit
- tree structure of the session build hierarchy is \<^emph>\<open>not\<close> relevant
+ As a side-effect of building sessions, Isabelle is able to generate theory
+ browsing information, including HTML documents that show the theory sources
+ and the relationship with its ancestors and descendants. Besides the HTML
+ file that is generated for every theory, Isabelle stores links to all
+ theories of a session in an index file. As a second hierarchy, groups of
+ sessions are organized as \<^emph>\<open>chapters\<close>, with a separate index. Note that the
+ implicit tree structure of the session build hierarchy is \<^emph>\<open>not\<close> relevant
for the presentation.
- Isabelle also generates graph files that represent the theory
- dependencies within a session. There is a graph browser Java applet
- embedded in the generated HTML pages, and also a stand-alone
- application that allows browsing theory graphs without having to
- start a WWW client first. The latter version also includes features
- such as generating Postscript files, which are not available in the
- applet version. See \secref{sec:browse} for further information.
+ Isabelle also generates graph files that represent the theory dependencies
+ within a session. There is a graph browser Java applet embedded in the
+ generated HTML pages, and also a stand-alone application that allows
+ browsing theory graphs without having to start a WWW client first. The
+ latter version also includes features such as generating Postscript files,
+ which are not available in the applet version. See \secref{sec:browse} for
+ further information.
\<^medskip>
- The easiest way to let Isabelle generate theory browsing information
- for existing sessions is to invoke @{tool build} with suitable
- options:
+ The easiest way to let Isabelle generate theory browsing information for
+ existing sessions is to invoke @{tool build} with suitable options:
@{verbatim [display] \<open>isabelle build -o browser_info -v -c FOL\<close>}
- The presentation output will appear in \<^verbatim>\<open>$ISABELLE_BROWSER_INFO/FOL/FOL\<close>
- as reported by the above verbose invocation of the build process.
+ The presentation output will appear in \<^verbatim>\<open>$ISABELLE_BROWSER_INFO/FOL/FOL\<close> as
+ reported by the above verbose invocation of the build process.
Many Isabelle sessions (such as \<^verbatim>\<open>HOL-Library\<close> in @{file
- "~~/src/HOL/Library"}) also provide actual printable documents.
- These are prepared automatically as well if enabled like this:
+ "~~/src/HOL/Library"}) also provide actual printable documents. These are
+ prepared automatically as well if enabled like this:
@{verbatim [display] \<open>isabelle build -o browser_info -o document=pdf -v -c HOL-Library\<close>}
- Enabling both browser info and document preparation simultaneously
- causes an appropriate ``document'' link to be included in the HTML
- index. Documents may be generated independently of browser
- information as well, see \secref{sec:tool-document} for further
- details.
+ Enabling both browser info and document preparation simultaneously causes an
+ appropriate ``document'' link to be included in the HTML index. Documents
+ may be generated independently of browser information as well, see
+ \secref{sec:tool-document} for further details.
\<^bigskip>
- The theory browsing information is stored in a
- sub-directory directory determined by the @{setting_ref
- ISABELLE_BROWSER_INFO} setting plus a prefix corresponding to the
- session chapter and identifier. In order to present Isabelle
- applications on the web, the corresponding subdirectory from
- @{setting ISABELLE_BROWSER_INFO} can be put on a WWW server.\<close>
+ The theory browsing information is stored in a sub-directory directory
+ determined by the @{setting_ref ISABELLE_BROWSER_INFO} setting plus a prefix
+ corresponding to the session chapter and identifier. In order to present
+ Isabelle applications on the web, the corresponding subdirectory from
+ @{setting ISABELLE_BROWSER_INFO} can be put on a WWW server.
+\<close>
section \<open>Preparing session root directories \label{sec:tool-mkroot}\<close>
-text \<open>The @{tool_def mkroot} tool configures a given directory as
- session root, with some \<^verbatim>\<open>ROOT\<close> file and optional document
- source directory. Its usage is:
+text \<open>
+ The @{tool_def mkroot} tool configures a given directory as session root,
+ with some \<^verbatim>\<open>ROOT\<close> file and optional document source directory. Its usage is:
@{verbatim [display]
\<open>Usage: isabelle mkroot [OPTIONS] [DIR]
@@ -114,46 +114,45 @@
Prepare session root DIR (default: current directory).\<close>}
- The results are placed in the given directory \<open>dir\<close>, which
- refers to the current directory by default. The @{tool mkroot} tool
- is conservative in the sense that it does not overwrite existing
- files or directories. Earlier attempts to generate a session root
- need to be deleted manually.
+ The results are placed in the given directory \<open>dir\<close>, which refers to the
+ current directory by default. The @{tool mkroot} tool is conservative in the
+ sense that it does not overwrite existing files or directories. Earlier
+ attempts to generate a session root need to be deleted manually.
\<^medskip>
- Option \<^verbatim>\<open>-d\<close> indicates that the session shall be
- accompanied by a formal document, with \<open>DIR\<close>\<^verbatim>\<open>/document/root.tex\<close>
- as its {\LaTeX} entry point (see also \chref{ch:present}).
+ Option \<^verbatim>\<open>-d\<close> indicates that the session shall be accompanied by a formal
+ document, with \<open>DIR\<close>\<^verbatim>\<open>/document/root.tex\<close> as its {\LaTeX} entry point (see
+ also \chref{ch:present}).
- Option \<^verbatim>\<open>-n\<close> allows to specify an alternative session
- name; otherwise the base name of the given directory is used.
+ Option \<^verbatim>\<open>-n\<close> allows to specify an alternative session name; otherwise the
+ base name of the given directory is used.
\<^medskip>
- The implicit Isabelle settings variable @{setting
- ISABELLE_LOGIC} specifies the parent session, and @{setting
- ISABELLE_DOCUMENT_FORMAT} the document format to be filled filled
- into the generated \<^verbatim>\<open>ROOT\<close> file.
+ The implicit Isabelle settings variable @{setting ISABELLE_LOGIC} specifies
+ the parent session, and @{setting ISABELLE_DOCUMENT_FORMAT} the document
+ format to be filled filled into the generated \<^verbatim>\<open>ROOT\<close> file.
\<close>
subsubsection \<open>Examples\<close>
-text \<open>Produce session \<^verbatim>\<open>Test\<close> (with document preparation)
- within a separate directory of the same name:
+text \<open>
+ Produce session \<^verbatim>\<open>Test\<close> (with document preparation) within a separate
+ directory of the same name:
@{verbatim [display] \<open>isabelle mkroot -d Test && isabelle build -D Test\<close>}
\<^medskip>
- Upgrade the current directory into a session ROOT with
- document preparation, and build it:
+ Upgrade the current directory into a session ROOT with document preparation,
+ and build it:
@{verbatim [display] \<open>isabelle mkroot -d && isabelle build -D .\<close>}
\<close>
section \<open>Preparing Isabelle session documents \label{sec:tool-document}\<close>
-text \<open>The @{tool_def document} tool prepares logic session
- documents, processing the sources as provided by the user and
- generated by Isabelle. Its usage is:
+text \<open>
+ The @{tool_def document} tool prepares logic session documents, processing
+ the sources as provided by the user and generated by Isabelle. Its usage is:
@{verbatim [display]
\<open>Usage: isabelle document [OPTIONS] [DIR]
@@ -168,90 +167,83 @@
This tool is usually run automatically as part of the Isabelle build
process, provided document preparation has been enabled via suitable
- options. It may be manually invoked on the generated browser
- information document output as well, e.g.\ in case of errors
- encountered in the batch run.
+ options. It may be manually invoked on the generated browser information
+ document output as well, e.g.\ in case of errors encountered in the batch
+ run.
\<^medskip>
- The \<^verbatim>\<open>-c\<close> option tells @{tool document} to
- dispose the document sources after successful operation! This is
- the right thing to do for sources generated by an Isabelle process,
- but take care of your files in manual document preparation!
+ The \<^verbatim>\<open>-c\<close> option tells @{tool document} to dispose the document sources
+ after successful operation! This is the right thing to do for sources
+ generated by an Isabelle process, but take care of your files in manual
+ document preparation!
\<^medskip>
- The \<^verbatim>\<open>-n\<close> and \<^verbatim>\<open>-o\<close> option specify
- the final output file name and format, the default is ``\<^verbatim>\<open>document.dvi\<close>''.
- Note that the result will appear in the parent of the target \<^verbatim>\<open>DIR\<close>.
+ The \<^verbatim>\<open>-n\<close> and \<^verbatim>\<open>-o\<close> option specify the final output file name and format,
+ the default is ``\<^verbatim>\<open>document.dvi\<close>''. Note that the result will appear in the
+ parent of the target \<^verbatim>\<open>DIR\<close>.
\<^medskip>
- The \<^verbatim>\<open>-t\<close> option tells {\LaTeX} how to interpret
- tagged Isabelle command regions. Tags are specified as a comma
- separated list of modifier/name pairs: ``\<^verbatim>\<open>+\<close>\<open>foo\<close>'' (or just ``\<open>foo\<close>'')
- means to keep, ``\<^verbatim>\<open>-\<close>\<open>foo\<close>'' to drop, and ``\<^verbatim>\<open>/\<close>\<open>foo\<close>'' to
- fold text tagged as \<open>foo\<close>. The builtin default is equivalent
- to the tag specification ``\<^verbatim>\<open>+theory,+proof,+ML,+visible,-invisible\<close>'';
- see also the {\LaTeX} macros \<^verbatim>\<open>\isakeeptag\<close>, \<^verbatim>\<open>\isadroptag\<close>, and
- \<^verbatim>\<open>\isafoldtag\<close>, in @{file "~~/lib/texinputs/isabelle.sty"}.
+ The \<^verbatim>\<open>-t\<close> option tells {\LaTeX} how to interpret tagged Isabelle command
+ regions. Tags are specified as a comma separated list of modifier/name
+ pairs: ``\<^verbatim>\<open>+\<close>\<open>foo\<close>'' (or just ``\<open>foo\<close>'') means to keep, ``\<^verbatim>\<open>-\<close>\<open>foo\<close>'' to
+ drop, and ``\<^verbatim>\<open>/\<close>\<open>foo\<close>'' to fold text tagged as \<open>foo\<close>. The builtin default is
+ equivalent to the tag specification
+ ``\<^verbatim>\<open>+theory,+proof,+ML,+visible,-invisible\<close>''; see also the {\LaTeX} macros
+ \<^verbatim>\<open>\isakeeptag\<close>, \<^verbatim>\<open>\isadroptag\<close>, and \<^verbatim>\<open>\isafoldtag\<close>, in @{file
+ "~~/lib/texinputs/isabelle.sty"}.
\<^medskip>
- Document preparation requires a \<^verbatim>\<open>document\<close>
- directory within the session sources. This directory is supposed to
- contain all the files needed to produce the final document --- apart
- from the actual theories which are generated by Isabelle.
+ Document preparation requires a \<^verbatim>\<open>document\<close> directory within the session
+ sources. This directory is supposed to contain all the files needed to
+ produce the final document --- apart from the actual theories which are
+ generated by Isabelle.
\<^medskip>
- For most practical purposes, @{tool document} is smart
- enough to create any of the specified output formats, taking
- \<^verbatim>\<open>root.tex\<close> supplied by the user as a starting point. This
- even includes multiple runs of {\LaTeX} to accommodate references
- and bibliographies (the latter assumes \<^verbatim>\<open>root.bib\<close> within
- the same directory).
+ For most practical purposes, @{tool document} is smart enough to create any
+ of the specified output formats, taking \<^verbatim>\<open>root.tex\<close> supplied by the user as
+ a starting point. This even includes multiple runs of {\LaTeX} to
+ accommodate references and bibliographies (the latter assumes \<^verbatim>\<open>root.bib\<close>
+ within the same directory).
- In more complex situations, a separate \<^verbatim>\<open>build\<close> script for
- the document sources may be given. It is invoked with command-line
- arguments for the document format and the document variant name.
- The script needs to produce corresponding output files, e.g.\
- \<^verbatim>\<open>root.pdf\<close> for target format \<^verbatim>\<open>pdf\<close> (and default
- variants). The main work can be again delegated to @{tool latex},
- but it is also possible to harvest generated {\LaTeX} sources and
- copy them elsewhere.
+ In more complex situations, a separate \<^verbatim>\<open>build\<close> script for the document
+ sources may be given. It is invoked with command-line arguments for the
+ document format and the document variant name. The script needs to produce
+ corresponding output files, e.g.\ \<^verbatim>\<open>root.pdf\<close> for target format \<^verbatim>\<open>pdf\<close> (and
+ default variants). The main work can be again delegated to @{tool latex},
+ but it is also possible to harvest generated {\LaTeX} sources and copy them
+ elsewhere.
\<^medskip>
- When running the session, Isabelle copies the content of
- the original \<^verbatim>\<open>document\<close> directory into its proper place
- within @{setting ISABELLE_BROWSER_INFO}, according to the session
- path and document variant. Then, for any processed theory \<open>A\<close>
- some {\LaTeX} source is generated and put there as \<open>A\<close>\<^verbatim>\<open>.tex\<close>.
- Furthermore, a list of all generated theory
- files is put into \<^verbatim>\<open>session.tex\<close>. Typically, the root
- {\LaTeX} file provided by the user would include \<^verbatim>\<open>session.tex\<close>
- to get a document containing all the theories.
+ When running the session, Isabelle copies the content of the original
+ \<^verbatim>\<open>document\<close> directory into its proper place within @{setting
+ ISABELLE_BROWSER_INFO}, according to the session path and document variant.
+ Then, for any processed theory \<open>A\<close> some {\LaTeX} source is generated and put
+ there as \<open>A\<close>\<^verbatim>\<open>.tex\<close>. Furthermore, a list of all generated theory files is
+ put into \<^verbatim>\<open>session.tex\<close>. Typically, the root {\LaTeX} file provided by the
+ user would include \<^verbatim>\<open>session.tex\<close> to get a document containing all the
+ theories.
- The {\LaTeX} versions of the theories require some macros defined in
- @{file "~~/lib/texinputs/isabelle.sty"}. Doing \<^verbatim>\<open>\usepackage{isabelle}\<close>
- in \<^verbatim>\<open>root.tex\<close> should be fine; the underlying @{tool latex} already
- includes an appropriate path specification for {\TeX} inputs.
+ The {\LaTeX} versions of the theories require some macros defined in @{file
+ "~~/lib/texinputs/isabelle.sty"}. Doing \<^verbatim>\<open>\usepackage{isabelle}\<close> in
+ \<^verbatim>\<open>root.tex\<close> should be fine; the underlying @{tool latex} already includes an
+ appropriate path specification for {\TeX} inputs.
- If the text contains any references to Isabelle symbols (such as
- \<^verbatim>\<open>\<forall>\<close>) then \<^verbatim>\<open>isabellesym.sty\<close> should be included as well.
- This package contains a standard set of {\LaTeX} macro definitions
- \<^verbatim>\<open>\isasym\<close>\<open>foo\<close> corresponding to \<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>foo\<close>\<^verbatim>\<open>>\<close>,
- see @{cite "isabelle-implementation"} for a
- complete list of predefined Isabelle symbols. Users may invent
- further symbols as well, just by providing {\LaTeX} macros in a
- similar fashion as in @{file "~~/lib/texinputs/isabellesym.sty"} of
- the Isabelle distribution.
+ If the text contains any references to Isabelle symbols (such as \<^verbatim>\<open>\<forall>\<close>) then
+ \<^verbatim>\<open>isabellesym.sty\<close> should be included as well. This package contains a
+ standard set of {\LaTeX} macro definitions \<^verbatim>\<open>\isasym\<close>\<open>foo\<close> corresponding to
+ \<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>foo\<close>\<^verbatim>\<open>>\<close>, see @{cite "isabelle-implementation"} for a complete list
+ of predefined Isabelle symbols. Users may invent further symbols as well,
+ just by providing {\LaTeX} macros in a similar fashion as in @{file
+ "~~/lib/texinputs/isabellesym.sty"} of the Isabelle distribution.
- For proper setup of DVI and PDF documents (with hyperlinks and
- bookmarks), we recommend to include @{file
- "~~/lib/texinputs/pdfsetup.sty"} as well.
+ For proper setup of DVI and PDF documents (with hyperlinks and bookmarks),
+ we recommend to include @{file "~~/lib/texinputs/pdfsetup.sty"} as well.
\<^medskip>
- As a final step of Isabelle document preparation, @{tool
- document}~\<^verbatim>\<open>-c\<close> is run on the resulting copy of the
- \<^verbatim>\<open>document\<close> directory. Thus the actual output document is
- built and installed in its proper place. The generated sources are
- deleted after successful run of {\LaTeX} and friends.
+ As a final step of Isabelle document preparation, @{tool document}~\<^verbatim>\<open>-c\<close> is
+ run on the resulting copy of the \<^verbatim>\<open>document\<close> directory. Thus the actual
+ output document is built and installed in its proper place. The generated
+ sources are deleted after successful run of {\LaTeX} and friends.
Some care is needed if the document output location is configured
differently, say within a directory whose content is still required
@@ -262,8 +254,9 @@
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:
+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]
@@ -274,32 +267,30 @@
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 latex},
- @{executable pdflatex}, @{executable dvips}, @{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.).
+ Appropriate {\LaTeX}-related programs are run on the input file, according
+ to the given output format: @{executable latex}, @{executable pdflatex},
+ @{executable dvips}, @{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.
+ 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.
- The \<^verbatim>\<open>syms\<close> output is for internal use; it generates lists
- of symbols that are available without loading additional {\LaTeX}
- packages.
+ The \<^verbatim>\<open>syms\<close> output is for internal use; it generates lists of symbols that
+ are available without loading additional {\LaTeX} packages.
\<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:
+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"
--- a/src/Doc/System/Scala.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/System/Scala.thy Thu Nov 05 08:27:22 2015 +0100
@@ -1,43 +1,47 @@
+(*:wrap=hard:maxLineLen=78:*)
+
theory Scala
imports Base
begin
chapter \<open>Isabelle/Scala development tools\<close>
-text \<open>Isabelle/ML and Isabelle/Scala are the two main language
-environments for Isabelle tool implementations. There are some basic
-command-line tools to work with the underlying Java Virtual Machine,
-the Scala toplevel and compiler. Note that Isabelle/jEdit
-@{cite "isabelle-jedit"} provides a Scala Console for interactive
-experimentation within the running application.\<close>
+text \<open>
+ Isabelle/ML and Isabelle/Scala are the two main language environments for
+ Isabelle tool implementations. There are some basic command-line tools to
+ work with the underlying Java Virtual Machine, the Scala toplevel and
+ compiler. Note that Isabelle/jEdit @{cite "isabelle-jedit"} provides a Scala
+ Console for interactive experimentation within the running application.
+\<close>
section \<open>Java Runtime Environment within Isabelle \label{sec:tool-java}\<close>
-text \<open>The @{tool_def java} tool is a direct wrapper for the Java
- Runtime Environment, within the regular Isabelle settings
- environment (\secref{sec:settings}). The command line arguments are
- that of the underlying Java version. It is run in \<^verbatim>\<open>-server\<close> mode
- if possible, to improve performance (at the cost of extra startup time).
+text \<open>
+ The @{tool_def java} tool is a direct wrapper for the Java Runtime
+ Environment, within the regular Isabelle settings environment
+ (\secref{sec:settings}). The command line arguments are that of the
+ underlying Java version. It is run in \<^verbatim>\<open>-server\<close> mode if possible, to
+ improve performance (at the cost of extra startup time).
- The \<^verbatim>\<open>java\<close> executable is the one within @{setting
- ISABELLE_JDK_HOME}, according to the standard directory layout for
- official JDK distributions. The class loader is augmented such that
- the name space of \<^verbatim>\<open>Isabelle/Pure.jar\<close> is available,
- which is the main Isabelle/Scala module.
+ The \<^verbatim>\<open>java\<close> executable is the one within @{setting ISABELLE_JDK_HOME},
+ according to the standard directory layout for official JDK distributions.
+ The class loader is augmented such that the name space of
+ \<^verbatim>\<open>Isabelle/Pure.jar\<close> is available, which is the main Isabelle/Scala module.
- For example, the following command-line invokes the main method of
- class \<^verbatim>\<open>isabelle.GUI_Setup\<close>, which opens a windows with
- some diagnostic information about the Isabelle environment:
+ For example, the following command-line invokes the main method of class
+ \<^verbatim>\<open>isabelle.GUI_Setup\<close>, which opens a windows with some diagnostic
+ information about the Isabelle environment:
@{verbatim [display] \<open>isabelle java isabelle.GUI_Setup\<close>}
\<close>
section \<open>Scala toplevel \label{sec:tool-scala}\<close>
-text \<open>The @{tool_def scala} tool is a direct wrapper for the Scala
- toplevel; see also @{tool java} above. The command line arguments
- are that of the underlying Scala version.
+text \<open>
+ The @{tool_def scala} tool is a direct wrapper for the Scala toplevel; see
+ also @{tool java} above. The command line arguments are that of the
+ underlying Scala version.
This allows to interact with Isabelle/Scala in TTY mode like this:
@{verbatim [display]
@@ -51,32 +55,33 @@
section \<open>Scala compiler \label{sec:tool-scalac}\<close>
-text \<open>The @{tool_def scalac} tool is a direct wrapper for the Scala
- compiler; see also @{tool scala} above. The command line arguments
- are that of the underlying Scala version.
+text \<open>
+ The @{tool_def scalac} tool is a direct wrapper for the Scala compiler; see
+ also @{tool scala} above. The command line arguments are that of the
+ underlying Scala version.
This allows to compile further Scala modules, depending on existing
- Isabelle/Scala functionality. The resulting class or jar files can
- be added to the Java classpath using the \<^verbatim>\<open>classpath\<close> Bash
- function that is provided by the Isabelle process environment. Thus
- add-on components can register themselves in a modular manner, see
- also \secref{sec:components}.
+ Isabelle/Scala functionality. The resulting class or jar files can be added
+ to the Java classpath using the \<^verbatim>\<open>classpath\<close> Bash function that is provided
+ by the Isabelle process environment. Thus add-on components can register
+ themselves in a modular manner, see also \secref{sec:components}.
- Note that jEdit @{cite "isabelle-jedit"} has its own mechanisms for
- adding plugin components, which needs special attention since
- it overrides the standard Java class loader.\<close>
+ Note that jEdit @{cite "isabelle-jedit"} has its own mechanisms for adding
+ plugin components, which needs special attention since it overrides the
+ standard Java class loader.
+\<close>
section \<open>Scala script wrapper\<close>
-text \<open>The executable @{executable
- "$ISABELLE_HOME/bin/isabelle_scala_script"} allows to run
- Isabelle/Scala source files stand-alone programs, by using a
+text \<open>
+ The executable @{executable "$ISABELLE_HOME/bin/isabelle_scala_script"}
+ allows to run Isabelle/Scala source files stand-alone programs, by using a
suitable ``hash-bang'' line and executable file permissions.
- The subsequent example assumes that the main Isabelle binaries have
- been installed in some directory that is included in @{setting PATH}
- (see also @{tool "install"}):
+ The subsequent example assumes that the main Isabelle binaries have been
+ installed in some directory that is included in @{setting PATH} (see also
+ @{tool "install"}):
@{verbatim [display]
\<open>#!/usr/bin/env isabelle_scala_script
@@ -84,8 +89,8 @@
Console.println("browser_info = " + options.bool("browser_info"))
Console.println("document = " + options.string("document"))\<close>}
- Alternatively the full @{file
- "$ISABELLE_HOME/bin/isabelle_scala_script"} may be specified in
- expanded form.\<close>
+ Alternatively the full @{file "$ISABELLE_HOME/bin/isabelle_scala_script"}
+ may be specified in expanded form.
+\<close>
end
--- a/src/Doc/System/Sessions.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Doc/System/Sessions.thy Thu Nov 05 08:27:22 2015 +0100
@@ -1,52 +1,52 @@
+(*:wrap=hard:maxLineLen=78:*)
+
theory Sessions
imports Base
begin
chapter \<open>Isabelle sessions and build management \label{ch:session}\<close>
-text \<open>An Isabelle \<^emph>\<open>session\<close> consists of a collection of related
- theories that may be associated with formal documents
- (\chref{ch:present}). There is also a notion of \<^emph>\<open>persistent
- heap\<close> image to capture the state of a session, similar to
- object-code in compiled programming languages. Thus the concept of
- session resembles that of a ``project'' in common IDE environments,
- but the specific name emphasizes the connection to interactive
- theorem proving: the session wraps-up the results of
- user-interaction with the prover in a persistent form.
+text \<open>
+ An Isabelle \<^emph>\<open>session\<close> consists of a collection of related theories that may
+ be associated with formal documents (\chref{ch:present}). There is also a
+ notion of \<^emph>\<open>persistent heap\<close> image to capture the state of a session,
+ similar to object-code in compiled programming languages. Thus the concept
+ of session resembles that of a ``project'' in common IDE environments, but
+ the specific name emphasizes the connection to interactive theorem proving:
+ the session wraps-up the results of user-interaction with the prover in a
+ persistent form.
- Application sessions are built on a given parent session, which may
- be built recursively on other parents. Following this path in the
- hierarchy eventually leads to some major object-logic session like
- \<open>HOL\<close>, which itself is based on \<open>Pure\<close> as the common
- root of all sessions.
+ Application sessions are built on a given parent session, which may be built
+ recursively on other parents. Following this path in the hierarchy
+ eventually leads to some major object-logic session like \<open>HOL\<close>, which itself
+ is based on \<open>Pure\<close> as the common root of all sessions.
- Processing sessions may take considerable time. Isabelle build
- management helps to organize this efficiently. This includes
- support for parallel build jobs, in addition to the multithreaded
- theory and proof checking that is already provided by the prover
- process itself.\<close>
+ Processing sessions may take considerable time. Isabelle build management
+ helps to organize this efficiently. This includes support for parallel build
+ jobs, in addition to the multithreaded theory and proof checking that is
+ already provided by the prover process itself.
+\<close>
section \<open>Session ROOT specifications \label{sec:session-root}\<close>
-text \<open>Session specifications reside in files called \<^verbatim>\<open>ROOT\<close>
- within certain directories, such as the home locations of registered
- Isabelle components or additional project directories given by the
- user.
+text \<open>
+ Session specifications reside in files called \<^verbatim>\<open>ROOT\<close> within certain
+ directories, such as the home locations of registered Isabelle components or
+ additional project directories given by the user.
- The ROOT file format follows the lexical conventions of the
- \<^emph>\<open>outer syntax\<close> of Isabelle/Isar, see also
- @{cite "isabelle-isar-ref"}. This defines common forms like
- identifiers, names, quoted strings, verbatim text, nested comments
- etc. The grammar for @{syntax session_chapter} and @{syntax
- session_entry} is given as syntax diagram below; each ROOT file may
- contain multiple specifications like this. Chapters help to
- organize browser info (\secref{sec:info}), but have no formal
- meaning. The default chapter is ``\<open>Unsorted\<close>''.
+ The ROOT file format follows the lexical conventions of the \<^emph>\<open>outer syntax\<close>
+ of Isabelle/Isar, see also @{cite "isabelle-isar-ref"}. This defines common
+ forms like identifiers, names, quoted strings, verbatim text, nested
+ comments etc. The grammar for @{syntax session_chapter} and @{syntax
+ session_entry} is given as syntax diagram below; each ROOT file may contain
+ multiple specifications like this. Chapters help to organize browser info
+ (\secref{sec:info}), but have no formal meaning. The default chapter is
+ ``\<open>Unsorted\<close>''.
- Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing
- mode \<^verbatim>\<open>isabelle-root\<close> for session ROOT files, which is
- enabled by default for any file of that name.
+ Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing mode
+ \<^verbatim>\<open>isabelle-root\<close> for session ROOT files, which is enabled by default for any
+ file of that name.
@{rail \<open>
@{syntax_def session_chapter}: @'chapter' @{syntax name}
@@ -77,151 +77,138 @@
document_files: @'document_files' ('(' dir ')')? (@{syntax name}+)
\<close>}
- \<^descr> \isakeyword{session}~\<open>A = B + body\<close> defines a new
- session \<open>A\<close> based on parent session \<open>B\<close>, with its
- content given in \<open>body\<close> (theories and auxiliary source files).
- Note that a parent (like \<open>HOL\<close>) is mandatory in practical
+ \<^descr> \isakeyword{session}~\<open>A = B + body\<close> defines a new session \<open>A\<close> based on
+ parent session \<open>B\<close>, with its content given in \<open>body\<close> (theories and auxiliary
+ source files). Note that a parent (like \<open>HOL\<close>) is mandatory in practical
applications: only Isabelle/Pure can bootstrap itself from nothing.
- All such session specifications together describe a hierarchy (tree)
- of sessions, with globally unique names. The new session name
- \<open>A\<close> should be sufficiently long and descriptive to stand on
- its own in a potentially large library.
+ All such session specifications together describe a hierarchy (tree) of
+ sessions, with globally unique names. The new session name \<open>A\<close> should be
+ sufficiently long and descriptive to stand on its own in a potentially large
+ library.
- \<^descr> \isakeyword{session}~\<open>A (groups)\<close> indicates a
- collection of groups where the new session is a member. Group names
- are uninterpreted and merely follow certain conventions. For
- example, the Isabelle distribution tags some important sessions by
- the group name called ``\<open>main\<close>''. Other projects may invent
- their own conventions, but this requires some care to avoid clashes
+ \<^descr> \isakeyword{session}~\<open>A (groups)\<close> indicates a collection of groups where
+ the new session is a member. Group names are uninterpreted and merely follow
+ certain conventions. For example, the Isabelle distribution tags some
+ important sessions by the group name called ``\<open>main\<close>''. Other projects may
+ invent their own conventions, but this requires some care to avoid clashes
within this unchecked name space.
- \<^descr> \isakeyword{session}~\<open>A\<close>~\isakeyword{in}~\<open>dir\<close>
- specifies an explicit directory for this session; by default this is
- the current directory of the \<^verbatim>\<open>ROOT\<close> file.
+ \<^descr> \isakeyword{session}~\<open>A\<close>~\isakeyword{in}~\<open>dir\<close> specifies an explicit
+ directory for this session; by default this is the current directory of the
+ \<^verbatim>\<open>ROOT\<close> file.
- All theories and auxiliary source files are located relatively to
- the session directory. The prover process is run within the same as
- its current working directory.
-
- \<^descr> \isakeyword{description}~\<open>text\<close> is a free-form
- annotation for this session.
+ All theories and auxiliary source files are located relatively to the
+ session directory. The prover process is run within the same as its current
+ working directory.
- \<^descr> \isakeyword{options}~\<open>[x = a, y = b, z]\<close> defines
- separate options (\secref{sec:system-options}) that are used when
- processing this session, but \<^emph>\<open>without\<close> propagation to child
- sessions. Note that \<open>z\<close> abbreviates \<open>z = true\<close> for
- Boolean options.
+ \<^descr> \isakeyword{description}~\<open>text\<close> is a free-form annotation for this
+ session.
- \<^descr> \isakeyword{theories}~\<open>options names\<close> specifies a
- block of theories that are processed within an environment that is
- augmented by the given options, in addition to the global session
- options given before. Any number of blocks of \isakeyword{theories}
- may be given. Options are only active for each
+ \<^descr> \isakeyword{options}~\<open>[x = a, y = b, z]\<close> defines separate options
+ (\secref{sec:system-options}) that are used when processing this session,
+ but \<^emph>\<open>without\<close> propagation to child sessions. Note that \<open>z\<close> abbreviates \<open>z =
+ true\<close> for Boolean options.
+
+ \<^descr> \isakeyword{theories}~\<open>options names\<close> specifies a block of theories that
+ are processed within an environment that is augmented by the given options,
+ in addition to the global session options given before. Any number of blocks
+ of \isakeyword{theories} may be given. Options are only active for each
\isakeyword{theories} block separately.
- \<^descr> \isakeyword{files}~\<open>files\<close> lists additional source
- files that are involved in the processing of this session. This
- should cover anything outside the formal content of the theory
- sources. In contrast, files that are loaded formally
- within a theory, e.g.\ via @{command "ML_file"}, need not be
+ \<^descr> \isakeyword{files}~\<open>files\<close> lists additional source files that are involved
+ in the processing of this session. This should cover anything outside the
+ formal content of the theory sources. In contrast, files that are loaded
+ formally within a theory, e.g.\ via @{command "ML_file"}, need not be
declared again.
- \<^descr> \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>base_dir) files\<close> lists source files for document preparation,
- typically \<^verbatim>\<open>.tex\<close> and \<^verbatim>\<open>.sty\<close> for {\LaTeX}.
- Only these explicitly given files are copied from the base directory
- to the document output directory, before formal document processing
- is started (see also \secref{sec:tool-document}). The local path
- structure of the \<open>files\<close> is preserved, which allows to
- reconstruct the original directory hierarchy of \<open>base_dir\<close>.
+ \<^descr> \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>base_dir) files\<close> lists
+ source files for document preparation, typically \<^verbatim>\<open>.tex\<close> and \<^verbatim>\<open>.sty\<close> for
+ {\LaTeX}. Only these explicitly given files are copied from the base
+ directory to the document output directory, before formal document
+ processing is started (see also \secref{sec:tool-document}). The local path
+ structure of the \<open>files\<close> is preserved, which allows to reconstruct the
+ original directory hierarchy of \<open>base_dir\<close>.
\<^descr> \isakeyword{document_files}~\<open>files\<close> abbreviates
- \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>document) files\<close>, i.e.\ document sources are taken from the base
- directory \<^verbatim>\<open>document\<close> within the session root directory.
+ \isakeyword{document_files}~\<open>(\<close>\isakeyword{in}~\<open>document) files\<close>, i.e.\
+ document sources are taken from the base directory \<^verbatim>\<open>document\<close> within the
+ session root directory.
\<close>
subsubsection \<open>Examples\<close>
-text \<open>See @{file "~~/src/HOL/ROOT"} for a diversity of practically
- relevant situations, although it uses relatively complex
- quasi-hierarchic naming conventions like \<^verbatim>\<open>HOL-SPARK\<close>,
- \<^verbatim>\<open>HOL-SPARK-Examples\<close>. An alternative is to use
- unqualified names that are relatively long and descriptive, as in
- the Archive of Formal Proofs (@{url "http://afp.sourceforge.net"}), for
- example.\<close>
+text \<open>
+ See @{file "~~/src/HOL/ROOT"} for a diversity of practically relevant
+ situations, although it uses relatively complex quasi-hierarchic naming
+ conventions like \<^verbatim>\<open>HOL-SPARK\<close>, \<^verbatim>\<open>HOL-SPARK-Examples\<close>. An alternative is to
+ use unqualified names that are relatively long and descriptive, as in the
+ Archive of Formal Proofs (@{url "http://afp.sourceforge.net"}), for
+ example.
+\<close>
section \<open>System build options \label{sec:system-options}\<close>
-text \<open>See @{file "~~/etc/options"} for the main defaults provided by
- the Isabelle distribution. Isabelle/jEdit @{cite "isabelle-jedit"}
- includes a simple editing mode \<^verbatim>\<open>isabelle-options\<close> for
- this file-format.
+text \<open>
+ See @{file "~~/etc/options"} for the main defaults provided by the Isabelle
+ distribution. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple
+ editing mode \<^verbatim>\<open>isabelle-options\<close> for this file-format.
- The following options are particularly relevant to build Isabelle
- sessions, in particular with document preparation
- (\chref{ch:present}).
+ The following options are particularly relevant to build Isabelle sessions,
+ in particular with document preparation (\chref{ch:present}).
- \<^item> @{system_option_def "browser_info"} controls output of HTML
- browser info, see also \secref{sec:info}.
+ \<^item> @{system_option_def "browser_info"} controls output of HTML browser
+ info, see also \secref{sec:info}.
- \<^item> @{system_option_def "document"} specifies the document output
- format, see @{tool document} option \<^verbatim>\<open>-o\<close> in
- \secref{sec:tool-document}. In practice, the most relevant values
- are \<^verbatim>\<open>document=false\<close> or \<^verbatim>\<open>document=pdf\<close>.
+ \<^item> @{system_option_def "document"} specifies the document output format,
+ see @{tool document} option \<^verbatim>\<open>-o\<close> in \secref{sec:tool-document}. In
+ practice, the most relevant values are \<^verbatim>\<open>document=false\<close> or
+ \<^verbatim>\<open>document=pdf\<close>.
- \<^item> @{system_option_def "document_output"} specifies an
- alternative directory for generated output of the document
- preparation system; the default is within the @{setting
- "ISABELLE_BROWSER_INFO"} hierarchy as explained in
- \secref{sec:info}. See also @{tool mkroot}, which generates a
- default configuration with output readily available to the author of
- the document.
+ \<^item> @{system_option_def "document_output"} specifies an alternative
+ directory for generated output of the document preparation system; the
+ default is within the @{setting "ISABELLE_BROWSER_INFO"} hierarchy as
+ explained in \secref{sec:info}. See also @{tool mkroot}, which generates a
+ default configuration with output readily available to the author of the
+ document.
- \<^item> @{system_option_def "document_variants"} specifies document
- variants as a colon-separated list of \<open>name=tags\<close> entries,
- corresponding to @{tool document} options \<^verbatim>\<open>-n\<close> and
- \<^verbatim>\<open>-t\<close>.
+ \<^item> @{system_option_def "document_variants"} specifies document variants as
+ a colon-separated list of \<open>name=tags\<close> entries, corresponding to @{tool
+ document} options \<^verbatim>\<open>-n\<close> and \<^verbatim>\<open>-t\<close>.
- For example, \<^verbatim>\<open>document_variants=document:outline=/proof,/ML\<close> indicates
- two documents: the one called \<^verbatim>\<open>document\<close> with default tags,
- and the other called \<^verbatim>\<open>outline\<close> where proofs and ML
- sections are folded.
+ For example, \<^verbatim>\<open>document_variants=document:outline=/proof,/ML\<close> indicates
+ two documents: the one called \<^verbatim>\<open>document\<close> with default tags, and the other
+ called \<^verbatim>\<open>outline\<close> where proofs and ML sections are folded.
- Document variant names are just a matter of conventions. It is also
- possible to use different document variant names (without tags) for
- different document root entries, see also
- \secref{sec:tool-document}.
+ Document variant names are just a matter of conventions. It is also
+ possible to use different document variant names (without tags) for
+ different document root entries, see also \secref{sec:tool-document}.
- \<^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 machines with many cores or with
- hyperthreading, this is often requires manual adjustment (on the
- command-line or within personal settings or preferences, not within
- a session \<^verbatim>\<open>ROOT\<close>).
+ \<^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
+ machines with many cores or with hyperthreading, this is often requires
+ manual adjustment (on the command-line or within personal settings or
+ preferences, not within a session \<^verbatim>\<open>ROOT\<close>).
- \<^item> @{system_option_def "condition"} specifies a comma-separated
- list of process environment variables (or Isabelle settings) that
- are required for the subsequent theories to be processed.
- Conditions are considered ``true'' if the corresponding environment
- value is defined and non-empty.
+ \<^item> @{system_option_def "condition"} specifies a comma-separated list of
+ process environment variables (or Isabelle settings) that are required for
+ the subsequent theories to be processed. Conditions are considered
+ ``true'' if the corresponding environment value is defined and non-empty.
- For example, the \<^verbatim>\<open>condition=ISABELLE_FULL_TEST\<close> may be
- used to guard extraordinary theories, which are meant to be enabled
- explicitly via some shell prefix \<^verbatim>\<open>env ISABELLE_FULL_TEST=true\<close>
- before invoking @{tool build}.
+ For example, the \<^verbatim>\<open>condition=ISABELLE_FULL_TEST\<close> may be used to guard
+ extraordinary theories, which are meant to be enabled explicitly via some
+ shell prefix \<^verbatim>\<open>env ISABELLE_FULL_TEST=true\<close> before invoking @{tool build}.
- \<^item> @{system_option_def "timeout"} specifies a real wall-clock
- timeout (in seconds) for the session as a whole. The timer is
- controlled outside the ML process by the JVM that runs
- Isabelle/Scala. Thus it is relatively reliable in canceling
- processes that get out of control, even if there is a deadlock
- without CPU time usage.
+ \<^item> @{system_option_def "timeout"} specifies a real wall-clock timeout (in
+ seconds) for the session as a whole. The timer is controlled outside the
+ ML process by the JVM that runs Isabelle/Scala. Thus it is relatively
+ reliable in canceling processes that get out of control, even if there is
+ a deadlock without CPU time usage.
-
- The @{tool_def options} tool prints Isabelle system options. Its
+ The @{tool_def options} tool prints Isabelle system options. Its
command-line usage is:
@{verbatim [display]
\<open>Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...]
@@ -235,32 +222,29 @@
Report Isabelle system options, augmented by MORE_OPTIONS given as
arguments NAME=VAL or NAME.\<close>}
- The command line arguments provide additional system options of the
- form \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close>
- for Boolean options.
+ The command line arguments provide additional system options of the form
+ \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close> for Boolean options.
+
+ Option \<^verbatim>\<open>-b\<close> augments the implicit environment of system options by the ones
+ of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ \secref{sec:tool-build}.
- Option \<^verbatim>\<open>-b\<close> augments the implicit environment of system
- options by the ones of @{setting ISABELLE_BUILD_OPTIONS}, cf.\
- \secref{sec:tool-build}.
+ Option \<^verbatim>\<open>-g\<close> prints the value of the given option. Option \<^verbatim>\<open>-l\<close> lists all
+ options with their declaration and current value.
- Option \<^verbatim>\<open>-g\<close> prints the value of the given option.
- Option \<^verbatim>\<open>-l\<close> lists all options with their declaration and
- current value.
-
- Option \<^verbatim>\<open>-x\<close> specifies a file to export the result in
- YXML format, instead of printing it in human-readable form.
+ Option \<^verbatim>\<open>-x\<close> specifies a file to export the result in YXML format, instead
+ of printing it in human-readable form.
\<close>
section \<open>Invoking the build process \label{sec:tool-build}\<close>
-text \<open>The @{tool_def build} tool invokes the build process for
- Isabelle sessions. It manages dependencies between sessions,
- related sources of theories and auxiliary files, and target heap
- images. Accordingly, it runs instances of the prover process with
- optional document preparation. Its command-line usage
- is:\footnote{Isabelle/Scala provides the same functionality via
- \<^verbatim>\<open>isabelle.Build.build\<close>.}
+text \<open>
+ The @{tool_def build} tool invokes the build process for Isabelle sessions.
+ It manages dependencies between sessions, related sources of theories and
+ auxiliary files, and target heap images. Accordingly, it runs instances of
+ the prover process with optional document preparation. Its command-line
+ usage is:\<^footnote>\<open>Isabelle/Scala provides the same functionality via
+ \<^verbatim>\<open>isabelle.Build.build\<close>.\<close>
@{verbatim [display]
\<open>Usage: isabelle build [OPTIONS] [SESSIONS ...]
@@ -291,98 +275,87 @@
ML_OPTIONS="..."\<close>}
\<^medskip>
- Isabelle sessions are defined via session ROOT files as
- described in (\secref{sec:session-root}). The totality of sessions
- is determined by collecting such specifications from all Isabelle
- component directories (\secref{sec:components}), augmented by more
- directories given via options \<^verbatim>\<open>-d\<close>~\<open>DIR\<close> on the
- command line. Each such directory may contain a session
+ Isabelle sessions are defined via session ROOT files as described in
+ (\secref{sec:session-root}). The totality of sessions is determined by
+ collecting such specifications from all Isabelle component directories
+ (\secref{sec:components}), augmented by more directories given via options
+ \<^verbatim>\<open>-d\<close>~\<open>DIR\<close> on the command line. Each such directory may contain a session
\<^verbatim>\<open>ROOT\<close> file with several session specifications.
- Any session root directory may refer recursively to further
- directories of the same kind, by listing them in a catalog file
- \<^verbatim>\<open>ROOTS\<close> line-by-line. This helps to organize large
- collections of session specifications, or to make \<^verbatim>\<open>-d\<close>
- command line options persistent (say within
+ Any session root directory may refer recursively to further directories of
+ the same kind, by listing them in a catalog file \<^verbatim>\<open>ROOTS\<close> line-by-line. This
+ helps to organize large collections of session specifications, or to make
+ \<^verbatim>\<open>-d\<close> command line options persistent (say within
\<^verbatim>\<open>$ISABELLE_HOME_USER/ROOTS\<close>).
\<^medskip>
- The subset of sessions to be managed is determined via
- individual \<open>SESSIONS\<close> given as command-line arguments, or
- session groups that are given via one or more options \<^verbatim>\<open>-g\<close>~\<open>NAME\<close>.
- Option \<^verbatim>\<open>-a\<close> selects all sessions.
- The build tool takes session dependencies into account: the set of
- selected sessions is completed by including all ancestors.
+ The subset of sessions to be managed is determined via individual \<open>SESSIONS\<close>
+ given as command-line arguments, or session groups that are given via one or
+ more options \<^verbatim>\<open>-g\<close>~\<open>NAME\<close>. Option \<^verbatim>\<open>-a\<close> selects all sessions. The build tool
+ takes session dependencies into account: the set of selected sessions is
+ completed by including all ancestors.
\<^medskip>
- One or more options \<^verbatim>\<open>-x\<close>~\<open>NAME\<close> specify
- sessions to be excluded. All descendents of excluded sessions are removed
- from the selection as specified above. Option \<^verbatim>\<open>-X\<close> is
- analogous to this, but excluded sessions are specified by session group
- membership.
+ One or more options \<^verbatim>\<open>-x\<close>~\<open>NAME\<close> specify sessions to be excluded. All
+ descendents of excluded sessions are removed from the selection as specified
+ above. Option \<^verbatim>\<open>-X\<close> is analogous to this, but excluded sessions are
+ specified by session group membership.
\<^medskip>
- Option \<^verbatim>\<open>-R\<close> reverses the selection in the sense
- that it refers to its requirements: all ancestor sessions excluding
- the original selection. This allows to prepare the stage for some
- build process with different options, before running the main build
- itself (without option \<^verbatim>\<open>-R\<close>).
+ Option \<^verbatim>\<open>-R\<close> reverses the selection in the sense that it refers to its
+ requirements: all ancestor sessions excluding the original selection. This
+ allows to prepare the stage for some build process with different options,
+ before running the main build itself (without option \<^verbatim>\<open>-R\<close>).
\<^medskip>
- Option \<^verbatim>\<open>-D\<close> is similar to \<^verbatim>\<open>-d\<close>, but
- selects all sessions that are defined in the given directories.
+ Option \<^verbatim>\<open>-D\<close> is similar to \<^verbatim>\<open>-d\<close>, but selects all sessions that are defined
+ in the given directories.
\<^medskip>
The build process depends on additional options
- (\secref{sec:system-options}) that are passed to the prover
- eventually. The settings variable @{setting_ref
- ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\
- \<^verbatim>\<open>ISABELLE_BUILD_OPTIONS="document=pdf threads=4"\<close>. Moreover,
- the environment of system build options may be augmented on the
- command line via \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<^verbatim>\<open>-o\<close>~\<open>name\<close>, which
- abbreviates \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=true\<close> for
- Boolean options. Multiple occurrences of \<^verbatim>\<open>-o\<close> on the
- command-line are applied in the given order.
+ (\secref{sec:system-options}) that are passed to the prover eventually. The
+ settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide
+ additional defaults, e.g.\ \<^verbatim>\<open>ISABELLE_BUILD_OPTIONS="document=pdf
+ threads=4"\<close>. Moreover, the environment of system build options may be
+ augmented on the command line via \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<^verbatim>\<open>-o\<close>~\<open>name\<close>,
+ which abbreviates \<^verbatim>\<open>-o\<close>~\<open>name\<close>\<^verbatim>\<open>=true\<close> for Boolean options. Multiple
+ occurrences of \<^verbatim>\<open>-o\<close> on the command-line are applied in the given order.
\<^medskip>
- Option \<^verbatim>\<open>-b\<close> ensures that heap images are
- produced for all selected sessions. By default, images are only
- saved for inner nodes of the hierarchy of sessions, as required for
- other sessions to continue later on.
+ Option \<^verbatim>\<open>-b\<close> ensures that heap images are produced for all selected
+ sessions. By default, images are only saved for inner nodes of the hierarchy
+ of sessions, as required for other sessions to continue later on.
\<^medskip>
- Option \<^verbatim>\<open>-c\<close> cleans all descendants of the
- selected sessions before performing the specified build operation.
+ Option \<^verbatim>\<open>-c\<close> cleans all descendants of the selected sessions before
+ performing the specified build operation.
\<^medskip>
- Option \<^verbatim>\<open>-n\<close> omits the actual build process
- after the preparatory stage (including optional cleanup). Note that
- the return code always indicates the status of the set of selected
- sessions.
+ Option \<^verbatim>\<open>-n\<close> omits the actual build process after the preparatory stage
+ (including optional cleanup). Note that the return code always indicates the
+ status of the set of selected sessions.
\<^medskip>
- Option \<^verbatim>\<open>-j\<close> specifies the maximum number of
- parallel build jobs (prover processes). Each prover process is
- subject to a separate limit of parallel worker threads, cf.\ system
- option @{system_option_ref threads}.
+ Option \<^verbatim>\<open>-j\<close> specifies the maximum number of parallel build jobs (prover
+ processes). Each prover process is subject to a separate limit of parallel
+ worker threads, cf.\ system option @{system_option_ref threads}.
\<^medskip>
- Option \<^verbatim>\<open>-s\<close> enables \<^emph>\<open>system mode\<close>, which
- means that resulting heap images and log files are stored in
- @{file_unchecked "$ISABELLE_HOME/heaps"} instead of the default location
- @{setting ISABELLE_OUTPUT} (which is normally in @{setting
- ISABELLE_HOME_USER}, i.e.\ the user's home directory).
+ Option \<^verbatim>\<open>-s\<close> enables \<^emph>\<open>system mode\<close>, which means that resulting heap images
+ and log files are stored in @{file_unchecked "$ISABELLE_HOME/heaps"} instead
+ of the default location @{setting ISABELLE_OUTPUT} (which is normally in
+ @{setting ISABELLE_HOME_USER}, i.e.\ the user's home directory).
\<^medskip>
- Option \<^verbatim>\<open>-v\<close> increases the general level of
- verbosity. Option \<^verbatim>\<open>-l\<close> lists the source files that
- contribute to a session.
+ Option \<^verbatim>\<open>-v\<close> increases the general level of verbosity. Option \<^verbatim>\<open>-l\<close> lists
+ the source files that contribute to a session.
\<^medskip>
- Option \<^verbatim>\<open>-k\<close> specifies a newly proposed keyword for
- outer syntax (multiple uses allowed). The theory sources are checked for
- conflicts wrt.\ this hypothetical change of syntax, e.g.\ to reveal
- occurrences of identifiers that need to be quoted.\<close>
+ Option \<^verbatim>\<open>-k\<close> specifies a newly proposed keyword for outer syntax (multiple
+ uses allowed). The theory sources are checked for conflicts wrt.\ this
+ hypothetical change of syntax, e.g.\ to reveal occurrences of identifiers
+ that need to be quoted.
+\<close>
subsubsection \<open>Examples\<close>
@@ -396,23 +369,22 @@
@{verbatim [display] \<open>isabelle build -b -g main\<close>}
\<^smallskip>
- Provide a general overview of the status of all Isabelle
- sessions, without building anything:
+ Provide a general overview of the status of all Isabelle sessions, without
+ building anything:
@{verbatim [display] \<open>isabelle build -a -n -v\<close>}
\<^smallskip>
- Build all sessions with HTML browser info and PDF
- document preparation:
+ Build all sessions with HTML browser info and PDF document preparation:
@{verbatim [display] \<open>isabelle build -a -o browser_info -o document=pdf\<close>}
\<^smallskip>
- Build all sessions with a maximum of 8 parallel prover
- processes and 4 worker threads each (on a machine with many cores):
+ Build all sessions with a maximum of 8 parallel prover processes and 4
+ worker threads each (on a machine with many cores):
@{verbatim [display] \<open>isabelle build -a -j8 -o threads=4\<close>}
\<^smallskip>
- Build some session images with cleanup of their
- descendants, while retaining their ancestry:
+ Build some session images with cleanup of their descendants, while retaining
+ their ancestry:
@{verbatim [display] \<open>isabelle build -b -c HOL-Algebra HOL-Word\<close>}
\<^smallskip>
@@ -420,14 +392,14 @@
@{verbatim [display] \<open>isabelle build -a -n -c\<close>}
\<^smallskip>
- Build all sessions from some other directory hierarchy,
- according to the settings variable \<^verbatim>\<open>AFP\<close> that happens to
- be defined inside the Isabelle environment:
+ Build all sessions from some other directory hierarchy, according to the
+ settings variable \<^verbatim>\<open>AFP\<close> that happens to be defined inside the Isabelle
+ environment:
@{verbatim [display] \<open>isabelle build -D '$AFP'\<close>}
\<^smallskip>
- Inform about the status of all sessions required for AFP,
- without building anything yet:
+ Inform about the status of all sessions required for AFP, without building
+ anything yet:
@{verbatim [display] \<open>isabelle build -D '$AFP' -R -v -n\<close>}
\<close>
--- a/src/Pure/General/symbol.ML Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/General/symbol.ML Thu Nov 05 08:27:22 2015 +0100
@@ -10,6 +10,7 @@
val STX: symbol
val DEL: symbol
val space: symbol
+ val comment: symbol
val is_char: symbol -> bool
val is_utf8: symbol -> bool
val is_symbolic: symbol -> bool
@@ -93,6 +94,8 @@
val space = chr 32;
+val comment = "\\<comment>";
+
fun is_char s = size s = 1;
fun is_utf8 s = size s > 0 andalso forall_string (fn c => ord c >= 128) s;
--- a/src/Pure/General/symbol.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/General/symbol.scala Thu Nov 05 08:27:22 2015 +0100
@@ -433,6 +433,11 @@
val symbolic = recode_set((for { (sym, _) <- symbols; if raw_symbolic(sym) } yield sym): _*)
+ /* comment */
+
+ val comment_decoded = decode(comment)
+
+
/* cartouches */
val open_decoded = decode(open)
@@ -496,10 +501,16 @@
def is_blank(sym: Symbol): Boolean = symbols.blanks.contains(sym)
+ /* comment */
+
+ val comment: Symbol = "\\<comment>"
+ def comment_decoded: Symbol = symbols.comment_decoded
+
+
/* cartouches */
- val open = "\\<open>"
- val close = "\\<close>"
+ val open: Symbol = "\\<open>"
+ val close: Symbol = "\\<close>"
def open_decoded: Symbol = symbols.open_decoded
def close_decoded: Symbol = symbols.close_decoded
--- a/src/Pure/Isar/outer_syntax.ML Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/Isar/outer_syntax.ML Thu Nov 05 08:27:22 2015 +0100
@@ -197,7 +197,7 @@
in msg ^ quote (Markup.markup Markup.keyword1 name) end))
end);
-val parse_cmt = Parse.$$$ "--" -- Parse.!!! Parse.document_source;
+val parse_cmt = (Parse.$$$ "--" || Parse.$$$ Symbol.comment) -- Parse.!!! Parse.document_source;
fun commands_source thy =
Token.source_proper #>
@@ -261,7 +261,7 @@
(* side-comments *)
fun cmts (t1 :: t2 :: toks) =
- if Token.keyword_with (fn s => s = "--") t1 then t2 :: cmts toks
+ if Token.keyword_with (fn s => s = "--" orelse s = Symbol.comment) t1 then t2 :: cmts toks
else cmts (t2 :: toks)
| cmts _ = [];
--- a/src/Pure/PIDE/command.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/PIDE/command.scala Thu Nov 05 08:27:22 2015 +0100
@@ -308,10 +308,11 @@
private def clean_tokens(tokens: List[Token]): List[(Token, Int)] =
{
+ val markers = Set("%", "--", Symbol.comment, Symbol.comment_decoded)
def clean(toks: List[(Token, Int)]): List[(Token, Int)] =
toks match {
case (t1, i1) :: (t2, i2) :: rest =>
- if (t1.is_keyword && (t1.source == "%" || t1.source == "--")) clean(rest)
+ if (t1.is_keyword && markers(t1.source)) clean(rest)
else (t1, i1) :: clean((t2, i2) :: rest)
case _ => toks
}
--- a/src/Pure/Pure.thy Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/Pure.thy Thu Nov 05 08:27:22 2015 +0100
@@ -6,7 +6,7 @@
theory Pure
keywords
- "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<equiv>"
+ "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<comment>" "\<equiv>"
"\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
"\<subseteq>" "]" "assumes" "attach" "binder" "constrains"
"defines" "fixes" "for" "identifier" "if" "in" "includes" "infix"
--- a/src/Pure/System/options.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/System/options.scala Thu Nov 05 08:27:22 2015 +0100
@@ -75,7 +75,7 @@
private val PREFS = PREFS_DIR + Path.basic("preferences")
lazy val options_syntax =
- Outer_Syntax.init() + ":" + "=" + "--" +
+ Outer_Syntax.init() + ":" + "=" + "--" + Symbol.comment + Symbol.comment_decoded +
(SECTION, Keyword.DOCUMENT_HEADING) + PUBLIC + (OPTION, Keyword.THY_DECL)
lazy val prefs_syntax = Outer_Syntax.init() + "="
@@ -89,12 +89,15 @@
{ case s ~ n => if (s.isDefined) "-" + n else n } |
atom("option value", tok => tok.is_name || tok.is_float)
+ def comment_marker: Parser[String] =
+ $$$("--") | $$$(Symbol.comment) | $$$(Symbol.comment_decoded)
+
val option_entry: Parser[Options => Options] =
{
command(SECTION) ~! text ^^
{ case _ ~ a => (options: Options) => options.set_section(a) } |
opt($$$(PUBLIC)) ~ command(OPTION) ~! (position(option_name) ~ $$$(":") ~ option_type ~
- $$$("=") ~ option_value ~ ($$$("--") ~! text ^^ { case _ ~ x => x } | success(""))) ^^
+ $$$("=") ~ option_value ~ (comment_marker ~! text ^^ { case _ ~ x => x } | success(""))) ^^
{ case a ~ _ ~ ((b, pos) ~ _ ~ c ~ _ ~ d ~ e) =>
(options: Options) => options.declare(a.isDefined, pos, b, c, d, e) }
}
--- a/src/Pure/Thy/latex.ML Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/Thy/latex.ML Thu Nov 05 08:27:22 2015 +0100
@@ -36,7 +36,7 @@
| "\t" => "\\ "
| "\n" => "\\isanewline\n"
| s =>
- if exists_string (fn s' => s = s') "\"#$%&'<>\\^_{}~"
+ if exists_string (fn s' => s = s') "\"#$%&',-<>\\^_`{}~"
then enclose "{\\char`\\" "}" s else s);
--- a/src/Pure/Thy/thy_output.ML Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/Thy/thy_output.ML Thu Nov 05 08:27:22 2015 +0100
@@ -430,7 +430,8 @@
(Basic_Token cmd, (markup_false, d)))]));
val cmt = Scan.peek (fn d =>
- Parse.$$$ "--" |-- Parse.!!!! (improper |-- Parse.document_source) >>
+ (Parse.$$$ "--" || Parse.$$$ Symbol.comment) |--
+ Parse.!!!! (improper |-- Parse.document_source) >>
(fn source => (NONE, (Markup_Token ("cmt", source), ("", d)))));
val other = Scan.peek (fn d =>
@@ -615,7 +616,8 @@
val _ =
Theory.setup
- (control_antiquotation @{binding emph} "\\emph{" "}" #>
+ (control_antiquotation @{binding footnote} "\\footnote{" "}" #>
+ control_antiquotation @{binding emph} "\\emph{" "}" #>
control_antiquotation @{binding bold} "\\textbf{" "}");
end;
--- a/src/Pure/Tools/update_cartouches.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Pure/Tools/update_cartouches.scala Thu Nov 05 08:27:22 2015 +0100
@@ -1,7 +1,7 @@
/* Title: Pure/Tools/update_cartouches.scala
Author: Makarius
-Update theory syntax to use cartouches.
+Update theory syntax to use cartouches etc.
*/
package isabelle
@@ -37,11 +37,11 @@
}
}
- def update_cartouches(replace_text: Boolean, path: Path)
+ def update_cartouches(replace_comment: Boolean, replace_text: Boolean, path: Path)
{
val text0 = File.read(path)
- // outer syntax cartouches
+ // outer syntax cartouches and comment markers
val text1 =
(for (tok <- Token.explode(Keyword.Keywords.empty, text0).iterator)
yield {
@@ -52,6 +52,7 @@
case s => tok.source
}
}
+ else if (replace_comment && tok.source == "--") Symbol.comment
else tok.source
}
).mkString
@@ -87,8 +88,10 @@
{
Command_Line.tool0 {
args.toList match {
- case Properties.Value.Boolean(replace_text) :: files =>
- files.foreach(file => update_cartouches(replace_text, Path.explode(file)))
+ case Properties.Value.Boolean(replace_comment) ::
+ Properties.Value.Boolean(replace_text) :: files =>
+ files.foreach(file =>
+ update_cartouches(replace_comment, replace_text, Path.explode(file)))
case _ => error("Bad arguments:\n" + cat_lines(args))
}
}
--- a/src/Tools/jEdit/src/pretty_text_area.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Tools/jEdit/src/pretty_text_area.scala Thu Nov 05 08:27:22 2015 +0100
@@ -18,9 +18,11 @@
import scala.swing.{Label, Component}
import scala.util.matching.Regex
-import org.gjt.sp.jedit.{jEdit, View, Registers}
+import org.gjt.sp.jedit.{jEdit, View, Registers, JEditBeanShellAction}
+import org.gjt.sp.jedit.input.{DefaultInputHandlerProvider, TextAreaInputHandler}
import org.gjt.sp.jedit.textarea.{AntiAlias, JEditEmbeddedTextArea}
import org.gjt.sp.jedit.syntax.SyntaxStyle
+import org.gjt.sp.jedit.gui.KeyEventTranslator
import org.gjt.sp.util.{SyntaxUtilities, Log}
@@ -139,11 +141,9 @@
current_rendering = rendering
JEdit_Lib.buffer_edit(getBuffer) {
rich_text_area.active_reset()
- getBuffer.setReadOnly(false)
getBuffer.setFoldHandler(new Fold_Handling.Document_Fold_Handler(rendering))
setText(text)
setCaretPosition(0)
- getBuffer.setReadOnly(true)
}
}
})
@@ -224,6 +224,14 @@
/* key handling */
+ inputHandlerProvider =
+ new DefaultInputHandlerProvider(new TextAreaInputHandler(text_area) {
+ override def getAction(action: String): JEditBeanShellAction =
+ text_area.getActionContext.getAction(action)
+ override def processKeyEvent(evt: KeyEvent, from: Int, global: Boolean) {}
+ override def handleKey(key: KeyEventTranslator.Key, dry_run: Boolean): Boolean = false
+ })
+
addKeyListener(JEdit_Lib.key_listener(
key_pressed = (evt: KeyEvent) =>
{
@@ -260,9 +268,7 @@
getPainter.setLineHighlightEnabled(false)
getBuffer.setTokenMarker(Isabelle.mode_token_marker("isabelle-output").get)
- getBuffer.setReadOnly(true)
getBuffer.setStringProperty("noWordSep", "_'.?")
rich_text_area.activate()
}
-
--- a/src/Tools/jEdit/src/rendering.scala Thu Nov 05 08:27:14 2015 +0100
+++ b/src/Tools/jEdit/src/rendering.scala Thu Nov 05 08:27:22 2015 +0100
@@ -94,6 +94,7 @@
def token_markup(syntax: Outer_Syntax, token: Token): Byte =
if (token.is_command) command_style(syntax.keywords.command_kind(token.content).getOrElse(""))
+ else if (token.is_keyword && token.source == Symbol.comment_decoded) JEditToken.NULL
else if (token.is_delimiter) JEditToken.OPERATOR
else token_style(token.kind)