# HG changeset patch # User wenzelm # Date 1254421245 -7200 # Node ID 4c6e3e7ac2bf1a9a4109fe4cf757fe73a3fc6dd1 # Parent 00c14c4a6b4fdf4768b52e0a0cbbba57f95ff32f updated generated files; diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/IsarImplementation/Thy/document/Integration.tex --- a/doc-src/IsarImplementation/Thy/document/Integration.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Integration.tex Thu Oct 01 20:20:45 2009 +0200 @@ -86,9 +86,9 @@ \indexdef{}{ML}{Toplevel.is\_toplevel}\verb|Toplevel.is_toplevel: Toplevel.state -> bool| \\ \indexdef{}{ML}{Toplevel.theory\_of}\verb|Toplevel.theory_of: Toplevel.state -> theory| \\ \indexdef{}{ML}{Toplevel.proof\_of}\verb|Toplevel.proof_of: Toplevel.state -> Proof.state| \\ - \indexdef{}{ML}{Toplevel.debug}\verb|Toplevel.debug: bool ref| \\ - \indexdef{}{ML}{Toplevel.timing}\verb|Toplevel.timing: bool ref| \\ - \indexdef{}{ML}{Toplevel.profiling}\verb|Toplevel.profiling: int ref| \\ + \indexdef{}{ML}{Toplevel.debug}\verb|Toplevel.debug: bool Unsynchronized.ref| \\ + \indexdef{}{ML}{Toplevel.timing}\verb|Toplevel.timing: bool Unsynchronized.ref| \\ + \indexdef{}{ML}{Toplevel.profiling}\verb|Toplevel.profiling: int Unsynchronized.ref| \\ \end{mldecls} \begin{description} @@ -112,11 +112,11 @@ \item \verb|Toplevel.proof_of|~\isa{state} selects the Isar proof state if available, otherwise raises \verb|Toplevel.UNDEF|. - \item \verb|set Toplevel.debug| makes the toplevel print further + \item \verb|Toplevel.debug := true| makes the toplevel print further details about internal error conditions, exceptions being raised etc. - \item \verb|set Toplevel.timing| makes the toplevel print timing + \item \verb|Toplevel.timing := true| makes the toplevel print timing information for each Isar command being executed. \item \verb|Toplevel.profiling|~\verb|:=|~\isa{n} controls diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/IsarImplementation/Thy/document/Logic.tex --- a/doc-src/IsarImplementation/Thy/document/Logic.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Thu Oct 01 20:20:45 2009 +0200 @@ -546,7 +546,7 @@ \end{mldecls} \begin{mldecls} \indexdef{}{ML type}{thm}\verb|type thm| \\ - \indexdef{}{ML}{proofs}\verb|proofs: int ref| \\ + \indexdef{}{ML}{proofs}\verb|proofs: int Unsynchronized.ref| \\ \indexdef{}{ML}{Thm.assume}\verb|Thm.assume: cterm -> thm| \\ \indexdef{}{ML}{Thm.forall\_intr}\verb|Thm.forall_intr: cterm -> thm -> thm| \\ \indexdef{}{ML}{Thm.forall\_elim}\verb|Thm.forall_elim: cterm -> thm -> thm| \\ diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/IsarImplementation/Thy/document/ML.tex --- a/doc-src/IsarImplementation/Thy/document/ML.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/IsarImplementation/Thy/document/ML.tex Thu Oct 01 20:20:45 2009 +0200 @@ -277,7 +277,7 @@ \begin{mldecls} \indexdef{}{ML}{NAMED\_CRITICAL}\verb|NAMED_CRITICAL: string -> (unit -> 'a) -> 'a| \\ \indexdef{}{ML}{CRITICAL}\verb|CRITICAL: (unit -> 'a) -> 'a| \\ - \indexdef{}{ML}{setmp}\verb|setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\ + \indexdef{}{ML}{setmp}\verb|setmp: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\ \end{mldecls} \begin{description} diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/IsarOverview/Isar/document/Logic.tex --- a/doc-src/IsarOverview/Isar/document/Logic.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/IsarOverview/Isar/document/Logic.tex Thu Oct 01 20:20:45 2009 +0200 @@ -1228,19 +1228,6 @@ decompose goals and facts. This can lead to very tedious proofs:% \end{isamarkuptext}% \isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML \isacommand{lemma}\isamarkupfalse% \ {\isachardoublequoteopen}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline % diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/IsarRef/Thy/document/Inner_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Thu Oct 01 20:20:45 2009 +0200 @@ -118,19 +118,19 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML}{show\_types}\verb|show_types: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_consts}\verb|show_consts: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{long\_names}\verb|long_names: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{short\_names}\verb|short_names: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{unique\_names}\verb|unique_names: bool ref| & default \verb|true| \\ - \indexdef{}{ML}{show\_brackets}\verb|show_brackets: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{eta\_contract}\verb|eta_contract: bool ref| & default \verb|true| \\ - \indexdef{}{ML}{goals\_limit}\verb|goals_limit: int ref| & default \verb|10| \\ - \indexdef{}{ML}{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_hyps}\verb|show_hyps: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_tags}\verb|show_tags: bool ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_question\_marks}\verb|show_question_marks: bool ref| & default \verb|true| \\ + \indexdef{}{ML}{show\_types}\verb|show_types: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_consts}\verb|show_consts: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{long\_names}\verb|long_names: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{short\_names}\verb|short_names: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{unique\_names}\verb|unique_names: bool Unsynchronized.ref| & default \verb|true| \\ + \indexdef{}{ML}{show\_brackets}\verb|show_brackets: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{eta\_contract}\verb|eta_contract: bool Unsynchronized.ref| & default \verb|true| \\ + \indexdef{}{ML}{goals\_limit}\verb|goals_limit: int Unsynchronized.ref| & default \verb|10| \\ + \indexdef{}{ML}{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_hyps}\verb|show_hyps: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_tags}\verb|show_tags: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_question\_marks}\verb|show_question_marks: bool Unsynchronized.ref| & default \verb|true| \\ \end{mldecls} These global ML variables control the detail of information that is diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/Locales/Locales/document/Examples3.tex --- a/doc-src/Locales/Locales/document/Examples3.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/Locales/Locales/document/Examples3.tex Thu Oct 01 20:20:45 2009 +0200 @@ -42,7 +42,7 @@ \isatagvisible \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ nat{\isacharunderscore}less{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline \isacommand{proof}\isamarkupfalse% \ {\isacharminus}\isanewline \ \ \isacommand{show}\isamarkupfalse% @@ -101,13 +101,12 @@ \isatagvisible \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline +\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline \isacommand{proof}\isamarkupfalse% \ {\isacharminus}\isanewline \ \ \isacommand{show}\isamarkupfalse% -\ lattice{\isacharcolon}\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}% +\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}% \begin{isamarkuptxt}% We have already shown that this is a partial order,% \end{isamarkuptxt}% @@ -134,21 +133,12 @@ \ \ \ \ \isacommand{by}\isamarkupfalse% \ arith{\isacharplus}% \begin{isamarkuptxt}% -For the first of the equations, we refer to the theorem - shown in the previous interpretation.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq{\isacharparenright}% -\begin{isamarkuptxt}% -In order to show the remaining equations, we put ourselves in a +In order to show the equations, we put ourselves in a situation where the lattice theorems can be used in a convenient way.% \end{isamarkuptxt}% \isamarkuptrue% -\ \ \isacommand{from}\isamarkupfalse% -\ lattice\ \isacommand{interpret}\isamarkupfalse% +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{interpret}\isamarkupfalse% \ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% \isanewline \ \ \isacommand{show}\isamarkupfalse% @@ -180,45 +170,8 @@ \isatagvisible \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}total{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isacharunderscore}locales\ arith\isanewline -\isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -Since the locale hierarchy reflects that total - orders are distributive lattices, an explicit interpretation of - distributive lattices for the order relation on natural numbers is - only necessary for mapping the definitions to the right operators on - \isa{nat}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -% -\endisadelimvisible -% -\isatagvisible -\isacommand{interpretation}\isamarkupfalse% -\ nat{\isacharcolon}\ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline \ \ \isacommand{by}\isamarkupfalse% -\ unfold{\isacharunderscore}locales\ {\isacharbrackleft}{\isadigit{1}}{\isacharbrackright}\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}% +\ unfold{\isacharunderscore}locales\ arith% \endisatagvisible {\isafoldvisible}% % @@ -248,7 +201,30 @@ \hrule \caption{Interpreted theorems for \isa{{\isasymle}} on the natural numbers.} \label{tab:nat-lattice} -\end{table}% +\end{table} + + Note that since the locale hierarchy reflects that total orders are + distributive lattices, an explicit interpretation of distributive + lattices for the order relation on natural numbers is not neccessary. + + Why not push this idea further and just give the last interpretation + as a single interpretation instead of the sequence of three? The + reasons for this are twofold: +\begin{itemize} +\item + Often it is easier to work in an incremental fashion, because later + interpretations require theorems provided by earlier + interpretations. +\item + Assume that a definition is made in some locale $l_1$, and that $l_2$ + imports $l_1$. Let an equation for the definition be + proved in an interpretation of $l_2$. The equation will be unfolded + in interpretations of theorems added to $l_2$ or below in the import + hierarchy, but not for theorems added above $l_2$. + Hence, an equation interpreting a definition should always be given in + an interpretation of the locale where the definition is made, not in + an interpretation of a locale further down the hierarchy. +\end{itemize}% \end{isamarkuptext}% \isamarkuptrue% % @@ -264,8 +240,7 @@ \isamarkuptrue% \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline % \isadelimproof % @@ -306,8 +281,7 @@ \isamarkuptrue% \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline +\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline \ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline % \isadelimproof @@ -338,10 +312,6 @@ \ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% \isanewline \ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}rule\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq{\isacharparenright}\isanewline -\ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{apply}\isamarkupfalse% \ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline @@ -408,14 +378,7 @@ \isacommand{interpretation}\isamarkupfalse% \ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline \ \ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline -\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline -\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isachardoublequoteopen}distrib{\isacharunderscore}lattice\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ \ \isacommand{apply}\isamarkupfalse% \ unfold{\isacharunderscore}locales% \begin{isamarkuptxt}% \begin{isabelle}% @@ -426,7 +389,7 @@ \end{isabelle}% \end{isamarkuptxt}% \isamarkuptrue% -\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ \ \isacommand{apply}\isamarkupfalse% \ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}% \begin{isamarkuptxt}% \begin{isabelle}% @@ -434,12 +397,9 @@ \end{isabelle}% \end{isamarkuptxt}% \isamarkuptrue% -\ \ \ \ \isacommand{apply}\isamarkupfalse% -\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\isanewline -\ \ \ \ \isacommand{done}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}rule\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}% +\ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\ \isacommand{done}\isamarkupfalse% +% \endisatagvisible {\isafoldvisible}% % diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/Main/Docs/document/Main_Doc.tex --- a/doc-src/Main/Docs/document/Main_Doc.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/Main/Docs/document/Main_Doc.tex Thu Oct 01 20:20:45 2009 +0200 @@ -519,7 +519,7 @@ \isa{takeWhile} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\ \isa{tl} & \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\ \isa{upt} & \isa{nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat\ list}\\ -\isa{upto} & \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\ +\isa{upto} & \isa{int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ int\ list}\\ \isa{zip} & \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}b\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ list}\\ \end{supertabular} diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/TutorialI/Inductive/document/Advanced.tex --- a/doc-src/TutorialI/Inductive/document/Advanced.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/TutorialI/Inductive/document/Advanced.tex Thu Oct 01 20:20:45 2009 +0200 @@ -362,7 +362,7 @@ subgoal may look uninviting, but fortunately \isa{lists} distributes over intersection: \begin{isabelle}% -lists\ {\isacharparenleft}A\ {\isasyminter}\ B{\isacharparenright}\ {\isacharequal}\ lists\ A\ {\isasyminter}\ lists\ B\rulename{lists{\isacharunderscore}Int{\isacharunderscore}eq}% +listsp\ {\isacharparenleft}{\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ A{\isacharparenright}\ {\isasyminter}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ B{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ lists\ A{\isacharparenright}\ {\isasyminter}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ lists\ B{\isacharparenright}\rulename{lists{\isacharunderscore}Int{\isacharunderscore}eq}% \end{isabelle} Thanks to this default simplification rule, the induction hypothesis is quickly replaced by its two parts: diff -r 00c14c4a6b4f -r 4c6e3e7ac2bf doc-src/TutorialI/Types/document/Numbers.tex --- a/doc-src/TutorialI/Types/document/Numbers.tex Thu Oct 01 20:13:32 2009 +0200 +++ b/doc-src/TutorialI/Types/document/Numbers.tex Thu Oct 01 20:20:45 2009 +0200 @@ -550,44 +550,18 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isadelimML -% -\endisadelimML -% -\isatagML -\isacommand{ML}\isamarkupfalse% -{\isacharverbatimopen}set\ show{\isacharunderscore}sorts{\isacharverbatimclose}% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% \begin{isamarkuptext}% effect of show sorts on the above \begin{isabelle}% -{\isacharparenleft}c\ {\isacharasterisk}\ a\ {\isacharequal}\ c\ {\isacharasterisk}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}c\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isasymor}\ a\ {\isacharequal}\ b{\isacharparenright}% +{\isacharparenleft}{\isacharparenleft}c{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isacharequal}\isanewline +\isaindent{{\isacharparenleft}}c\ {\isacharasterisk}\ {\isacharparenleft}b{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +{\isacharparenleft}c\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isasymor}\ a\ {\isacharequal}\ b{\isacharparenright}% \end{isabelle} \rulename{mult_cancel_left}% \end{isamarkuptext}% \isamarkuptrue% % -\isadelimML -% -\endisadelimML -% -\isatagML -\isacommand{ML}\isamarkupfalse% -{\isacharverbatimopen}reset\ show{\isacharunderscore}sorts{\isacharverbatimclose}% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% \begin{isamarkuptext}% absolute value