--- a/doc-src/TutorialI/CTL/CTL.thy Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/CTL/CTL.thy Wed Jan 10 10:40:34 2001 +0100
@@ -62,7 +62,7 @@
"mc(Neg f) = -mc f"
"mc(And f g) = mc f \<inter> mc g"
"mc(AX f) = {s. \<forall>t. (s,t) \<in> M \<longrightarrow> t \<in> mc f}"
-"mc(EF f) = lfp(\<lambda>T. mc f \<union> M\<inverse> ``` T)"(*>*)
+"mc(EF f) = lfp(\<lambda>T. mc f \<union> M\<inverse> `` T)"(*>*)
"mc(AF f) = lfp(af(mc f))";
text{*\noindent
@@ -75,12 +75,12 @@
apply blast;
done
(*<*)
-lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> ``` T)";
+lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> `` T)";
apply(rule monoI);
by(blast);
lemma EF_lemma:
- "lfp(\<lambda>T. A \<union> M\<inverse> ``` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}";
+ "lfp(\<lambda>T. A \<union> M\<inverse> `` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}";
apply(rule equalityI);
apply(rule subsetI);
apply(simp);
@@ -366,7 +366,7 @@
Note that @{term EU} is not definable in terms of the other operators!
Model checking @{term EU} is again a least fixed point construction:
-@{text[display]"mc(EU f g) = lfp(\<lambda>T. mc g \<union> mc f \<inter> (M\<inverse> ``` T))"}
+@{text[display]"mc(EU f g) = lfp(\<lambda>T. mc g \<union> mc f \<inter> (M\<inverse> `` T))"}
\begin{exercise}
Extend the datatype of formulae by the above until operator
@@ -382,7 +382,7 @@
(*<*)
constdefs
eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set"
-"eufix A B T \<equiv> B \<union> A \<inter> (M\<inverse> ``` T)"
+"eufix A B T \<equiv> B \<union> A \<inter> (M\<inverse> `` T)"
lemma "lfp(eufix A B) \<subseteq> eusem A B"
apply(rule lfp_lowerbound)
--- a/doc-src/TutorialI/CTL/PDL.thy Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/CTL/PDL.thy Wed Jan 10 10:40:34 2001 +0100
@@ -58,14 +58,14 @@
"mc(Neg f) = -mc f"
"mc(And f g) = mc f \<inter> mc g"
"mc(AX f) = {s. \<forall>t. (s,t) \<in> M \<longrightarrow> t \<in> mc f}"
-"mc(EF f) = lfp(\<lambda>T. mc f \<union> M\<inverse> ``` T)"
+"mc(EF f) = lfp(\<lambda>T. mc f \<union> M\<inverse> `` T)"
text{*\noindent
Only the equation for @{term EF} deserves some comments. Remember that the
-postfix @{text"\<inverse>"} and the infix @{text"```"} are predefined and denote the
+postfix @{text"\<inverse>"} and the infix @{text"``"} are predefined and denote the
converse of a relation and the application of a relation to a set. Thus
-@{term "M\<inverse> ``` T"} is the set of all predecessors of @{term T} and the least
-fixed point (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M\<inverse> ``` T"} is the least set
+@{term "M\<inverse> `` T"} is the set of all predecessors of @{term T} and the least
+fixed point (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M\<inverse> `` T"} is the least set
@{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you
find it hard to see that @{term"mc(EF f)"} contains exactly those states from
which there is a path to a state where @{term f} is true, do not worry---that
@@ -74,7 +74,7 @@
First we prove monotonicity of the function inside @{term lfp}
*}
-lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> ``` T)"
+lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> `` T)"
apply(rule monoI)
apply blast
done
@@ -87,7 +87,7 @@
*}
lemma EF_lemma:
- "lfp(\<lambda>T. A \<union> M\<inverse> ``` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
+ "lfp(\<lambda>T. A \<union> M\<inverse> `` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
txt{*\noindent
The equality is proved in the canonical fashion by proving that each set
--- a/doc-src/TutorialI/CTL/document/CTL.tex Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/CTL/document/CTL.tex Wed Jan 10 10:40:34 2001 +0100
@@ -300,7 +300,7 @@
Model checking \isa{EU} is again a least fixed point construction:
\begin{isabelle}%
-\ \ \ \ \ mc{\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ g\ {\isasymunion}\ mc\ f\ {\isasyminter}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}%
+\ \ \ \ \ mc{\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ g\ {\isasymunion}\ mc\ f\ {\isasyminter}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}%
\end{isabelle}
\begin{exercise}
--- a/doc-src/TutorialI/CTL/document/PDL.tex Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/CTL/document/PDL.tex Wed Jan 10 10:40:34 2001 +0100
@@ -58,14 +58,14 @@
{\isachardoublequote}mc{\isacharparenleft}Neg\ f{\isacharparenright}\ \ \ {\isacharequal}\ {\isacharminus}mc\ f{\isachardoublequote}\isanewline
{\isachardoublequote}mc{\isacharparenleft}And\ f\ g{\isacharparenright}\ {\isacharequal}\ mc\ f\ {\isasyminter}\ mc\ g{\isachardoublequote}\isanewline
{\isachardoublequote}mc{\isacharparenleft}AX\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ \ {\isasymlongrightarrow}\ t\ {\isasymin}\ mc\ f{\isacharbraceright}{\isachardoublequote}\isanewline
-{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}%
+{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}%
\begin{isamarkuptext}%
\noindent
Only the equation for \isa{EF} deserves some comments. Remember that the
-postfix \isa{{\isasyminverse}} and the infix \isa{{\isacharbackquote}{\isacharbackquote}{\isacharbackquote}} are predefined and denote the
+postfix \isa{{\isasyminverse}} and the infix \isa{{\isacharbackquote}{\isacharbackquote}} are predefined and denote the
converse of a relation and the application of a relation to a set. Thus
-\isa{M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T} is the set of all predecessors of \isa{T} and the least
-fixed point (\isa{lfp}) of \isa{{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T} is the least set
+\isa{M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T} is the set of all predecessors of \isa{T} and the least
+fixed point (\isa{lfp}) of \isa{{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T} is the least set
\isa{T} containing \isa{mc\ f} and all predecessors of \isa{T}. If you
find it hard to see that \isa{mc\ {\isacharparenleft}EF\ f{\isacharparenright}} contains exactly those states from
which there is a path to a state where \isa{f} is true, do not worry---that
@@ -73,7 +73,7 @@
First we prove monotonicity of the function inside \isa{lfp}%
\end{isamarkuptext}%
-\isacommand{lemma}\ mono{\isacharunderscore}ef{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}\isanewline
+\isacommand{lemma}\ mono{\isacharunderscore}ef{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}\isanewline
\isacommand{apply}{\isacharparenleft}rule\ monoI{\isacharparenright}\isanewline
\isacommand{apply}\ blast\isanewline
\isacommand{done}%
@@ -85,7 +85,7 @@
a separate lemma:%
\end{isamarkuptext}%
\isacommand{lemma}\ EF{\isacharunderscore}lemma{\isacharcolon}\isanewline
-\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}%
+\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}%
\begin{isamarkuptxt}%
\noindent
The equality is proved in the canonical fashion by proving that each set
@@ -98,7 +98,7 @@
\noindent
Simplification leaves us with the following first subgoal
\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A%
\end{isabelle}
which is proved by \isa{lfp}-induction:%
\end{isamarkuptxt}%
@@ -127,7 +127,7 @@
\noindent
After simplification and clarification we are left with
\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}{\isacharsemicolon}\ t\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}{\isacharsemicolon}\ t\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
\end{isabelle}
This goal is proved by induction on \isa{{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}}. But since the model
checker works backwards (from \isa{t} to \isa{s}), we cannot use the
@@ -148,14 +148,14 @@
\noindent
The base case
\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
\end{isabelle}
is solved by unrolling \isa{lfp} once%
\end{isamarkuptxt}%
\ \isacommand{apply}{\isacharparenleft}rule\ ssubst{\isacharbrackleft}OF\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharbrackright}{\isacharparenright}%
\begin{isamarkuptxt}%
\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}%
\end{isabelle}
and disposing of the resulting trivial subgoal automatically:%
\end{isamarkuptxt}%
--- a/doc-src/TutorialI/Misc/document/pairs.tex Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/Misc/document/pairs.tex Wed Jan 10 10:40:34 2001 +0100
@@ -24,7 +24,7 @@
Products, like type \isa{nat}, are datatypes, which means
in particular that \isa{induct{\isacharunderscore}tac} and \isa{case{\isacharunderscore}tac} are applicable to
terms of product type.
-Both replace the term by tuple of variables.
+Both replace the term by a pair of variables.
\item
Tuples with more than two or three components become unwieldy;
records are preferable.
--- a/doc-src/TutorialI/Misc/pairs.thy Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/Misc/pairs.thy Wed Jan 10 10:40:34 2001 +0100
@@ -22,7 +22,7 @@
Products, like type @{typ nat}, are datatypes, which means
in particular that @{text induct_tac} and @{text case_tac} are applicable to
terms of product type.
-Both replace the term by tuple of variables.
+Both replace the term by a pair of variables.
\item
Tuples with more than two or three components become unwieldy;
records are preferable.
--- a/doc-src/TutorialI/Sets/Functions.thy Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/Sets/Functions.thy Wed Jan 10 10:40:34 2001 +0100
@@ -119,17 +119,17 @@
text{*
illustrates Union as well as image
*}
-lemma "f``A \<union> g``A = (\<Union>x\<in>A. {f x, g x})"
+lemma "f`A \<union> g`A = (\<Union>x\<in>A. {f x, g x})"
apply (blast)
done
-lemma "f `` {(x,y). P x y} = {f(x,y) | x y. P x y}"
+lemma "f ` {(x,y). P x y} = {f(x,y) | x y. P x y}"
apply (blast)
done
text{*actually a macro!*}
-lemma "range f = f``UNIV"
+lemma "range f = f`UNIV"
apply (blast)
done
--- a/doc-src/TutorialI/isabelle.sty Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/isabelle.sty Wed Jan 10 10:40:34 2001 +0100
@@ -1,5 +1,4 @@
%%
-%% $Id$
%% Author: Markus Wenzel, TU Muenchen
%% License: GPL (GNU GENERAL PUBLIC LICENSE)
%%
--- a/doc-src/TutorialI/isabellesym.sty Wed Jan 10 00:15:33 2001 +0100
+++ b/doc-src/TutorialI/isabellesym.sty Wed Jan 10 10:40:34 2001 +0100
@@ -1,5 +1,4 @@
%%
-%% $Id$
%% Author: Markus Wenzel, TU Muenchen
%% License: GPL (GNU GENERAL PUBLIC LICENSE)
%%