--- a/NEWS Tue Mar 09 15:42:23 2010 +0100
+++ b/NEWS Fri Mar 12 16:02:42 2010 +0100
@@ -76,9 +76,16 @@
within a local theory context, with explicit checking of the
constructors involved (in contrast to the raw 'syntax' versions).
+* Command 'typedecl' now works within a local theory context --
+without introducing dependencies on parameters or assumptions, which
+is not possible in Isabelle/Pure.
+
*** HOL ***
+* Split off theory Big_Operators containing setsum, setprod, Inf_fin, Sup_fin,
+Min, Max from theory Finite_Set. INCOMPATIBILITY.
+
* Theory "Rational" renamed to "Rat", for consistency with "Nat", "Int" etc.
INCOMPATIBILITY.
@@ -98,7 +105,7 @@
INCOMPATIBILITY.
-* Class division ring also requires proof of fact divide_inverse. However instantiation
+* Class division_ring also requires proof of fact divide_inverse. However instantiation
of parameter divide has also been required previously. INCOMPATIBILITY.
* More consistent naming of type classes involving orderings (and lattices):
--- a/doc-src/IsarRef/Thy/Spec.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/doc-src/IsarRef/Thy/Spec.thy Fri Mar 12 16:02:42 2010 +0100
@@ -954,7 +954,7 @@
text {*
\begin{matharray}{rcll}
@{command_def "types"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def "typedecl"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def "typedecl"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
@{command_def "arities"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
\end{matharray}
--- a/doc-src/IsarRef/Thy/document/Spec.tex Tue Mar 09 15:42:23 2010 +0100
+++ b/doc-src/IsarRef/Thy/document/Spec.tex Fri Mar 12 16:02:42 2010 +0100
@@ -991,7 +991,7 @@
\begin{isamarkuptext}%
\begin{matharray}{rcll}
\indexdef{}{command}{types}\hypertarget{command.types}{\hyperlink{command.types}{\mbox{\isa{\isacommand{types}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{}{command}{typedecl}\hypertarget{command.typedecl}{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{}{command}{typedecl}\hypertarget{command.typedecl}{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
\indexdef{}{command}{arities}\hypertarget{command.arities}{\hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} & (axiomatic!) \\
\end{matharray}
--- a/doc-src/Nitpick/nitpick.tex Tue Mar 09 15:42:23 2010 +0100
+++ b/doc-src/Nitpick/nitpick.tex Fri Mar 12 16:02:42 2010 +0100
@@ -4,6 +4,7 @@
\usepackage{amssymb}
\usepackage[english,french]{babel}
\usepackage{color}
+\usepackage{footmisc}
\usepackage{graphicx}
%\usepackage{mathpazo}
\usepackage{multicol}
@@ -136,8 +137,8 @@
suggesting several textual improvements.
% and Perry James for reporting a typo.
-\section{Overview}
-\label{overview}
+\section{First Steps}
+\label{first-steps}
This section introduces Nitpick by presenting small examples. If possible, you
should try out the examples on your workstation. Your theory file should start
@@ -149,12 +150,12 @@
\textbf{begin}
\postw
-The results presented here were obtained using the JNI version of MiniSat and
-with multithreading disabled to reduce nondeterminism and a time limit of
-15~seconds (instead of 30~seconds). This was done by adding the line
+The results presented here were obtained using the JNI (Java Native Interface)
+version of MiniSat and with multithreading disabled to reduce nondeterminism.
+This was done by adding the line
\prew
-\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1, \,\textit{timeout} = 15$\,s$]
+\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1]
\postw
after the \textbf{begin} keyword. The JNI version of MiniSat is bundled with
@@ -471,7 +472,7 @@
\prew
\textbf{lemma} ``$\forall n.\; \textit{Suc}~n \mathbin{\not=} n \,\Longrightarrow\, P$'' \\
-\textbf{nitpick} [\textit{card~nat}~= 100, \textit{check\_potential}] \\[2\smallskipamount]
+\textbf{nitpick} [\textit{card~nat}~= 50, \textit{check\_potential}] \\[2\smallskipamount]
\slshape Warning: The conjecture either trivially holds for the given scopes or lies outside Nitpick's supported
fragment. Only potential counterexamples may be found. \\[2\smallskipamount]
Nitpick found a potential counterexample: \\[2\smallskipamount]
@@ -628,7 +629,7 @@
unlikely that one could be found for smaller cardinalities.
\subsection{Typedefs, Quotient Types, Records, Rationals, and Reals}
-\label{typedefs-records-rationals-and-reals}
+\label{typedefs-quotient-types-records-rationals-and-reals}
Nitpick generally treats types declared using \textbf{typedef} as datatypes
whose single constructor is the corresponding \textit{Abs\_\kern.1ex} function.
@@ -683,7 +684,26 @@
In the counterexample, $\Abs{(0,\, 0)}$ and $\Abs{(1,\, 0)}$ represent the
integers $0$ and $1$, respectively. Other representants would have been
-possible---e.g., $\Abs{(5,\, 5)}$ and $\Abs{(12,\, 11)}$.
+possible---e.g., $\Abs{(5,\, 5)}$ and $\Abs{(12,\, 11)}$. If we are going to
+use \textit{my\_int} extensively, it pays off to install a term postprocessor
+that converts the pair notation to the standard mathematical notation:
+
+\prew
+$\textbf{ML}~\,\{{*} \\
+\!\begin{aligned}[t]
+%& ({*}~\,\textit{Proof.context} \rightarrow \textit{string} \rightarrow (\textit{typ} \rightarrow \textit{term~list\/}) \rightarrow \textit{typ} \rightarrow \textit{term} \\[-2pt]
+%& \phantom{(*}~\,{\rightarrow}\;\textit{term}~\,{*}) \\[-2pt]
+& \textbf{fun}\,~\textit{my\_int\_postproc}~\_~\_~\_~T~(\textit{Const}~\_~\$~(\textit{Const}~\_~\$~\textit{t1}~\$~\textit{t2\/})) = {} \\[-2pt]
+& \phantom{fun}\,~\textit{HOLogic.mk\_number}~T~(\textit{snd}~(\textit{HOLogic.dest\_number~t1}) \\[-2pt]
+& \phantom{fun\,~\textit{HOLogic.mk\_number}~T~(}{-}~\textit{snd}~(\textit{HOLogic.dest\_number~t2\/})) \\[-2pt]
+& \phantom{fun}\!{\mid}\,~\textit{my\_int\_postproc}~\_~\_~\_~\_~t = t \\[-2pt]
+{*}\}\end{aligned}$ \\[2\smallskipamount]
+$\textbf{setup}~\,\{{*} \\
+\!\begin{aligned}[t]
+& \textit{Nitpick.register\_term\_postprocessor}~\!\begin{aligned}[t]
+ & @\{\textrm{typ}~\textit{my\_int}\}~\textit{my\_int\_postproc}\end{aligned} \\[-2pt]
+{*}\}\end{aligned}$
+\postw
Records are also handled as datatypes with a single constructor:
@@ -770,25 +790,25 @@
\prew
\textbf{lemma} ``$\exists n.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
-\textbf{nitpick}~[\textit{card nat}~= 100, \textit{unary\_ints}, \textit{verbose}] \\[2\smallskipamount]
+\textbf{nitpick}~[\textit{card nat}~= 50, \textit{unary\_ints}, \textit{verbose}] \\[2\smallskipamount]
\slshape The inductive predicate ``\textit{even}'' was proved well-founded.
Nitpick can compute it efficiently. \\[2\smallskipamount]
Trying 1 scope: \\
-\hbox{}\qquad \textit{card nat}~= 100. \\[2\smallskipamount]
-Nitpick found a potential counterexample for \textit{card nat}~= 100: \\[2\smallskipamount]
+\hbox{}\qquad \textit{card nat}~= 50. \\[2\smallskipamount]
+Nitpick found a potential counterexample for \textit{card nat}~= 50: \\[2\smallskipamount]
\hbox{}\qquad Empty assignment \\[2\smallskipamount]
Nitpick could not find a better counterexample. \\[2\smallskipamount]
Total time: 2274 ms.
\postw
No genuine counterexample is possible because Nitpick cannot rule out the
-existence of a natural number $n \ge 100$ such that both $\textit{even}~n$ and
+existence of a natural number $n \ge 50$ such that both $\textit{even}~n$ and
$\textit{even}~(\textit{Suc}~n)$ are true. To help Nitpick, we can bound the
existential quantifier:
\prew
-\textbf{lemma} ``$\exists n \mathbin{\le} 99.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
-\textbf{nitpick}~[\textit{card nat}~= 100, \textit{unary\_ints}] \\[2\smallskipamount]
+\textbf{lemma} ``$\exists n \mathbin{\le} 49.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
+\textbf{nitpick}~[\textit{card nat}~= 50, \textit{unary\_ints}] \\[2\smallskipamount]
\slshape Nitpick found a counterexample: \\[2\smallskipamount]
\hbox{}\qquad Empty assignment
\postw
@@ -1019,7 +1039,7 @@
can be disabled by setting the \textit{bisim\_depth} option to $-1$. The
bisimilarity check is then performed \textsl{after} the counterexample has been
found to ensure correctness. If this after-the-fact check fails, the
-counterexample is tagged as ``likely genuine'' and Nitpick recommends to try
+counterexample is tagged as ``quasi genuine'' and Nitpick recommends to try
again with \textit{bisim\_depth} set to a nonnegative integer. Disabling the
check for the previous example saves approximately 150~milli\-seconds; the speed
gains can be more significant for larger scopes.
@@ -1031,7 +1051,7 @@
\textbf{lemma} ``$\lbrakk xs = \textit{LCons}~a~\textit{xs};\>\, \textit{ys} = \textit{LCons}~a~\textit{ys}\rbrakk
\,\Longrightarrow\, \textit{xs} = \textit{ys}$'' \\
\textbf{nitpick} [\textit{bisim\_depth} = $-1$, \textit{show\_datatypes}] \\[2\smallskipamount]
-\slshape Nitpick found a likely genuine counterexample for $\textit{card}~'a$ = 2: \\[2\smallskipamount]
+\slshape Nitpick found a quasi genuine counterexample for $\textit{card}~'a$ = 2: \\[2\smallskipamount]
\hbox{}\qquad Free variables: \nopagebreak \\
\hbox{}\qquad\qquad $a = a_1$ \\
\hbox{}\qquad\qquad $\textit{xs} = \textsl{THE}~\omega.\; \omega =
@@ -1210,26 +1230,26 @@
Trying 8 scopes: \\
\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} $'b$~= 1,
\textit{card} \textit{nat}~= 1, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 1, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 1, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 1. \\
+\textit{list\/}''~= 1, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 1, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 1. \\
\hbox{}\qquad \textit{card} $'a$~= 2, \textit{card} $'b$~= 2,
\textit{card} \textit{nat}~= 2, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 2, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 2, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 2. \\
+\textit{list\/}''~= 2, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 2, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 2. \\
\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount]
\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} $'b$~= 8,
\textit{card} \textit{nat}~= 8, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 8, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 8, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 8.
+\textit{list\/}''~= 8, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 8, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 8.
\\[2\smallskipamount]
Nitpick found a counterexample for
\textit{card} $'a$~= 5, \textit{card} $'b$~= 5,
\textit{card} \textit{nat}~= 5, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 5, \textit{card} ``\kern1pt$'a$ \textit{list}''~= 5, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 5:
+\textit{list\/}''~= 5, \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 5, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 5:
\\[2\smallskipamount]
\hbox{}\qquad Free variables: \nopagebreak \\
\hbox{}\qquad\qquad $\textit{xs} = [a_1, a_2]$ \\
@@ -1850,7 +1870,7 @@
The number of options can be overwhelming at first glance. Do not let that worry
you: Nitpick's defaults have been chosen so that it almost always does the right
thing, and the most important options have been covered in context in
-\S\ref{overview}.
+\S\ref{first-steps}.
The descriptions below refer to the following syntactic quantities:
@@ -1901,7 +1921,7 @@
\textbf{Warning:} If the option is set to \textit{true}, Nitpick might
nonetheless ignore some polymorphic axioms. Counterexamples generated under
-these conditions are tagged as ``likely genuine.'' The \textit{debug}
+these conditions are tagged as ``quasi genuine.'' The \textit{debug}
(\S\ref{output-format}) option can be used to find out which axioms were
considered.
@@ -2002,7 +2022,7 @@
\begin{enum}
\item[$\bullet$] \textbf{\textit{true}}: Tentatively treat the (co)in\-duc\-tive
predicate as if it were well-founded. Since this is generally not sound when the
-predicate is not well-founded, the counterexamples are tagged as ``likely
+predicate is not well-founded, the counterexamples are tagged as ``quasi
genuine.''
\item[$\bullet$] \textbf{\textit{false}}: Treat the (co)in\-duc\-tive predicate
@@ -2050,7 +2070,7 @@
of $-1$ means that no predicate is generated, in which case Nitpick performs an
after-the-fact check to see if the known coinductive datatype values are
bidissimilar. If two values are found to be bisimilar, the counterexample is
-tagged as ``likely genuine.'' The iteration counts are automatically bounded by
+tagged as ``quasi genuine.'' The iteration counts are automatically bounded by
the sum of the cardinalities of the coinductive datatypes occurring in the
formula to falsify.
@@ -2100,7 +2120,7 @@
\begin{enum}
\item[$\bullet$] \textbf{\textit{true}}: Finitize the datatype. Since this is
-unsound, counterexamples generated under these conditions are tagged as ``likely
+unsound, counterexamples generated under these conditions are tagged as ``quasi
genuine.''
\item[$\bullet$] \textbf{\textit{false}}: Don't attempt to finitize the datatype.
\item[$\bullet$] \textbf{\textit{smart}}: Perform a monotonicity analysis to
@@ -2202,11 +2222,9 @@
Specifies the maximum number of potential counterexamples to display. Setting
this option to 0 speeds up the search for a genuine counterexample. This option
is implicitly set to 0 for automatic runs. If you set this option to a value
-greater than 1, you will need an incremental SAT solver: For efficiency, it is
-recommended to install the JNI version of MiniSat and set \textit{sat\_solver} =
-\textit{MiniSat\_JNI}. Also be aware that many of the counterexamples may look
-identical, unless the \textit{show\_all} (\S\ref{output-format}) option is
-enabled.
+greater than 1, you will need an incremental SAT solver, such as
+\textit{MiniSat\_JNI} (recommended) and \textit{SAT4J}. Be aware that many of
+the counterexamples may be identical.
\nopagebreak
{\small See also \textit{check\_potential} (\S\ref{authentication}) and
@@ -2214,11 +2232,9 @@
\opdefault{max\_genuine}{int}{$\mathbf{1}$}
Specifies the maximum number of genuine counterexamples to display. If you set
-this option to a value greater than 1, you will need an incremental SAT solver:
-For efficiency, it is recommended to install the JNI version of MiniSat and set
-\textit{sat\_solver} = \textit{MiniSat\_JNI}. Also be aware that many of the
-counterexamples may look identical, unless the \textit{show\_all}
-(\S\ref{output-format}) option is enabled.
+this option to a value greater than 1, you will need an incremental SAT solver,
+such as \textit{MiniSat\_JNI} (recommended) and \textit{SAT4J}. Be aware that
+many of the counterexamples may be identical.
\nopagebreak
{\small See also \textit{check\_genuine} (\S\ref{authentication}) and
@@ -2268,7 +2284,7 @@
{\small See also \textit{max\_potential} (\S\ref{output-format}).}
\opfalse{check\_genuine}{trust\_genuine}
-Specifies whether genuine and likely genuine counterexamples should be given to
+Specifies whether genuine and quasi genuine counterexamples should be given to
Isabelle's \textit{auto} tactic to assess their validity. If a ``genuine''
counterexample is shown to be spurious, the user is kindly asked to send a bug
report to the author at
@@ -2282,7 +2298,7 @@
\begin{enum}
\item[$\bullet$] \textbf{\textit{genuine}}: Nitpick found a genuine counterexample.
-\item[$\bullet$] \textbf{\textit{likely\_genuine}}: Nitpick found a ``likely
+\item[$\bullet$] \textbf{\textit{quasi\_genuine}}: Nitpick found a ``quasi
genuine'' counterexample (i.e., a counterexample that is genuine unless
it contradicts a missing axiom or a dangerous option was used inappropriately).
\item[$\bullet$] \textbf{\textit{potential}}: Nitpick found a potential counterexample.
@@ -2313,63 +2329,77 @@
The supported solvers are listed below:
+\let\foo
+
\begin{enum}
\item[$\bullet$] \textbf{\textit{MiniSat}}: MiniSat is an efficient solver
written in \cpp{}. To use MiniSat, set the environment variable
\texttt{MINISAT\_HOME} to the directory that contains the \texttt{minisat}
-executable. The \cpp{} sources and executables for MiniSat are available at
+executable.%
+\footnote{Important note for Cygwin users: The path must be specified using
+native Windows syntax. Make sure to escape backslashes properly.%
+\label{cygwin-paths}}
+The \cpp{} sources and executables for MiniSat are available at
\url{http://minisat.se/MiniSat.html}. Nitpick has been tested with versions 1.14
and 2.0 beta (2007-07-21).
\item[$\bullet$] \textbf{\textit{MiniSat\_JNI}}: The JNI (Java Native Interface)
-version of MiniSat is bundled in \texttt{nativesolver.\allowbreak tgz}, which
-you will find on Kodkod's web site \cite{kodkod-2009}. Unlike the standard
+version of MiniSat is bundled with Kodkodi and is precompiled for the major
+platforms. It is also available from \texttt{native\-solver.\allowbreak tgz},
+which you will find on Kodkod's web site \cite{kodkod-2009}. Unlike the standard
version of MiniSat, the JNI version can be used incrementally.
-%%% No longer true:
-%%% "It is bundled with Kodkodi and requires no further installation or
-%%% configuration steps. Alternatively,"
\item[$\bullet$] \textbf{\textit{PicoSAT}}: PicoSAT is an efficient solver
written in C. You can install a standard version of
PicoSAT and set the environment variable \texttt{PICOSAT\_HOME} to the directory
-that contains the \texttt{picosat} executable. The C sources for PicoSAT are
+that contains the \texttt{picosat} executable.%
+\footref{cygwin-paths}
+The C sources for PicoSAT are
available at \url{http://fmv.jku.at/picosat/} and are also bundled with Kodkodi.
Nitpick has been tested with version 913.
\item[$\bullet$] \textbf{\textit{zChaff}}: zChaff is an efficient solver written
in \cpp{}. To use zChaff, set the environment variable \texttt{ZCHAFF\_HOME} to
-the directory that contains the \texttt{zchaff} executable. The \cpp{} sources
-and executables for zChaff are available at
+the directory that contains the \texttt{zchaff} executable.%
+\footref{cygwin-paths}
+The \cpp{} sources and executables for zChaff are available at
\url{http://www.princeton.edu/~chaff/zchaff.html}. Nitpick has been tested with
versions 2004-05-13, 2004-11-15, and 2007-03-12.
\item[$\bullet$] \textbf{\textit{zChaff\_JNI}}: The JNI version of zChaff is
-bundled in \texttt{native\-solver.\allowbreak tgz}, which you will find on
-Kodkod's web site \cite{kodkod-2009}.
+bundled with Kodkodi and is precompiled for the major
+platforms. It is also available from \texttt{native\-solver.\allowbreak tgz},
+which you will find on Kodkod's web site \cite{kodkod-2009}.
\item[$\bullet$] \textbf{\textit{RSat}}: RSat is an efficient solver written in
\cpp{}. To use RSat, set the environment variable \texttt{RSAT\_HOME} to the
-directory that contains the \texttt{rsat} executable. The \cpp{} sources for
-RSat are available at \url{http://reasoning.cs.ucla.edu/rsat/}. Nitpick has been
-tested with version 2.01.
+directory that contains the \texttt{rsat} executable.%
+\footref{cygwin-paths}
+The \cpp{} sources for RSat are available at
+\url{http://reasoning.cs.ucla.edu/rsat/}. Nitpick has been tested with version
+2.01.
\item[$\bullet$] \textbf{\textit{BerkMin}}: BerkMin561 is an efficient solver
written in C. To use BerkMin, set the environment variable
\texttt{BERKMIN\_HOME} to the directory that contains the \texttt{BerkMin561}
-executable. The BerkMin executables are available at
+executable.\footref{cygwin-paths}
+The BerkMin executables are available at
\url{http://eigold.tripod.com/BerkMin.html}.
\item[$\bullet$] \textbf{\textit{BerkMin\_Alloy}}: Variant of BerkMin that is
included with Alloy 4 and calls itself ``sat56'' in its banner text. To use this
version of BerkMin, set the environment variable
\texttt{BERKMINALLOY\_HOME} to the directory that contains the \texttt{berkmin}
-executable.
+executable.%
+\footref{cygwin-paths}
\item[$\bullet$] \textbf{\textit{Jerusat}}: Jerusat 1.3 is an efficient solver
written in C. To use Jerusat, set the environment variable
\texttt{JERUSAT\_HOME} to the directory that contains the \texttt{Jerusat1.3}
-executable. The C sources for Jerusat are available at
+executable.%
+\footref{cygwin-paths}
+The C sources for Jerusat are available at
\url{http://www.cs.tau.ac.il/~ale1/Jerusat1.3.tgz}.
\item[$\bullet$] \textbf{\textit{SAT4J}}: SAT4J is a reasonably efficient solver
@@ -2384,7 +2414,9 @@
\item[$\bullet$] \textbf{\textit{HaifaSat}}: HaifaSat 1.0 beta is an
experimental solver written in \cpp. To use HaifaSat, set the environment
variable \texttt{HAIFASAT\_\allowbreak HOME} to the directory that contains the
-\texttt{HaifaSat} executable. The \cpp{} sources for HaifaSat are available at
+\texttt{HaifaSat} executable.%
+\footref{cygwin-paths}
+The \cpp{} sources for HaifaSat are available at
\url{http://cs.technion.ac.il/~gershman/HaifaSat.htm}.
\item[$\bullet$] \textbf{\textit{smart}}: If \textit{sat\_solver} is set to
@@ -2609,7 +2641,7 @@
then Nitpick assumes that the definition was made using an inductive package and
based on the introduction rules marked with \textit{nitpick\_\allowbreak
-ind\_\allowbreak intros} tries to determine whether the definition is
+\allowbreak intros} tries to determine whether the definition is
well-founded.
\end{enum}
\end{enum}
@@ -2651,7 +2683,8 @@
Nitpick provides a rich Standard ML interface used mainly for internal purposes
and debugging. Among the most interesting functions exported by Nitpick are
those that let you invoke the tool programmatically and those that let you
-register and unregister custom coinductive datatypes.
+register and unregister custom coinductive datatypes as well as term
+postprocessors.
\subsection{Invocation of Nitpick}
\label{invocation-of-nitpick}
@@ -2668,7 +2701,7 @@
\postw
The return value is a new proof state paired with an outcome string
-(``genuine'', ``likely\_genuine'', ``potential'', ``none'', or ``unknown''). The
+(``genuine'', ``quasi\_genuine'', ``potential'', ``none'', or ``unknown''). The
\textit{params} type is a large record that lets you set Nitpick's options. The
current default options can be retrieved by calling the following function
defined in the \textit{Nitpick\_Isar} structure:
@@ -2682,7 +2715,7 @@
put into a \textit{params} record. Here is an example:
\prew
-$\textbf{val}\,~\textit{params} = \textit{Nitpick\_Isar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
+$\textbf{val}\,~\textit{params} = \textit{Nitpick\_Isar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout\/}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
$\textbf{val}\,~(\textit{outcome},\, \textit{state}') = \textit{Nitpick.pick\_nits\_in\_subgoal}~\begin{aligned}[t]
& \textit{state}~\textit{params}~\textit{false} \\[-2pt]
& \textit{subgoal}\end{aligned}$
@@ -2713,7 +2746,7 @@
\prew
$\textbf{setup}~\,\{{*}\,~\!\begin{aligned}[t]
& \textit{Nitpick.register\_codatatype} \\[-2pt]
-& \qquad @\{\antiq{typ}~``\kern1pt'a~\textit{llist}\textrm{''}\}~@\{\antiq{const\_name}~ \textit{llist\_case}\} \\[-2pt] %% TYPESETTING
+& \qquad @\{\antiq{typ}~``\kern1pt'a~\textit{llist\/}\textrm{''}\}~@\{\antiq{const\_name}~ \textit{llist\_case}\} \\[-2pt] %% TYPESETTING
& \qquad (\textit{map}~\textit{dest\_Const}~[@\{\antiq{term}~\textit{LNil}\},\, @\{\antiq{term}~\textit{LCons}\}])\,\ {*}\}\end{aligned}$
\postw
@@ -2727,7 +2760,7 @@
\prew
$\textbf{setup}~\,\{{*}\,~\textit{Nitpick.unregister\_codatatype}~@\{\antiq{typ}~``
-\kern1pt'a~\textit{list}\textrm{''}\}\ \,{*}\}$
+\kern1pt'a~\textit{list\/}\textrm{''}\}\ \,{*}\}$
\postw
Inductive datatypes can be registered as coinductive datatypes, given
@@ -2735,6 +2768,26 @@
the use of the inductive constructors---Nitpick will generate an error if they
are needed.
+\subsection{Registration of Term Postprocessors}
+\label{registration-of-term-postprocessors}
+
+It is possible to change the output of any term that Nitpick considers a
+datatype by registering a term postprocessor. The interface for registering and
+unregistering postprocessors consists of the following pair of functions defined
+in the \textit{Nitpick} structure:
+
+\prew
+$\textbf{type}\,~\textit{term\_postprocessor}\,~{=} {}$ \\
+$\hbox{}\quad\textit{Proof.context} \rightarrow \textit{string} \rightarrow (\textit{typ} \rightarrow \textit{term~list\/}) \rightarrow \textit{typ} \rightarrow \textit{term} \rightarrow \textit{term}$ \\
+$\textbf{val}\,~\textit{register\_term\_postprocessors} : {}$ \\
+$\hbox{}\quad\textit{typ} \rightarrow \textit{term\_postprocessor} \rightarrow \textit{theory} \rightarrow \textit{theory}$ \\
+$\textbf{val}\,~\textit{unregister\_term\_postprocessors} :\,
+\textit{typ} \rightarrow \textit{theory} \rightarrow \textit{theory}$
+\postw
+
+\S\ref{typedefs-quotient-types-records-rationals-and-reals} and
+\texttt{src/HOL/Library/Multiset.thy} illustrate this feature in context.
+
\section{Known Bugs and Limitations}
\label{known-bugs-and-limitations}
--- a/src/HOL/Auth/TLS.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Auth/TLS.thy Fri Mar 12 16:02:42 2010 +0100
@@ -41,7 +41,7 @@
header{*The TLS Protocol: Transport Layer Security*}
-theory TLS imports Public Nat_Int_Bij begin
+theory TLS imports Public Nat_Bijection begin
definition certificate :: "[agent,key] => msg" where
"certificate A KA == Crypt (priSK Server) {|Agent A, Key KA|}"
@@ -74,19 +74,17 @@
specification (PRF)
inj_PRF: "inj PRF"
--{*the pseudo-random function is collision-free*}
- apply (rule exI [of _ "%(x,y,z). nat2_to_nat(x, nat2_to_nat(y,z))"])
- apply (simp add: inj_on_def)
- apply (blast dest!: nat2_to_nat_inj [THEN injD])
+ apply (rule exI [of _ "%(x,y,z). prod_encode(x, prod_encode(y,z))"])
+ apply (simp add: inj_on_def prod_encode_eq)
done
specification (sessionK)
inj_sessionK: "inj sessionK"
--{*sessionK is collision-free; also, no clientK clashes with any serverK.*}
apply (rule exI [of _
- "%((x,y,z), r). nat2_to_nat(role_case 0 1 r,
- nat2_to_nat(x, nat2_to_nat(y,z)))"])
- apply (simp add: inj_on_def split: role.split)
- apply (blast dest!: nat2_to_nat_inj [THEN injD])
+ "%((x,y,z), r). prod_encode(role_case 0 1 r,
+ prod_encode(x, prod_encode(y,z)))"])
+ apply (simp add: inj_on_def prod_encode_eq split: role.split)
done
axioms
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Big_Operators.thy Fri Mar 12 16:02:42 2010 +0100
@@ -0,0 +1,1711 @@
+(* Title: HOL/Big_Operators.thy
+ Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
+ with contributions by Jeremy Avigad
+*)
+
+header {* Big operators and finite (non-empty) sets *}
+
+theory Big_Operators
+imports Plain
+begin
+
+subsection {* Generalized summation over a set *}
+
+interpretation comm_monoid_add: comm_monoid_mult "op +" "0::'a::comm_monoid_add"
+ proof qed (auto intro: add_assoc add_commute)
+
+definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
+where "setsum f A == if finite A then fold_image (op +) f 0 A else 0"
+
+abbreviation
+ Setsum ("\<Sum>_" [1000] 999) where
+ "\<Sum>A == setsum (%x. x) A"
+
+text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
+written @{text"\<Sum>x\<in>A. e"}. *}
+
+syntax
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10)
+syntax (xsymbols)
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
+syntax (HTML output)
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
+
+translations -- {* Beware of argument permutation! *}
+ "SUM i:A. b" == "CONST setsum (%i. b) A"
+ "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
+
+text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Sum>x|P. e"}. *}
+
+syntax
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
+syntax (xsymbols)
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
+syntax (HTML output)
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
+
+translations
+ "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
+ "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
+
+print_translation {*
+let
+ fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
+ if x <> y then raise Match
+ else
+ let
+ val x' = Syntax.mark_bound x;
+ val t' = subst_bound (x', t);
+ val P' = subst_bound (x', P);
+ in Syntax.const @{syntax_const "_qsetsum"} $ Syntax.mark_bound x $ P' $ t' end
+ | setsum_tr' _ = raise Match;
+in [(@{const_syntax setsum}, setsum_tr')] end
+*}
+
+
+lemma setsum_empty [simp]: "setsum f {} = 0"
+by (simp add: setsum_def)
+
+lemma setsum_insert [simp]:
+ "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
+by (simp add: setsum_def)
+
+lemma setsum_infinite [simp]: "~ finite A ==> setsum f A = 0"
+by (simp add: setsum_def)
+
+lemma setsum_reindex:
+ "inj_on f B ==> setsum h (f ` B) = setsum (h \<circ> f) B"
+by(auto simp add: setsum_def comm_monoid_add.fold_image_reindex dest!:finite_imageD)
+
+lemma setsum_reindex_id:
+ "inj_on f B ==> setsum f B = setsum id (f ` B)"
+by (auto simp add: setsum_reindex)
+
+lemma setsum_reindex_nonzero:
+ assumes fS: "finite S"
+ and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
+ shows "setsum h (f ` S) = setsum (h o f) S"
+using nz
+proof(induct rule: finite_induct[OF fS])
+ case 1 thus ?case by simp
+next
+ case (2 x F)
+ {assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
+ then obtain y where y: "y \<in> F" "f x = f y" by auto
+ from "2.hyps" y have xy: "x \<noteq> y" by auto
+
+ from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
+ have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
+ also have "\<dots> = setsum (h o f) (insert x F)"
+ unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+ using h0
+ apply simp
+ apply (rule "2.hyps"(3))
+ apply (rule_tac y="y" in "2.prems")
+ apply simp_all
+ done
+ finally have ?case .}
+ moreover
+ {assume fxF: "f x \<notin> f ` F"
+ have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"
+ using fxF "2.hyps" by simp
+ also have "\<dots> = setsum (h o f) (insert x F)"
+ unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+ apply simp
+ apply (rule cong[OF refl[of "op + (h (f x))"]])
+ apply (rule "2.hyps"(3))
+ apply (rule_tac y="y" in "2.prems")
+ apply simp_all
+ done
+ finally have ?case .}
+ ultimately show ?case by blast
+qed
+
+lemma setsum_cong:
+ "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
+by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong)
+
+lemma strong_setsum_cong[cong]:
+ "A = B ==> (!!x. x:B =simp=> f x = g x)
+ ==> setsum (%x. f x) A = setsum (%x. g x) B"
+by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong)
+
+lemma setsum_cong2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setsum f A = setsum g A"
+by (rule setsum_cong[OF refl], auto)
+
+lemma setsum_reindex_cong:
+ "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]
+ ==> setsum h B = setsum g A"
+by (simp add: setsum_reindex cong: setsum_cong)
+
+
+lemma setsum_0[simp]: "setsum (%i. 0) A = 0"
+apply (clarsimp simp: setsum_def)
+apply (erule finite_induct, auto)
+done
+
+lemma setsum_0': "ALL a:A. f a = 0 ==> setsum f A = 0"
+by(simp add:setsum_cong)
+
+lemma setsum_Un_Int: "finite A ==> finite B ==>
+ setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
+ -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
+by(simp add: setsum_def comm_monoid_add.fold_image_Un_Int [symmetric])
+
+lemma setsum_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
+by (subst setsum_Un_Int [symmetric], auto)
+
+lemma setsum_mono_zero_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 0"
+ shows "setsum f S = setsum f T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_right:
+ "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. f i = 0 \<Longrightarrow> setsum f T = setsum f S"
+by(blast intro!: setsum_mono_zero_left[symmetric])
+
+lemma setsum_mono_zero_cong_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. g i = 0"
+ and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+ shows "setsum f S = setsum g T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_cong_right:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 0"
+ and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+ shows "setsum f T = setsum g S"
+using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto
+
+lemma setsum_delta:
+ assumes fS: "finite S"
+ shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else 0)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = 0" by simp
+ hence ?thesis using a by simp}
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
+ using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a by simp}
+ ultimately show ?thesis by blast
+qed
+lemma setsum_delta':
+ assumes fS: "finite S" shows
+ "setsum (\<lambda>k. if a = k then b k else 0) S =
+ (if a\<in> S then b a else 0)"
+ using setsum_delta[OF fS, of a b, symmetric]
+ by (auto intro: setsum_cong)
+
+lemma setsum_restrict_set:
+ assumes fA: "finite A"
+ shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
+proof-
+ from fA have fab: "finite (A \<inter> B)" by auto
+ have aba: "A \<inter> B \<subseteq> A" by blast
+ let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
+ from setsum_mono_zero_left[OF fA aba, of ?g]
+ show ?thesis by simp
+qed
+
+lemma setsum_cases:
+ assumes fA: "finite A"
+ shows "setsum (\<lambda>x. if P x then f x else g x) A =
+ setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
+proof-
+ have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}"
+ "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
+ by blast+
+ from fA
+ have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
+ let ?g = "\<lambda>x. if P x then f x else g x"
+ from setsum_Un_disjoint[OF f a(2), of ?g] a(1)
+ show ?thesis by simp
+qed
+
+
+(*But we can't get rid of finite I. If infinite, although the rhs is 0,
+ the lhs need not be, since UNION I A could still be finite.*)
+lemma setsum_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
+ setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
+by(simp add: setsum_def comm_monoid_add.fold_image_UN_disjoint cong: setsum_cong)
+
+text{*No need to assume that @{term C} is finite. If infinite, the rhs is
+directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*}
+lemma setsum_Union_disjoint:
+ "[| (ALL A:C. finite A);
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
+ ==> setsum f (Union C) = setsum (setsum f) C"
+apply (cases "finite C")
+ prefer 2 apply (force dest: finite_UnionD simp add: setsum_def)
+ apply (frule setsum_UN_disjoint [of C id f])
+ apply (unfold Union_def id_def, assumption+)
+done
+
+(*But we can't get rid of finite A. If infinite, although the lhs is 0,
+ the rhs need not be, since SIGMA A B could still be finite.*)
+lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
+ (\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"
+by(simp add:setsum_def comm_monoid_add.fold_image_Sigma split_def cong:setsum_cong)
+
+text{*Here we can eliminate the finiteness assumptions, by cases.*}
+lemma setsum_cartesian_product:
+ "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
+apply (cases "finite A")
+ apply (cases "finite B")
+ apply (simp add: setsum_Sigma)
+ apply (cases "A={}", simp)
+ apply (simp)
+apply (auto simp add: setsum_def
+ dest: finite_cartesian_productD1 finite_cartesian_productD2)
+done
+
+lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
+by(simp add:setsum_def comm_monoid_add.fold_image_distrib)
+
+
+subsubsection {* Properties in more restricted classes of structures *}
+
+lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
+apply (case_tac "finite A")
+ prefer 2 apply (simp add: setsum_def)
+apply (erule rev_mp)
+apply (erule finite_induct, auto)
+done
+
+lemma setsum_eq_0_iff [simp]:
+ "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
+by (induct set: finite) auto
+
+lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>
+ (setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"
+apply(erule finite_induct)
+apply (auto simp add:add_is_1)
+done
+
+lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
+
+lemma setsum_Un_nat: "finite A ==> finite B ==>
+ (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
+ -- {* For the natural numbers, we have subtraction. *}
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
+
+lemma setsum_Un: "finite A ==> finite B ==>
+ (setsum f (A Un B) :: 'a :: ab_group_add) =
+ setsum f A + setsum f B - setsum f (A Int B)"
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
+
+lemma (in comm_monoid_mult) fold_image_1: "finite S \<Longrightarrow> (\<forall>x\<in>S. f x = 1) \<Longrightarrow> fold_image op * f 1 S = 1"
+ apply (induct set: finite)
+ apply simp by auto
+
+lemma (in comm_monoid_mult) fold_image_Un_one:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
+ shows "fold_image (op *) f 1 (S \<union> T) = fold_image (op *) f 1 S * fold_image (op *) f 1 T"
+proof-
+ have "fold_image op * f 1 (S \<inter> T) = 1"
+ apply (rule fold_image_1)
+ using fS fT I0 by auto
+ with fold_image_Un_Int[OF fS fT] show ?thesis by simp
+qed
+
+lemma setsum_eq_general_reverses:
+ assumes fS: "finite S" and fT: "finite T"
+ and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
+ and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"
+ shows "setsum f S = setsum g T"
+ apply (simp add: setsum_def fS fT)
+ apply (rule comm_monoid_add.fold_image_eq_general_inverses[OF fS])
+ apply (erule kh)
+ apply (erule hk)
+ done
+
+
+
+lemma setsum_Un_zero:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 0"
+ shows "setsum f (S \<union> T) = setsum f S + setsum f T"
+ using fS fT
+ apply (simp add: setsum_def)
+ apply (rule comm_monoid_add.fold_image_Un_one)
+ using I0 by auto
+
+
+lemma setsum_UNION_zero:
+ assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
+ and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"
+ shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
+ using fSS f0
+proof(induct rule: finite_induct[OF fS])
+ case 1 thus ?case by simp
+next
+ case (2 T F)
+ then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"
+ and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
+ from fTF have fUF: "finite (\<Union>F)" by auto
+ from "2.prems" TF fTF
+ show ?case
+ by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f])
+qed
+
+
+lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
+ (if a:A then setsum f A - f a else setsum f A)"
+apply (case_tac "finite A")
+ prefer 2 apply (simp add: setsum_def)
+apply (erule finite_induct)
+ apply (auto simp add: insert_Diff_if)
+apply (drule_tac a = a in mk_disjoint_insert, auto)
+done
+
+lemma setsum_diff1: "finite A \<Longrightarrow>
+ (setsum f (A - {a}) :: ('a::ab_group_add)) =
+ (if a:A then setsum f A - f a else setsum f A)"
+by (erule finite_induct) (auto simp add: insert_Diff_if)
+
+lemma setsum_diff1'[rule_format]:
+ "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
+apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
+apply (auto simp add: insert_Diff_if add_ac)
+done
+
+lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
+ shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
+unfolding setsum_diff1'[OF assms] by auto
+
+(* By Jeremy Siek: *)
+
+lemma setsum_diff_nat:
+assumes "finite B" and "B \<subseteq> A"
+shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"
+using assms
+proof induct
+ show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
+next
+ fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"
+ and xFinA: "insert x F \<subseteq> A"
+ and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"
+ from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp
+ from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
+ by (simp add: setsum_diff1_nat)
+ from xFinA have "F \<subseteq> A" by simp
+ with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
+ with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
+ by simp
+ from xnotinF have "A - insert x F = (A - F) - {x}" by auto
+ with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
+ by simp
+ from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp
+ with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
+ by simp
+ thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp
+qed
+
+lemma setsum_diff:
+ assumes le: "finite A" "B \<subseteq> A"
+ shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
+proof -
+ from le have finiteB: "finite B" using finite_subset by auto
+ show ?thesis using finiteB le
+ proof induct
+ case empty
+ thus ?case by auto
+ next
+ case (insert x F)
+ thus ?case using le finiteB
+ by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
+ qed
+qed
+
+lemma setsum_mono:
+ assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
+ shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
+proof (cases "finite K")
+ case True
+ thus ?thesis using le
+ proof induct
+ case empty
+ thus ?case by simp
+ next
+ case insert
+ thus ?case using add_mono by fastsimp
+ qed
+next
+ case False
+ thus ?thesis
+ by (simp add: setsum_def)
+qed
+
+lemma setsum_strict_mono:
+ fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
+ assumes "finite A" "A \<noteq> {}"
+ and "!!x. x:A \<Longrightarrow> f x < g x"
+ shows "setsum f A < setsum g A"
+ using prems
+proof (induct rule: finite_ne_induct)
+ case singleton thus ?case by simp
+next
+ case insert thus ?case by (auto simp: add_strict_mono)
+qed
+
+lemma setsum_negf:
+ "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"
+proof (cases "finite A")
+ case True thus ?thesis by (induct set: finite) auto
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_subtractf:
+ "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
+ setsum f A - setsum g A"
+proof (cases "finite A")
+ case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_nonneg:
+ assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
+ shows "0 \<le> setsum f A"
+proof (cases "finite A")
+ case True thus ?thesis using nn
+ proof induct
+ case empty then show ?case by simp
+ next
+ case (insert x F)
+ then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
+ with insert show ?case by simp
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_nonpos:
+ assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
+ shows "setsum f A \<le> 0"
+proof (cases "finite A")
+ case True thus ?thesis using np
+ proof induct
+ case empty then show ?case by simp
+ next
+ case (insert x F)
+ then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
+ with insert show ?case by simp
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_mono2:
+fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_semigroup_add_imp_le,comm_monoid_add}"
+assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
+shows "setsum f A \<le> setsum f B"
+proof -
+ have "setsum f A \<le> setsum f A + setsum f (B-A)"
+ by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
+ also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
+ by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
+ also have "A \<union> (B-A) = B" using sub by blast
+ finally show ?thesis .
+qed
+
+lemma setsum_mono3: "finite B ==> A <= B ==>
+ ALL x: B - A.
+ 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
+ setsum f A <= setsum f B"
+ apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")
+ apply (erule ssubst)
+ apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")
+ apply simp
+ apply (rule add_left_mono)
+ apply (erule setsum_nonneg)
+ apply (subst setsum_Un_disjoint [THEN sym])
+ apply (erule finite_subset, assumption)
+ apply (rule finite_subset)
+ prefer 2
+ apply assumption
+ apply (auto simp add: sup_absorb2)
+done
+
+lemma setsum_right_distrib:
+ fixes f :: "'a => ('b::semiring_0)"
+ shows "r * setsum f A = setsum (%n. r * f n) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: right_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_left_distrib:
+ "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"
+proof (cases "finite A")
+ case True
+ then show ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: left_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_divide_distrib:
+ "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"
+proof (cases "finite A")
+ case True
+ then show ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: add_divide_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_abs[iff]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A)
+ thus ?case by (auto intro: abs_triangle_ineq order_trans)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_abs_ge_zero[iff]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "0 \<le> setsum (%i. abs(f i)) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma abs_setsum_abs[simp]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert a A)
+ hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp
+ also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp
+ also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"
+ by (simp del: abs_of_nonneg)
+ also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp
+ finally show ?case .
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+
+lemma setsum_Plus:
+ fixes A :: "'a set" and B :: "'b set"
+ assumes fin: "finite A" "finite B"
+ shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
+proof -
+ have "A <+> B = Inl ` A \<union> Inr ` B" by auto
+ moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
+ by(auto intro: finite_imageI)
+ moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
+ moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
+ ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex)
+qed
+
+
+text {* Commuting outer and inner summation *}
+
+lemma swap_inj_on:
+ "inj_on (%(i, j). (j, i)) (A \<times> B)"
+ by (unfold inj_on_def) fast
+
+lemma swap_product:
+ "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
+ by (simp add: split_def image_def) blast
+
+lemma setsum_commute:
+ "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
+proof (simp add: setsum_cartesian_product)
+ have "(\<Sum>(x,y) \<in> A <*> B. f x y) =
+ (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"
+ (is "?s = _")
+ apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on)
+ apply (simp add: split_def)
+ done
+ also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"
+ (is "_ = ?t")
+ apply (simp add: swap_product)
+ done
+ finally show "?s = ?t" .
+qed
+
+lemma setsum_product:
+ fixes f :: "'a => ('b::semiring_0)"
+ shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
+ by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
+
+lemma setsum_mult_setsum_if_inj:
+fixes f :: "'a => ('b::semiring_0)"
+shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
+ setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
+by(auto simp: setsum_product setsum_cartesian_product
+ intro!: setsum_reindex_cong[symmetric])
+
+lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"
+apply (cases "finite A")
+apply (erule finite_induct)
+apply (auto simp add: algebra_simps)
+done
+
+lemma setsum_bounded:
+ assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
+ shows "setsum f A \<le> of_nat(card A) * K"
+proof (cases "finite A")
+ case True
+ thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+
+subsubsection {* Cardinality as special case of @{const setsum} *}
+
+lemma card_eq_setsum:
+ "card A = setsum (\<lambda>x. 1) A"
+ by (simp only: card_def setsum_def)
+
+lemma card_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {})
+ ==> card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
+apply (simp add: card_eq_setsum del: setsum_constant)
+apply (subgoal_tac
+ "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I")
+apply (simp add: setsum_UN_disjoint del: setsum_constant)
+apply (simp cong: setsum_cong)
+done
+
+lemma card_Union_disjoint:
+ "finite C ==> (ALL A:C. finite A) ==>
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
+ ==> card (Union C) = setsum card C"
+apply (frule card_UN_disjoint [of C id])
+apply (unfold Union_def id_def, assumption+)
+done
+
+text{*The image of a finite set can be expressed using @{term fold_image}.*}
+lemma image_eq_fold_image:
+ "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A"
+proof (induct rule: finite_induct)
+ case empty then show ?case by simp
+next
+ interpret ab_semigroup_mult "op Un"
+ proof qed auto
+ case insert
+ then show ?case by simp
+qed
+
+subsubsection {* Cardinality of products *}
+
+lemma card_SigmaI [simp]:
+ "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
+ \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
+by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant)
+
+(*
+lemma SigmaI_insert: "y \<notin> A ==>
+ (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
+ by auto
+*)
+
+lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"
+ by (cases "finite A \<and> finite B")
+ (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2)
+
+lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)"
+by (simp add: card_cartesian_product)
+
+
+subsection {* Generalized product over a set *}
+
+definition setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult"
+where "setprod f A == if finite A then fold_image (op *) f 1 A else 1"
+
+abbreviation
+ Setprod ("\<Prod>_" [1000] 999) where
+ "\<Prod>A == setprod (%x. x) A"
+
+syntax
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10)
+syntax (xsymbols)
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
+syntax (HTML output)
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
+
+translations -- {* Beware of argument permutation! *}
+ "PROD i:A. b" == "CONST setprod (%i. b) A"
+ "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
+
+text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Prod>x|P. e"}. *}
+
+syntax
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
+syntax (xsymbols)
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
+syntax (HTML output)
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
+
+translations
+ "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
+ "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
+
+
+lemma setprod_empty [simp]: "setprod f {} = 1"
+by (auto simp add: setprod_def)
+
+lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
+ setprod f (insert a A) = f a * setprod f A"
+by (simp add: setprod_def)
+
+lemma setprod_infinite [simp]: "~ finite A ==> setprod f A = 1"
+by (simp add: setprod_def)
+
+lemma setprod_reindex:
+ "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
+by(auto simp: setprod_def fold_image_reindex dest!:finite_imageD)
+
+lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)"
+by (auto simp add: setprod_reindex)
+
+lemma setprod_cong:
+ "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
+by(fastsimp simp: setprod_def intro: fold_image_cong)
+
+lemma strong_setprod_cong[cong]:
+ "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
+by(fastsimp simp: simp_implies_def setprod_def intro: fold_image_cong)
+
+lemma setprod_reindex_cong: "inj_on f A ==>
+ B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
+by (frule setprod_reindex, simp)
+
+lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"
+ and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
+ shows "setprod h B = setprod g A"
+proof-
+ have "setprod h B = setprod (h o f) A"
+ by (simp add: B setprod_reindex[OF i, of h])
+ then show ?thesis apply simp
+ apply (rule setprod_cong)
+ apply simp
+ by (simp add: eq)
+qed
+
+lemma setprod_Un_one:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
+ shows "setprod f (S \<union> T) = setprod f S * setprod f T"
+ using fS fT
+ apply (simp add: setprod_def)
+ apply (rule fold_image_Un_one)
+ using I0 by auto
+
+
+lemma setprod_1: "setprod (%i. 1) A = 1"
+apply (case_tac "finite A")
+apply (erule finite_induct, auto simp add: mult_ac)
+done
+
+lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1"
+apply (subgoal_tac "setprod f F = setprod (%x. 1) F")
+apply (erule ssubst, rule setprod_1)
+apply (rule setprod_cong, auto)
+done
+
+lemma setprod_Un_Int: "finite A ==> finite B
+ ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
+by(simp add: setprod_def fold_image_Un_Int[symmetric])
+
+lemma setprod_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
+by (subst setprod_Un_Int [symmetric], auto)
+
+lemma setprod_mono_one_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 1"
+ shows "setprod f S = setprod f T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ by (simp add: setprod_Un_disjoint[OF f d, unfolded eq[symmetric]] setprod_1'[OF z])
+qed
+
+lemmas setprod_mono_one_right = setprod_mono_one_left [THEN sym]
+
+lemma setprod_delta:
+ assumes fS: "finite S"
+ shows "setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else 1)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = 1" by simp
+ hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) }
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto
+ have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+ using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)}
+ ultimately show ?thesis by blast
+qed
+
+lemma setprod_delta':
+ assumes fS: "finite S" shows
+ "setprod (\<lambda>k. if a = k then b k else 1) S =
+ (if a\<in> S then b a else 1)"
+ using setprod_delta[OF fS, of a b, symmetric]
+ by (auto intro: setprod_cong)
+
+
+lemma setprod_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
+ setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
+by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong)
+
+lemma setprod_Union_disjoint:
+ "[| (ALL A:C. finite A);
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
+ ==> setprod f (Union C) = setprod (setprod f) C"
+apply (cases "finite C")
+ prefer 2 apply (force dest: finite_UnionD simp add: setprod_def)
+ apply (frule setprod_UN_disjoint [of C id f])
+ apply (unfold Union_def id_def, assumption+)
+done
+
+lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
+ (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
+ (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
+by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong)
+
+text{*Here we can eliminate the finiteness assumptions, by cases.*}
+lemma setprod_cartesian_product:
+ "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
+apply (cases "finite A")
+ apply (cases "finite B")
+ apply (simp add: setprod_Sigma)
+ apply (cases "A={}", simp)
+ apply (simp add: setprod_1)
+apply (auto simp add: setprod_def
+ dest: finite_cartesian_productD1 finite_cartesian_productD2)
+done
+
+lemma setprod_timesf:
+ "setprod (%x. f x * g x) A = (setprod f A * setprod g A)"
+by(simp add:setprod_def fold_image_distrib)
+
+
+subsubsection {* Properties in more restricted classes of structures *}
+
+lemma setprod_eq_1_iff [simp]:
+ "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"
+by (induct set: finite) auto
+
+lemma setprod_zero:
+ "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"
+apply (induct set: finite, force, clarsimp)
+apply (erule disjE, auto)
+done
+
+lemma setprod_nonneg [rule_format]:
+ "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
+by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)
+
+lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)
+ --> 0 < setprod f A"
+by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)
+
+lemma setprod_zero_iff[simp]: "finite A ==>
+ (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
+ (EX x: A. f x = 0)"
+by (erule finite_induct, auto simp:no_zero_divisors)
+
+lemma setprod_pos_nat:
+ "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
+using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
+
+lemma setprod_pos_nat_iff[simp]:
+ "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"
+using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
+
+lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
+ (setprod f (A Un B) :: 'a ::{field})
+ = setprod f A * setprod f B / setprod f (A Int B)"
+by (subst setprod_Un_Int [symmetric], auto)
+
+lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
+ (setprod f (A - {a}) :: 'a :: {field}) =
+ (if a:A then setprod f A / f a else setprod f A)"
+by (erule finite_induct) (auto simp add: insert_Diff_if)
+
+lemma setprod_inversef:
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
+by (erule finite_induct) auto
+
+lemma setprod_dividef:
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ shows "finite A
+ ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
+apply (subgoal_tac
+ "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
+apply (erule ssubst)
+apply (subst divide_inverse)
+apply (subst setprod_timesf)
+apply (subst setprod_inversef, assumption+, rule refl)
+apply (rule setprod_cong, rule refl)
+apply (subst divide_inverse, auto)
+done
+
+lemma setprod_dvd_setprod [rule_format]:
+ "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply (auto simp add: dvd_def)
+ apply (rule_tac x = "k * ka" in exI)
+ apply (simp add: algebra_simps)
+done
+
+lemma setprod_dvd_setprod_subset:
+ "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
+ apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
+ apply (unfold dvd_def, blast)
+ apply (subst setprod_Un_disjoint [symmetric])
+ apply (auto elim: finite_subset intro: setprod_cong)
+done
+
+lemma setprod_dvd_setprod_subset2:
+ "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>
+ setprod f A dvd setprod g B"
+ apply (rule dvd_trans)
+ apply (rule setprod_dvd_setprod, erule (1) bspec)
+ apply (erule (1) setprod_dvd_setprod_subset)
+done
+
+lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>
+ (f i ::'a::comm_semiring_1) dvd setprod f A"
+by (induct set: finite) (auto intro: dvd_mult)
+
+lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>
+ (d::'a::comm_semiring_1) dvd (SUM x : A. f x)"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply auto
+done
+
+lemma setprod_mono:
+ fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"
+ assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"
+ shows "setprod f A \<le> setprod g A"
+proof (cases "finite A")
+ case True
+ hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]
+ proof (induct A rule: finite_subset_induct)
+ case (insert a F)
+ thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
+ unfolding setprod_insert[OF insert(1,3)]
+ using assms[rule_format,OF insert(2)] insert
+ by (auto intro: mult_mono mult_nonneg_nonneg)
+ qed auto
+ thus ?thesis by simp
+qed auto
+
+lemma abs_setprod:
+ fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
+ shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"
+proof (cases "finite A")
+ case True thus ?thesis
+ by induct (auto simp add: field_simps abs_mult)
+qed auto
+
+lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
+apply (erule finite_induct)
+apply auto
+done
+
+lemma setprod_gen_delta:
+ assumes fS: "finite S"
+ shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult}) * c^ (card S - 1) else c^ card S)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else c)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = c" by simp
+ hence ?thesis using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) }
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
+ apply (rule setprod_cong) by auto
+ have cA: "card ?A = card S - 1" using fS a by auto
+ have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto
+ have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+ using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a cA
+ by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
+ ultimately show ?thesis by blast
+qed
+
+
+subsubsection {* Fold1 in lattices with @{const inf} and @{const sup} *}
+
+text{*
+ As an application of @{text fold1} we define infimum
+ and supremum in (not necessarily complete!) lattices
+ over (non-empty) sets by means of @{text fold1}.
+*}
+
+context semilattice_inf
+begin
+
+lemma below_fold1_iff:
+ assumes "finite A" "A \<noteq> {}"
+ shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
+qed
+
+lemma fold1_belowI:
+ assumes "finite A"
+ and "a \<in> A"
+ shows "fold1 inf A \<le> a"
+proof -
+ from assms have "A \<noteq> {}" by auto
+ from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
+ proof (induct rule: finite_ne_induct)
+ case singleton thus ?case by simp
+ next
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ case (insert x F)
+ from insert(5) have "a = x \<or> a \<in> F" by simp
+ thus ?case
+ proof
+ assume "a = x" thus ?thesis using insert
+ by (simp add: mult_ac)
+ next
+ assume "a \<in> F"
+ hence bel: "fold1 inf F \<le> a" by (rule insert)
+ have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"
+ using insert by (simp add: mult_ac)
+ also have "inf (fold1 inf F) a = fold1 inf F"
+ using bel by (auto intro: antisym)
+ also have "inf x \<dots> = fold1 inf (insert x F)"
+ using insert by (simp add: mult_ac)
+ finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .
+ moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp
+ ultimately show ?thesis by simp
+ qed
+ qed
+qed
+
+end
+
+context lattice
+begin
+
+definition
+ Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900)
+where
+ "Inf_fin = fold1 inf"
+
+definition
+ Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900)
+where
+ "Sup_fin = fold1 sup"
+
+lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A"
+apply(unfold Sup_fin_def Inf_fin_def)
+apply(subgoal_tac "EX a. a:A")
+prefer 2 apply blast
+apply(erule exE)
+apply(rule order_trans)
+apply(erule (1) fold1_belowI)
+apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice])
+done
+
+lemma sup_Inf_absorb [simp]:
+ "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a"
+apply(subst sup_commute)
+apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI)
+done
+
+lemma inf_Sup_absorb [simp]:
+ "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"
+by (simp add: Sup_fin_def inf_absorb1
+ semilattice_inf.fold1_belowI [OF dual_semilattice])
+
+end
+
+context distrib_lattice
+begin
+
+lemma sup_Inf1_distrib:
+ assumes "finite A"
+ and "A \<noteq> {}"
+ shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ from assms show ?thesis
+ by (simp add: Inf_fin_def image_def
+ hom_fold1_commute [where h="sup x", OF sup_inf_distrib1])
+ (rule arg_cong [where f="fold1 inf"], blast)
+qed
+
+lemma sup_Inf2_distrib:
+ assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
+ shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}"
+using A proof (induct rule: finite_ne_induct)
+ case singleton thus ?case
+ by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
+next
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ case (insert x A)
+ have finB: "finite {sup x b |b. b \<in> B}"
+ by(rule finite_surj[where f = "sup x", OF B(1)], auto)
+ have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
+ proof -
+ have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
+ by blast
+ thus ?thesis by(simp add: insert(1) B(1))
+ qed
+ have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
+ have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)"
+ using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def])
+ also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2)
+ also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})"
+ using insert by(simp add:sup_Inf1_distrib[OF B])
+ also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
+ (is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M")
+ using B insert
+ by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne])
+ also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
+ by blast
+ finally show ?case .
+qed
+
+lemma inf_Sup1_distrib:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
+proof -
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ from assms show ?thesis
+ by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
+ (rule arg_cong [where f="fold1 sup"], blast)
+qed
+
+lemma inf_Sup2_distrib:
+ assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
+ shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}"
+using A proof (induct rule: finite_ne_induct)
+ case singleton thus ?case
+ by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def])
+next
+ case (insert x A)
+ have finB: "finite {inf x b |b. b \<in> B}"
+ by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)
+ have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
+ proof -
+ have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
+ by blast
+ thus ?thesis by(simp add: insert(1) B(1))
+ qed
+ have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
+ using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
+ also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2)
+ also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})"
+ using insert by(simp add:inf_Sup1_distrib[OF B])
+ also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
+ (is "_ = \<Squnion>\<^bsub>fin\<^esub>?M")
+ using B insert
+ by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne])
+ also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
+ by blast
+ finally show ?case .
+qed
+
+end
+
+context complete_lattice
+begin
+
+lemma Inf_fin_Inf:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+ moreover with `finite A` have "finite B" by simp
+ ultimately show ?thesis
+ by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
+ (simp add: Inf_fold_inf)
+qed
+
+lemma Sup_fin_Sup:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
+proof -
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+ moreover with `finite A` have "finite B" by simp
+ ultimately show ?thesis
+ by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])
+ (simp add: Sup_fold_sup)
+qed
+
+end
+
+
+subsubsection {* Fold1 in linear orders with @{const min} and @{const max} *}
+
+text{*
+ As an application of @{text fold1} we define minimum
+ and maximum in (not necessarily complete!) linear orders
+ over (non-empty) sets by means of @{text fold1}.
+*}
+
+context linorder
+begin
+
+lemma ab_semigroup_idem_mult_min:
+ "ab_semigroup_idem_mult min"
+ proof qed (auto simp add: min_def)
+
+lemma ab_semigroup_idem_mult_max:
+ "ab_semigroup_idem_mult max"
+ proof qed (auto simp add: max_def)
+
+lemma max_lattice:
+ "semilattice_inf (op \<ge>) (op >) max"
+ by (fact min_max.dual_semilattice)
+
+lemma dual_max:
+ "ord.max (op \<ge>) = min"
+ by (auto simp add: ord.max_def_raw min_def expand_fun_eq)
+
+lemma dual_min:
+ "ord.min (op \<ge>) = max"
+ by (auto simp add: ord.min_def_raw max_def expand_fun_eq)
+
+lemma strict_below_fold1_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert)
+qed
+
+lemma fold1_below_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert min_le_iff_disj)
+qed
+
+lemma fold1_strict_below_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert min_less_iff_disj)
+qed
+
+lemma fold1_antimono:
+ assumes "A \<noteq> {}" and "A \<subseteq> B" and "finite B"
+ shows "fold1 min B \<le> fold1 min A"
+proof cases
+ assume "A = B" thus ?thesis by simp
+next
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ assume "A \<noteq> B"
+ have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
+ have "fold1 min B = fold1 min (A \<union> (B-A))" by(subst B)(rule refl)
+ also have "\<dots> = min (fold1 min A) (fold1 min (B-A))"
+ proof -
+ have "finite A" by(rule finite_subset[OF `A \<subseteq> B` `finite B`])
+ moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *)
+ moreover have "(B-A) \<noteq> {}" using prems by blast
+ moreover have "A Int (B-A) = {}" using prems by blast
+ ultimately show ?thesis using `A \<noteq> {}` by (rule_tac fold1_Un)
+ qed
+ also have "\<dots> \<le> fold1 min A" by (simp add: min_le_iff_disj)
+ finally show ?thesis .
+qed
+
+definition
+ Min :: "'a set \<Rightarrow> 'a"
+where
+ "Min = fold1 min"
+
+definition
+ Max :: "'a set \<Rightarrow> 'a"
+where
+ "Max = fold1 max"
+
+lemmas Min_singleton [simp] = fold1_singleton_def [OF Min_def]
+lemmas Max_singleton [simp] = fold1_singleton_def [OF Max_def]
+
+lemma Min_insert [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min (insert x A) = min x (Min A)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
+qed
+
+lemma Max_insert [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max (insert x A) = max x (Max A)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
+qed
+
+lemma Min_in [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A \<in> A"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
+qed
+
+lemma Max_in [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A \<in> A"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
+qed
+
+lemma Min_Un:
+ assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
+ shows "Min (A \<union> B) = min (Min A) (Min B)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (simp add: Min_def fold1_Un2)
+qed
+
+lemma Max_Un:
+ assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
+ shows "Max (A \<union> B) = max (Max A) (Max B)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis
+ by (simp add: Max_def fold1_Un2)
+qed
+
+lemma hom_Min_commute:
+ assumes "\<And>x y. h (min x y) = min (h x) (h y)"
+ and "finite N" and "N \<noteq> {}"
+ shows "h (Min N) = Min (h ` N)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (simp add: Min_def hom_fold1_commute)
+qed
+
+lemma hom_Max_commute:
+ assumes "\<And>x y. h (max x y) = max (h x) (h y)"
+ and "finite N" and "N \<noteq> {}"
+ shows "h (Max N) = Max (h ` N)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis
+ by (simp add: Max_def hom_fold1_commute [of h])
+qed
+
+lemma Min_le [simp]:
+ assumes "finite A" and "x \<in> A"
+ shows "Min A \<le> x"
+ using assms by (simp add: Min_def min_max.fold1_belowI)
+
+lemma Max_ge [simp]:
+ assumes "finite A" and "x \<in> A"
+ shows "x \<le> Max A"
+proof -
+ interpret semilattice_inf "op \<ge>" "op >" max
+ by (rule max_lattice)
+ from assms show ?thesis by (simp add: Max_def fold1_belowI)
+qed
+
+lemma Min_ge_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
+ using assms by (simp add: Min_def min_max.below_fold1_iff)
+
+lemma Max_le_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
+proof -
+ interpret semilattice_inf "op \<ge>" "op >" max
+ by (rule max_lattice)
+ from assms show ?thesis by (simp add: Max_def below_fold1_iff)
+qed
+
+lemma Min_gr_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
+ using assms by (simp add: Min_def strict_below_fold1_iff)
+
+lemma Max_less_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.strict_below_fold1_iff [folded dual.dual_max])
+qed
+
+lemma Min_le_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
+ using assms by (simp add: Min_def fold1_below_iff)
+
+lemma Max_ge_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_below_iff [folded dual.dual_max])
+qed
+
+lemma Min_less_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
+ using assms by (simp add: Min_def fold1_strict_below_iff)
+
+lemma Max_gr_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_strict_below_iff [folded dual.dual_max])
+qed
+
+lemma Min_eqI:
+ assumes "finite A"
+ assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"
+ and "x \<in> A"
+ shows "Min A = x"
+proof (rule antisym)
+ from `x \<in> A` have "A \<noteq> {}" by auto
+ with assms show "Min A \<ge> x" by simp
+next
+ from assms show "x \<ge> Min A" by simp
+qed
+
+lemma Max_eqI:
+ assumes "finite A"
+ assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
+ and "x \<in> A"
+ shows "Max A = x"
+proof (rule antisym)
+ from `x \<in> A` have "A \<noteq> {}" by auto
+ with assms show "Max A \<le> x" by simp
+next
+ from assms show "x \<le> Max A" by simp
+qed
+
+lemma Min_antimono:
+ assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
+ shows "Min N \<le> Min M"
+ using assms by (simp add: Min_def fold1_antimono)
+
+lemma Max_mono:
+ assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
+ shows "Max M \<le> Max N"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_antimono [folded dual.dual_max])
+qed
+
+lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:
+ "finite A \<Longrightarrow> P {} \<Longrightarrow>
+ (!!b A. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A))
+ \<Longrightarrow> P A"
+proof (induct rule: finite_psubset_induct)
+ fix A :: "'a set"
+ assume IH: "!! B. finite B \<Longrightarrow> B < A \<Longrightarrow> P {} \<Longrightarrow>
+ (!!b A. finite A \<Longrightarrow> (\<forall>a\<in>A. a<b) \<Longrightarrow> P A \<Longrightarrow> P (insert b A))
+ \<Longrightarrow> P B"
+ and "finite A" and "P {}"
+ and step: "!!b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)"
+ show "P A"
+ proof (cases "A = {}")
+ assume "A = {}" thus "P A" using `P {}` by simp
+ next
+ let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B"
+ assume "A \<noteq> {}"
+ with `finite A` have "Max A : A" by auto
+ hence A: "?A = A" using insert_Diff_single insert_absorb by auto
+ moreover have "finite ?B" using `finite A` by simp
+ ultimately have "P ?B" using `P {}` step IH[of ?B] by blast
+ moreover have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp
+ ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp
+ qed
+qed
+
+lemma finite_linorder_min_induct[consumes 1, case_names empty insert]:
+ "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
+by(rule linorder.finite_linorder_max_induct[OF dual_linorder])
+
+end
+
+context linordered_ab_semigroup_add
+begin
+
+lemma add_Min_commute:
+ fixes k
+ assumes "finite N" and "N \<noteq> {}"
+ shows "k + Min N = Min {k + m | m. m \<in> N}"
+proof -
+ have "\<And>x y. k + min x y = min (k + x) (k + y)"
+ by (simp add: min_def not_le)
+ (blast intro: antisym less_imp_le add_left_mono)
+ with assms show ?thesis
+ using hom_Min_commute [of "plus k" N]
+ by simp (blast intro: arg_cong [where f = Min])
+qed
+
+lemma add_Max_commute:
+ fixes k
+ assumes "finite N" and "N \<noteq> {}"
+ shows "k + Max N = Max {k + m | m. m \<in> N}"
+proof -
+ have "\<And>x y. k + max x y = max (k + x) (k + y)"
+ by (simp add: max_def not_le)
+ (blast intro: antisym less_imp_le add_left_mono)
+ with assms show ?thesis
+ using hom_Max_commute [of "plus k" N]
+ by simp (blast intro: arg_cong [where f = Max])
+qed
+
+end
+
+context linordered_ab_group_add
+begin
+
+lemma minus_Max_eq_Min [simp]:
+ "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Max S) = Min (uminus ` S)"
+ by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)
+
+lemma minus_Min_eq_Max [simp]:
+ "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Min S) = Max (uminus ` S)"
+ by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)
+
+end
+
+end
--- a/src/HOL/Boogie/Tools/boogie_loader.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML Fri Mar 12 16:02:42 2010 +0100
@@ -88,7 +88,7 @@
let
val args = Name.variant_list [] (replicate arity "'")
val (T, thy') =
- Typedecl.typedecl (Binding.name isa_name, args, mk_syntax name arity) thy
+ Typedecl.typedecl_global (Binding.name isa_name, args, mk_syntax name arity) thy
val type_name = fst (Term.dest_Type T)
in (((name, type_name), log_new name type_name), thy') end)
end
--- a/src/HOL/Code_Numeral.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Code_Numeral.thy Fri Mar 12 16:02:42 2010 +0100
@@ -247,7 +247,7 @@
"nat_of i = nat_of_aux i 0"
by (simp add: nat_of_aux_def)
-definition div_mod_code_numeral :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
+definition div_mod_code_numeral :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
[code del]: "div_mod_code_numeral n m = (n div m, n mod m)"
lemma [code]:
--- a/src/HOL/Divides.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Divides.thy Fri Mar 12 16:02:42 2010 +0100
@@ -376,7 +376,7 @@
end
-class ring_div = semiring_div + idom
+class ring_div = semiring_div + comm_ring_1
begin
text {* Negation respects modular equivalence. *}
@@ -2353,47 +2353,6 @@
apply (rule Divides.div_less_dividend, simp_all)
done
-text {* code generator setup *}
-
-context ring_1
-begin
-
-lemma of_int_num [code]:
- "of_int k = (if k = 0 then 0 else if k < 0 then
- - of_int (- k) else let
- (l, m) = divmod_int k 2;
- l' = of_int l
- in if m = 0 then l' + l' else l' + l' + 1)"
-proof -
- have aux1: "k mod (2\<Colon>int) \<noteq> (0\<Colon>int) \<Longrightarrow>
- of_int k = of_int (k div 2 * 2 + 1)"
- proof -
- have "k mod 2 < 2" by (auto intro: pos_mod_bound)
- moreover have "0 \<le> k mod 2" by (auto intro: pos_mod_sign)
- moreover assume "k mod 2 \<noteq> 0"
- ultimately have "k mod 2 = 1" by arith
- moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
- ultimately show ?thesis by auto
- qed
- have aux2: "\<And>x. of_int 2 * x = x + x"
- proof -
- fix x
- have int2: "(2::int) = 1 + 1" by arith
- show "of_int 2 * x = x + x"
- unfolding int2 of_int_add left_distrib by simp
- qed
- have aux3: "\<And>x. x * of_int 2 = x + x"
- proof -
- fix x
- have int2: "(2::int) = 1 + 1" by arith
- show "x * of_int 2 = x + x"
- unfolding int2 of_int_add right_distrib by simp
- qed
- from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3)
-qed
-
-end
-
lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y"
proof
assume H: "x mod n = y mod n"
@@ -2482,6 +2441,7 @@
mod_div_equality' [THEN eq_reflection]
zmod_zdiv_equality' [THEN eq_reflection]
+
subsubsection {* Code generation *}
definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
@@ -2515,6 +2475,45 @@
then show ?thesis by (simp add: divmod_int_pdivmod)
qed
+context ring_1
+begin
+
+lemma of_int_num [code]:
+ "of_int k = (if k = 0 then 0 else if k < 0 then
+ - of_int (- k) else let
+ (l, m) = divmod_int k 2;
+ l' = of_int l
+ in if m = 0 then l' + l' else l' + l' + 1)"
+proof -
+ have aux1: "k mod (2\<Colon>int) \<noteq> (0\<Colon>int) \<Longrightarrow>
+ of_int k = of_int (k div 2 * 2 + 1)"
+ proof -
+ have "k mod 2 < 2" by (auto intro: pos_mod_bound)
+ moreover have "0 \<le> k mod 2" by (auto intro: pos_mod_sign)
+ moreover assume "k mod 2 \<noteq> 0"
+ ultimately have "k mod 2 = 1" by arith
+ moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
+ ultimately show ?thesis by auto
+ qed
+ have aux2: "\<And>x. of_int 2 * x = x + x"
+ proof -
+ fix x
+ have int2: "(2::int) = 1 + 1" by arith
+ show "of_int 2 * x = x + x"
+ unfolding int2 of_int_add left_distrib by simp
+ qed
+ have aux3: "\<And>x. x * of_int 2 = x + x"
+ proof -
+ fix x
+ have int2: "(2::int) = 1 + 1" by arith
+ show "x * of_int 2 = x + x"
+ unfolding int2 of_int_add right_distrib by simp
+ qed
+ from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3)
+qed
+
+end
+
code_modulename SML
Divides Arith
--- a/src/HOL/Equiv_Relations.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Equiv_Relations.thy Fri Mar 12 16:02:42 2010 +0100
@@ -5,7 +5,7 @@
header {* Equivalence Relations in Higher-Order Set Theory *}
theory Equiv_Relations
-imports Finite_Set Relation Plain
+imports Big_Operators Relation Plain
begin
subsection {* Equivalence relations *}
--- a/src/HOL/Finite_Set.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Finite_Set.thy Fri Mar 12 16:02:42 2010 +0100
@@ -6,7 +6,7 @@
header {* Finite sets *}
theory Finite_Set
-imports Power Product_Type Sum_Type
+imports Power Option
begin
subsection {* Definition and basic properties *}
@@ -527,17 +527,24 @@
lemma UNIV_unit [noatp]:
"UNIV = {()}" by auto
-instance unit :: finite
- by default (simp add: UNIV_unit)
+instance unit :: finite proof
+qed (simp add: UNIV_unit)
lemma UNIV_bool [noatp]:
"UNIV = {False, True}" by auto
-instance bool :: finite
- by default (simp add: UNIV_bool)
+instance bool :: finite proof
+qed (simp add: UNIV_bool)
+
+instance * :: (finite, finite) finite proof
+qed (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
-instance * :: (finite, finite) finite
- by default (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
+lemma finite_option_UNIV [simp]:
+ "finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"
+ by (auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some)
+
+instance option :: (finite) finite proof
+qed (simp add: UNIV_option_conv)
lemma inj_graph: "inj (%f. {(x, y). y = f x})"
by (rule inj_onI, auto simp add: expand_set_eq expand_fun_eq)
@@ -556,8 +563,8 @@
qed
qed
-instance "+" :: (finite, finite) finite
- by default (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
+instance "+" :: (finite, finite) finite proof
+qed (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
subsection {* A fold functional for finite sets *}
@@ -1053,1470 +1060,6 @@
end
-subsection {* Generalized summation over a set *}
-
-interpretation comm_monoid_add: comm_monoid_mult "op +" "0::'a::comm_monoid_add"
- proof qed (auto intro: add_assoc add_commute)
-
-definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
-where "setsum f A == if finite A then fold_image (op +) f 0 A else 0"
-
-abbreviation
- Setsum ("\<Sum>_" [1000] 999) where
- "\<Sum>A == setsum (%x. x) A"
-
-text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
-written @{text"\<Sum>x\<in>A. e"}. *}
-
-syntax
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10)
-syntax (xsymbols)
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
-syntax (HTML output)
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
-
-translations -- {* Beware of argument permutation! *}
- "SUM i:A. b" == "CONST setsum (%i. b) A"
- "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
-
-text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Sum>x|P. e"}. *}
-
-syntax
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
-syntax (xsymbols)
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
-syntax (HTML output)
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
-
-translations
- "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
- "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
-
-print_translation {*
-let
- fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
- if x <> y then raise Match
- else
- let
- val x' = Syntax.mark_bound x;
- val t' = subst_bound (x', t);
- val P' = subst_bound (x', P);
- in Syntax.const @{syntax_const "_qsetsum"} $ Syntax.mark_bound x $ P' $ t' end
- | setsum_tr' _ = raise Match;
-in [(@{const_syntax setsum}, setsum_tr')] end
-*}
-
-
-lemma setsum_empty [simp]: "setsum f {} = 0"
-by (simp add: setsum_def)
-
-lemma setsum_insert [simp]:
- "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
-by (simp add: setsum_def)
-
-lemma setsum_infinite [simp]: "~ finite A ==> setsum f A = 0"
-by (simp add: setsum_def)
-
-lemma setsum_reindex:
- "inj_on f B ==> setsum h (f ` B) = setsum (h \<circ> f) B"
-by(auto simp add: setsum_def comm_monoid_add.fold_image_reindex dest!:finite_imageD)
-
-lemma setsum_reindex_id:
- "inj_on f B ==> setsum f B = setsum id (f ` B)"
-by (auto simp add: setsum_reindex)
-
-lemma setsum_reindex_nonzero:
- assumes fS: "finite S"
- and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
- shows "setsum h (f ` S) = setsum (h o f) S"
-using nz
-proof(induct rule: finite_induct[OF fS])
- case 1 thus ?case by simp
-next
- case (2 x F)
- {assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
- then obtain y where y: "y \<in> F" "f x = f y" by auto
- from "2.hyps" y have xy: "x \<noteq> y" by auto
-
- from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
- have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
- also have "\<dots> = setsum (h o f) (insert x F)"
- unfolding setsum_insert[OF `finite F` `x\<notin>F`]
- using h0
- apply simp
- apply (rule "2.hyps"(3))
- apply (rule_tac y="y" in "2.prems")
- apply simp_all
- done
- finally have ?case .}
- moreover
- {assume fxF: "f x \<notin> f ` F"
- have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"
- using fxF "2.hyps" by simp
- also have "\<dots> = setsum (h o f) (insert x F)"
- unfolding setsum_insert[OF `finite F` `x\<notin>F`]
- apply simp
- apply (rule cong[OF refl[of "op + (h (f x))"]])
- apply (rule "2.hyps"(3))
- apply (rule_tac y="y" in "2.prems")
- apply simp_all
- done
- finally have ?case .}
- ultimately show ?case by blast
-qed
-
-lemma setsum_cong:
- "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
-by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong)
-
-lemma strong_setsum_cong[cong]:
- "A = B ==> (!!x. x:B =simp=> f x = g x)
- ==> setsum (%x. f x) A = setsum (%x. g x) B"
-by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong)
-
-lemma setsum_cong2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setsum f A = setsum g A"
-by (rule setsum_cong[OF refl], auto)
-
-lemma setsum_reindex_cong:
- "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]
- ==> setsum h B = setsum g A"
-by (simp add: setsum_reindex cong: setsum_cong)
-
-
-lemma setsum_0[simp]: "setsum (%i. 0) A = 0"
-apply (clarsimp simp: setsum_def)
-apply (erule finite_induct, auto)
-done
-
-lemma setsum_0': "ALL a:A. f a = 0 ==> setsum f A = 0"
-by(simp add:setsum_cong)
-
-lemma setsum_Un_Int: "finite A ==> finite B ==>
- setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
- -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
-by(simp add: setsum_def comm_monoid_add.fold_image_Un_Int [symmetric])
-
-lemma setsum_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
-by (subst setsum_Un_Int [symmetric], auto)
-
-lemma setsum_mono_zero_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 0"
- shows "setsum f S = setsum f T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
-qed
-
-lemma setsum_mono_zero_right:
- "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. f i = 0 \<Longrightarrow> setsum f T = setsum f S"
-by(blast intro!: setsum_mono_zero_left[symmetric])
-
-lemma setsum_mono_zero_cong_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. g i = 0"
- and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
- shows "setsum f S = setsum g T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
-qed
-
-lemma setsum_mono_zero_cong_right:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 0"
- and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
- shows "setsum f T = setsum g S"
-using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto
-
-lemma setsum_delta:
- assumes fS: "finite S"
- shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else 0)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = 0" by simp
- hence ?thesis using a by simp}
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
- using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a by simp}
- ultimately show ?thesis by blast
-qed
-lemma setsum_delta':
- assumes fS: "finite S" shows
- "setsum (\<lambda>k. if a = k then b k else 0) S =
- (if a\<in> S then b a else 0)"
- using setsum_delta[OF fS, of a b, symmetric]
- by (auto intro: setsum_cong)
-
-lemma setsum_restrict_set:
- assumes fA: "finite A"
- shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
-proof-
- from fA have fab: "finite (A \<inter> B)" by auto
- have aba: "A \<inter> B \<subseteq> A" by blast
- let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
- from setsum_mono_zero_left[OF fA aba, of ?g]
- show ?thesis by simp
-qed
-
-lemma setsum_cases:
- assumes fA: "finite A"
- shows "setsum (\<lambda>x. if P x then f x else g x) A =
- setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
-proof-
- have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}"
- "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
- by blast+
- from fA
- have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
- let ?g = "\<lambda>x. if P x then f x else g x"
- from setsum_Un_disjoint[OF f a(2), of ?g] a(1)
- show ?thesis by simp
-qed
-
-
-(*But we can't get rid of finite I. If infinite, although the rhs is 0,
- the lhs need not be, since UNION I A could still be finite.*)
-lemma setsum_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
- setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
-by(simp add: setsum_def comm_monoid_add.fold_image_UN_disjoint cong: setsum_cong)
-
-text{*No need to assume that @{term C} is finite. If infinite, the rhs is
-directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*}
-lemma setsum_Union_disjoint:
- "[| (ALL A:C. finite A);
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
- ==> setsum f (Union C) = setsum (setsum f) C"
-apply (cases "finite C")
- prefer 2 apply (force dest: finite_UnionD simp add: setsum_def)
- apply (frule setsum_UN_disjoint [of C id f])
- apply (unfold Union_def id_def, assumption+)
-done
-
-(*But we can't get rid of finite A. If infinite, although the lhs is 0,
- the rhs need not be, since SIGMA A B could still be finite.*)
-lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
- (\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"
-by(simp add:setsum_def comm_monoid_add.fold_image_Sigma split_def cong:setsum_cong)
-
-text{*Here we can eliminate the finiteness assumptions, by cases.*}
-lemma setsum_cartesian_product:
- "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
-apply (cases "finite A")
- apply (cases "finite B")
- apply (simp add: setsum_Sigma)
- apply (cases "A={}", simp)
- apply (simp)
-apply (auto simp add: setsum_def
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
-by(simp add:setsum_def comm_monoid_add.fold_image_distrib)
-
-
-subsubsection {* Properties in more restricted classes of structures *}
-
-lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
-apply (case_tac "finite A")
- prefer 2 apply (simp add: setsum_def)
-apply (erule rev_mp)
-apply (erule finite_induct, auto)
-done
-
-lemma setsum_eq_0_iff [simp]:
- "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
-by (induct set: finite) auto
-
-lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>
- (setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"
-apply(erule finite_induct)
-apply (auto simp add:add_is_1)
-done
-
-lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
-
-lemma setsum_Un_nat: "finite A ==> finite B ==>
- (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
- -- {* For the natural numbers, we have subtraction. *}
-by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
-
-lemma setsum_Un: "finite A ==> finite B ==>
- (setsum f (A Un B) :: 'a :: ab_group_add) =
- setsum f A + setsum f B - setsum f (A Int B)"
-by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
-
-lemma (in comm_monoid_mult) fold_image_1: "finite S \<Longrightarrow> (\<forall>x\<in>S. f x = 1) \<Longrightarrow> fold_image op * f 1 S = 1"
- apply (induct set: finite)
- apply simp by auto
-
-lemma (in comm_monoid_mult) fold_image_Un_one:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
- shows "fold_image (op *) f 1 (S \<union> T) = fold_image (op *) f 1 S * fold_image (op *) f 1 T"
-proof-
- have "fold_image op * f 1 (S \<inter> T) = 1"
- apply (rule fold_image_1)
- using fS fT I0 by auto
- with fold_image_Un_Int[OF fS fT] show ?thesis by simp
-qed
-
-lemma setsum_eq_general_reverses:
- assumes fS: "finite S" and fT: "finite T"
- and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
- and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"
- shows "setsum f S = setsum g T"
- apply (simp add: setsum_def fS fT)
- apply (rule comm_monoid_add.fold_image_eq_general_inverses[OF fS])
- apply (erule kh)
- apply (erule hk)
- done
-
-
-
-lemma setsum_Un_zero:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 0"
- shows "setsum f (S \<union> T) = setsum f S + setsum f T"
- using fS fT
- apply (simp add: setsum_def)
- apply (rule comm_monoid_add.fold_image_Un_one)
- using I0 by auto
-
-
-lemma setsum_UNION_zero:
- assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
- and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"
- shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
- using fSS f0
-proof(induct rule: finite_induct[OF fS])
- case 1 thus ?case by simp
-next
- case (2 T F)
- then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"
- and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
- from fTF have fUF: "finite (\<Union>F)" by auto
- from "2.prems" TF fTF
- show ?case
- by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f])
-qed
-
-
-lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
- (if a:A then setsum f A - f a else setsum f A)"
-apply (case_tac "finite A")
- prefer 2 apply (simp add: setsum_def)
-apply (erule finite_induct)
- apply (auto simp add: insert_Diff_if)
-apply (drule_tac a = a in mk_disjoint_insert, auto)
-done
-
-lemma setsum_diff1: "finite A \<Longrightarrow>
- (setsum f (A - {a}) :: ('a::ab_group_add)) =
- (if a:A then setsum f A - f a else setsum f A)"
-by (erule finite_induct) (auto simp add: insert_Diff_if)
-
-lemma setsum_diff1'[rule_format]:
- "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
-apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
-apply (auto simp add: insert_Diff_if add_ac)
-done
-
-lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
- shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
-unfolding setsum_diff1'[OF assms] by auto
-
-(* By Jeremy Siek: *)
-
-lemma setsum_diff_nat:
-assumes "finite B" and "B \<subseteq> A"
-shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"
-using assms
-proof induct
- show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
-next
- fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"
- and xFinA: "insert x F \<subseteq> A"
- and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"
- from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp
- from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
- by (simp add: setsum_diff1_nat)
- from xFinA have "F \<subseteq> A" by simp
- with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
- with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
- by simp
- from xnotinF have "A - insert x F = (A - F) - {x}" by auto
- with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
- by simp
- from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp
- with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
- by simp
- thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp
-qed
-
-lemma setsum_diff:
- assumes le: "finite A" "B \<subseteq> A"
- shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
-proof -
- from le have finiteB: "finite B" using finite_subset by auto
- show ?thesis using finiteB le
- proof induct
- case empty
- thus ?case by auto
- next
- case (insert x F)
- thus ?case using le finiteB
- by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
- qed
-qed
-
-lemma setsum_mono:
- assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
- shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
-proof (cases "finite K")
- case True
- thus ?thesis using le
- proof induct
- case empty
- thus ?case by simp
- next
- case insert
- thus ?case using add_mono by fastsimp
- qed
-next
- case False
- thus ?thesis
- by (simp add: setsum_def)
-qed
-
-lemma setsum_strict_mono:
- fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
- assumes "finite A" "A \<noteq> {}"
- and "!!x. x:A \<Longrightarrow> f x < g x"
- shows "setsum f A < setsum g A"
- using prems
-proof (induct rule: finite_ne_induct)
- case singleton thus ?case by simp
-next
- case insert thus ?case by (auto simp: add_strict_mono)
-qed
-
-lemma setsum_negf:
- "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"
-proof (cases "finite A")
- case True thus ?thesis by (induct set: finite) auto
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_subtractf:
- "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
- setsum f A - setsum g A"
-proof (cases "finite A")
- case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_nonneg:
- assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
- shows "0 \<le> setsum f A"
-proof (cases "finite A")
- case True thus ?thesis using nn
- proof induct
- case empty then show ?case by simp
- next
- case (insert x F)
- then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
- with insert show ?case by simp
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_nonpos:
- assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
- shows "setsum f A \<le> 0"
-proof (cases "finite A")
- case True thus ?thesis using np
- proof induct
- case empty then show ?case by simp
- next
- case (insert x F)
- then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
- with insert show ?case by simp
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_mono2:
-fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_semigroup_add_imp_le,comm_monoid_add}"
-assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
-shows "setsum f A \<le> setsum f B"
-proof -
- have "setsum f A \<le> setsum f A + setsum f (B-A)"
- by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
- also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
- by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
- also have "A \<union> (B-A) = B" using sub by blast
- finally show ?thesis .
-qed
-
-lemma setsum_mono3: "finite B ==> A <= B ==>
- ALL x: B - A.
- 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
- setsum f A <= setsum f B"
- apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")
- apply (erule ssubst)
- apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")
- apply simp
- apply (rule add_left_mono)
- apply (erule setsum_nonneg)
- apply (subst setsum_Un_disjoint [THEN sym])
- apply (erule finite_subset, assumption)
- apply (rule finite_subset)
- prefer 2
- apply assumption
- apply (auto simp add: sup_absorb2)
-done
-
-lemma setsum_right_distrib:
- fixes f :: "'a => ('b::semiring_0)"
- shows "r * setsum f A = setsum (%n. r * f n) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: right_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_left_distrib:
- "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"
-proof (cases "finite A")
- case True
- then show ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: left_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_divide_distrib:
- "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"
-proof (cases "finite A")
- case True
- then show ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: add_divide_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_abs[iff]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A)
- thus ?case by (auto intro: abs_triangle_ineq order_trans)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_abs_ge_zero[iff]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "0 \<le> setsum (%i. abs(f i)) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma abs_setsum_abs[simp]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert a A)
- hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp
- also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp
- also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"
- by (simp del: abs_of_nonneg)
- also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp
- finally show ?case .
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-
-lemma setsum_Plus:
- fixes A :: "'a set" and B :: "'b set"
- assumes fin: "finite A" "finite B"
- shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
-proof -
- have "A <+> B = Inl ` A \<union> Inr ` B" by auto
- moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
- by(auto intro: finite_imageI)
- moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
- moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
- ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex)
-qed
-
-
-text {* Commuting outer and inner summation *}
-
-lemma swap_inj_on:
- "inj_on (%(i, j). (j, i)) (A \<times> B)"
- by (unfold inj_on_def) fast
-
-lemma swap_product:
- "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
- by (simp add: split_def image_def) blast
-
-lemma setsum_commute:
- "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
-proof (simp add: setsum_cartesian_product)
- have "(\<Sum>(x,y) \<in> A <*> B. f x y) =
- (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"
- (is "?s = _")
- apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on)
- apply (simp add: split_def)
- done
- also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"
- (is "_ = ?t")
- apply (simp add: swap_product)
- done
- finally show "?s = ?t" .
-qed
-
-lemma setsum_product:
- fixes f :: "'a => ('b::semiring_0)"
- shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
- by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
-
-lemma setsum_mult_setsum_if_inj:
-fixes f :: "'a => ('b::semiring_0)"
-shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
- setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
-by(auto simp: setsum_product setsum_cartesian_product
- intro!: setsum_reindex_cong[symmetric])
-
-
-subsection {* Generalized product over a set *}
-
-definition setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult"
-where "setprod f A == if finite A then fold_image (op *) f 1 A else 1"
-
-abbreviation
- Setprod ("\<Prod>_" [1000] 999) where
- "\<Prod>A == setprod (%x. x) A"
-
-syntax
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10)
-syntax (xsymbols)
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
-syntax (HTML output)
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
-
-translations -- {* Beware of argument permutation! *}
- "PROD i:A. b" == "CONST setprod (%i. b) A"
- "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
-
-text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Prod>x|P. e"}. *}
-
-syntax
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
-syntax (xsymbols)
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
-syntax (HTML output)
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
-
-translations
- "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
- "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
-
-
-lemma setprod_empty [simp]: "setprod f {} = 1"
-by (auto simp add: setprod_def)
-
-lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
- setprod f (insert a A) = f a * setprod f A"
-by (simp add: setprod_def)
-
-lemma setprod_infinite [simp]: "~ finite A ==> setprod f A = 1"
-by (simp add: setprod_def)
-
-lemma setprod_reindex:
- "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
-by(auto simp: setprod_def fold_image_reindex dest!:finite_imageD)
-
-lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)"
-by (auto simp add: setprod_reindex)
-
-lemma setprod_cong:
- "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
-by(fastsimp simp: setprod_def intro: fold_image_cong)
-
-lemma strong_setprod_cong[cong]:
- "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
-by(fastsimp simp: simp_implies_def setprod_def intro: fold_image_cong)
-
-lemma setprod_reindex_cong: "inj_on f A ==>
- B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
-by (frule setprod_reindex, simp)
-
-lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"
- and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
- shows "setprod h B = setprod g A"
-proof-
- have "setprod h B = setprod (h o f) A"
- by (simp add: B setprod_reindex[OF i, of h])
- then show ?thesis apply simp
- apply (rule setprod_cong)
- apply simp
- by (simp add: eq)
-qed
-
-lemma setprod_Un_one:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
- shows "setprod f (S \<union> T) = setprod f S * setprod f T"
- using fS fT
- apply (simp add: setprod_def)
- apply (rule fold_image_Un_one)
- using I0 by auto
-
-
-lemma setprod_1: "setprod (%i. 1) A = 1"
-apply (case_tac "finite A")
-apply (erule finite_induct, auto simp add: mult_ac)
-done
-
-lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1"
-apply (subgoal_tac "setprod f F = setprod (%x. 1) F")
-apply (erule ssubst, rule setprod_1)
-apply (rule setprod_cong, auto)
-done
-
-lemma setprod_Un_Int: "finite A ==> finite B
- ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
-by(simp add: setprod_def fold_image_Un_Int[symmetric])
-
-lemma setprod_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
-by (subst setprod_Un_Int [symmetric], auto)
-
-lemma setprod_mono_one_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 1"
- shows "setprod f S = setprod f T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- by (simp add: setprod_Un_disjoint[OF f d, unfolded eq[symmetric]] setprod_1'[OF z])
-qed
-
-lemmas setprod_mono_one_right = setprod_mono_one_left [THEN sym]
-
-lemma setprod_delta:
- assumes fS: "finite S"
- shows "setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else 1)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = 1" by simp
- hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) }
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto
- have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
- using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)}
- ultimately show ?thesis by blast
-qed
-
-lemma setprod_delta':
- assumes fS: "finite S" shows
- "setprod (\<lambda>k. if a = k then b k else 1) S =
- (if a\<in> S then b a else 1)"
- using setprod_delta[OF fS, of a b, symmetric]
- by (auto intro: setprod_cong)
-
-
-lemma setprod_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
- setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
-by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong)
-
-lemma setprod_Union_disjoint:
- "[| (ALL A:C. finite A);
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
- ==> setprod f (Union C) = setprod (setprod f) C"
-apply (cases "finite C")
- prefer 2 apply (force dest: finite_UnionD simp add: setprod_def)
- apply (frule setprod_UN_disjoint [of C id f])
- apply (unfold Union_def id_def, assumption+)
-done
-
-lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
- (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
- (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
-by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong)
-
-text{*Here we can eliminate the finiteness assumptions, by cases.*}
-lemma setprod_cartesian_product:
- "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
-apply (cases "finite A")
- apply (cases "finite B")
- apply (simp add: setprod_Sigma)
- apply (cases "A={}", simp)
- apply (simp add: setprod_1)
-apply (auto simp add: setprod_def
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma setprod_timesf:
- "setprod (%x. f x * g x) A = (setprod f A * setprod g A)"
-by(simp add:setprod_def fold_image_distrib)
-
-
-subsubsection {* Properties in more restricted classes of structures *}
-
-lemma setprod_eq_1_iff [simp]:
- "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"
-by (induct set: finite) auto
-
-lemma setprod_zero:
- "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"
-apply (induct set: finite, force, clarsimp)
-apply (erule disjE, auto)
-done
-
-lemma setprod_nonneg [rule_format]:
- "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
-by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)
-
-lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)
- --> 0 < setprod f A"
-by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)
-
-lemma setprod_zero_iff[simp]: "finite A ==>
- (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
- (EX x: A. f x = 0)"
-by (erule finite_induct, auto simp:no_zero_divisors)
-
-lemma setprod_pos_nat:
- "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
-using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
-
-lemma setprod_pos_nat_iff[simp]:
- "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"
-using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
-
-lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
- (setprod f (A Un B) :: 'a ::{field})
- = setprod f A * setprod f B / setprod f (A Int B)"
-by (subst setprod_Un_Int [symmetric], auto)
-
-lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
- (setprod f (A - {a}) :: 'a :: {field}) =
- (if a:A then setprod f A / f a else setprod f A)"
-by (erule finite_induct) (auto simp add: insert_Diff_if)
-
-lemma setprod_inversef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
- shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
-by (erule finite_induct) auto
-
-lemma setprod_dividef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
- shows "finite A
- ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
-apply (subgoal_tac
- "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
-apply (erule ssubst)
-apply (subst divide_inverse)
-apply (subst setprod_timesf)
-apply (subst setprod_inversef, assumption+, rule refl)
-apply (rule setprod_cong, rule refl)
-apply (subst divide_inverse, auto)
-done
-
-lemma setprod_dvd_setprod [rule_format]:
- "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"
- apply (cases "finite A")
- apply (induct set: finite)
- apply (auto simp add: dvd_def)
- apply (rule_tac x = "k * ka" in exI)
- apply (simp add: algebra_simps)
-done
-
-lemma setprod_dvd_setprod_subset:
- "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
- apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
- apply (unfold dvd_def, blast)
- apply (subst setprod_Un_disjoint [symmetric])
- apply (auto elim: finite_subset intro: setprod_cong)
-done
-
-lemma setprod_dvd_setprod_subset2:
- "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>
- setprod f A dvd setprod g B"
- apply (rule dvd_trans)
- apply (rule setprod_dvd_setprod, erule (1) bspec)
- apply (erule (1) setprod_dvd_setprod_subset)
-done
-
-lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>
- (f i ::'a::comm_semiring_1) dvd setprod f A"
-by (induct set: finite) (auto intro: dvd_mult)
-
-lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>
- (d::'a::comm_semiring_1) dvd (SUM x : A. f x)"
- apply (cases "finite A")
- apply (induct set: finite)
- apply auto
-done
-
-lemma setprod_mono:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"
- assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"
- shows "setprod f A \<le> setprod g A"
-proof (cases "finite A")
- case True
- hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]
- proof (induct A rule: finite_subset_induct)
- case (insert a F)
- thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
- unfolding setprod_insert[OF insert(1,3)]
- using assms[rule_format,OF insert(2)] insert
- by (auto intro: mult_mono mult_nonneg_nonneg)
- qed auto
- thus ?thesis by simp
-qed auto
-
-lemma abs_setprod:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
- shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"
-proof (cases "finite A")
- case True thus ?thesis
- by induct (auto simp add: field_simps abs_mult)
-qed auto
-
-
-subsection {* Finite cardinality *}
-
-text {* This definition, although traditional, is ugly to work with:
-@{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}.
-But now that we have @{text setsum} things are easy:
-*}
-
-definition card :: "'a set \<Rightarrow> nat" where
- "card A = setsum (\<lambda>x. 1) A"
-
-lemmas card_eq_setsum = card_def
-
-lemma card_empty [simp]: "card {} = 0"
- by (simp add: card_def)
-
-lemma card_insert_disjoint [simp]:
- "finite A ==> x \<notin> A ==> card (insert x A) = Suc(card A)"
- by (simp add: card_def)
-
-lemma card_insert_if:
- "finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))"
- by (simp add: insert_absorb)
-
-lemma card_infinite [simp]: "~ finite A ==> card A = 0"
- by (simp add: card_def)
-
-lemma card_ge_0_finite:
- "card A > 0 \<Longrightarrow> finite A"
- by (rule ccontr) simp
-
-lemma card_0_eq [simp,noatp]: "finite A ==> (card A = 0) = (A = {})"
- apply auto
- apply (drule_tac a = x in mk_disjoint_insert, clarify, auto)
- done
-
-lemma finite_UNIV_card_ge_0:
- "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
- by (rule ccontr) simp
-
-lemma card_eq_0_iff: "(card A = 0) = (A = {} | ~ finite A)"
- by auto
-
-lemma card_gt_0_iff: "(0 < card A) = (A \<noteq> {} & finite A)"
- by (simp add: neq0_conv [symmetric] card_eq_0_iff)
-
-lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
-apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
-apply(simp del:insert_Diff_single)
-done
-
-lemma card_Diff_singleton:
- "finite A ==> x: A ==> card (A - {x}) = card A - 1"
-by (simp add: card_Suc_Diff1 [symmetric])
-
-lemma card_Diff_singleton_if:
- "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
-by (simp add: card_Diff_singleton)
-
-lemma card_Diff_insert[simp]:
-assumes "finite A" and "a:A" and "a ~: B"
-shows "card(A - insert a B) = card(A - B) - 1"
-proof -
- have "A - insert a B = (A - B) - {a}" using assms by blast
- then show ?thesis using assms by(simp add:card_Diff_singleton)
-qed
-
-lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
-by (simp add: card_insert_if card_Suc_Diff1 del:card_Diff_insert)
-
-lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
-by (simp add: card_insert_if)
-
-lemma card_mono: "\<lbrakk> finite B; A \<subseteq> B \<rbrakk> \<Longrightarrow> card A \<le> card B"
-by (simp add: card_def setsum_mono2)
-
-lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
-apply (induct set: finite, simp, clarify)
-apply (subgoal_tac "finite A & A - {x} <= F")
- prefer 2 apply (blast intro: finite_subset, atomize)
-apply (drule_tac x = "A - {x}" in spec)
-apply (simp add: card_Diff_singleton_if split add: split_if_asm)
-apply (case_tac "card A", auto)
-done
-
-lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
-apply (simp add: psubset_eq linorder_not_le [symmetric])
-apply (blast dest: card_seteq)
-done
-
-lemma card_Un_Int: "finite A ==> finite B
- ==> card A + card B = card (A Un B) + card (A Int B)"
-by(simp add:card_def setsum_Un_Int)
-
-lemma card_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> card (A Un B) = card A + card B"
-by (simp add: card_Un_Int)
-
-lemma card_Diff_subset:
- "finite B ==> B <= A ==> card (A - B) = card A - card B"
-by(simp add:card_def setsum_diff_nat)
-
-lemma card_Diff_subset_Int:
- assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
-proof -
- have "A - B = A - A \<inter> B" by auto
- thus ?thesis
- by (simp add: card_Diff_subset AB)
-qed
-
-lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
-apply (rule Suc_less_SucD)
-apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
-done
-
-lemma card_Diff2_less:
- "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
-apply (case_tac "x = y")
- apply (simp add: card_Diff1_less del:card_Diff_insert)
-apply (rule less_trans)
- prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
-done
-
-lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
-apply (case_tac "x : A")
- apply (simp_all add: card_Diff1_less less_imp_le)
-done
-
-lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
-by (erule psubsetI, blast)
-
-lemma insert_partition:
- "\<lbrakk> x \<notin> F; \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<rbrakk>
- \<Longrightarrow> x \<inter> \<Union> F = {}"
-by auto
-
-lemma finite_psubset_induct[consumes 1, case_names psubset]:
- assumes "finite A" and "!!A. finite A \<Longrightarrow> (!!B. finite B \<Longrightarrow> B \<subset> A \<Longrightarrow> P(B)) \<Longrightarrow> P(A)" shows "P A"
-using assms(1)
-proof (induct A rule: measure_induct_rule[where f=card])
- case (less A)
- show ?case
- proof(rule assms(2)[OF less(2)])
- fix B assume "finite B" "B \<subset> A"
- show "P B" by(rule less(1)[OF psubset_card_mono[OF less(2) `B \<subset> A`] `finite B`])
- qed
-qed
-
-text{* main cardinality theorem *}
-lemma card_partition [rule_format]:
- "finite C ==>
- finite (\<Union> C) -->
- (\<forall>c\<in>C. card c = k) -->
- (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
- k * card(C) = card (\<Union> C)"
-apply (erule finite_induct, simp)
-apply (simp add: card_Un_disjoint insert_partition
- finite_subset [of _ "\<Union> (insert x F)"])
-done
-
-lemma card_eq_UNIV_imp_eq_UNIV:
- assumes fin: "finite (UNIV :: 'a set)"
- and card: "card A = card (UNIV :: 'a set)"
- shows "A = (UNIV :: 'a set)"
-proof
- show "A \<subseteq> UNIV" by simp
- show "UNIV \<subseteq> A"
- proof
- fix x
- show "x \<in> A"
- proof (rule ccontr)
- assume "x \<notin> A"
- then have "A \<subset> UNIV" by auto
- with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
- with card show False by simp
- qed
- qed
-qed
-
-text{*The form of a finite set of given cardinality*}
-
-lemma card_eq_SucD:
-assumes "card A = Suc k"
-shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
-proof -
- have fin: "finite A" using assms by (auto intro: ccontr)
- moreover have "card A \<noteq> 0" using assms by auto
- ultimately obtain b where b: "b \<in> A" by auto
- show ?thesis
- proof (intro exI conjI)
- show "A = insert b (A-{b})" using b by blast
- show "b \<notin> A - {b}" by blast
- show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
- using assms b fin by(fastsimp dest:mk_disjoint_insert)+
- qed
-qed
-
-lemma card_Suc_eq:
- "(card A = Suc k) =
- (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
-apply(rule iffI)
- apply(erule card_eq_SucD)
-apply(auto)
-apply(subst card_insert)
- apply(auto intro:ccontr)
-done
-
-lemma finite_fun_UNIVD2:
- assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
- shows "finite (UNIV :: 'b set)"
-proof -
- from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
- by(rule finite_imageI)
- moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
- by(rule UNIV_eq_I) auto
- ultimately show "finite (UNIV :: 'b set)" by simp
-qed
-
-lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"
-apply (cases "finite A")
-apply (erule finite_induct)
-apply (auto simp add: algebra_simps)
-done
-
-lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
-apply (erule finite_induct)
-apply auto
-done
-
-lemma setprod_gen_delta:
- assumes fS: "finite S"
- shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult}) * c^ (card S - 1) else c^ card S)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else c)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = c" by simp
- hence ?thesis using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) }
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
- apply (rule setprod_cong) by auto
- have cA: "card ?A = card S - 1" using fS a by auto
- have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto
- have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
- using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a cA
- by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
- ultimately show ?thesis by blast
-qed
-
-
-lemma setsum_bounded:
- assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
- shows "setsum f A \<le> of_nat(card A) * K"
-proof (cases "finite A")
- case True
- thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-
-lemma card_UNIV_unit: "card (UNIV :: unit set) = 1"
- unfolding UNIV_unit by simp
-
-
-subsubsection {* Cardinality of unions *}
-
-lemma card_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {})
- ==> card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
-apply (simp add: card_def del: setsum_constant)
-apply (subgoal_tac
- "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I")
-apply (simp add: setsum_UN_disjoint del: setsum_constant)
-apply (simp cong: setsum_cong)
-done
-
-lemma card_Union_disjoint:
- "finite C ==> (ALL A:C. finite A) ==>
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
- ==> card (Union C) = setsum card C"
-apply (frule card_UN_disjoint [of C id])
-apply (unfold Union_def id_def, assumption+)
-done
-
-
-subsubsection {* Cardinality of image *}
-
-text{*The image of a finite set can be expressed using @{term fold_image}.*}
-lemma image_eq_fold_image:
- "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A"
-proof (induct rule: finite_induct)
- case empty then show ?case by simp
-next
- interpret ab_semigroup_mult "op Un"
- proof qed auto
- case insert
- then show ?case by simp
-qed
-
-lemma card_image_le: "finite A ==> card (f ` A) <= card A"
-apply (induct set: finite)
- apply simp
-apply (simp add: le_SucI card_insert_if)
-done
-
-lemma card_image: "inj_on f A ==> card (f ` A) = card A"
-by(simp add:card_def setsum_reindex o_def del:setsum_constant)
-
-lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
-by(auto simp: card_image bij_betw_def)
-
-lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
-by (simp add: card_seteq card_image)
-
-lemma eq_card_imp_inj_on:
- "[| finite A; card(f ` A) = card A |] ==> inj_on f A"
-apply (induct rule:finite_induct)
-apply simp
-apply(frule card_image_le[where f = f])
-apply(simp add:card_insert_if split:if_splits)
-done
-
-lemma inj_on_iff_eq_card:
- "finite A ==> inj_on f A = (card(f ` A) = card A)"
-by(blast intro: card_image eq_card_imp_inj_on)
-
-
-lemma card_inj_on_le:
- "[|inj_on f A; f ` A \<subseteq> B; finite B |] ==> card A \<le> card B"
-apply (subgoal_tac "finite A")
- apply (force intro: card_mono simp add: card_image [symmetric])
-apply (blast intro: finite_imageD dest: finite_subset)
-done
-
-lemma card_bij_eq:
- "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
- finite A; finite B |] ==> card A = card B"
-by (auto intro: le_antisym card_inj_on_le)
-
-
-subsubsection {* Cardinality of products *}
-
-(*
-lemma SigmaI_insert: "y \<notin> A ==>
- (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
- by auto
-*)
-
-lemma card_SigmaI [simp]:
- "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
- \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
-by(simp add:card_def setsum_Sigma del:setsum_constant)
-
-lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"
-apply (cases "finite A")
-apply (cases "finite B")
-apply (auto simp add: card_eq_0_iff
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)"
-by (simp add: card_cartesian_product)
-
-
-subsubsection {* Cardinality of sums *}
-
-lemma card_Plus:
- assumes "finite A" and "finite B"
- shows "card (A <+> B) = card A + card B"
-proof -
- have "Inl`A \<inter> Inr`B = {}" by fast
- with assms show ?thesis
- unfolding Plus_def
- by (simp add: card_Un_disjoint card_image)
-qed
-
-lemma card_Plus_conv_if:
- "card (A <+> B) = (if finite A \<and> finite B then card(A) + card(B) else 0)"
-by(auto simp: card_def setsum_Plus simp del: setsum_constant)
-
-
-subsubsection {* Cardinality of the Powerset *}
-
-lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *)
-apply (induct set: finite)
- apply (simp_all add: Pow_insert)
-apply (subst card_Un_disjoint, blast)
- apply (blast intro: finite_imageI, blast)
-apply (subgoal_tac "inj_on (insert x) (Pow F)")
- apply (simp add: card_image Pow_insert)
-apply (unfold inj_on_def)
-apply (blast elim!: equalityE)
-done
-
-text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
-
-lemma dvd_partition:
- "finite (Union C) ==>
- ALL c : C. k dvd card c ==>
- (ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
- k dvd card (Union C)"
-apply(frule finite_UnionD)
-apply(rotate_tac -1)
-apply (induct set: finite, simp_all, clarify)
-apply (subst card_Un_disjoint)
- apply (auto simp add: disjoint_eq_subset_Compl)
-done
-
-
-subsubsection {* Relating injectivity and surjectivity *}
-
-lemma finite_surj_inj: "finite(A) \<Longrightarrow> A <= f`A \<Longrightarrow> inj_on f A"
-apply(rule eq_card_imp_inj_on, assumption)
-apply(frule finite_imageI)
-apply(drule (1) card_seteq)
- apply(erule card_image_le)
-apply simp
-done
-
-lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
-shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
-by (blast intro: finite_surj_inj subset_UNIV dest:surj_range)
-
-lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
-shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
-by(fastsimp simp:surj_def dest!: endo_inj_surj)
-
-corollary infinite_UNIV_nat[iff]: "~finite(UNIV::nat set)"
-proof
- assume "finite(UNIV::nat set)"
- with finite_UNIV_inj_surj[of Suc]
- show False by simp (blast dest: Suc_neq_Zero surjD)
-qed
-
-(* Often leads to bogus ATP proofs because of reduced type information, hence noatp *)
-lemma infinite_UNIV_char_0[noatp]:
- "\<not> finite (UNIV::'a::semiring_char_0 set)"
-proof
- assume "finite (UNIV::'a set)"
- with subset_UNIV have "finite (range of_nat::'a set)"
- by (rule finite_subset)
- moreover have "inj (of_nat::nat \<Rightarrow> 'a)"
- by (simp add: inj_on_def)
- ultimately have "finite (UNIV::nat set)"
- by (rule finite_imageD)
- then show "False"
- by simp
-qed
subsection{* A fold functional for non-empty sets *}
@@ -2811,561 +1354,6 @@
qed
-subsubsection {* Fold1 in lattices with @{const inf} and @{const sup} *}
-
-text{*
- As an application of @{text fold1} we define infimum
- and supremum in (not necessarily complete!) lattices
- over (non-empty) sets by means of @{text fold1}.
-*}
-
-context semilattice_inf
-begin
-
-lemma below_fold1_iff:
- assumes "finite A" "A \<noteq> {}"
- shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
-qed
-
-lemma fold1_belowI:
- assumes "finite A"
- and "a \<in> A"
- shows "fold1 inf A \<le> a"
-proof -
- from assms have "A \<noteq> {}" by auto
- from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
- proof (induct rule: finite_ne_induct)
- case singleton thus ?case by simp
- next
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- case (insert x F)
- from insert(5) have "a = x \<or> a \<in> F" by simp
- thus ?case
- proof
- assume "a = x" thus ?thesis using insert
- by (simp add: mult_ac)
- next
- assume "a \<in> F"
- hence bel: "fold1 inf F \<le> a" by (rule insert)
- have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"
- using insert by (simp add: mult_ac)
- also have "inf (fold1 inf F) a = fold1 inf F"
- using bel by (auto intro: antisym)
- also have "inf x \<dots> = fold1 inf (insert x F)"
- using insert by (simp add: mult_ac)
- finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .
- moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp
- ultimately show ?thesis by simp
- qed
- qed
-qed
-
-end
-
-context lattice
-begin
-
-definition
- Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900)
-where
- "Inf_fin = fold1 inf"
-
-definition
- Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900)
-where
- "Sup_fin = fold1 sup"
-
-lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A"
-apply(unfold Sup_fin_def Inf_fin_def)
-apply(subgoal_tac "EX a. a:A")
-prefer 2 apply blast
-apply(erule exE)
-apply(rule order_trans)
-apply(erule (1) fold1_belowI)
-apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice])
-done
-
-lemma sup_Inf_absorb [simp]:
- "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a"
-apply(subst sup_commute)
-apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI)
-done
-
-lemma inf_Sup_absorb [simp]:
- "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"
-by (simp add: Sup_fin_def inf_absorb1
- semilattice_inf.fold1_belowI [OF dual_semilattice])
-
-end
-
-context distrib_lattice
-begin
-
-lemma sup_Inf1_distrib:
- assumes "finite A"
- and "A \<noteq> {}"
- shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- from assms show ?thesis
- by (simp add: Inf_fin_def image_def
- hom_fold1_commute [where h="sup x", OF sup_inf_distrib1])
- (rule arg_cong [where f="fold1 inf"], blast)
-qed
-
-lemma sup_Inf2_distrib:
- assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
- shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}"
-using A proof (induct rule: finite_ne_induct)
- case singleton thus ?case
- by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
-next
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- case (insert x A)
- have finB: "finite {sup x b |b. b \<in> B}"
- by(rule finite_surj[where f = "sup x", OF B(1)], auto)
- have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
- proof -
- have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
- by blast
- thus ?thesis by(simp add: insert(1) B(1))
- qed
- have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
- have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)"
- using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def])
- also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2)
- also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})"
- using insert by(simp add:sup_Inf1_distrib[OF B])
- also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
- (is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M")
- using B insert
- by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne])
- also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
- by blast
- finally show ?case .
-qed
-
-lemma inf_Sup1_distrib:
- assumes "finite A" and "A \<noteq> {}"
- shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
-proof -
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- from assms show ?thesis
- by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
- (rule arg_cong [where f="fold1 sup"], blast)
-qed
-
-lemma inf_Sup2_distrib:
- assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
- shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}"
-using A proof (induct rule: finite_ne_induct)
- case singleton thus ?case
- by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def])
-next
- case (insert x A)
- have finB: "finite {inf x b |b. b \<in> B}"
- by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)
- have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
- proof -
- have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
- by blast
- thus ?thesis by(simp add: insert(1) B(1))
- qed
- have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
- using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
- also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2)
- also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})"
- using insert by(simp add:inf_Sup1_distrib[OF B])
- also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
- (is "_ = \<Squnion>\<^bsub>fin\<^esub>?M")
- using B insert
- by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne])
- also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
- by blast
- finally show ?case .
-qed
-
-end
-
-
-subsubsection {* Fold1 in linear orders with @{const min} and @{const max} *}
-
-text{*
- As an application of @{text fold1} we define minimum
- and maximum in (not necessarily complete!) linear orders
- over (non-empty) sets by means of @{text fold1}.
-*}
-
-context linorder
-begin
-
-lemma ab_semigroup_idem_mult_min:
- "ab_semigroup_idem_mult min"
- proof qed (auto simp add: min_def)
-
-lemma ab_semigroup_idem_mult_max:
- "ab_semigroup_idem_mult max"
- proof qed (auto simp add: max_def)
-
-lemma max_lattice:
- "semilattice_inf (op \<ge>) (op >) max"
- by (fact min_max.dual_semilattice)
-
-lemma dual_max:
- "ord.max (op \<ge>) = min"
- by (auto simp add: ord.max_def_raw min_def expand_fun_eq)
-
-lemma dual_min:
- "ord.min (op \<ge>) = max"
- by (auto simp add: ord.min_def_raw max_def expand_fun_eq)
-
-lemma strict_below_fold1_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert)
-qed
-
-lemma fold1_below_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert min_le_iff_disj)
-qed
-
-lemma fold1_strict_below_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert min_less_iff_disj)
-qed
-
-lemma fold1_antimono:
- assumes "A \<noteq> {}" and "A \<subseteq> B" and "finite B"
- shows "fold1 min B \<le> fold1 min A"
-proof cases
- assume "A = B" thus ?thesis by simp
-next
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- assume "A \<noteq> B"
- have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
- have "fold1 min B = fold1 min (A \<union> (B-A))" by(subst B)(rule refl)
- also have "\<dots> = min (fold1 min A) (fold1 min (B-A))"
- proof -
- have "finite A" by(rule finite_subset[OF `A \<subseteq> B` `finite B`])
- moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *)
- moreover have "(B-A) \<noteq> {}" using prems by blast
- moreover have "A Int (B-A) = {}" using prems by blast
- ultimately show ?thesis using `A \<noteq> {}` by (rule_tac fold1_Un)
- qed
- also have "\<dots> \<le> fold1 min A" by (simp add: min_le_iff_disj)
- finally show ?thesis .
-qed
-
-definition
- Min :: "'a set \<Rightarrow> 'a"
-where
- "Min = fold1 min"
-
-definition
- Max :: "'a set \<Rightarrow> 'a"
-where
- "Max = fold1 max"
-
-lemmas Min_singleton [simp] = fold1_singleton_def [OF Min_def]
-lemmas Max_singleton [simp] = fold1_singleton_def [OF Max_def]
-
-lemma Min_insert [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min (insert x A) = min x (Min A)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
-qed
-
-lemma Max_insert [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max (insert x A) = max x (Max A)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
-qed
-
-lemma Min_in [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A \<in> A"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
-qed
-
-lemma Max_in [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A \<in> A"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
-qed
-
-lemma Min_Un:
- assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
- shows "Min (A \<union> B) = min (Min A) (Min B)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (simp add: Min_def fold1_Un2)
-qed
-
-lemma Max_Un:
- assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
- shows "Max (A \<union> B) = max (Max A) (Max B)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis
- by (simp add: Max_def fold1_Un2)
-qed
-
-lemma hom_Min_commute:
- assumes "\<And>x y. h (min x y) = min (h x) (h y)"
- and "finite N" and "N \<noteq> {}"
- shows "h (Min N) = Min (h ` N)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (simp add: Min_def hom_fold1_commute)
-qed
-
-lemma hom_Max_commute:
- assumes "\<And>x y. h (max x y) = max (h x) (h y)"
- and "finite N" and "N \<noteq> {}"
- shows "h (Max N) = Max (h ` N)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis
- by (simp add: Max_def hom_fold1_commute [of h])
-qed
-
-lemma Min_le [simp]:
- assumes "finite A" and "x \<in> A"
- shows "Min A \<le> x"
- using assms by (simp add: Min_def min_max.fold1_belowI)
-
-lemma Max_ge [simp]:
- assumes "finite A" and "x \<in> A"
- shows "x \<le> Max A"
-proof -
- interpret semilattice_inf "op \<ge>" "op >" max
- by (rule max_lattice)
- from assms show ?thesis by (simp add: Max_def fold1_belowI)
-qed
-
-lemma Min_ge_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
- using assms by (simp add: Min_def min_max.below_fold1_iff)
-
-lemma Max_le_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
-proof -
- interpret semilattice_inf "op \<ge>" "op >" max
- by (rule max_lattice)
- from assms show ?thesis by (simp add: Max_def below_fold1_iff)
-qed
-
-lemma Min_gr_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
- using assms by (simp add: Min_def strict_below_fold1_iff)
-
-lemma Max_less_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.strict_below_fold1_iff [folded dual.dual_max])
-qed
-
-lemma Min_le_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
- using assms by (simp add: Min_def fold1_below_iff)
-
-lemma Max_ge_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_below_iff [folded dual.dual_max])
-qed
-
-lemma Min_less_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
- using assms by (simp add: Min_def fold1_strict_below_iff)
-
-lemma Max_gr_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_strict_below_iff [folded dual.dual_max])
-qed
-
-lemma Min_eqI:
- assumes "finite A"
- assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"
- and "x \<in> A"
- shows "Min A = x"
-proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
- with assms show "Min A \<ge> x" by simp
-next
- from assms show "x \<ge> Min A" by simp
-qed
-
-lemma Max_eqI:
- assumes "finite A"
- assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
- and "x \<in> A"
- shows "Max A = x"
-proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
- with assms show "Max A \<le> x" by simp
-next
- from assms show "x \<le> Max A" by simp
-qed
-
-lemma Min_antimono:
- assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
- shows "Min N \<le> Min M"
- using assms by (simp add: Min_def fold1_antimono)
-
-lemma Max_mono:
- assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
- shows "Max M \<le> Max N"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_antimono [folded dual.dual_max])
-qed
-
-lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:
- "finite A \<Longrightarrow> P {} \<Longrightarrow>
- (!!b A. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A))
- \<Longrightarrow> P A"
-proof (induct rule: finite_psubset_induct)
- fix A :: "'a set"
- assume IH: "!! B. finite B \<Longrightarrow> B < A \<Longrightarrow> P {} \<Longrightarrow>
- (!!b A. finite A \<Longrightarrow> (\<forall>a\<in>A. a<b) \<Longrightarrow> P A \<Longrightarrow> P (insert b A))
- \<Longrightarrow> P B"
- and "finite A" and "P {}"
- and step: "!!b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)"
- show "P A"
- proof (cases "A = {}")
- assume "A = {}" thus "P A" using `P {}` by simp
- next
- let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B"
- assume "A \<noteq> {}"
- with `finite A` have "Max A : A" by auto
- hence A: "?A = A" using insert_Diff_single insert_absorb by auto
- moreover have "finite ?B" using `finite A` by simp
- ultimately have "P ?B" using `P {}` step IH[of ?B] by blast
- moreover have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp
- ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp
- qed
-qed
-
-lemma finite_linorder_min_induct[consumes 1, case_names empty insert]:
- "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
-by(rule linorder.finite_linorder_max_induct[OF dual_linorder])
-
-end
-
-context linordered_ab_semigroup_add
-begin
-
-lemma add_Min_commute:
- fixes k
- assumes "finite N" and "N \<noteq> {}"
- shows "k + Min N = Min {k + m | m. m \<in> N}"
-proof -
- have "\<And>x y. k + min x y = min (k + x) (k + y)"
- by (simp add: min_def not_le)
- (blast intro: antisym less_imp_le add_left_mono)
- with assms show ?thesis
- using hom_Min_commute [of "plus k" N]
- by simp (blast intro: arg_cong [where f = Min])
-qed
-
-lemma add_Max_commute:
- fixes k
- assumes "finite N" and "N \<noteq> {}"
- shows "k + Max N = Max {k + m | m. m \<in> N}"
-proof -
- have "\<And>x y. k + max x y = max (k + x) (k + y)"
- by (simp add: max_def not_le)
- (blast intro: antisym less_imp_le add_left_mono)
- with assms show ?thesis
- using hom_Max_commute [of "plus k" N]
- by simp (blast intro: arg_cong [where f = Max])
-qed
-
-end
-
-context linordered_ab_group_add
-begin
-
-lemma minus_Max_eq_Min [simp]:
- "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Max S) = Min (uminus ` S)"
- by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)
-
-lemma minus_Min_eq_Max [simp]:
- "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Min S) = Max (uminus ` S)"
- by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)
-
-end
-
-
subsection {* Expressing set operations via @{const fold} *}
lemma (in fun_left_comm) fun_left_comm_apply:
@@ -3445,32 +1433,6 @@
shows "Sup A = fold sup bot A"
using assms sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
-lemma Inf_fin_Inf:
- assumes "finite A" and "A \<noteq> {}"
- shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
- moreover with `finite A` have "finite B" by simp
- ultimately show ?thesis
- by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
- (simp add: Inf_fold_inf)
-qed
-
-lemma Sup_fin_Sup:
- assumes "finite A" and "A \<noteq> {}"
- shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
-proof -
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
- moreover with `finite A` have "finite B" by simp
- ultimately show ?thesis
- by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])
- (simp add: Sup_fold_sup)
-qed
-
lemma inf_INFI_fold_inf:
assumes "finite A"
shows "inf B (INFI A f) = fold (\<lambda>A. inf (f A)) B A" (is "?inf = ?fold")
@@ -3505,4 +1467,661 @@
end
+
+subsection {* Locales as mini-packages *}
+
+locale folding =
+ fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ fixes F :: "'a set \<Rightarrow> 'b \<Rightarrow> 'b"
+ assumes commute_comp: "f x \<circ> f y = f y \<circ> f x"
+ assumes eq_fold: "finite A \<Longrightarrow> F A s = fold f s A"
+begin
+
+lemma fun_left_commute:
+ "f x (f y s) = f y (f x s)"
+ using commute_comp [of x y] by (simp add: expand_fun_eq)
+
+lemma fun_left_comm:
+ "fun_left_comm f"
+proof
+qed (fact fun_left_commute)
+
+lemma empty [simp]:
+ "F {} = id"
+ by (simp add: eq_fold expand_fun_eq)
+
+lemma insert [simp]:
+ assumes "finite A" and "x \<notin> A"
+ shows "F (insert x A) = F A \<circ> f x"
+proof -
+ interpret fun_left_comm f by (fact fun_left_comm)
+ from fold_insert2 assms
+ have "\<And>s. fold f s (insert x A) = fold f (f x s) A" .
+ with `finite A` show ?thesis by (simp add: eq_fold expand_fun_eq)
+qed
+
+lemma remove:
+ assumes "finite A" and "x \<in> A"
+ shows "F A = F (A - {x}) \<circ> f x"
+proof -
+ from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ by (auto dest: mk_disjoint_insert)
+ moreover from `finite A` this have "finite B" by simp
+ ultimately show ?thesis by simp
+qed
+
+lemma insert_remove:
+ assumes "finite A"
+ shows "F (insert x A) = F (A - {x}) \<circ> f x"
+ using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
+
+lemma commute_comp':
+ assumes "finite A"
+ shows "f x \<circ> F A = F A \<circ> f x"
+proof (rule ext)
+ fix s
+ from assms show "(f x \<circ> F A) s = (F A \<circ> f x) s"
+ by (induct A arbitrary: s) (simp_all add: fun_left_commute)
+qed
+
+lemma fun_left_commute':
+ assumes "finite A"
+ shows "f x (F A s) = F A (f x s)"
+ using commute_comp' assms by (simp add: expand_fun_eq)
+
+lemma union:
+ assumes "finite A" and "finite B"
+ and "A \<inter> B = {}"
+ shows "F (A \<union> B) = F A \<circ> F B"
+using `finite A` `A \<inter> B = {}` proof (induct A)
+ case empty show ?case by simp
+next
+ case (insert x A)
+ then have "A \<inter> B = {}" by auto
+ with insert(3) have "F (A \<union> B) = F A \<circ> F B" .
+ moreover from insert have "x \<notin> B" by simp
+ moreover from `finite A` `finite B` have fin: "finite (A \<union> B)" by simp
+ moreover from `x \<notin> A` `x \<notin> B` have "x \<notin> A \<union> B" by simp
+ ultimately show ?case by (simp add: fun_left_commute')
+qed
+
end
+
+locale folding_idem = folding +
+ assumes idem_comp: "f x \<circ> f x = f x"
+begin
+
+declare insert [simp del]
+
+lemma fun_idem:
+ "f x (f x s) = f x s"
+ using idem_comp [of x] by (simp add: expand_fun_eq)
+
+lemma fun_left_comm_idem:
+ "fun_left_comm_idem f"
+proof
+qed (fact fun_left_commute fun_idem)+
+
+lemma insert_idem [simp]:
+ assumes "finite A"
+ shows "F (insert x A) = F A \<circ> f x"
+proof -
+ interpret fun_left_comm_idem f by (fact fun_left_comm_idem)
+ from fold_insert_idem2 assms
+ have "\<And>s. fold f s (insert x A) = fold f (f x s) A" .
+ with assms show ?thesis by (simp add: eq_fold expand_fun_eq)
+qed
+
+lemma union_idem:
+ assumes "finite A" and "finite B"
+ shows "F (A \<union> B) = F A \<circ> F B"
+using `finite A` proof (induct A)
+ case empty show ?case by simp
+next
+ case (insert x A)
+ from insert(3) have "F (A \<union> B) = F A \<circ> F B" .
+ moreover from `finite A` `finite B` have fin: "finite (A \<union> B)" by simp
+ ultimately show ?case by (simp add: fun_left_commute')
+qed
+
+end
+
+no_notation (in times) times (infixl "*" 70)
+no_notation (in one) Groups.one ("1")
+
+locale folding_image_simple = comm_monoid +
+ fixes g :: "('b \<Rightarrow> 'a)"
+ fixes F :: "'b set \<Rightarrow> 'a"
+ assumes eq_fold: "finite A \<Longrightarrow> F A = fold_image f g 1 A"
+begin
+
+lemma empty [simp]:
+ "F {} = 1"
+ by (simp add: eq_fold)
+
+lemma insert [simp]:
+ assumes "finite A" and "x \<notin> A"
+ shows "F (insert x A) = g x * F A"
+proof -
+ interpret fun_left_comm "%x y. (g x) * y" proof
+ qed (simp add: ac_simps)
+ with assms have "fold_image (op *) g 1 (insert x A) = g x * fold_image (op *) g 1 A"
+ by (simp add: fold_image_def)
+ with `finite A` show ?thesis by (simp add: eq_fold)
+qed
+
+lemma remove:
+ assumes "finite A" and "x \<in> A"
+ shows "F A = g x * F (A - {x})"
+proof -
+ from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ by (auto dest: mk_disjoint_insert)
+ moreover from `finite A` this have "finite B" by simp
+ ultimately show ?thesis by simp
+qed
+
+lemma insert_remove:
+ assumes "finite A"
+ shows "F (insert x A) = g x * F (A - {x})"
+ using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
+
+lemma union_inter:
+ assumes "finite A" and "finite B"
+ shows "F A * F B = F (A \<union> B) * F (A \<inter> B)"
+using assms proof (induct A)
+ case empty then show ?case by simp
+next
+ case (insert x A) then show ?case
+ by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute)
+qed
+
+corollary union_disjoint:
+ assumes "finite A" and "finite B"
+ assumes "A \<inter> B = {}"
+ shows "F (A \<union> B) = F A * F B"
+ using assms by (simp add: union_inter)
+
+end
+
+locale folding_image = comm_monoid +
+ fixes F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
+ assumes eq_fold: "\<And>g. finite A \<Longrightarrow> F g A = fold_image f g 1 A"
+
+sublocale folding_image < folding_image_simple "op *" 1 g "F g" proof
+qed (fact eq_fold)
+
+context folding_image
+begin
+
+lemma reindex:
+ assumes "finite A" and "inj_on h A"
+ shows "F g (h ` A) = F (g \<circ> h) A"
+ using assms by (induct A) auto
+
+lemma cong:
+ assumes "finite A" and "\<And>x. x \<in> A \<Longrightarrow> g x = h x"
+ shows "F g A = F h A"
+proof -
+ from assms have "ALL C. C <= A --> (ALL x:C. g x = h x) --> F g C = F h C"
+ apply - apply (erule finite_induct) apply simp
+ apply (simp add: subset_insert_iff, clarify)
+ apply (subgoal_tac "finite C")
+ prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl])
+ apply (subgoal_tac "C = insert x (C - {x})")
+ prefer 2 apply blast
+ apply (erule ssubst)
+ apply (drule spec)
+ apply (erule (1) notE impE)
+ apply (simp add: Ball_def del: insert_Diff_single)
+ done
+ with assms show ?thesis by simp
+qed
+
+lemma UNION_disjoint:
+ assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
+ and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
+ shows "F g (UNION I A) = F (F g \<circ> A) I"
+apply (insert assms)
+apply (induct set: finite, simp, atomize)
+apply (subgoal_tac "\<forall>i\<in>Fa. x \<noteq> i")
+ prefer 2 apply blast
+apply (subgoal_tac "A x Int UNION Fa A = {}")
+ prefer 2 apply blast
+apply (simp add: union_disjoint)
+done
+
+lemma distrib:
+ assumes "finite A"
+ shows "F (\<lambda>x. g x * h x) A = F g A * F h A"
+ using assms by (rule finite_induct) (simp_all add: assoc commute left_commute)
+
+lemma related:
+ assumes Re: "R 1 1"
+ and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 * y1) (x2 * y2)"
+ and fS: "finite S" and Rfg: "\<forall>x\<in>S. R (h x) (g x)"
+ shows "R (F h S) (F g S)"
+ using fS by (rule finite_subset_induct) (insert assms, auto)
+
+lemma eq_general:
+ assumes fS: "finite S"
+ and h: "\<forall>y\<in>S'. \<exists>!x. x \<in> S \<and> h x = y"
+ and f12: "\<forall>x\<in>S. h x \<in> S' \<and> f2 (h x) = f1 x"
+ shows "F f1 S = F f2 S'"
+proof-
+ from h f12 have hS: "h ` S = S'" by blast
+ {fix x y assume H: "x \<in> S" "y \<in> S" "h x = h y"
+ from f12 h H have "x = y" by auto }
+ hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast
+ from f12 have th: "\<And>x. x \<in> S \<Longrightarrow> (f2 \<circ> h) x = f1 x" by auto
+ from hS have "F f2 S' = F f2 (h ` S)" by simp
+ also have "\<dots> = F (f2 o h) S" using reindex [OF fS hinj, of f2] .
+ also have "\<dots> = F f1 S " using th cong [OF fS, of "f2 o h" f1]
+ by blast
+ finally show ?thesis ..
+qed
+
+lemma eq_general_inverses:
+ assumes fS: "finite S"
+ and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
+ and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = j x"
+ shows "F j S = F g T"
+ (* metis solves it, but not yet available here *)
+ apply (rule eq_general [OF fS, of T h g j])
+ apply (rule ballI)
+ apply (frule kh)
+ apply (rule ex1I[])
+ apply blast
+ apply clarsimp
+ apply (drule hk) apply simp
+ apply (rule sym)
+ apply (erule conjunct1[OF conjunct2[OF hk]])
+ apply (rule ballI)
+ apply (drule hk)
+ apply blast
+ done
+
+end
+
+notation (in times) times (infixl "*" 70)
+notation (in one) Groups.one ("1")
+
+
+subsection {* Finite cardinality *}
+
+text {* This definition, although traditional, is ugly to work with:
+@{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}.
+But now that we have @{text fold_image} things are easy:
+*}
+
+definition card :: "'a set \<Rightarrow> nat" where
+ "card A = (if finite A then fold_image (op +) (\<lambda>x. 1) 0 A else 0)"
+
+interpretation card!: folding_image_simple "op +" 0 "\<lambda>x. 1" card proof
+qed (simp add: card_def)
+
+lemma card_infinite [simp]:
+ "\<not> finite A \<Longrightarrow> card A = 0"
+ by (simp add: card_def)
+
+lemma card_empty:
+ "card {} = 0"
+ by (fact card.empty)
+
+lemma card_insert_disjoint:
+ "finite A ==> x \<notin> A ==> card (insert x A) = Suc (card A)"
+ by simp
+
+lemma card_insert_if:
+ "finite A ==> card (insert x A) = (if x \<in> A then card A else Suc (card A))"
+ by auto (simp add: card.insert_remove card.remove)
+
+lemma card_ge_0_finite:
+ "card A > 0 \<Longrightarrow> finite A"
+ by (rule ccontr) simp
+
+lemma card_0_eq [simp, noatp]:
+ "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
+ by (auto dest: mk_disjoint_insert)
+
+lemma finite_UNIV_card_ge_0:
+ "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
+ by (rule ccontr) simp
+
+lemma card_eq_0_iff:
+ "card A = 0 \<longleftrightarrow> A = {} \<or> \<not> finite A"
+ by auto
+
+lemma card_gt_0_iff:
+ "0 < card A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
+ by (simp add: neq0_conv [symmetric] card_eq_0_iff)
+
+lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
+apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
+apply(simp del:insert_Diff_single)
+done
+
+lemma card_Diff_singleton:
+ "finite A ==> x: A ==> card (A - {x}) = card A - 1"
+by (simp add: card_Suc_Diff1 [symmetric])
+
+lemma card_Diff_singleton_if:
+ "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
+by (simp add: card_Diff_singleton)
+
+lemma card_Diff_insert[simp]:
+assumes "finite A" and "a:A" and "a ~: B"
+shows "card(A - insert a B) = card(A - B) - 1"
+proof -
+ have "A - insert a B = (A - B) - {a}" using assms by blast
+ then show ?thesis using assms by(simp add:card_Diff_singleton)
+qed
+
+lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
+by (simp add: card_insert_if card_Suc_Diff1 del:card_Diff_insert)
+
+lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
+by (simp add: card_insert_if)
+
+lemma card_mono:
+ assumes "finite B" and "A \<subseteq> B"
+ shows "card A \<le> card B"
+proof -
+ from assms have "finite A" by (auto intro: finite_subset)
+ then show ?thesis using assms proof (induct A arbitrary: B)
+ case empty then show ?case by simp
+ next
+ case (insert x A)
+ then have "x \<in> B" by simp
+ from insert have "A \<subseteq> B - {x}" and "finite (B - {x})" by auto
+ with insert.hyps have "card A \<le> card (B - {x})" by auto
+ with `finite A` `x \<notin> A` `finite B` `x \<in> B` show ?case by simp (simp only: card.remove)
+ qed
+qed
+
+lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
+apply (induct set: finite, simp, clarify)
+apply (subgoal_tac "finite A & A - {x} <= F")
+ prefer 2 apply (blast intro: finite_subset, atomize)
+apply (drule_tac x = "A - {x}" in spec)
+apply (simp add: card_Diff_singleton_if split add: split_if_asm)
+apply (case_tac "card A", auto)
+done
+
+lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
+apply (simp add: psubset_eq linorder_not_le [symmetric])
+apply (blast dest: card_seteq)
+done
+
+lemma card_Un_Int: "finite A ==> finite B
+ ==> card A + card B = card (A Un B) + card (A Int B)"
+ by (fact card.union_inter)
+
+lemma card_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> card (A Un B) = card A + card B"
+ by (fact card.union_disjoint)
+
+lemma card_Diff_subset:
+ assumes "finite B" and "B \<subseteq> A"
+ shows "card (A - B) = card A - card B"
+proof (cases "finite A")
+ case False with assms show ?thesis by simp
+next
+ case True with assms show ?thesis by (induct B arbitrary: A) simp_all
+qed
+
+lemma card_Diff_subset_Int:
+ assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
+proof -
+ have "A - B = A - A \<inter> B" by auto
+ thus ?thesis
+ by (simp add: card_Diff_subset AB)
+qed
+
+lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
+apply (rule Suc_less_SucD)
+apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
+done
+
+lemma card_Diff2_less:
+ "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
+apply (case_tac "x = y")
+ apply (simp add: card_Diff1_less del:card_Diff_insert)
+apply (rule less_trans)
+ prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
+done
+
+lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
+apply (case_tac "x : A")
+ apply (simp_all add: card_Diff1_less less_imp_le)
+done
+
+lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
+by (erule psubsetI, blast)
+
+lemma insert_partition:
+ "\<lbrakk> x \<notin> F; \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<rbrakk>
+ \<Longrightarrow> x \<inter> \<Union> F = {}"
+by auto
+
+lemma finite_psubset_induct[consumes 1, case_names psubset]:
+ assumes "finite A" and "!!A. finite A \<Longrightarrow> (!!B. finite B \<Longrightarrow> B \<subset> A \<Longrightarrow> P(B)) \<Longrightarrow> P(A)" shows "P A"
+using assms(1)
+proof (induct A rule: measure_induct_rule[where f=card])
+ case (less A)
+ show ?case
+ proof(rule assms(2)[OF less(2)])
+ fix B assume "finite B" "B \<subset> A"
+ show "P B" by(rule less(1)[OF psubset_card_mono[OF less(2) `B \<subset> A`] `finite B`])
+ qed
+qed
+
+text{* main cardinality theorem *}
+lemma card_partition [rule_format]:
+ "finite C ==>
+ finite (\<Union> C) -->
+ (\<forall>c\<in>C. card c = k) -->
+ (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
+ k * card(C) = card (\<Union> C)"
+apply (erule finite_induct, simp)
+apply (simp add: card_Un_disjoint insert_partition
+ finite_subset [of _ "\<Union> (insert x F)"])
+done
+
+lemma card_eq_UNIV_imp_eq_UNIV:
+ assumes fin: "finite (UNIV :: 'a set)"
+ and card: "card A = card (UNIV :: 'a set)"
+ shows "A = (UNIV :: 'a set)"
+proof
+ show "A \<subseteq> UNIV" by simp
+ show "UNIV \<subseteq> A"
+ proof
+ fix x
+ show "x \<in> A"
+ proof (rule ccontr)
+ assume "x \<notin> A"
+ then have "A \<subset> UNIV" by auto
+ with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
+ with card show False by simp
+ qed
+ qed
+qed
+
+text{*The form of a finite set of given cardinality*}
+
+lemma card_eq_SucD:
+assumes "card A = Suc k"
+shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
+proof -
+ have fin: "finite A" using assms by (auto intro: ccontr)
+ moreover have "card A \<noteq> 0" using assms by auto
+ ultimately obtain b where b: "b \<in> A" by auto
+ show ?thesis
+ proof (intro exI conjI)
+ show "A = insert b (A-{b})" using b by blast
+ show "b \<notin> A - {b}" by blast
+ show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
+ using assms b fin by(fastsimp dest:mk_disjoint_insert)+
+ qed
+qed
+
+lemma card_Suc_eq:
+ "(card A = Suc k) =
+ (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
+apply(rule iffI)
+ apply(erule card_eq_SucD)
+apply(auto)
+apply(subst card_insert)
+ apply(auto intro:ccontr)
+done
+
+lemma finite_fun_UNIVD2:
+ assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
+ shows "finite (UNIV :: 'b set)"
+proof -
+ from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
+ by(rule finite_imageI)
+ moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
+ by(rule UNIV_eq_I) auto
+ ultimately show "finite (UNIV :: 'b set)" by simp
+qed
+
+lemma card_UNIV_unit: "card (UNIV :: unit set) = 1"
+ unfolding UNIV_unit by simp
+
+
+subsubsection {* Cardinality of image *}
+
+lemma card_image_le: "finite A ==> card (f ` A) <= card A"
+apply (induct set: finite)
+ apply simp
+apply (simp add: le_SucI card_insert_if)
+done
+
+lemma card_image:
+ assumes "inj_on f A"
+ shows "card (f ` A) = card A"
+proof (cases "finite A")
+ case True then show ?thesis using assms by (induct A) simp_all
+next
+ case False then have "\<not> finite (f ` A)" using assms by (auto dest: finite_imageD)
+ with False show ?thesis by simp
+qed
+
+lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
+by(auto simp: card_image bij_betw_def)
+
+lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
+by (simp add: card_seteq card_image)
+
+lemma eq_card_imp_inj_on:
+ "[| finite A; card(f ` A) = card A |] ==> inj_on f A"
+apply (induct rule:finite_induct)
+apply simp
+apply(frule card_image_le[where f = f])
+apply(simp add:card_insert_if split:if_splits)
+done
+
+lemma inj_on_iff_eq_card:
+ "finite A ==> inj_on f A = (card(f ` A) = card A)"
+by(blast intro: card_image eq_card_imp_inj_on)
+
+
+lemma card_inj_on_le:
+ "[|inj_on f A; f ` A \<subseteq> B; finite B |] ==> card A \<le> card B"
+apply (subgoal_tac "finite A")
+ apply (force intro: card_mono simp add: card_image [symmetric])
+apply (blast intro: finite_imageD dest: finite_subset)
+done
+
+lemma card_bij_eq:
+ "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
+ finite A; finite B |] ==> card A = card B"
+by (auto intro: le_antisym card_inj_on_le)
+
+
+subsubsection {* Cardinality of sums *}
+
+lemma card_Plus:
+ assumes "finite A" and "finite B"
+ shows "card (A <+> B) = card A + card B"
+proof -
+ have "Inl`A \<inter> Inr`B = {}" by fast
+ with assms show ?thesis
+ unfolding Plus_def
+ by (simp add: card_Un_disjoint card_image)
+qed
+
+lemma card_Plus_conv_if:
+ "card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
+ by (auto simp add: card_Plus)
+
+
+subsubsection {* Cardinality of the Powerset *}
+
+lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *)
+apply (induct set: finite)
+ apply (simp_all add: Pow_insert)
+apply (subst card_Un_disjoint, blast)
+ apply (blast intro: finite_imageI, blast)
+apply (subgoal_tac "inj_on (insert x) (Pow F)")
+ apply (simp add: card_image Pow_insert)
+apply (unfold inj_on_def)
+apply (blast elim!: equalityE)
+done
+
+text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
+
+lemma dvd_partition:
+ "finite (Union C) ==>
+ ALL c : C. k dvd card c ==>
+ (ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
+ k dvd card (Union C)"
+apply(frule finite_UnionD)
+apply(rotate_tac -1)
+apply (induct set: finite, simp_all, clarify)
+apply (subst card_Un_disjoint)
+ apply (auto simp add: disjoint_eq_subset_Compl)
+done
+
+
+subsubsection {* Relating injectivity and surjectivity *}
+
+lemma finite_surj_inj: "finite(A) \<Longrightarrow> A <= f`A \<Longrightarrow> inj_on f A"
+apply(rule eq_card_imp_inj_on, assumption)
+apply(frule finite_imageI)
+apply(drule (1) card_seteq)
+ apply(erule card_image_le)
+apply simp
+done
+
+lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
+shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
+by (blast intro: finite_surj_inj subset_UNIV dest:surj_range)
+
+lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
+shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
+by(fastsimp simp:surj_def dest!: endo_inj_surj)
+
+corollary infinite_UNIV_nat[iff]: "~finite(UNIV::nat set)"
+proof
+ assume "finite(UNIV::nat set)"
+ with finite_UNIV_inj_surj[of Suc]
+ show False by simp (blast dest: Suc_neq_Zero surjD)
+qed
+
+(* Often leads to bogus ATP proofs because of reduced type information, hence noatp *)
+lemma infinite_UNIV_char_0[noatp]:
+ "\<not> finite (UNIV::'a::semiring_char_0 set)"
+proof
+ assume "finite (UNIV::'a set)"
+ with subset_UNIV have "finite (range of_nat::'a set)"
+ by (rule finite_subset)
+ moreover have "inj (of_nat::nat \<Rightarrow> 'a)"
+ by (simp add: inj_on_def)
+ ultimately have "finite (UNIV::nat set)"
+ by (rule finite_imageD)
+ then show "False"
+ by simp
+qed
+
+end
--- a/src/HOL/GCD.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/GCD.thy Fri Mar 12 16:02:42 2010 +0100
@@ -1358,10 +1358,10 @@
done
lemma lcm_dvd2_nat: "(n::nat) dvd lcm m n"
- using lcm_dvd1_nat [of n m] by (simp only: lcm_nat_def times.commute gcd_nat.commute)
+ using lcm_dvd1_nat [of n m] by (simp only: lcm_nat_def mult.commute gcd_nat.commute)
lemma lcm_dvd2_int: "(n::int) dvd lcm m n"
- using lcm_dvd1_int [of n m] by (simp only: lcm_int_def lcm_nat_def times.commute gcd_nat.commute)
+ using lcm_dvd1_int [of n m] by (simp only: lcm_int_def lcm_nat_def mult.commute gcd_nat.commute)
lemma dvd_lcm_I1_nat[simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
by(metis lcm_dvd1_nat dvd_trans)
--- a/src/HOL/Groups.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Groups.thy Fri Mar 12 16:02:42 2010 +0100
@@ -67,6 +67,18 @@
end
+locale monoid = semigroup +
+ fixes z :: 'a ("1")
+ assumes left_neutral [simp]: "1 * a = a"
+ assumes right_neutral [simp]: "a * 1 = a"
+
+locale comm_monoid = abel_semigroup +
+ fixes z :: 'a ("1")
+ assumes comm_neutral: "a * 1 = a"
+
+sublocale comm_monoid < monoid proof
+qed (simp_all add: commute comm_neutral)
+
subsection {* Generic operations *}
@@ -129,19 +141,19 @@
class semigroup_add = plus +
assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
-sublocale semigroup_add < plus!: semigroup plus proof
+sublocale semigroup_add < add!: semigroup plus proof
qed (fact add_assoc)
class ab_semigroup_add = semigroup_add +
assumes add_commute [algebra_simps]: "a + b = b + a"
-sublocale ab_semigroup_add < plus!: abel_semigroup plus proof
+sublocale ab_semigroup_add < add!: abel_semigroup plus proof
qed (fact add_commute)
context ab_semigroup_add
begin
-lemmas add_left_commute [algebra_simps] = plus.left_commute
+lemmas add_left_commute [algebra_simps] = add.left_commute
theorems add_ac = add_assoc add_commute add_left_commute
@@ -152,19 +164,19 @@
class semigroup_mult = times +
assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)"
-sublocale semigroup_mult < times!: semigroup times proof
+sublocale semigroup_mult < mult!: semigroup times proof
qed (fact mult_assoc)
class ab_semigroup_mult = semigroup_mult +
assumes mult_commute [algebra_simps]: "a * b = b * a"
-sublocale ab_semigroup_mult < times!: abel_semigroup times proof
+sublocale ab_semigroup_mult < mult!: abel_semigroup times proof
qed (fact mult_commute)
context ab_semigroup_mult
begin
-lemmas mult_left_commute [algebra_simps] = times.left_commute
+lemmas mult_left_commute [algebra_simps] = mult.left_commute
theorems mult_ac = mult_assoc mult_commute mult_left_commute
@@ -173,36 +185,42 @@
theorems mult_ac = mult_assoc mult_commute mult_left_commute
class monoid_add = zero + semigroup_add +
- assumes add_0_left [simp]: "0 + a = a"
- and add_0_right [simp]: "a + 0 = a"
+ assumes add_0_left: "0 + a = a"
+ and add_0_right: "a + 0 = a"
+
+sublocale monoid_add < add!: monoid plus 0 proof
+qed (fact add_0_left add_0_right)+
lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"
by (rule eq_commute)
class comm_monoid_add = zero + ab_semigroup_add +
assumes add_0: "0 + a = a"
-begin
-subclass monoid_add
- proof qed (insert add_0, simp_all add: add_commute)
+sublocale comm_monoid_add < add!: comm_monoid plus 0 proof
+qed (insert add_0, simp add: ac_simps)
-end
+subclass (in comm_monoid_add) monoid_add proof
+qed (fact add.left_neutral add.right_neutral)+
class monoid_mult = one + semigroup_mult +
- assumes mult_1_left [simp]: "1 * a = a"
- assumes mult_1_right [simp]: "a * 1 = a"
+ assumes mult_1_left: "1 * a = a"
+ and mult_1_right: "a * 1 = a"
+
+sublocale monoid_mult < mult!: monoid times 1 proof
+qed (fact mult_1_left mult_1_right)+
lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"
by (rule eq_commute)
class comm_monoid_mult = one + ab_semigroup_mult +
assumes mult_1: "1 * a = a"
-begin
-subclass monoid_mult
- proof qed (insert mult_1, simp_all add: mult_commute)
+sublocale comm_monoid_mult < mult!: comm_monoid times 1 proof
+qed (insert mult_1, simp add: ac_simps)
-end
+subclass (in comm_monoid_mult) monoid_mult proof
+qed (fact mult.left_neutral mult.right_neutral)+
class cancel_semigroup_add = semigroup_add +
assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
--- a/src/HOL/IMP/Hoare.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/IMP/Hoare.thy Fri Mar 12 16:02:42 2010 +0100
@@ -36,76 +36,62 @@
wrt denotational semantics
*)
-lemma hoare_conseq1: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
-apply (erule hoare.conseq)
-apply assumption
-apply fast
-done
+lemma strengthen_pre: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
+by (blast intro: conseq)
-lemma hoare_conseq2: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
-apply (rule hoare.conseq)
-prefer 2 apply (assumption)
-apply fast
-apply fast
-done
+lemma weaken_post: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
+by (blast intro: conseq)
lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}"
-apply (unfold hoare_valid_def)
-apply (induct set: hoare)
- apply (simp_all (no_asm_simp))
- apply fast
- apply fast
-apply (rule allI, rule allI, rule impI)
-apply (erule lfp_induct2)
- apply (rule Gamma_mono)
-apply (unfold Gamma_def)
-apply fast
-done
+proof(induct rule: hoare.induct)
+ case (While P b c)
+ { fix s t
+ let ?G = "Gamma b (C c)"
+ assume "(s,t) \<in> lfp ?G"
+ hence "P s \<longrightarrow> P t \<and> \<not> b t"
+ proof(rule lfp_induct2)
+ show "mono ?G" by(rule Gamma_mono)
+ next
+ fix s t assume "(s,t) \<in> ?G (lfp ?G \<inter> {(s,t). P s \<longrightarrow> P t \<and> \<not> b t})"
+ thus "P s \<longrightarrow> P t \<and> \<not> b t" using While.hyps
+ by(auto simp: hoare_valid_def Gamma_def)
+ qed
+ }
+ thus ?case by(simp add:hoare_valid_def)
+qed (auto simp: hoare_valid_def)
+
lemma wp_SKIP: "wp \<SKIP> Q = Q"
-apply (unfold wp_def)
-apply (simp (no_asm))
-done
+by (simp add: wp_def)
lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\<mapsto>a s]))"
-apply (unfold wp_def)
-apply (simp (no_asm))
-done
+by (simp add: wp_def)
lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)"
-apply (unfold wp_def)
-apply (simp (no_asm))
-apply (rule ext)
-apply fast
-done
+by (rule ext) (auto simp: wp_def)
lemma wp_If:
"wp (\<IF> b \<THEN> c \<ELSE> d) Q = (%s. (b s --> wp c Q s) & (~b s --> wp d Q s))"
-apply (unfold wp_def)
-apply (simp (no_asm))
-apply (rule ext)
-apply fast
-done
+by (rule ext) (auto simp: wp_def)
-lemma wp_While_True:
- "b s ==> wp (\<WHILE> b \<DO> c) Q s = wp (c;\<WHILE> b \<DO> c) Q s"
-apply (unfold wp_def)
-apply (subst C_While_If)
-apply (simp (no_asm_simp))
-done
-
-lemma wp_While_False: "~b s ==> wp (\<WHILE> b \<DO> c) Q s = Q s"
-apply (unfold wp_def)
-apply (subst C_While_If)
-apply (simp (no_asm_simp))
-done
-
-lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False
+lemma wp_While_If:
+ "wp (\<WHILE> b \<DO> c) Q s =
+ wp (IF b THEN c;\<WHILE> b \<DO> c ELSE SKIP) Q s"
+by(simp only: wp_def C_While_If)
(*Not suitable for rewriting: LOOPS!*)
lemma wp_While_if:
"wp (\<WHILE> b \<DO> c) Q s = (if b s then wp (c;\<WHILE> b \<DO> c) Q s else Q s)"
- by simp
+by(simp add:wp_While_If wp_If wp_SKIP)
+
+lemma wp_While_True: "b s ==>
+ wp (\<WHILE> b \<DO> c) Q s = wp (c;\<WHILE> b \<DO> c) Q s"
+by(simp add: wp_While_if)
+
+lemma wp_While_False: "~b s ==> wp (\<WHILE> b \<DO> c) Q s = Q s"
+by(simp add: wp_While_if)
+
+lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False
lemma wp_While: "wp (\<WHILE> b \<DO> c) Q s =
(s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))"
@@ -132,23 +118,48 @@
lemmas [intro!] = hoare.skip hoare.ass hoare.semi hoare.If
lemma wp_is_pre: "|- {wp c Q} c {Q}"
-apply (induct c arbitrary: Q)
- apply (simp_all (no_asm))
- apply fast+
- apply (blast intro: hoare_conseq1)
-apply (rule hoare_conseq2)
- apply (rule hoare.While)
- apply (rule hoare_conseq1)
- prefer 2 apply fast
- apply safe
- apply simp
-apply simp
-done
+proof(induct c arbitrary: Q)
+ case SKIP show ?case by auto
+next
+ case Assign show ?case by auto
+next
+ case Semi thus ?case by auto
+next
+ case (Cond b c1 c2)
+ let ?If = "IF b THEN c1 ELSE c2"
+ show ?case
+ proof(rule If)
+ show "|- {\<lambda>s. wp ?If Q s \<and> b s} c1 {Q}"
+ proof(rule strengthen_pre[OF _ Cond(1)])
+ show "\<forall>s. wp ?If Q s \<and> b s \<longrightarrow> wp c1 Q s" by auto
+ qed
+ show "|- {\<lambda>s. wp ?If Q s \<and> \<not> b s} c2 {Q}"
+ proof(rule strengthen_pre[OF _ Cond(2)])
+ show "\<forall>s. wp ?If Q s \<and> \<not> b s \<longrightarrow> wp c2 Q s" by auto
+ qed
+ qed
+next
+ case (While b c)
+ let ?w = "WHILE b DO c"
+ have "|- {wp ?w Q} ?w {\<lambda>s. wp ?w Q s \<and> \<not> b s}"
+ proof(rule hoare.While)
+ show "|- {\<lambda>s. wp ?w Q s \<and> b s} c {wp ?w Q}"
+ proof(rule strengthen_pre[OF _ While(1)])
+ show "\<forall>s. wp ?w Q s \<and> b s \<longrightarrow> wp c (wp ?w Q) s" by auto
+ qed
+ qed
+ thus ?case
+ proof(rule weaken_post)
+ show "\<forall>s. wp ?w Q s \<and> \<not> b s \<longrightarrow> Q s" by auto
+ qed
+qed
-lemma hoare_relative_complete: "|= {P}c{Q} ==> |- {P}c{Q}"
-apply (rule hoare_conseq1 [OF _ wp_is_pre])
-apply (unfold hoare_valid_def wp_def)
-apply fast
-done
+lemma hoare_relative_complete: assumes "|= {P}c{Q}" shows "|- {P}c{Q}"
+proof(rule conseq)
+ show "\<forall>s. P s \<longrightarrow> wp c Q s" using assms
+ by (auto simp: hoare_valid_def wp_def)
+ show "|- {wp c Q} c {Q}" by(rule wp_is_pre)
+ show "\<forall>s. Q s \<longrightarrow> Q s" by auto
+qed
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Hoare_Op.thy Fri Mar 12 16:02:42 2010 +0100
@@ -0,0 +1,130 @@
+(* Title: HOL/IMP/Hoare_Op.thy
+ ID: $Id$
+ Author: Tobias Nipkow
+*)
+
+header "Hoare Logic (justified wrt operational semantics)"
+
+theory Hoare_Op imports Natural begin
+
+types assn = "state => bool"
+
+definition
+ hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where
+ "|= {P}c{Q} = (!s t. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t --> P s --> Q t)"
+
+inductive
+ hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
+where
+ skip: "|- {P}\<SKIP>{P}"
+| ass: "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
+| semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
+| If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==>
+ |- {P} \<IF> b \<THEN> c \<ELSE> d {Q}"
+| While: "|- {%s. P s & b s} c {P} ==>
+ |- {P} \<WHILE> b \<DO> c {%s. P s & ~b s}"
+| conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==>
+ |- {P'}c{Q'}"
+
+lemmas [simp] = skip ass semi If
+
+lemma strengthen_pre: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}"
+by (blast intro: conseq)
+
+lemma weaken_post: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}"
+by (blast intro: conseq)
+
+lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}"
+proof(induct rule: hoare.induct)
+ case (While P b c)
+ { fix s t
+ assume "\<langle>WHILE b DO c,s\<rangle> \<longrightarrow>\<^sub>c t"
+ hence "P s \<longrightarrow> P t \<and> \<not> b t"
+ proof(induct "WHILE b DO c" s t)
+ case WhileFalse thus ?case by blast
+ next
+ case WhileTrue thus ?case
+ using While(2) unfolding hoare_valid_def by blast
+ qed
+
+ }
+ thus ?case unfolding hoare_valid_def by blast
+qed (auto simp: hoare_valid_def)
+
+
+definition
+ wp :: "com => assn => assn" where
+ "wp c Q = (%s. !t. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t --> Q t)"
+
+lemma wp_SKIP: "wp \<SKIP> Q = Q"
+by (simp add: wp_def)
+
+lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\<mapsto>a s]))"
+by (simp add: wp_def)
+
+lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)"
+by (rule ext) (auto simp: wp_def)
+
+lemma wp_If:
+ "wp (\<IF> b \<THEN> c \<ELSE> d) Q = (%s. (b s --> wp c Q s) & (~b s --> wp d Q s))"
+by (rule ext) (auto simp: wp_def)
+
+lemma wp_While_If:
+ "wp (\<WHILE> b \<DO> c) Q s =
+ wp (IF b THEN c;\<WHILE> b \<DO> c ELSE SKIP) Q s"
+unfolding wp_def by (metis equivD1 equivD2 unfold_while)
+
+lemma wp_While_True: "b s ==>
+ wp (\<WHILE> b \<DO> c) Q s = wp (c;\<WHILE> b \<DO> c) Q s"
+by(simp add: wp_While_If wp_If wp_SKIP)
+
+lemma wp_While_False: "~b s ==> wp (\<WHILE> b \<DO> c) Q s = Q s"
+by(simp add: wp_While_If wp_If wp_SKIP)
+
+lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False
+
+lemma wp_is_pre: "|- {wp c Q} c {Q}"
+proof(induct c arbitrary: Q)
+ case SKIP show ?case by auto
+next
+ case Assign show ?case by auto
+next
+ case Semi thus ?case by(auto intro: semi)
+next
+ case (Cond b c1 c2)
+ let ?If = "IF b THEN c1 ELSE c2"
+ show ?case
+ proof(rule If)
+ show "|- {\<lambda>s. wp ?If Q s \<and> b s} c1 {Q}"
+ proof(rule strengthen_pre[OF _ Cond(1)])
+ show "\<forall>s. wp ?If Q s \<and> b s \<longrightarrow> wp c1 Q s" by auto
+ qed
+ show "|- {\<lambda>s. wp ?If Q s \<and> \<not> b s} c2 {Q}"
+ proof(rule strengthen_pre[OF _ Cond(2)])
+ show "\<forall>s. wp ?If Q s \<and> \<not> b s \<longrightarrow> wp c2 Q s" by auto
+ qed
+ qed
+next
+ case (While b c)
+ let ?w = "WHILE b DO c"
+ have "|- {wp ?w Q} ?w {\<lambda>s. wp ?w Q s \<and> \<not> b s}"
+ proof(rule hoare.While)
+ show "|- {\<lambda>s. wp ?w Q s \<and> b s} c {wp ?w Q}"
+ proof(rule strengthen_pre[OF _ While(1)])
+ show "\<forall>s. wp ?w Q s \<and> b s \<longrightarrow> wp c (wp ?w Q) s" by auto
+ qed
+ qed
+ thus ?case
+ proof(rule weaken_post)
+ show "\<forall>s. wp ?w Q s \<and> \<not> b s \<longrightarrow> Q s" by auto
+ qed
+qed
+
+lemma hoare_relative_complete: assumes "|= {P}c{Q}" shows "|- {P}c{Q}"
+proof(rule strengthen_pre)
+ show "\<forall>s. P s \<longrightarrow> wp c Q s" using assms
+ by (auto simp: hoare_valid_def wp_def)
+ show "|- {wp c Q} c {Q}" by(rule wp_is_pre)
+qed
+
+end
--- a/src/HOL/IMP/ROOT.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/IMP/ROOT.ML Fri Mar 12 16:02:42 2010 +0100
@@ -6,4 +6,4 @@
Caveat: HOLCF/IMP depends on HOL/IMP
*)
-use_thys ["Expr", "Transition", "VC", "Examples", "Compiler0", "Compiler", "Live"];
+use_thys ["Expr", "Transition", "Hoare_Op", "VC", "Examples", "Compiler0", "Compiler", "Live"];
--- a/src/HOL/IsaMakefile Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/IsaMakefile Fri Mar 12 16:02:42 2010 +0100
@@ -247,6 +247,7 @@
MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
ATP_Linkup.thy \
+ Big_Operators.thy \
Code_Evaluation.thy \
Code_Numeral.thy \
Divides.thy \
@@ -393,7 +394,7 @@
Library/Inner_Product.thy Library/Kleene_Algebra.thy \
Library/Lattice_Algebras.thy Library/Lattice_Syntax.thy \
Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy \
- Library/State_Monad.thy Library/Nat_Int_Bij.thy Library/Multiset.thy \
+ Library/State_Monad.thy Library/Multiset.thy \
Library/Permutation.thy Library/Quotient_Type.thy \
Library/Quicksort.thy Library/Nat_Infinity.thy Library/Word.thy \
Library/README.html Library/Continuity.thy \
@@ -415,6 +416,7 @@
Library/Enum.thy Library/Float.thy Library/Quotient_List.thy \
Library/Quotient_Option.thy Library/Quotient_Product.thy \
Library/Quotient_Sum.thy Library/Quotient_Syntax.thy \
+ Library/Nat_Bijection.thy \
$(SRC)/Tools/float.ML \
$(SRC)/HOL/Tools/float_arith.ML Library/positivstellensatz.ML \
Library/reify_data.ML Library/reflection.ML Library/LaTeXsugar.thy \
--- a/src/HOL/Lattices.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Lattices.thy Fri Mar 12 16:02:42 2010 +0100
@@ -43,7 +43,7 @@
end
-subsection {* Conrete lattices *}
+subsection {* Concrete lattices *}
notation
less_eq (infix "\<sqsubseteq>" 50) and
--- a/src/HOL/Library/Countable.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Library/Countable.thy Fri Mar 12 16:02:42 2010 +0100
@@ -5,7 +5,7 @@
header {* Encoding (almost) everything into natural numbers *}
theory Countable
-imports Main Rat Nat_Int_Bij
+imports Main Rat Nat_Bijection
begin
subsection {* The class of countable types *}
@@ -48,7 +48,7 @@
subsection {* Countable types *}
instance nat :: countable
- by (rule countable_classI [of "id"]) simp
+ by (rule countable_classI [of "id"]) simp
subclass (in finite) countable
proof
@@ -63,47 +63,9 @@
text {* Pairs *}
-primrec sum :: "nat \<Rightarrow> nat"
-where
- "sum 0 = 0"
-| "sum (Suc n) = Suc n + sum n"
-
-lemma sum_arith: "sum n = n * Suc n div 2"
- by (induct n) auto
-
-lemma sum_mono: "n \<ge> m \<Longrightarrow> sum n \<ge> sum m"
- by (induct n m rule: diff_induct) auto
-
-definition
- "pair_encode = (\<lambda>(m, n). sum (m + n) + m)"
-
-lemma inj_pair_cencode: "inj pair_encode"
- unfolding pair_encode_def
-proof (rule injI, simp only: split_paired_all split_conv)
- fix a b c d
- assume eq: "sum (a + b) + a = sum (c + d) + c"
- have "a + b = c + d \<or> a + b \<ge> Suc (c + d) \<or> c + d \<ge> Suc (a + b)" by arith
- then
- show "(a, b) = (c, d)"
- proof (elim disjE)
- assume sumeq: "a + b = c + d"
- then have "a = c" using eq by auto
- moreover from sumeq this have "b = d" by auto
- ultimately show ?thesis by simp
- next
- assume "a + b \<ge> Suc (c + d)"
- from sum_mono[OF this] eq
- show ?thesis by auto
- next
- assume "c + d \<ge> Suc (a + b)"
- from sum_mono[OF this] eq
- show ?thesis by auto
- qed
-qed
-
instance "*" :: (countable, countable) countable
-by (rule countable_classI [of "\<lambda>(x, y). pair_encode (to_nat x, to_nat y)"])
- (auto dest: injD [OF inj_pair_cencode] injD [OF inj_to_nat])
+ by (rule countable_classI [of "\<lambda>(x, y). prod_encode (to_nat x, to_nat y)"])
+ (auto simp add: prod_encode_eq)
text {* Sums *}
@@ -111,76 +73,28 @@
instance "+":: (countable, countable) countable
by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a)
| Inr b \<Rightarrow> to_nat (True, to_nat b))"])
- (auto split:sum.splits)
+ (simp split: sum.split_asm)
text {* Integers *}
-lemma int_cases: "(i::int) = 0 \<or> i < 0 \<or> i > 0"
-by presburger
-
-lemma int_pos_neg_zero:
- obtains (zero) "(z::int) = 0" "sgn z = 0" "abs z = 0"
- | (pos) n where "z = of_nat n" "sgn z = 1" "abs z = of_nat n"
- | (neg) n where "z = - (of_nat n)" "sgn z = -1" "abs z = of_nat n"
-apply atomize_elim
-apply (insert int_cases[of z])
-apply (auto simp:zsgn_def)
-apply (rule_tac x="nat (-z)" in exI, simp)
-apply (rule_tac x="nat z" in exI, simp)
-done
-
instance int :: countable
-proof (rule countable_classI [of "(\<lambda>i. to_nat (nat (sgn i + 1), nat (abs i)))"],
- auto dest: injD [OF inj_to_nat])
- fix x y
- assume a: "nat (sgn x + 1) = nat (sgn y + 1)" "nat (abs x) = nat (abs y)"
- show "x = y"
- proof (cases rule: int_pos_neg_zero[of x])
- case zero
- with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto
- next
- case (pos n)
- with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto
- next
- case (neg n)
- with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto
- qed
-qed
+ by (rule countable_classI [of "int_encode"])
+ (simp add: int_encode_eq)
text {* Options *}
instance option :: (countable) countable
-by (rule countable_classI[of "\<lambda>x. case x of None \<Rightarrow> 0
- | Some y \<Rightarrow> Suc (to_nat y)"])
- (auto split:option.splits)
+ by (rule countable_classI [of "option_case 0 (Suc \<circ> to_nat)"])
+ (simp split: option.split_asm)
text {* Lists *}
-lemma from_nat_to_nat_map [simp]: "map from_nat (map to_nat xs) = xs"
- by (simp add: comp_def)
-
-primrec
- list_encode :: "'a\<Colon>countable list \<Rightarrow> nat"
-where
- "list_encode [] = 0"
-| "list_encode (x#xs) = Suc (to_nat (x, list_encode xs))"
-
instance list :: (countable) countable
-proof (rule countable_classI [of "list_encode"])
- fix xs ys :: "'a list"
- assume cenc: "list_encode xs = list_encode ys"
- then show "xs = ys"
- proof (induct xs arbitrary: ys)
- case (Nil ys)
- with cenc show ?case by (cases ys, auto)
- next
- case (Cons x xs' ys)
- thus ?case by (cases ys) auto
- qed
-qed
+ by (rule countable_classI [of "list_encode \<circ> map to_nat"])
+ (simp add: list_encode_eq)
text {* Functions *}
@@ -200,8 +114,8 @@
subsection {* The Rationals are Countably Infinite *}
definition nat_to_rat_surj :: "nat \<Rightarrow> rat" where
-"nat_to_rat_surj n = (let (a,b) = nat_to_nat2 n
- in Fract (nat_to_int_bij a) (nat_to_int_bij b))"
+"nat_to_rat_surj n = (let (a,b) = prod_decode n
+ in Fract (int_decode a) (int_decode b))"
lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj"
unfolding surj_def
@@ -210,10 +124,9 @@
show "\<exists>n. r = nat_to_rat_surj n"
proof (cases r)
fix i j assume [simp]: "r = Fract i j" and "j > 0"
- have "r = (let m = inv nat_to_int_bij i; n = inv nat_to_int_bij j
- in nat_to_rat_surj(nat2_to_nat (m,n)))"
- using nat2_to_nat_inj surj_f_inv_f[OF surj_nat_to_int_bij]
- by(simp add:Let_def nat_to_rat_surj_def nat_to_nat2_def)
+ have "r = (let m = int_encode i; n = int_encode j
+ in nat_to_rat_surj(prod_encode (m,n)))"
+ by (simp add: Let_def nat_to_rat_surj_def)
thus "\<exists>n. r = nat_to_rat_surj n" by(auto simp:Let_def)
qed
qed
--- a/src/HOL/Library/Dlist.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Library/Dlist.thy Fri Mar 12 16:02:42 2010 +0100
@@ -123,8 +123,6 @@
"list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"
by (simp add: filter_def)
-declare null_def [code] member_def [code] length_def [code] fold_def [code] -- {* explicit is better than implicit *}
-
section {* Implementation of sets by distinct lists -- canonical! *}
--- a/src/HOL/Library/Efficient_Nat.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy Fri Mar 12 16:02:42 2010 +0100
@@ -400,7 +400,7 @@
(SML "IntInf.max/ (/0,/ _)")
(OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
-text {* For Haskell ans Scala, things are slightly different again. *}
+text {* For Haskell and Scala, things are slightly different again. *}
code_const int and nat
(Haskell "toInteger" and "fromInteger")
--- a/src/HOL/Library/Multiset.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Library/Multiset.thy Fri Mar 12 16:02:42 2010 +0100
@@ -1729,4 +1729,33 @@
"M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P"
by (fact multiset_order.less_asym)
+ML {*
+(* Proof.context -> string -> (typ -> term list) -> typ -> term -> term *)
+fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T]))
+ (Const _ $ t') =
+ let
+ val (maybe_opt, ps) =
+ Nitpick_Model.dest_plain_fun t' ||> op ~~
+ ||> map (apsnd (snd o HOLogic.dest_number))
+ fun elems_for t =
+ case AList.lookup (op =) ps t of
+ SOME n => replicate n t
+ | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]
+ in
+ case maps elems_for (all_values elem_T) @
+ (if maybe_opt then [Const (Nitpick_Model.unrep, elem_T)] else []) of
+ [] => Const (@{const_name zero_class.zero}, T)
+ | ts => foldl1 (fn (t1, t2) =>
+ Const (@{const_name plus_class.plus}, T --> T --> T)
+ $ t1 $ t2)
+ (map (curry (op $) (Const (@{const_name single},
+ elem_T --> T))) ts)
+ end
+ | multiset_postproc _ _ _ _ t = t
+*}
+
+setup {*
+Nitpick.register_term_postprocessor @{typ "'a multiset"} multiset_postproc
+*}
+
end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Nat_Bijection.thy Fri Mar 12 16:02:42 2010 +0100
@@ -0,0 +1,370 @@
+(* Title: HOL/Nat_Bijection.thy
+ Author: Brian Huffman
+ Author: Florian Haftmann
+ Author: Stefan Richter
+ Author: Tobias Nipkow
+ Author: Alexander Krauss
+*)
+
+header {* Bijections between natural numbers and other types *}
+
+theory Nat_Bijection
+imports Main Parity
+begin
+
+subsection {* Type @{typ "nat \<times> nat"} *}
+
+text "Triangle numbers: 0, 1, 3, 6, 10, 15, ..."
+
+definition
+ triangle :: "nat \<Rightarrow> nat"
+where
+ "triangle n = n * Suc n div 2"
+
+lemma triangle_0 [simp]: "triangle 0 = 0"
+unfolding triangle_def by simp
+
+lemma triangle_Suc [simp]: "triangle (Suc n) = triangle n + Suc n"
+unfolding triangle_def by simp
+
+definition
+ prod_encode :: "nat \<times> nat \<Rightarrow> nat"
+where
+ "prod_encode = (\<lambda>(m, n). triangle (m + n) + m)"
+
+text {* In this auxiliary function, @{term "triangle k + m"} is an invariant. *}
+
+fun
+ prod_decode_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat"
+where
+ "prod_decode_aux k m =
+ (if m \<le> k then (m, k - m) else prod_decode_aux (Suc k) (m - Suc k))"
+
+declare prod_decode_aux.simps [simp del]
+
+definition
+ prod_decode :: "nat \<Rightarrow> nat \<times> nat"
+where
+ "prod_decode = prod_decode_aux 0"
+
+lemma prod_encode_prod_decode_aux:
+ "prod_encode (prod_decode_aux k m) = triangle k + m"
+apply (induct k m rule: prod_decode_aux.induct)
+apply (subst prod_decode_aux.simps)
+apply (simp add: prod_encode_def)
+done
+
+lemma prod_decode_inverse [simp]: "prod_encode (prod_decode n) = n"
+unfolding prod_decode_def by (simp add: prod_encode_prod_decode_aux)
+
+lemma prod_decode_triangle_add:
+ "prod_decode (triangle k + m) = prod_decode_aux k m"
+apply (induct k arbitrary: m)
+apply (simp add: prod_decode_def)
+apply (simp only: triangle_Suc add_assoc)
+apply (subst prod_decode_aux.simps, simp)
+done
+
+lemma prod_encode_inverse [simp]: "prod_decode (prod_encode x) = x"
+unfolding prod_encode_def
+apply (induct x)
+apply (simp add: prod_decode_triangle_add)
+apply (subst prod_decode_aux.simps, simp)
+done
+
+lemma inj_prod_encode: "inj_on prod_encode A"
+by (rule inj_on_inverseI, rule prod_encode_inverse)
+
+lemma inj_prod_decode: "inj_on prod_decode A"
+by (rule inj_on_inverseI, rule prod_decode_inverse)
+
+lemma surj_prod_encode: "surj prod_encode"
+by (rule surjI, rule prod_decode_inverse)
+
+lemma surj_prod_decode: "surj prod_decode"
+by (rule surjI, rule prod_encode_inverse)
+
+lemma bij_prod_encode: "bij prod_encode"
+by (rule bijI [OF inj_prod_encode surj_prod_encode])
+
+lemma bij_prod_decode: "bij prod_decode"
+by (rule bijI [OF inj_prod_decode surj_prod_decode])
+
+lemma prod_encode_eq: "prod_encode x = prod_encode y \<longleftrightarrow> x = y"
+by (rule inj_prod_encode [THEN inj_eq])
+
+lemma prod_decode_eq: "prod_decode x = prod_decode y \<longleftrightarrow> x = y"
+by (rule inj_prod_decode [THEN inj_eq])
+
+text {* Ordering properties *}
+
+lemma le_prod_encode_1: "a \<le> prod_encode (a, b)"
+unfolding prod_encode_def by simp
+
+lemma le_prod_encode_2: "b \<le> prod_encode (a, b)"
+unfolding prod_encode_def by (induct b, simp_all)
+
+
+subsection {* Type @{typ "nat + nat"} *}
+
+definition
+ sum_encode :: "nat + nat \<Rightarrow> nat"
+where
+ "sum_encode x = (case x of Inl a \<Rightarrow> 2 * a | Inr b \<Rightarrow> Suc (2 * b))"
+
+definition
+ sum_decode :: "nat \<Rightarrow> nat + nat"
+where
+ "sum_decode n = (if even n then Inl (n div 2) else Inr (n div 2))"
+
+lemma sum_encode_inverse [simp]: "sum_decode (sum_encode x) = x"
+unfolding sum_decode_def sum_encode_def
+by (induct x) simp_all
+
+lemma sum_decode_inverse [simp]: "sum_encode (sum_decode n) = n"
+unfolding sum_decode_def sum_encode_def numeral_2_eq_2
+by (simp add: even_nat_div_two_times_two odd_nat_div_two_times_two_plus_one
+ del: mult_Suc)
+
+lemma inj_sum_encode: "inj_on sum_encode A"
+by (rule inj_on_inverseI, rule sum_encode_inverse)
+
+lemma inj_sum_decode: "inj_on sum_decode A"
+by (rule inj_on_inverseI, rule sum_decode_inverse)
+
+lemma surj_sum_encode: "surj sum_encode"
+by (rule surjI, rule sum_decode_inverse)
+
+lemma surj_sum_decode: "surj sum_decode"
+by (rule surjI, rule sum_encode_inverse)
+
+lemma bij_sum_encode: "bij sum_encode"
+by (rule bijI [OF inj_sum_encode surj_sum_encode])
+
+lemma bij_sum_decode: "bij sum_decode"
+by (rule bijI [OF inj_sum_decode surj_sum_decode])
+
+lemma sum_encode_eq: "sum_encode x = sum_encode y \<longleftrightarrow> x = y"
+by (rule inj_sum_encode [THEN inj_eq])
+
+lemma sum_decode_eq: "sum_decode x = sum_decode y \<longleftrightarrow> x = y"
+by (rule inj_sum_decode [THEN inj_eq])
+
+
+subsection {* Type @{typ "int"} *}
+
+definition
+ int_encode :: "int \<Rightarrow> nat"
+where
+ "int_encode i = sum_encode (if 0 \<le> i then Inl (nat i) else Inr (nat (- i - 1)))"
+
+definition
+ int_decode :: "nat \<Rightarrow> int"
+where
+ "int_decode n = (case sum_decode n of Inl a \<Rightarrow> int a | Inr b \<Rightarrow> - int b - 1)"
+
+lemma int_encode_inverse [simp]: "int_decode (int_encode x) = x"
+unfolding int_decode_def int_encode_def by simp
+
+lemma int_decode_inverse [simp]: "int_encode (int_decode n) = n"
+unfolding int_decode_def int_encode_def using sum_decode_inverse [of n]
+by (cases "sum_decode n", simp_all)
+
+lemma inj_int_encode: "inj_on int_encode A"
+by (rule inj_on_inverseI, rule int_encode_inverse)
+
+lemma inj_int_decode: "inj_on int_decode A"
+by (rule inj_on_inverseI, rule int_decode_inverse)
+
+lemma surj_int_encode: "surj int_encode"
+by (rule surjI, rule int_decode_inverse)
+
+lemma surj_int_decode: "surj int_decode"
+by (rule surjI, rule int_encode_inverse)
+
+lemma bij_int_encode: "bij int_encode"
+by (rule bijI [OF inj_int_encode surj_int_encode])
+
+lemma bij_int_decode: "bij int_decode"
+by (rule bijI [OF inj_int_decode surj_int_decode])
+
+lemma int_encode_eq: "int_encode x = int_encode y \<longleftrightarrow> x = y"
+by (rule inj_int_encode [THEN inj_eq])
+
+lemma int_decode_eq: "int_decode x = int_decode y \<longleftrightarrow> x = y"
+by (rule inj_int_decode [THEN inj_eq])
+
+
+subsection {* Type @{typ "nat list"} *}
+
+fun
+ list_encode :: "nat list \<Rightarrow> nat"
+where
+ "list_encode [] = 0"
+| "list_encode (x # xs) = Suc (prod_encode (x, list_encode xs))"
+
+function
+ list_decode :: "nat \<Rightarrow> nat list"
+where
+ "list_decode 0 = []"
+| "list_decode (Suc n) = (case prod_decode n of (x, y) \<Rightarrow> x # list_decode y)"
+by pat_completeness auto
+
+termination list_decode
+apply (relation "measure id", simp_all)
+apply (drule arg_cong [where f="prod_encode"])
+apply (simp add: le_imp_less_Suc le_prod_encode_2)
+done
+
+lemma list_encode_inverse [simp]: "list_decode (list_encode x) = x"
+by (induct x rule: list_encode.induct) simp_all
+
+lemma list_decode_inverse [simp]: "list_encode (list_decode n) = n"
+apply (induct n rule: list_decode.induct, simp)
+apply (simp split: prod.split)
+apply (simp add: prod_decode_eq [symmetric])
+done
+
+lemma inj_list_encode: "inj_on list_encode A"
+by (rule inj_on_inverseI, rule list_encode_inverse)
+
+lemma inj_list_decode: "inj_on list_decode A"
+by (rule inj_on_inverseI, rule list_decode_inverse)
+
+lemma surj_list_encode: "surj list_encode"
+by (rule surjI, rule list_decode_inverse)
+
+lemma surj_list_decode: "surj list_decode"
+by (rule surjI, rule list_encode_inverse)
+
+lemma bij_list_encode: "bij list_encode"
+by (rule bijI [OF inj_list_encode surj_list_encode])
+
+lemma bij_list_decode: "bij list_decode"
+by (rule bijI [OF inj_list_decode surj_list_decode])
+
+lemma list_encode_eq: "list_encode x = list_encode y \<longleftrightarrow> x = y"
+by (rule inj_list_encode [THEN inj_eq])
+
+lemma list_decode_eq: "list_decode x = list_decode y \<longleftrightarrow> x = y"
+by (rule inj_list_decode [THEN inj_eq])
+
+
+subsection {* Finite sets of naturals *}
+
+subsubsection {* Preliminaries *}
+
+lemma finite_vimage_Suc_iff: "finite (Suc -` F) \<longleftrightarrow> finite F"
+apply (safe intro!: finite_vimageI inj_Suc)
+apply (rule finite_subset [where B="insert 0 (Suc ` Suc -` F)"])
+apply (rule subsetI, case_tac x, simp, simp)
+apply (rule finite_insert [THEN iffD2])
+apply (erule finite_imageI)
+done
+
+lemma vimage_Suc_insert_0: "Suc -` insert 0 A = Suc -` A"
+by auto
+
+lemma vimage_Suc_insert_Suc:
+ "Suc -` insert (Suc n) A = insert n (Suc -` A)"
+by auto
+
+lemma even_nat_Suc_div_2: "even x \<Longrightarrow> Suc x div 2 = x div 2"
+by (simp only: numeral_2_eq_2 even_nat_plus_one_div_two)
+
+lemma div2_even_ext_nat:
+ "\<lbrakk>x div 2 = y div 2; even x = even y\<rbrakk> \<Longrightarrow> (x::nat) = y"
+apply (rule mod_div_equality [of x 2, THEN subst])
+apply (rule mod_div_equality [of y 2, THEN subst])
+apply (case_tac "even x")
+apply (simp add: numeral_2_eq_2 even_nat_equiv_def)
+apply (simp add: numeral_2_eq_2 odd_nat_equiv_def)
+done
+
+subsubsection {* From sets to naturals *}
+
+definition
+ set_encode :: "nat set \<Rightarrow> nat"
+where
+ "set_encode = setsum (op ^ 2)"
+
+lemma set_encode_empty [simp]: "set_encode {} = 0"
+by (simp add: set_encode_def)
+
+lemma set_encode_insert [simp]:
+ "\<lbrakk>finite A; n \<notin> A\<rbrakk> \<Longrightarrow> set_encode (insert n A) = 2^n + set_encode A"
+by (simp add: set_encode_def)
+
+lemma even_set_encode_iff: "finite A \<Longrightarrow> even (set_encode A) \<longleftrightarrow> 0 \<notin> A"
+unfolding set_encode_def by (induct set: finite, auto)
+
+lemma set_encode_vimage_Suc: "set_encode (Suc -` A) = set_encode A div 2"
+apply (cases "finite A")
+apply (erule finite_induct, simp)
+apply (case_tac x)
+apply (simp add: even_nat_Suc_div_2 even_set_encode_iff vimage_Suc_insert_0)
+apply (simp add: finite_vimageI add_commute vimage_Suc_insert_Suc)
+apply (simp add: set_encode_def finite_vimage_Suc_iff)
+done
+
+lemmas set_encode_div_2 = set_encode_vimage_Suc [symmetric]
+
+subsubsection {* From naturals to sets *}
+
+definition
+ set_decode :: "nat \<Rightarrow> nat set"
+where
+ "set_decode x = {n. odd (x div 2 ^ n)}"
+
+lemma set_decode_0 [simp]: "0 \<in> set_decode x \<longleftrightarrow> odd x"
+by (simp add: set_decode_def)
+
+lemma set_decode_Suc [simp]:
+ "Suc n \<in> set_decode x \<longleftrightarrow> n \<in> set_decode (x div 2)"
+by (simp add: set_decode_def div_mult2_eq)
+
+lemma set_decode_zero [simp]: "set_decode 0 = {}"
+by (simp add: set_decode_def)
+
+lemma set_decode_div_2: "set_decode (x div 2) = Suc -` set_decode x"
+by auto
+
+lemma set_decode_plus_power_2:
+ "n \<notin> set_decode z \<Longrightarrow> set_decode (2 ^ n + z) = insert n (set_decode z)"
+ apply (induct n arbitrary: z, simp_all)
+ apply (rule set_ext, induct_tac x, simp, simp add: even_nat_Suc_div_2)
+ apply (rule set_ext, induct_tac x, simp, simp add: add_commute)
+done
+
+lemma finite_set_decode [simp]: "finite (set_decode n)"
+apply (induct n rule: nat_less_induct)
+apply (case_tac "n = 0", simp)
+apply (drule_tac x="n div 2" in spec, simp)
+apply (simp add: set_decode_div_2)
+apply (simp add: finite_vimage_Suc_iff)
+done
+
+subsubsection {* Proof of isomorphism *}
+
+lemma set_decode_inverse [simp]: "set_encode (set_decode n) = n"
+apply (induct n rule: nat_less_induct)
+apply (case_tac "n = 0", simp)
+apply (drule_tac x="n div 2" in spec, simp)
+apply (simp add: set_decode_div_2 set_encode_vimage_Suc)
+apply (erule div2_even_ext_nat)
+apply (simp add: even_set_encode_iff)
+done
+
+lemma set_encode_inverse [simp]: "finite A \<Longrightarrow> set_decode (set_encode A) = A"
+apply (erule finite_induct, simp_all)
+apply (simp add: set_decode_plus_power_2)
+done
+
+lemma inj_on_set_encode: "inj_on set_encode (Collect finite)"
+by (rule inj_on_inverseI [where g="set_decode"], simp)
+
+lemma set_encode_eq:
+ "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> set_encode A = set_encode B \<longleftrightarrow> A = B"
+by (rule iffI, simp add: inj_onD [OF inj_on_set_encode], simp)
+
+end
--- a/src/HOL/Library/Nat_Int_Bij.thy Tue Mar 09 15:42:23 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,178 +0,0 @@
-(* Title: HOL/Nat_Int_Bij.thy
- Author: Stefan Richter, Tobias Nipkow
-*)
-
-header{* Bijections $\mathbb{N}\to\mathbb{N}^2$ and $\mathbb{N}\to\mathbb{Z}$*}
-
-theory Nat_Int_Bij
-imports Main
-begin
-
-subsection{* A bijection between @{text "\<nat>"} and @{text "\<nat>\<twosuperior>"} *}
-
-text{* Definition and proofs are from \cite[page 85]{Oberschelp:1993}. *}
-
-definition nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" where
-"nat2_to_nat pair = (let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n)"
-definition nat_to_nat2:: "nat \<Rightarrow> (nat * nat)" where
-"nat_to_nat2 = inv nat2_to_nat"
-
-lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)"
-proof (cases "2 dvd a")
- case True
- then show ?thesis by (rule dvd_mult2)
-next
- case False
- then have "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0)
- then have "Suc a mod 2 = 0" by (simp add: mod_Suc)
- then have "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0)
- then show ?thesis by (rule dvd_mult)
-qed
-
-lemma
- assumes eq: "nat2_to_nat (u,v) = nat2_to_nat (x,y)"
- shows nat2_to_nat_help: "u+v \<le> x+y"
-proof (rule classical)
- assume "\<not> ?thesis"
- then have contrapos: "x+y < u+v"
- by simp
- have "nat2_to_nat (x,y) < (x+y) * Suc (x+y) div 2 + Suc (x + y)"
- by (unfold nat2_to_nat_def) (simp add: Let_def)
- also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2"
- by (simp only: div_mult_self1_is_m)
- also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2
- + ((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2"
- proof -
- have "2 dvd (x+y)*Suc(x+y)"
- by (rule dvd2_a_x_suc_a)
- then have "(x+y)*Suc(x+y) mod 2 = 0"
- by (simp only: dvd_eq_mod_eq_0)
- also
- have "2 * Suc(x+y) mod 2 = 0"
- by (rule mod_mult_self1_is_0)
- ultimately have
- "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0"
- by simp
- then show ?thesis
- by simp
- qed
- also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2"
- by (rule div_add1_eq [symmetric])
- also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2"
- by (simp only: add_mult_distrib [symmetric])
- also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2"
- by (simp only: mult_le_mono div_le_mono)
- also have "\<dots> \<le> nat2_to_nat (u,v)"
- by (unfold nat2_to_nat_def) (simp add: Let_def)
- finally show ?thesis
- by (simp only: eq)
-qed
-
-theorem nat2_to_nat_inj: "inj nat2_to_nat"
-proof -
- {
- fix u v x y
- assume eq1: "nat2_to_nat (u,v) = nat2_to_nat (x,y)"
- then have "u+v \<le> x+y" by (rule nat2_to_nat_help)
- also from eq1 [symmetric] have "x+y \<le> u+v"
- by (rule nat2_to_nat_help)
- finally have eq2: "u+v = x+y" .
- with eq1 have ux: "u=x"
- by (simp add: nat2_to_nat_def Let_def)
- with eq2 have vy: "v=y" by simp
- with ux have "(u,v) = (x,y)" by simp
- }
- then have "\<And>x y. nat2_to_nat x = nat2_to_nat y \<Longrightarrow> x=y" by fast
- then show ?thesis unfolding inj_on_def by simp
-qed
-
-lemma nat_to_nat2_surj: "surj nat_to_nat2"
-by (simp only: nat_to_nat2_def nat2_to_nat_inj inj_imp_surj_inv)
-
-
-lemma gauss_sum_nat_upto: "2 * (\<Sum>i\<le>n::nat. i) = n * (n + 1)"
-using gauss_sum[where 'a = nat]
-by (simp add:atLeast0AtMost setsum_shift_lb_Suc0_0 numeral_2_eq_2)
-
-lemma nat2_to_nat_surj: "surj nat2_to_nat"
-proof (unfold surj_def)
- {
- fix z::nat
- def r \<equiv> "Max {r. (\<Sum>i\<le>r. i) \<le> z}"
- def x \<equiv> "z - (\<Sum>i\<le>r. i)"
-
- hence "finite {r. (\<Sum>i\<le>r. i) \<le> z}"
- by (simp add: lessThan_Suc_atMost[symmetric] lessThan_Suc finite_less_ub)
- also have "0 \<in> {r. (\<Sum>i\<le>r. i) \<le> z}" by simp
- hence "{r::nat. (\<Sum>i\<le>r. i) \<le> z} \<noteq> {}" by fast
- ultimately have a: "r \<in> {r. (\<Sum>i\<le>r. i) \<le> z} \<and> (\<forall>s \<in> {r. (\<Sum>i\<le>r. i) \<le> z}. s \<le> r)"
- by (simp add: r_def del:mem_Collect_eq)
- {
- assume "r<x"
- hence "r+1\<le>x" by simp
- hence "(\<Sum>i\<le>r. i)+(r+1)\<le>z" using x_def by arith
- hence "(r+1) \<in> {r. (\<Sum>i\<le>r. i) \<le> z}" by simp
- with a have "(r+1)\<le>r" by simp
- }
- hence b: "x\<le>r" by force
-
- def y \<equiv> "r-x"
- have "2*z=2*(\<Sum>i\<le>r. i)+2*x" using x_def a by simp arith
- also have "\<dots> = r * (r+1) + 2*x" using gauss_sum_nat_upto by simp
- also have "\<dots> = (x+y)*(x+y+1)+2*x" using y_def b by simp
- also { have "2 dvd ((x+y)*(x+y+1))" using dvd2_a_x_suc_a by simp }
- hence "\<dots> = 2 * nat2_to_nat(x,y)"
- using nat2_to_nat_def by (simp add: Let_def dvd_mult_div_cancel)
- finally have "z=nat2_to_nat (x, y)" by simp
- }
- thus "\<forall>y. \<exists>x. y = nat2_to_nat x" by fast
-qed
-
-lemma nat_to_nat2_inj: "inj nat_to_nat2"
- by (simp add: nat_to_nat2_def surj_imp_inj_inv nat2_to_nat_surj)
-
-
-subsection{* A bijection between @{text "\<nat>"} and @{text "\<int>"} *}
-
-definition nat_to_int_bij :: "nat \<Rightarrow> int" where
-"nat_to_int_bij n = (if 2 dvd n then int(n div 2) else -int(Suc n div 2))"
-
-definition int_to_nat_bij :: "int \<Rightarrow> nat" where
-"int_to_nat_bij i = (if 0<=i then 2*nat(i) else 2*nat(-i) - 1)"
-
-lemma i2n_n2i_id: "int_to_nat_bij (nat_to_int_bij n) = n"
-by (simp add: int_to_nat_bij_def nat_to_int_bij_def) presburger
-
-lemma n2i_i2n_id: "nat_to_int_bij(int_to_nat_bij i) = i"
-proof -
- have "ALL m n::nat. m>0 \<longrightarrow> 2 * m - Suc 0 \<noteq> 2 * n" by presburger
- thus ?thesis
- by(simp add: nat_to_int_bij_def int_to_nat_bij_def, simp add:dvd_def)
-qed
-
-lemma inv_nat_to_int_bij: "inv nat_to_int_bij = int_to_nat_bij"
-by (simp add: i2n_n2i_id inv_equality n2i_i2n_id)
-
-lemma inv_int_to_nat_bij: "inv int_to_nat_bij = nat_to_int_bij"
-by (simp add: i2n_n2i_id inv_equality n2i_i2n_id)
-
-lemma surj_nat_to_int_bij: "surj nat_to_int_bij"
-by (blast intro: n2i_i2n_id surjI)
-
-lemma surj_int_to_nat_bij: "surj int_to_nat_bij"
-by (blast intro: i2n_n2i_id surjI)
-
-lemma inj_nat_to_int_bij: "inj nat_to_int_bij"
-by(simp add:inv_int_to_nat_bij[symmetric] surj_int_to_nat_bij surj_imp_inj_inv)
-
-lemma inj_int_to_nat_bij: "inj int_to_nat_bij"
-by(simp add:inv_nat_to_int_bij[symmetric] surj_nat_to_int_bij surj_imp_inj_inv)
-
-lemma bij_nat_to_int_bij: "bij nat_to_int_bij"
-by(simp add:bij_def inj_nat_to_int_bij surj_nat_to_int_bij)
-
-lemma bij_int_to_nat_bij: "bij int_to_nat_bij"
-by(simp add:bij_def inj_int_to_nat_bij surj_int_to_nat_bij)
-
-
-end
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Mar 12 16:02:42 2010 +0100
@@ -94,8 +94,8 @@
using lem1[unfolded lem3 lem2 lem5] by auto
have even_minus_odd:"\<And>x y. even x \<Longrightarrow> odd (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
have odd_minus_even:"\<And>x y. odd x \<Longrightarrow> even (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
- show ?thesis unfolding even_nat_def unfolding card_def and lem4[THEN sym] and *[unfolded card_def]
- unfolding card_def[THEN sym] apply(rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)
+ show ?thesis unfolding even_nat_def unfolding card_eq_setsum and lem4[THEN sym] and *[unfolded card_eq_setsum]
+ unfolding card_eq_setsum[THEN sym] apply (rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)
apply(rule assms(7)[unfolded even_nat_def]) unfolding int_mult by auto qed
subsection {* The odd/even result for faces of complete vertices, generalized. *}
--- a/src/HOL/Multivariate_Analysis/Integration.cert Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.cert Fri Mar 12 16:02:42 2010 +0100
@@ -3294,3 +3294,215 @@
else -> val!7
}
sat
+qkVrmXMcHAG5MLuJ9d8jXQ 211 0
+#2 := false
+#33 := 0::real
+decl uf_11 :: (-> T5 T6 real)
+decl uf_15 :: T6
+#28 := uf_15
+decl uf_16 :: T5
+#30 := uf_16
+#31 := (uf_11 uf_16 uf_15)
+decl uf_12 :: (-> T7 T8 T5)
+decl uf_14 :: T8
+#26 := uf_14
+decl uf_13 :: (-> T1 T7)
+decl uf_8 :: T1
+#16 := uf_8
+#25 := (uf_13 uf_8)
+#27 := (uf_12 #25 uf_14)
+#29 := (uf_11 #27 uf_15)
+#73 := -1::real
+#84 := (* -1::real #29)
+#85 := (+ #84 #31)
+#74 := (* -1::real #31)
+#75 := (+ #29 #74)
+#112 := (>= #75 0::real)
+#119 := (ite #112 #75 #85)
+#127 := (* -1::real #119)
+decl uf_17 :: T5
+#37 := uf_17
+#38 := (uf_11 uf_17 uf_15)
+#102 := -1/3::real
+#103 := (* -1/3::real #38)
+#128 := (+ #103 #127)
+#100 := 1/3::real
+#101 := (* 1/3::real #31)
+#129 := (+ #101 #128)
+#130 := (<= #129 0::real)
+#131 := (not #130)
+#40 := 3::real
+#39 := (- #31 #38)
+#41 := (/ #39 3::real)
+#32 := (- #29 #31)
+#35 := (- #32)
+#34 := (< #32 0::real)
+#36 := (ite #34 #35 #32)
+#42 := (< #36 #41)
+#136 := (iff #42 #131)
+#104 := (+ #101 #103)
+#78 := (< #75 0::real)
+#90 := (ite #78 #85 #75)
+#109 := (< #90 #104)
+#134 := (iff #109 #131)
+#124 := (< #119 #104)
+#132 := (iff #124 #131)
+#133 := [rewrite]: #132
+#125 := (iff #109 #124)
+#122 := (= #90 #119)
+#113 := (not #112)
+#116 := (ite #113 #85 #75)
+#120 := (= #116 #119)
+#121 := [rewrite]: #120
+#117 := (= #90 #116)
+#114 := (iff #78 #113)
+#115 := [rewrite]: #114
+#118 := [monotonicity #115]: #117
+#123 := [trans #118 #121]: #122
+#126 := [monotonicity #123]: #125
+#135 := [trans #126 #133]: #134
+#110 := (iff #42 #109)
+#107 := (= #41 #104)
+#93 := (* -1::real #38)
+#94 := (+ #31 #93)
+#97 := (/ #94 3::real)
+#105 := (= #97 #104)
+#106 := [rewrite]: #105
+#98 := (= #41 #97)
+#95 := (= #39 #94)
+#96 := [rewrite]: #95
+#99 := [monotonicity #96]: #98
+#108 := [trans #99 #106]: #107
+#91 := (= #36 #90)
+#76 := (= #32 #75)
+#77 := [rewrite]: #76
+#88 := (= #35 #85)
+#81 := (- #75)
+#86 := (= #81 #85)
+#87 := [rewrite]: #86
+#82 := (= #35 #81)
+#83 := [monotonicity #77]: #82
+#89 := [trans #83 #87]: #88
+#79 := (iff #34 #78)
+#80 := [monotonicity #77]: #79
+#92 := [monotonicity #80 #89 #77]: #91
+#111 := [monotonicity #92 #108]: #110
+#137 := [trans #111 #135]: #136
+#72 := [asserted]: #42
+#138 := [mp #72 #137]: #131
+decl uf_1 :: T1
+#4 := uf_1
+#43 := (uf_13 uf_1)
+#44 := (uf_12 #43 uf_14)
+#45 := (uf_11 #44 uf_15)
+#149 := (* -1::real #45)
+#150 := (+ #38 #149)
+#140 := (+ #93 #45)
+#161 := (<= #150 0::real)
+#168 := (ite #161 #140 #150)
+#176 := (* -1::real #168)
+#177 := (+ #103 #176)
+#178 := (+ #101 #177)
+#179 := (<= #178 0::real)
+#180 := (not #179)
+#46 := (- #45 #38)
+#48 := (- #46)
+#47 := (< #46 0::real)
+#49 := (ite #47 #48 #46)
+#50 := (< #49 #41)
+#185 := (iff #50 #180)
+#143 := (< #140 0::real)
+#155 := (ite #143 #150 #140)
+#158 := (< #155 #104)
+#183 := (iff #158 #180)
+#173 := (< #168 #104)
+#181 := (iff #173 #180)
+#182 := [rewrite]: #181
+#174 := (iff #158 #173)
+#171 := (= #155 #168)
+#162 := (not #161)
+#165 := (ite #162 #150 #140)
+#169 := (= #165 #168)
+#170 := [rewrite]: #169
+#166 := (= #155 #165)
+#163 := (iff #143 #162)
+#164 := [rewrite]: #163
+#167 := [monotonicity #164]: #166
+#172 := [trans #167 #170]: #171
+#175 := [monotonicity #172]: #174
+#184 := [trans #175 #182]: #183
+#159 := (iff #50 #158)
+#156 := (= #49 #155)
+#141 := (= #46 #140)
+#142 := [rewrite]: #141
+#153 := (= #48 #150)
+#146 := (- #140)
+#151 := (= #146 #150)
+#152 := [rewrite]: #151
+#147 := (= #48 #146)
+#148 := [monotonicity #142]: #147
+#154 := [trans #148 #152]: #153
+#144 := (iff #47 #143)
+#145 := [monotonicity #142]: #144
+#157 := [monotonicity #145 #154 #142]: #156
+#160 := [monotonicity #157 #108]: #159
+#186 := [trans #160 #184]: #185
+#139 := [asserted]: #50
+#187 := [mp #139 #186]: #180
+#299 := (+ #140 #176)
+#300 := (<= #299 0::real)
+#290 := (= #140 #168)
+#329 := [hypothesis]: #162
+#191 := (+ #29 #149)
+#192 := (<= #191 0::real)
+#51 := (<= #29 #45)
+#193 := (iff #51 #192)
+#194 := [rewrite]: #193
+#188 := [asserted]: #51
+#195 := [mp #188 #194]: #192
+#298 := (+ #75 #127)
+#301 := (<= #298 0::real)
+#284 := (= #75 #119)
+#302 := [hypothesis]: #113
+#296 := (+ #85 #127)
+#297 := (<= #296 0::real)
+#285 := (= #85 #119)
+#288 := (or #112 #285)
+#289 := [def-axiom]: #288
+#303 := [unit-resolution #289 #302]: #285
+#304 := (not #285)
+#305 := (or #304 #297)
+#306 := [th-lemma]: #305
+#307 := [unit-resolution #306 #303]: #297
+#315 := (not #290)
+#310 := (not #300)
+#311 := (or #310 #112)
+#308 := [hypothesis]: #300
+#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
+#312 := [lemma #309]: #311
+#322 := [unit-resolution #312 #302]: #310
+#316 := (or #315 #300)
+#313 := [hypothesis]: #310
+#314 := [hypothesis]: #290
+#317 := [th-lemma]: #316
+#318 := [unit-resolution #317 #314 #313]: false
+#319 := [lemma #318]: #316
+#323 := [unit-resolution #319 #322]: #315
+#292 := (or #162 #290)
+#293 := [def-axiom]: #292
+#324 := [unit-resolution #293 #323]: #162
+#325 := [th-lemma #324 #307 #138 #302 #195]: false
+#326 := [lemma #325]: #112
+#286 := (or #113 #284)
+#287 := [def-axiom]: #286
+#330 := [unit-resolution #287 #326]: #284
+#331 := (not #284)
+#332 := (or #331 #301)
+#333 := [th-lemma]: #332
+#334 := [unit-resolution #333 #330]: #301
+#335 := [th-lemma #326 #334 #195 #329 #138]: false
+#336 := [lemma #335]: #161
+#327 := [unit-resolution #293 #336]: #290
+#328 := [unit-resolution #319 #327]: #300
+[th-lemma #326 #334 #195 #328 #187 #138]: false
+unsat
--- a/src/HOL/Nat_Transfer.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nat_Transfer.thy Fri Mar 12 16:02:42 2010 +0100
@@ -27,17 +27,17 @@
text {* set up transfer direction *}
lemma transfer_morphism_nat_int: "transfer_morphism nat (op <= (0::int))"
- by (simp add: transfer_morphism_def)
+ by (fact transfer_morphismI)
-declare transfer_morphism_nat_int [transfer
- add mode: manual
+declare transfer_morphism_nat_int [transfer add
+ mode: manual
return: nat_0_le
- labels: natint
+ labels: nat_int
]
text {* basic functions and relations *}
-lemma transfer_nat_int_numerals:
+lemma transfer_nat_int_numerals [transfer key: transfer_morphism_nat_int]:
"(0::nat) = nat 0"
"(1::nat) = nat 1"
"(2::nat) = nat 2"
@@ -52,8 +52,7 @@
lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y"
by (simp add: tsub_def)
-
-lemma transfer_nat_int_functions:
+lemma transfer_nat_int_functions [transfer key: transfer_morphism_nat_int]:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)"
@@ -61,7 +60,7 @@
by (auto simp add: eq_nat_nat_iff nat_mult_distrib
nat_power_eq tsub_def)
-lemma transfer_nat_int_function_closures:
+lemma transfer_nat_int_function_closures [transfer key: transfer_morphism_nat_int]:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0"
@@ -73,7 +72,7 @@
"int z >= 0"
by (auto simp add: zero_le_mult_iff tsub_def)
-lemma transfer_nat_int_relations:
+lemma transfer_nat_int_relations [transfer key: transfer_morphism_nat_int]:
"x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
(nat (x::int) = nat y) = (x = y)"
"x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
@@ -84,13 +83,6 @@
(nat (x::int) dvd nat y) = (x dvd y)"
by (auto simp add: zdvd_int)
-declare transfer_morphism_nat_int [transfer add return:
- transfer_nat_int_numerals
- transfer_nat_int_functions
- transfer_nat_int_function_closures
- transfer_nat_int_relations
-]
-
text {* first-order quantifiers *}
@@ -108,7 +100,7 @@
then show "\<exists>x. P x" by auto
qed
-lemma transfer_nat_int_quantifiers:
+lemma transfer_nat_int_quantifiers [transfer key: transfer_morphism_nat_int]:
"(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))"
"(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))"
by (rule all_nat, rule ex_nat)
@@ -123,18 +115,15 @@
by auto
declare transfer_morphism_nat_int [transfer add
- return: transfer_nat_int_quantifiers
cong: all_cong ex_cong]
text {* if *}
-lemma nat_if_cong: "(if P then (nat x) else (nat y)) =
- nat (if P then x else y)"
+lemma nat_if_cong [transfer key: transfer_morphism_nat_int]:
+ "(if P then (nat x) else (nat y)) = nat (if P then x else y)"
by auto
-declare transfer_morphism_nat_int [transfer add return: nat_if_cong]
-
text {* operations with sets *}
@@ -276,22 +265,18 @@
text {* set up transfer direction *}
-lemma transfer_morphism_int_nat: "transfer_morphism int (UNIV :: nat set)"
- by (simp add: transfer_morphism_def)
+lemma transfer_morphism_int_nat: "transfer_morphism int (\<lambda>n. True)"
+ by (fact transfer_morphismI)
declare transfer_morphism_int_nat [transfer add
mode: manual
-(* labels: int-nat *)
return: nat_int
+ labels: int_nat
]
text {* basic functions and relations *}
-lemma UNIV_apply:
- "UNIV x = True"
- by (simp add: top_fun_eq top_bool_eq)
-
definition
is_nat :: "int \<Rightarrow> bool"
where
@@ -335,7 +320,6 @@
transfer_int_nat_functions
transfer_int_nat_function_closures
transfer_int_nat_relations
- UNIV_apply
]
--- a/src/HOL/Nitpick.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nitpick.thy Fri Mar 12 16:02:42 2010 +0100
@@ -36,7 +36,8 @@
and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
and bisim_iterator_max :: bisim_iterator
and Quot :: "'a \<Rightarrow> 'b"
- and Tha :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
+ and safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
+ and safe_Eps :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
datatype ('a, 'b) fin_fun = FinFun "('a \<Rightarrow> 'b)"
datatype ('a, 'b) fun_box = FunBox "('a \<Rightarrow> 'b)"
@@ -95,10 +96,10 @@
else THE y. wfrec_rel R (%f x. F (Recdef.cut f R x) x) x y"
definition card' :: "('a \<Rightarrow> bool) \<Rightarrow> nat" where
-"card' X \<equiv> length (SOME xs. set xs = X \<and> distinct xs)"
+"card' A \<equiv> if finite A then length (safe_Eps (\<lambda>xs. set xs = A \<and> distinct xs)) else 0"
definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b" where
-"setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
+"setsum' f A \<equiv> if finite A then listsum (map f (safe_Eps (\<lambda>xs. set xs = A \<and> distinct xs))) else 0"
inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" where
"fold_graph' f z {} z" |
@@ -237,10 +238,11 @@
setup {* Nitpick_Isar.setup *}
hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim
- bisim_iterator_max Quot Tha FinFun FunBox PairBox Word refl' wf' wf_wfrec
- wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm
- Frac Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac plus_frac
- times_frac uminus_frac number_of_frac inverse_frac less_eq_frac of_frac
+ bisim_iterator_max Quot safe_The safe_Eps FinFun FunBox PairBox Word refl'
+ wf' wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm
+ int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom
+ norm_frac frac plus_frac times_frac uminus_frac number_of_frac inverse_frac
+ less_eq_frac of_frac
hide (open) type bisim_iterator fin_fun fun_box pair_box unsigned_bit signed_bit
word
hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def
--- a/src/HOL/Nitpick_Examples/Induct_Nits.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Induct_Nits.thy Fri Mar 12 16:02:42 2010 +0100
@@ -61,7 +61,7 @@
lemma "q2 = {}"
nitpick [expect = genuine]
nitpick [dont_star_linear_preds, expect = genuine]
-nitpick [wf, expect = likely_genuine]
+nitpick [wf, expect = quasi_genuine]
oops
lemma "p2 = UNIV"
@@ -72,7 +72,7 @@
lemma "q2 = UNIV"
nitpick [expect = none]
nitpick [dont_star_linear_preds, expect = none]
-nitpick [wf, expect = likely_genuine]
+nitpick [wf, expect = quasi_genuine]
sorry
lemma "p2 = q2"
--- a/src/HOL/Nitpick_Examples/Integer_Nits.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Integer_Nits.thy Fri Mar 12 16:02:42 2010 +0100
@@ -206,4 +206,39 @@
nitpick [binary_ints, expect = none]
sorry
+datatype tree = Null | Node nat tree tree
+
+primrec labels where
+"labels Null = {}" |
+"labels (Node x t u) = {x} \<union> labels t \<union> labels u"
+
+lemma "labels (Node x t u) \<noteq> labels (Node y v w)"
+nitpick [expect = genuine]
+nitpick [dont_finitize, expect = potential]
+oops
+
+lemma "labels (Node x t u) \<noteq> {}"
+nitpick [expect = none]
+oops
+
+lemma "card (labels t) > 0"
+nitpick [expect = genuine]
+nitpick [dont_finitize, expect = potential]
+oops
+
+lemma "(\<Sum>n \<in> labels t. n + 2) \<ge> 2"
+nitpick [expect = genuine]
+nitpick [dont_finitize, expect = potential]
+oops
+
+lemma "t \<noteq> Null \<Longrightarrow> (\<Sum>n \<in> labels t. n + 2) \<ge> 2"
+nitpick [expect = none]
+nitpick [dont_finitize, expect = none]
+sorry
+
+lemma "(\<Sum>i \<in> labels (Node x t u). f i\<Colon>nat) = f x + (\<Sum>i \<in> labels t. f i) + (\<Sum>i \<in> labels u. f i)"
+nitpick [expect = genuine]
+nitpick [dont_finitize, expect = none]
+oops
+
end
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Fri Mar 12 16:02:42 2010 +0100
@@ -13,34 +13,34 @@
chapter {* 3. First Steps *}
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 15 s]
+nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1]
subsection {* 3.1. Propositional Logic *}
lemma "P \<longleftrightarrow> Q"
-nitpick
+nitpick [expect = genuine]
apply auto
-nitpick 1
-nitpick 2
+nitpick [expect = genuine] 1
+nitpick [expect = genuine] 2
oops
subsection {* 3.2. Type Variables *}
lemma "P x \<Longrightarrow> P (THE y. P y)"
-nitpick [verbose]
+nitpick [verbose, expect = genuine]
oops
subsection {* 3.3. Constants *}
lemma "P x \<Longrightarrow> P (THE y. P y)"
-nitpick [show_consts]
-nitpick [full_descrs, show_consts]
-nitpick [dont_specialize, full_descrs, show_consts]
+nitpick [show_consts, expect = genuine]
+nitpick [full_descrs, show_consts, expect = genuine]
+nitpick [dont_specialize, full_descrs, show_consts, expect = genuine]
oops
lemma "\<exists>!x. P x \<Longrightarrow> P (THE y. P y)"
-nitpick
-nitpick [card 'a = 1-50]
+nitpick [expect = none]
+nitpick [card 'a = 1\<midarrow>50, expect = none]
(* sledgehammer *)
apply (metis the_equality)
done
@@ -48,45 +48,45 @@
subsection {* 3.4. Skolemization *}
lemma "\<exists>g. \<forall>x. g (f x) = x \<Longrightarrow> \<forall>y. \<exists>x. y = f x"
-nitpick
+nitpick [expect = genuine]
oops
lemma "\<exists>x. \<forall>f. f x = x"
-nitpick
+nitpick [expect = genuine]
oops
lemma "refl r \<Longrightarrow> sym r"
-nitpick
+nitpick [expect = genuine]
oops
subsection {* 3.5. Natural Numbers and Integers *}
lemma "\<lbrakk>i \<le> j; n \<le> (m\<Colon>int)\<rbrakk> \<Longrightarrow> i * n + j * m \<le> i * m + j * n"
-nitpick
+nitpick [expect = genuine]
oops
lemma "\<forall>n. Suc n \<noteq> n \<Longrightarrow> P"
-nitpick [card nat = 100, check_potential]
+nitpick [card nat = 100, check_potential, expect = genuine]
oops
lemma "P Suc"
-nitpick
+nitpick [expect = none]
oops
lemma "P (op +\<Colon>nat\<Rightarrow>nat\<Rightarrow>nat)"
-nitpick [card nat = 1]
-nitpick [card nat = 2]
+nitpick [card nat = 1, expect = genuine]
+nitpick [card nat = 2, expect = none]
oops
subsection {* 3.6. Inductive Datatypes *}
lemma "hd (xs @ [y, y]) = hd xs"
-nitpick
-nitpick [show_consts, show_datatypes]
+nitpick [expect = genuine]
+nitpick [show_consts, show_datatypes, expect = genuine]
oops
lemma "\<lbrakk>length xs = 1; length ys = 1\<rbrakk> \<Longrightarrow> xs = ys"
-nitpick [show_datatypes]
+nitpick [show_datatypes, expect = genuine]
oops
subsection {* 3.7. Typedefs, Records, Rationals, and Reals *}
@@ -99,7 +99,7 @@
definition C :: three where "C \<equiv> Abs_three 2"
lemma "\<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P x"
-nitpick [show_datatypes]
+nitpick [show_datatypes, expect = genuine]
oops
fun my_int_rel where
@@ -114,6 +114,20 @@
quotient_definition "add\<Colon>my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw
lemma "add x y = add x x"
+nitpick [show_datatypes, expect = genuine]
+oops
+
+ML {*
+(* Proof.context -> string -> (typ -> term list) -> typ -> term -> term *)
+fun my_int_postproc _ _ _ T (Const _ $ (Const _ $ t1 $ t2)) =
+ HOLogic.mk_number T (snd (HOLogic.dest_number t1)
+ - snd (HOLogic.dest_number t2))
+ | my_int_postproc _ _ _ _ t = t
+*}
+
+setup {* Nitpick.register_term_postprocessor @{typ my_int} my_int_postproc *}
+
+lemma "add x y = add x x"
nitpick [show_datatypes]
oops
@@ -122,11 +136,11 @@
Ycoord :: int
lemma "Xcoord (p\<Colon>point) = Xcoord (q\<Colon>point)"
-nitpick [show_datatypes]
+nitpick [show_datatypes, expect = genuine]
oops
lemma "4 * x + 3 * (y\<Colon>real) \<noteq> 1 / 2"
-nitpick [show_datatypes]
+nitpick [show_datatypes, expect = genuine]
oops
subsection {* 3.8. Inductive and Coinductive Predicates *}
@@ -136,11 +150,11 @@
"even n \<Longrightarrow> even (Suc (Suc n))"
lemma "\<exists>n. even n \<and> even (Suc n)"
-nitpick [card nat = 100, unary_ints, verbose]
+nitpick [card nat = 50, unary_ints, verbose, expect = potential]
oops
-lemma "\<exists>n \<le> 99. even n \<and> even (Suc n)"
-nitpick [card nat = 100, unary_ints, verbose]
+lemma "\<exists>n \<le> 49. even n \<and> even (Suc n)"
+nitpick [card nat = 50, unary_ints, verbose, expect = genuine]
oops
inductive even' where
@@ -149,18 +163,18 @@
"\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)"
lemma "\<exists>n \<in> {0, 2, 4, 6, 8}. \<not> even' n"
-nitpick [card nat = 10, unary_ints, verbose, show_consts]
+nitpick [card nat = 10, unary_ints, verbose, show_consts, expect = genuine]
oops
lemma "even' (n - 2) \<Longrightarrow> even' n"
-nitpick [card nat = 10, show_consts]
+nitpick [card nat = 10, show_consts, expect = genuine]
oops
coinductive nats where
"nats (x\<Colon>nat) \<Longrightarrow> nats x"
lemma "nats = {0, 1, 2, 3, 4}"
-nitpick [card nat = 10, show_consts]
+nitpick [card nat = 10, show_consts, expect = genuine]
oops
inductive odd where
@@ -168,7 +182,7 @@
"\<lbrakk>odd m; even n\<rbrakk> \<Longrightarrow> odd (m + n)"
lemma "odd n \<Longrightarrow> odd (n - 2)"
-nitpick [card nat = 10, show_consts]
+nitpick [card nat = 10, show_consts, expect = genuine]
oops
subsection {* 3.9. Coinductive Datatypes *}
@@ -179,7 +193,8 @@
typedef 'a llist = "UNIV\<Colon>('a list + (nat \<Rightarrow> 'a)) set" by auto
-definition LNil where "LNil = Abs_llist (Inl [])"
+definition LNil where
+"LNil = Abs_llist (Inl [])"
definition LCons where
"LCons y ys = Abs_llist (case Rep_llist ys of
Inl ys' \<Rightarrow> Inl (y # ys')
@@ -197,16 +212,16 @@
*}
lemma "xs \<noteq> LCons a xs"
-nitpick
+nitpick [expect = genuine]
oops
lemma "\<lbrakk>xs = LCons a xs; ys = iterates (\<lambda>b. a) b\<rbrakk> \<Longrightarrow> xs = ys"
-nitpick [verbose]
+nitpick [verbose, expect = genuine]
oops
lemma "\<lbrakk>xs = LCons a xs; ys = LCons a ys\<rbrakk> \<Longrightarrow> xs = ys"
-nitpick [bisim_depth = -1, show_datatypes]
-nitpick
+nitpick [bisim_depth = -1, show_datatypes, expect = quasi_genuine]
+nitpick [expect = none]
sorry
subsection {* 3.10. Boxing *}
@@ -230,9 +245,9 @@
"subst\<^isub>1 \<sigma> (App t u) = App (subst\<^isub>1 \<sigma> t) (subst\<^isub>1 \<sigma> u)"
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>1 \<sigma> t = t"
-nitpick [verbose]
-nitpick [eval = "subst\<^isub>1 \<sigma> t"]
-(* nitpick [dont_box] *)
+nitpick [verbose, expect = genuine]
+nitpick [eval = "subst\<^isub>1 \<sigma> t", expect = genuine]
+(* nitpick [dont_box, expect = unknown] *)
oops
primrec subst\<^isub>2 where
@@ -242,19 +257,19 @@
"subst\<^isub>2 \<sigma> (App t u) = App (subst\<^isub>2 \<sigma> t) (subst\<^isub>2 \<sigma> u)"
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>2 \<sigma> t = t"
-nitpick [card = 1\<midarrow>6]
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
subsection {* 3.11. Scope Monotonicity *}
lemma "length xs = length ys \<Longrightarrow> rev (zip xs ys) = zip xs (rev ys)"
-nitpick [verbose]
-nitpick [card = 8, verbose]
+nitpick [verbose, expect = genuine]
+nitpick [card = 8, verbose, expect = genuine]
oops
lemma "\<exists>g. \<forall>x\<Colon>'b. g (f x) = x \<Longrightarrow> \<forall>y\<Colon>'a. \<exists>x. y = f x"
-nitpick [mono]
-nitpick
+nitpick [mono, expect = none]
+nitpick [expect = genuine]
oops
subsection {* 3.12. Inductive Properties *}
@@ -265,21 +280,21 @@
"n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
lemma "n \<in> reach \<Longrightarrow> 2 dvd n"
-nitpick [unary_ints]
+nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (induct set: reach)
apply auto
- nitpick
+ nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (thin_tac "n \<in> reach")
- nitpick
+ nitpick [expect = genuine]
oops
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0"
-nitpick [unary_ints]
+nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (induct set: reach)
apply auto
- nitpick
+ nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (thin_tac "n \<in> reach")
- nitpick
+ nitpick [expect = genuine]
oops
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<ge> 4"
@@ -297,13 +312,13 @@
"swap (Branch t u) a b = Branch (swap t a b) (swap u a b)"
lemma "{a, b} \<subseteq> labels t \<Longrightarrow> labels (swap t a b) = labels t"
-nitpick
+(* nitpick *)
proof (induct t)
case Leaf thus ?case by simp
next
case (Branch t u) thus ?case
- nitpick
- nitpick [non_std, show_all]
+ (* nitpick *)
+ nitpick [non_std, show_all, expect = genuine]
oops
lemma "labels (swap t a b) =
@@ -316,7 +331,7 @@
case Leaf thus ?case by simp
next
case (Branch t u) thus ?case
- nitpick [non_std, show_all]
+ nitpick [non_std, card = 1\<midarrow>5, expect = none]
by auto
qed
@@ -338,7 +353,7 @@
theorem S\<^isub>1_sound:
"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick
+nitpick [expect = genuine]
oops
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
@@ -351,7 +366,7 @@
theorem S\<^isub>2_sound:
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick
+nitpick [expect = genuine]
oops
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
@@ -364,12 +379,12 @@
theorem S\<^isub>3_sound:
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick
+nitpick [card = 1\<midarrow>6, expect = none]
sorry
theorem S\<^isub>3_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
-nitpick
+nitpick [expect = genuine]
oops
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
@@ -383,19 +398,19 @@
theorem S\<^isub>4_sound:
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick
+nitpick [card = 1\<midarrow>6, expect = none]
sorry
theorem S\<^isub>4_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-nitpick
+nitpick [card = 1\<midarrow>6, expect = none]
sorry
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
-nitpick
+nitpick [card = 1\<midarrow>6, expect = none]
sorry
subsection {* 4.2. AA Trees *}
@@ -449,13 +464,13 @@
theorem dataset_skew_split:
"dataset (skew t) = dataset t"
"dataset (split t) = dataset t"
-nitpick
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
theorem wf_skew_split:
"wf t \<Longrightarrow> skew t = t"
"wf t \<Longrightarrow> split t = t"
-nitpick
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
primrec insort\<^isub>1 where
@@ -465,11 +480,11 @@
(if x > y then insort\<^isub>1 u x else u))"
theorem wf_insort\<^isub>1: "wf t \<Longrightarrow> wf (insort\<^isub>1 t x)"
-nitpick
+nitpick [expect = genuine]
oops
theorem wf_insort\<^isub>1_nat: "wf t \<Longrightarrow> wf (insort\<^isub>1 t (x\<Colon>nat))"
-nitpick [eval = "insort\<^isub>1 t x"]
+nitpick [eval = "insort\<^isub>1 t x", expect = genuine]
oops
primrec insort\<^isub>2 where
@@ -479,11 +494,11 @@
(if x > y then insort\<^isub>2 u x else u))"
theorem wf_insort\<^isub>2: "wf t \<Longrightarrow> wf (insort\<^isub>2 t x)"
-nitpick
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
theorem dataset_insort\<^isub>2: "dataset (insort\<^isub>2 t x) = {x} \<union> dataset t"
-nitpick
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
end
--- a/src/HOL/Nitpick_Examples/Mini_Nits.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Fri Mar 12 16:02:42 2010 +0100
@@ -24,8 +24,6 @@
fun unknown n t = (minipick n t = "unknown" orelse raise FAIL)
*}
-ML {* minipick 1 @{prop "\<forall>x\<Colon>'a. \<exists>y\<Colon>'b. f x = y"} *}
-
ML {* genuine 1 @{prop "x = Not"} *}
ML {* none 1 @{prop "\<exists>x. x = Not"} *}
ML {* none 1 @{prop "\<not> False"} *}
--- a/src/HOL/Nominal/nominal_fresh_fun.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML Fri Mar 12 16:02:42 2010 +0100
@@ -58,7 +58,7 @@
val goal = List.nth(prems_of thm, i-1);
val ps = Logic.strip_params goal;
val Ts = rev (map snd ps);
- fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]);
+ fun is_of_fs_name T = Sign.of_sort thy (T, [Sign.intern_class thy ("fs_"^atom_basename)]);
(* rebuild de bruijn indices *)
val bvs = map_index (Bound o fst) ps;
(* select variables of the right class *)
--- a/src/HOL/Number_Theory/Binomial.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Number_Theory/Binomial.thy Fri Mar 12 16:02:42 2010 +0100
@@ -364,7 +364,7 @@
finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
card F choose (k + 1) + (card F choose k)".
with iassms choose_plus_one_nat show ?thesis
- by auto
+ by (auto simp del: card.insert)
qed
qed
qed
--- a/src/HOL/Option.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Option.thy Fri Mar 12 16:02:42 2010 +0100
@@ -5,7 +5,7 @@
header {* Datatype option *}
theory Option
-imports Datatype Finite_Set
+imports Datatype
begin
datatype 'a option = None | Some 'a
@@ -33,13 +33,6 @@
lemma UNIV_option_conv: "UNIV = insert None (range Some)"
by(auto intro: classical)
-lemma finite_option_UNIV[simp]:
- "finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"
-by(auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some)
-
-instance option :: (finite) finite proof
-qed (simp add: UNIV_option_conv)
-
subsubsection {* Operations *}
--- a/src/HOL/Probability/Borel.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Probability/Borel.thy Fri Mar 12 16:02:42 2010 +0100
@@ -15,11 +15,6 @@
definition indicator_fn where
"indicator_fn s = (\<lambda>x. if x \<in> s then 1 else (0::real))"
-definition mono_convergent where
- "mono_convergent u f s \<equiv>
- (\<forall>x\<in>s. incseq (\<lambda>n. u n x)) \<and>
- (\<forall>x \<in> s. (\<lambda>i. u i x) ----> f x)"
-
lemma in_borel_measurable:
"f \<in> borel_measurable M \<longleftrightarrow>
sigma_algebra M \<and>
@@ -191,20 +186,20 @@
definition
nat_to_rat_surj :: "nat \<Rightarrow> rat" where
- "nat_to_rat_surj n = (let (i,j) = nat_to_nat2 n
- in Fract (nat_to_int_bij i) (nat_to_int_bij j))"
+ "nat_to_rat_surj n = (let (i,j) = prod_decode n
+ in Fract (int_decode i) (int_decode j))"
lemma nat_to_rat_surj: "surj nat_to_rat_surj"
proof (auto simp add: surj_def nat_to_rat_surj_def)
fix y
- show "\<exists>x. y = (\<lambda>(i, j). Fract (nat_to_int_bij i) (nat_to_int_bij j)) (nat_to_nat2 x)"
+ show "\<exists>x. y = (\<lambda>(i, j). Fract (int_decode i) (int_decode j)) (prod_decode x)"
proof (cases y)
case (Fract a b)
- obtain i where i: "nat_to_int_bij i = a" using surj_nat_to_int_bij
+ obtain i where i: "int_decode i = a" using surj_int_decode
by (metis surj_def)
- obtain j where j: "nat_to_int_bij j = b" using surj_nat_to_int_bij
+ obtain j where j: "int_decode j = b" using surj_int_decode
by (metis surj_def)
- obtain n where n: "nat_to_nat2 n = (i,j)" using nat_to_nat2_surj
+ obtain n where n: "prod_decode n = (i,j)" using surj_prod_decode
by (metis surj_def)
from Fract i j n show ?thesis
@@ -375,6 +370,103 @@
by (fast intro: borel_measurable_add_borel_measurable
borel_measurable_uminus_borel_measurable f g)
+lemma (in measure_space) borel_measurable_setsum_borel_measurable:
+ assumes s: "finite s"
+ shows "(!!i. i \<in> s ==> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) s) \<in> borel_measurable M" using s
+proof (induct s)
+ case empty
+ thus ?case
+ by (simp add: borel_measurable_const)
+next
+ case (insert x s)
+ thus ?case
+ by (auto simp add: borel_measurable_add_borel_measurable)
+qed
+
+lemma (in measure_space) borel_measurable_cong:
+ assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M"
+using assms unfolding in_borel_measurable by (simp cong: vimage_inter_cong)
+
+lemma (in measure_space) borel_measurable_inverse:
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
+ unfolding borel_measurable_ge_iff
+proof (safe, rule linorder_cases)
+ fix a :: real assume "0 < a"
+ { fix w
+ from `0 < a` have "a \<le> inverse (f w) \<longleftrightarrow> 0 < f w \<and> f w \<le> 1 / a"
+ by (metis inverse_eq_divide inverse_inverse_eq le_imp_inverse_le
+ linorder_not_le real_le_refl real_le_trans real_less_def
+ xt1(7) zero_less_divide_1_iff) }
+ hence "{w \<in> space M. a \<le> inverse (f w)} =
+ {w \<in> space M. 0 < f w} \<inter> {w \<in> space M. f w \<le> 1 / a}" by auto
+ with Int assms[unfolded borel_measurable_gr_iff]
+ assms[unfolded borel_measurable_le_iff]
+ show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
+next
+ fix a :: real assume "0 = a"
+ { fix w have "a \<le> inverse (f w) \<longleftrightarrow> 0 \<le> f w"
+ unfolding `0 = a`[symmetric] by auto }
+ thus "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M"
+ using assms[unfolded borel_measurable_ge_iff] by simp
+next
+ fix a :: real assume "a < 0"
+ { fix w
+ from `a < 0` have "a \<le> inverse (f w) \<longleftrightarrow> f w \<le> 1 / a \<or> 0 \<le> f w"
+ apply (cases "0 \<le> f w")
+ apply (metis inverse_eq_divide linorder_not_le xt1(8) xt1(9)
+ zero_le_divide_1_iff)
+ apply (metis inverse_eq_divide inverse_inverse_eq inverse_le_imp_le_neg
+ linorder_not_le real_le_refl real_le_trans)
+ done }
+ hence "{w \<in> space M. a \<le> inverse (f w)} =
+ {w \<in> space M. f w \<le> 1 / a} \<union> {w \<in> space M. 0 \<le> f w}" by auto
+ with Un assms[unfolded borel_measurable_ge_iff]
+ assms[unfolded borel_measurable_le_iff]
+ show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
+qed
+
+lemma (in measure_space) borel_measurable_divide:
+ assumes "f \<in> borel_measurable M"
+ and "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
+ unfolding field_divide_inverse
+ by (rule borel_measurable_inverse borel_measurable_times_borel_measurable assms)+
+
+lemma (in measure_space) borel_measurable_vimage:
+ assumes borel: "f \<in> borel_measurable M"
+ shows "f -` {X} \<inter> space M \<in> sets M"
+proof -
+ have "{w \<in> space M. f w = X} = {w. f w = X} \<inter> space M" by auto
+ with borel_measurable_eq_borel_measurable[OF borel borel_measurable_const, of X]
+ show ?thesis unfolding vimage_def by simp
+qed
+
+section "Monotone convergence"
+
+definition mono_convergent where
+ "mono_convergent u f s \<equiv>
+ (\<forall>x\<in>s. incseq (\<lambda>n. u n x)) \<and>
+ (\<forall>x \<in> s. (\<lambda>i. u i x) ----> f x)"
+
+definition "upclose f g x = max (f x) (g x)"
+
+primrec mon_upclose where
+"mon_upclose 0 u = u 0" |
+"mon_upclose (Suc n) u = upclose (u (Suc n)) (mon_upclose n u)"
+
+lemma mono_convergentD:
+ assumes "mono_convergent u f s" and "x \<in> s"
+ shows "incseq (\<lambda>n. u n x)" and "(\<lambda>i. u i x) ----> f x"
+ using assms unfolding mono_convergent_def by auto
+
+lemma mono_convergentI:
+ assumes "\<And>x. x \<in> s \<Longrightarrow> incseq (\<lambda>n. u n x)"
+ assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>i. u i x) ----> f x"
+ shows "mono_convergent u f s"
+ using assms unfolding mono_convergent_def by auto
+
lemma (in measure_space) mono_convergent_borel_measurable:
assumes u: "!!n. u n \<in> borel_measurable M"
assumes mc: "mono_convergent u f (space M)"
@@ -409,44 +501,11 @@
by (auto simp add: borel_measurable_le_iff)
qed
-lemma (in measure_space) borel_measurable_setsum_borel_measurable:
- assumes s: "finite s"
- shows "(!!i. i \<in> s ==> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) s) \<in> borel_measurable M" using s
-proof (induct s)
- case empty
- thus ?case
- by (simp add: borel_measurable_const)
-next
- case (insert x s)
- thus ?case
- by (auto simp add: borel_measurable_add_borel_measurable)
-qed
-
-section "Monotone convergence"
-
-lemma mono_convergentD:
- assumes "mono_convergent u f s" and "x \<in> s"
- shows "incseq (\<lambda>n. u n x)" and "(\<lambda>i. u i x) ----> f x"
- using assms unfolding mono_convergent_def by auto
-
-
-lemma mono_convergentI:
- assumes "\<And>x. x \<in> s \<Longrightarrow> incseq (\<lambda>n. u n x)"
- assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>i. u i x) ----> f x"
- shows "mono_convergent u f s"
- using assms unfolding mono_convergent_def by auto
-
lemma mono_convergent_le:
assumes "mono_convergent u f s" and "t \<in> s"
shows "u n t \<le> f t"
using mono_convergentD[OF assms] by (auto intro!: incseq_le)
-definition "upclose f g x = max (f x) (g x)"
-
-primrec mon_upclose where
-"mon_upclose 0 u = u 0" |
-"mon_upclose (Suc n) u = upclose (u (Suc n)) (mon_upclose n u)"
-
lemma mon_upclose_ex:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ('b\<Colon>linorder)"
shows "\<exists>n \<le> m. mon_upclose m u x = u n x"
@@ -561,4 +620,19 @@
qed
qed
+lemma mono_conv_outgrow:
+ assumes "incseq x" "x ----> y" "z < y"
+ shows "\<exists>b. \<forall> a \<ge> b. z < x a"
+using assms
+proof -
+ from assms have "y - z > 0" by simp
+ hence A: "\<exists>n. (\<forall> m \<ge> n. \<bar> x m + - y \<bar> < y - z)" using assms
+ unfolding incseq_def LIMSEQ_def dist_real_def real_diff_def
+ by simp
+ have "\<forall>m. x m \<le> y" using incseq_le assms by auto
+ hence B: "\<forall>m. \<bar> x m + - y \<bar> = y - x m"
+ by (metis abs_if abs_minus_add_cancel less_iff_diff_less_0 linorder_not_le real_diff_def)
+ from A B show ?thesis by auto
+qed
+
end
--- a/src/HOL/Probability/Caratheodory.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Probability/Caratheodory.thy Fri Mar 12 16:02:42 2010 +0100
@@ -816,9 +816,9 @@
by (simp add: eqe)
finally show ?thesis using Series.summable_le2 [OF 1 2] by auto
qed
- def C \<equiv> "(split BB) o nat_to_nat2"
+ def C \<equiv> "(split BB) o prod_decode"
have C: "!!n. C n \<in> sets M"
- apply (rule_tac p="nat_to_nat2 n" in PairE)
+ apply (rule_tac p="prod_decode n" in PairE)
apply (simp add: C_def)
apply (metis BB subsetD rangeI)
done
@@ -828,11 +828,10 @@
assume x: "x \<in> A i"
with sbBB [of i] obtain j where "x \<in> BB i j"
by blast
- thus "\<exists>i. x \<in> split BB (nat_to_nat2 i)"
- by (metis nat_to_nat2_surj internal_split_def prod.cases
- prod_case_split surj_f_inv_f)
+ thus "\<exists>i. x \<in> split BB (prod_decode i)"
+ by (metis prod_encode_inverse prod.cases prod_case_split)
qed
- have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> nat_to_nat2"
+ have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
by (rule ext) (auto simp add: C_def)
also have "... sums suminf ll"
proof (rule suminf_2dimen)
--- a/src/HOL/Probability/Lebesgue.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Probability/Lebesgue.thy Fri Mar 12 16:02:42 2010 +0100
@@ -25,9 +25,58 @@
shows "nonneg (neg_part f)"
unfolding nonneg_def neg_part_def min_def by auto
+lemma (in measure_space)
+ assumes "f \<in> borel_measurable M"
+ shows pos_part_borel_measurable: "pos_part f \<in> borel_measurable M"
+ and neg_part_borel_measurable: "neg_part f \<in> borel_measurable M"
+using assms
+proof -
+ { fix a :: real
+ { assume asm: "0 \<le> a"
+ from asm have pp: "\<And> w. (pos_part f w \<le> a) = (f w \<le> a)" unfolding pos_part_def by auto
+ have "{w | w. w \<in> space M \<and> f w \<le> a} \<in> sets M"
+ unfolding pos_part_def using assms borel_measurable_le_iff by auto
+ hence "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M" using pp by auto }
+ moreover have "a < 0 \<Longrightarrow> {w \<in> space M. pos_part f w \<le> a} \<in> sets M"
+ unfolding pos_part_def using empty_sets by auto
+ ultimately have "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M"
+ using le_less_linear by auto
+ } hence pos: "pos_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
+ { fix a :: real
+ { assume asm: "0 \<le> a"
+ from asm have pp: "\<And> w. (neg_part f w \<le> a) = (f w \<ge> - a)" unfolding neg_part_def by auto
+ have "{w | w. w \<in> space M \<and> f w \<ge> - a} \<in> sets M"
+ unfolding neg_part_def using assms borel_measurable_ge_iff by auto
+ hence "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M" using pp by auto }
+ moreover have "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} = {}" unfolding neg_part_def by auto
+ moreover hence "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} \<in> sets M" by (simp only: empty_sets)
+ ultimately have "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M"
+ using le_less_linear by auto
+ } hence neg: "neg_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
+ from pos neg show "pos_part f \<in> borel_measurable M" and "neg_part f \<in> borel_measurable M" by auto
+qed
+
+lemma (in measure_space) pos_part_neg_part_borel_measurable_iff:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ pos_part f \<in> borel_measurable M \<and> neg_part f \<in> borel_measurable M"
+proof -
+ { fix x
+ have "f x = pos_part f x - neg_part f x"
+ unfolding pos_part_def neg_part_def unfolding max_def min_def
+ by auto }
+ hence "(\<lambda> x. f x) = (\<lambda> x. pos_part f x - neg_part f x)" by auto
+ hence "f = (\<lambda> x. pos_part f x - neg_part f x)" by blast
+ from pos_part_borel_measurable[of f] neg_part_borel_measurable[of f]
+ borel_measurable_diff_borel_measurable[of "pos_part f" "neg_part f"]
+ this
+ show ?thesis by auto
+qed
+
context measure_space
begin
+section "Simple discrete step function"
+
definition
"pos_simple f =
{ (s :: nat set, a, x).
@@ -40,29 +89,6 @@
definition
"psfis f = pos_simple_integral ` (pos_simple f)"
-definition
- "nnfis f = { y. \<exists>u x. mono_convergent u f (space M) \<and>
- (\<forall>n. x n \<in> psfis (u n)) \<and> x ----> y }"
-
-definition
- "integrable f \<longleftrightarrow> (\<exists>x. x \<in> nnfis (pos_part f)) \<and> (\<exists>y. y \<in> nnfis (neg_part f))"
-
-definition
- "integral f = (THE i :: real. i \<in> nnfis (pos_part f)) - (THE j. j \<in> nnfis (neg_part f))"
-
-definition
- "enumerate s \<equiv> SOME f. bij_betw f UNIV s"
-
-definition
- "countable_space_integral f \<equiv>
- let e = enumerate (f ` space M) in
- suminf (\<lambda>r. e r * measure M (f -` {e r} \<inter> space M))"
-
-definition
- "RN_deriv v \<equiv> SOME f. measure_space (M\<lparr>measure := v\<rparr>) \<and>
- f \<in> borel_measurable M \<and>
- (\<forall>a \<in> sets M. (integral (\<lambda>x. f x * indicator_fn a x) = v a))"
-
lemma pos_simpleE[consumes 1]:
assumes ps: "(s, a, x) \<in> pos_simple f"
obtains "finite s" and "nonneg f" and "nonneg x"
@@ -105,6 +131,11 @@
shows "psfis f = psfis g"
unfolding psfis_def using pos_simple_cong[OF assms] by simp
+lemma psfis_0: "0 \<in> psfis (\<lambda>x. 0)"
+ unfolding psfis_def pos_simple_def pos_simple_integral_def
+ by (auto simp: nonneg_def
+ intro: image_eqI[where x="({0}, (\<lambda>i. space M), (\<lambda>i. 0))"])
+
lemma pos_simple_setsum_indicator_fn:
assumes ps: "(s, a, x) \<in> pos_simple f"
and "t \<in> space M"
@@ -121,28 +152,7 @@
using `i \<in> s` by simp
qed
-lemma (in measure_space) measure_setsum_split:
- assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
- assumes "(\<Union>i \<in> r. b i) = space M"
- assumes "disjoint_family_on b r"
- shows "(measure M) a = (\<Sum> i \<in> r. measure M (a \<inter> (b i)))"
-proof -
- have *: "measure M a = measure M (\<Union>i \<in> r. a \<inter> b i)"
- using assms by auto
- show ?thesis unfolding *
- proof (rule measure_finitely_additive''[symmetric])
- show "finite r" using `finite r` by auto
- { fix i assume "i \<in> r"
- hence "b i \<in> sets M" using br_in_M by auto
- thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto
- }
- show "disjoint_family_on (\<lambda>i. a \<inter> b i) r"
- using `disjoint_family_on b r`
- unfolding disjoint_family_on_def by auto
- qed
-qed
-
-lemma (in measure_space) pos_simple_common_partition:
+lemma pos_simple_common_partition:
assumes psf: "(s, a, x) \<in> pos_simple f"
assumes psg: "(s', b, y) \<in> pos_simple g"
obtains z z' c k where "(k, c, z) \<in> pos_simple f" "(k, c, z') \<in> pos_simple g"
@@ -229,7 +239,7 @@
qed
qed
-lemma (in measure_space) pos_simple_integral_equal:
+lemma pos_simple_integral_equal:
assumes psx: "(s, a, x) \<in> pos_simple f"
assumes psy: "(s', b, y) \<in> pos_simple f"
shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"
@@ -269,7 +279,7 @@
finally show "?left = ?right" .
qed
-lemma (in measure_space)psfis_present:
+lemma psfis_present:
assumes "A \<in> psfis f"
assumes "B \<in> psfis g"
obtains z z' c k where
@@ -295,7 +305,7 @@
qed
qed
-lemma (in measure_space) pos_simple_integral_add:
+lemma pos_simple_integral_add:
assumes "(s, a, x) \<in> pos_simple f"
assumes "(s', b, y) \<in> pos_simple g"
obtains s'' c z where
@@ -468,70 +478,6 @@
thus ?thesis using r unfolding psfis_def image_def by auto
qed
-lemma pos_simple_integral_setsum_image:
- assumes "finite P"
- assumes "\<And> i. i \<in> P \<Longrightarrow> (s i, a i, x i) \<in> pos_simple (f i)"
- obtains s' a' x' where
- "(s', a', x') \<in> pos_simple (\<lambda>t. (\<Sum> i \<in> P. f i t))"
- "pos_simple_integral (s', a', x') = (\<Sum> i \<in> P. pos_simple_integral (s i, a i, x i))"
-using assms that
-proof (induct P arbitrary:thesis rule:finite.induct)
- case emptyI note asms = this
- def s' == "{0 :: nat}"
- def a' == "\<lambda> x. if x = (0 :: nat) then space M else {}"
- def x' == "\<lambda> x :: nat. (0 :: real)"
- have "(s', a', x') \<in> pos_simple (\<lambda> t. 0)"
- unfolding s'_def a'_def x'_def pos_simple_def nonneg_def by auto
- moreover have "pos_simple_integral (s', a', x') = 0"
- unfolding s'_def a'_def x'_def pos_simple_integral_def by auto
- ultimately show ?case using asms by auto
-next
-
- case (insertI P c) note asms = this
- then obtain s' a' x' where
- sax: "(s', a', x') \<in> pos_simple (\<lambda>t. \<Sum>i\<in>P. f i t)"
- "pos_simple_integral (s', a', x') =
- (\<Sum>i\<in>P. pos_simple_integral (s i, a i, x i))"
- by auto
-
- { assume "c \<in> P"
- hence "P = insert c P" by auto
- hence thesis using asms sax by auto
- }
- moreover
- { assume "c \<notin> P"
- from asms obtain s'' a'' x'' where
- sax': "s'' = s c" "a'' = a c" "x'' = x c"
- "(s'', a'', x'') \<in> pos_simple (f c)" by auto
- from sax sax' obtain k :: "nat \<Rightarrow> bool" and b :: "nat \<Rightarrow> 'a \<Rightarrow> bool" and z :: "nat \<Rightarrow> real" where
- tybbie:
- "(k, b, z) \<in> pos_simple (\<lambda>t. ((\<Sum>i\<in>P. f i t) + f c t))"
- "(pos_simple_integral (s', a', x') +
- pos_simple_integral (s'', a'', x'') =
- pos_simple_integral (k, b, z))"
- using pos_simple_integral_add by blast
-
- from tybbie have
- "pos_simple_integral (k, b, z) =
- pos_simple_integral (s', a', x') +
- pos_simple_integral (s'', a'', x'')" by simp
- also have "\<dots> = (\<Sum> i \<in> P. pos_simple_integral (s i, a i, x i))
- + pos_simple_integral (s c, a c, x c)"
- using sax sax' by auto
- also have "\<dots> = (\<Sum> i \<in> insert c P. pos_simple_integral (s i, a i, x i))"
- using asms `c \<notin> P` by auto
- finally have A: "pos_simple_integral (k, b, z) = (\<Sum> i \<in> insert c P. pos_simple_integral (s i, a i, x i))"
- by simp
-
- have "\<And> t. (\<Sum> i \<in> P. f i t) + f c t = (\<Sum> i \<in> insert c P. f i t)"
- using `c \<notin> P` `finite P` by auto
- hence B: "(k, b, z) \<in> pos_simple (\<lambda> t. (\<Sum> i \<in> insert c P. f i t))"
- using tybbie by simp
-
- from A B have thesis using asms by auto
- } ultimately show thesis by blast
-qed
-
lemma psfis_setsum_image:
assumes "finite P"
assumes "\<And>i. i \<in> P \<Longrightarrow> a i \<in> psfis (f i)"
@@ -577,57 +523,6 @@
unfolding psfis_def pos_simple_integral_def image_def pos_simple_def nonneg_def
using positive[unfolded positive_def] by (auto intro!:setsum_nonneg mult_nonneg_nonneg)
-lemma pos_part_neg_part_borel_measurable:
- assumes "f \<in> borel_measurable M"
- shows "pos_part f \<in> borel_measurable M" and "neg_part f \<in> borel_measurable M"
-using assms
-proof -
- { fix a :: real
- { assume asm: "0 \<le> a"
- from asm have pp: "\<And> w. (pos_part f w \<le> a) = (f w \<le> a)" unfolding pos_part_def by auto
- have "{w | w. w \<in> space M \<and> f w \<le> a} \<in> sets M"
- unfolding pos_part_def using assms borel_measurable_le_iff by auto
- hence "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M" using pp by auto }
- moreover have "a < 0 \<Longrightarrow> {w \<in> space M. pos_part f w \<le> a} \<in> sets M"
- unfolding pos_part_def using empty_sets by auto
- ultimately have "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M"
- using le_less_linear by auto
- } hence pos: "pos_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
- { fix a :: real
- { assume asm: "0 \<le> a"
- from asm have pp: "\<And> w. (neg_part f w \<le> a) = (f w \<ge> - a)" unfolding neg_part_def by auto
- have "{w | w. w \<in> space M \<and> f w \<ge> - a} \<in> sets M"
- unfolding neg_part_def using assms borel_measurable_ge_iff by auto
- hence "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M" using pp by auto }
- moreover have "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} = {}" unfolding neg_part_def by auto
- moreover hence "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} \<in> sets M" by (simp only: empty_sets)
- ultimately have "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M"
- using le_less_linear by auto
- } hence neg: "neg_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
- from pos neg show "pos_part f \<in> borel_measurable M" and "neg_part f \<in> borel_measurable M" by auto
-qed
-
-lemma pos_part_neg_part_borel_measurable_iff:
- "f \<in> borel_measurable M \<longleftrightarrow>
- pos_part f \<in> borel_measurable M \<and> neg_part f \<in> borel_measurable M"
-proof -
- { fix x
- have "f x = pos_part f x - neg_part f x"
- unfolding pos_part_def neg_part_def unfolding max_def min_def
- by auto }
- hence "(\<lambda> x. f x) = (\<lambda> x. pos_part f x - neg_part f x)" by auto
- hence "f = (\<lambda> x. pos_part f x - neg_part f x)" by blast
- from pos_part_neg_part_borel_measurable[of f]
- borel_measurable_diff_borel_measurable[of "pos_part f" "neg_part f"]
- this
- show ?thesis by auto
-qed
-
-lemma borel_measurable_cong:
- assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
- shows "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M"
-using assms unfolding in_borel_measurable by (simp cong: vimage_inter_cong)
-
lemma psfis_borel_measurable:
assumes "a \<in> psfis f"
shows "f \<in> borel_measurable M"
@@ -654,21 +549,6 @@
show ?thesis by simp
qed
-lemma mono_conv_outgrow:
- assumes "incseq x" "x ----> y" "z < y"
- shows "\<exists>b. \<forall> a \<ge> b. z < x a"
-using assms
-proof -
- from assms have "y - z > 0" by simp
- hence A: "\<exists>n. (\<forall> m \<ge> n. \<bar> x m + - y \<bar> < y - z)" using assms
- unfolding incseq_def LIMSEQ_def dist_real_def real_diff_def
- by simp
- have "\<forall>m. x m \<le> y" using incseq_le assms by auto
- hence B: "\<forall>m. \<bar> x m + - y \<bar> = y - x m"
- by (metis abs_if abs_minus_add_cancel less_iff_diff_less_0 linorder_not_le real_diff_def)
- from A B show ?thesis by auto
-qed
-
lemma psfis_mono_conv_mono:
assumes mono: "mono_convergent u f (space M)"
and ps_u: "\<And>n. x n \<in> psfis (u n)"
@@ -679,7 +559,6 @@
proof (rule field_le_mult_one_interval)
fix z :: real assume "0 < z" and "z < 1"
hence "0 \<le> z" by auto
-(* def B' \<equiv> "\<lambda>n. {w \<in> space M. z * s w <= u n w}" *)
let "?B' n" = "{w \<in> space M. z * s w <= u n w}"
have "incseq x" unfolding incseq_def
@@ -783,11 +662,20 @@
qed
qed
+section "Continuous posititve integration"
+
+definition
+ "nnfis f = { y. \<exists>u x. mono_convergent u f (space M) \<and>
+ (\<forall>n. x n \<in> psfis (u n)) \<and> x ----> y }"
+
lemma psfis_nnfis:
"a \<in> psfis f \<Longrightarrow> a \<in> nnfis f"
unfolding nnfis_def mono_convergent_def incseq_def
by (auto intro!: exI[of _ "\<lambda>n. f"] exI[of _ "\<lambda>n. a"] LIMSEQ_const)
+lemma nnfis_0: "0 \<in> nnfis (\<lambda> x. 0)"
+ by (rule psfis_nnfis[OF psfis_0])
+
lemma nnfis_times:
assumes "a \<in> nnfis f" and "0 \<le> z"
shows "z * a \<in> nnfis (\<lambda>t. z * f t)"
@@ -951,7 +839,7 @@
using assms
proof -
from assms[unfolded nnfis_def] guess u y by auto note uy = this
- hence "\<And> n. 0 \<le> u n x"
+ hence "\<And> n. 0 \<le> u n x"
unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def
by auto
thus "0 \<le> f x" using uy[rule_format]
@@ -1136,11 +1024,6 @@
show ?thesis unfolding nnfis_def by auto
qed
-lemma psfis_0: "0 \<in> psfis (\<lambda>x. 0)"
- unfolding psfis_def pos_simple_def pos_simple_integral_def
- by (auto simp: nonneg_def
- intro: image_eqI[where x="({0}, (\<lambda>i. space M), (\<lambda>i. 0))"])
-
lemma the_nnfis[simp]: "a \<in> nnfis f \<Longrightarrow> (THE a. a \<in> nnfis f) = a"
by (auto intro: the_equality nnfis_unique)
@@ -1158,6 +1041,14 @@
show ?thesis by safe
qed
+section "Lebesgue Integral"
+
+definition
+ "integrable f \<longleftrightarrow> (\<exists>x. x \<in> nnfis (pos_part f)) \<and> (\<exists>y. y \<in> nnfis (neg_part f))"
+
+definition
+ "integral f = (THE i :: real. i \<in> nnfis (pos_part f)) - (THE j. j \<in> nnfis (neg_part f))"
+
lemma integral_cong:
assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
shows "integral f = integral g"
@@ -1203,7 +1094,7 @@
(is "\<exists>x. x \<in> nnfis ?pp \<and> _")
proof (rule nnfis_dom_conv)
show "?pp \<in> borel_measurable M"
- using borel by (rule pos_part_neg_part_borel_measurable)
+ using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)
show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)
fix t assume "t \<in> space M"
with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]
@@ -1217,7 +1108,7 @@
(is "\<exists>x. x \<in> nnfis ?np \<and> _")
proof (rule nnfis_dom_conv)
show "?np \<in> borel_measurable M"
- using borel by (rule pos_part_neg_part_borel_measurable)
+ using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)
show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)
fix t assume "t \<in> space M"
with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]
@@ -1419,6 +1310,192 @@
by (auto intro!: mult_imp_le_div_pos[OF `0 < a ^ n`], simp add: real_mult_commute)
qed
+section "Lebesgue integration on countable spaces"
+
+lemma nnfis_on_countable:
+ assumes borel: "f \<in> borel_measurable M"
+ and bij: "bij_betw enum S (f ` space M - {0})"
+ and enum_zero: "enum ` (-S) \<subseteq> {0}"
+ and nn_enum: "\<And>n. 0 \<le> enum n"
+ and sums: "(\<lambda>r. enum r * measure M (f -` {enum r} \<inter> space M)) sums x" (is "?sum sums x")
+ shows "x \<in> nnfis f"
+proof -
+ have inj_enum: "inj_on enum S"
+ and range_enum: "enum ` S = f ` space M - {0}"
+ using bij by (auto simp: bij_betw_def)
+
+ let "?x n z" = "\<Sum>i = 0..<n. enum i * indicator_fn (f -` {enum i} \<inter> space M) z"
+
+ show ?thesis
+ proof (rule nnfis_mon_conv)
+ show "(\<lambda>n. \<Sum>i = 0..<n. ?sum i) ----> x" using sums unfolding sums_def .
+ next
+ fix n
+ show "(\<Sum>i = 0..<n. ?sum i) \<in> nnfis (?x n)"
+ proof (induct n)
+ case 0 thus ?case by (simp add: nnfis_0)
+ next
+ case (Suc n) thus ?case using nn_enum
+ by (auto intro!: nnfis_add nnfis_times psfis_nnfis[OF psfis_indicator] borel_measurable_vimage[OF borel])
+ qed
+ next
+ show "mono_convergent ?x f (space M)"
+ proof (rule mono_convergentI)
+ fix x
+ show "incseq (\<lambda>n. ?x n x)"
+ by (rule incseq_SucI, auto simp: indicator_fn_def nn_enum)
+
+ have fin: "\<And>n. finite (enum ` ({0..<n} \<inter> S))" by auto
+
+ assume "x \<in> space M"
+ hence "f x \<in> enum ` S \<or> f x = 0" using range_enum by auto
+ thus "(\<lambda>n. ?x n x) ----> f x"
+ proof (rule disjE)
+ assume "f x \<in> enum ` S"
+ then obtain i where "i \<in> S" and "f x = enum i" by auto
+
+ { fix n
+ have sum_ranges:
+ "i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {enum i}"
+ "\<not> i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {}"
+ using `x \<in> space M` `i \<in> S` inj_enum[THEN inj_on_iff] by auto
+ have "?x n x =
+ (\<Sum>i \<in> {0..<n} \<inter> S. enum i * indicator_fn (f -` {enum i} \<inter> space M) x)"
+ using enum_zero by (auto intro!: setsum_mono_zero_cong_right)
+ also have "... =
+ (\<Sum>z \<in> enum`({0..<n} \<inter> S). z * indicator_fn (f -` {z} \<inter> space M) x)"
+ using inj_enum[THEN subset_inj_on] by (auto simp: setsum_reindex)
+ also have "... = (if i < n then f x else 0)"
+ unfolding indicator_fn_def if_distrib[where x=1 and y=0]
+ setsum_cases[OF fin]
+ using sum_ranges `f x = enum i`
+ by auto
+ finally have "?x n x = (if i < n then f x else 0)" . }
+ note sum_equals_if = this
+
+ show ?thesis unfolding sum_equals_if
+ by (rule LIMSEQ_offset[where k="i + 1"]) (auto intro!: LIMSEQ_const)
+ next
+ assume "f x = 0"
+ { fix n have "?x n x = 0"
+ unfolding indicator_fn_def if_distrib[where x=1 and y=0]
+ setsum_cases[OF finite_atLeastLessThan]
+ using `f x = 0` `x \<in> space M`
+ by (auto split: split_if) }
+ thus ?thesis using `f x = 0` by (auto intro!: LIMSEQ_const)
+ qed
+ qed
+ qed
+qed
+
+lemma integral_on_countable:
+ assumes borel: "f \<in> borel_measurable M"
+ and bij: "bij_betw enum S (f ` space M)"
+ and enum_zero: "enum ` (-S) \<subseteq> {0}"
+ and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
+ shows "integrable f"
+ and "integral f = (\<Sum>r. enum r * measure M (f -` {enum r} \<inter> space M))" (is "_ = suminf (?sum f enum)")
+proof -
+ { fix f enum
+ assume borel: "f \<in> borel_measurable M"
+ and bij: "bij_betw enum S (f ` space M)"
+ and enum_zero: "enum ` (-S) \<subseteq> {0}"
+ and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
+ have inj_enum: "inj_on enum S" and range_enum: "f ` space M = enum ` S"
+ using bij unfolding bij_betw_def by auto
+
+ have [simp, intro]: "\<And>X. 0 \<le> measure M (f -` {X} \<inter> space M)"
+ by (rule positive, rule borel_measurable_vimage[OF borel])
+
+ have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) \<in> nnfis (pos_part f) \<and>
+ summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"
+ proof (rule conjI, rule nnfis_on_countable)
+ have pos_f_image: "pos_part f ` space M - {0} = f ` space M \<inter> {0<..}"
+ unfolding pos_part_def max_def by auto
+
+ show "bij_betw (pos_part enum) {x \<in> S. 0 < enum x} (pos_part f ` space M - {0})"
+ unfolding bij_betw_def pos_f_image
+ unfolding pos_part_def max_def range_enum
+ by (auto intro!: inj_onI simp: inj_enum[THEN inj_on_eq_iff])
+
+ show "\<And>n. 0 \<le> pos_part enum n" unfolding pos_part_def max_def by auto
+
+ show "pos_part f \<in> borel_measurable M"
+ by (rule pos_part_borel_measurable[OF borel])
+
+ show "pos_part enum ` (- {x \<in> S. 0 < enum x}) \<subseteq> {0}"
+ unfolding pos_part_def max_def using enum_zero by auto
+
+ show "summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"
+ proof (rule summable_comparison_test[OF _ abs_summable], safe intro!: exI[of _ 0])
+ fix n :: nat
+ have "pos_part enum n \<noteq> 0 \<Longrightarrow> (pos_part f -` {enum n} \<inter> space M) =
+ (if 0 < enum n then (f -` {enum n} \<inter> space M) else {})"
+ unfolding pos_part_def max_def by (auto split: split_if_asm)
+ thus "norm (?sum (pos_part f) (pos_part enum) n) \<le> \<bar>?sum f enum n \<bar>"
+ by (cases "pos_part enum n = 0",
+ auto simp: pos_part_def max_def abs_mult not_le split: split_if_asm intro!: mult_nonpos_nonneg)
+ qed
+ thus "(\<lambda>r. ?sum (pos_part f) (pos_part enum) r) sums (\<Sum>r. ?sum (pos_part f) (pos_part enum) r)"
+ by (rule summable_sums)
+ qed }
+ note pos = this
+
+ note pos_part = pos[OF assms(1-4)]
+ moreover
+ have neg_part_to_pos_part:
+ "\<And>f :: _ \<Rightarrow> real. neg_part f = pos_part (uminus \<circ> f)"
+ by (auto simp: pos_part_def neg_part_def min_def max_def expand_fun_eq)
+
+ have neg_part: "(\<Sum>r. ?sum (neg_part f) (neg_part enum) r) \<in> nnfis (neg_part f) \<and>
+ summable (\<lambda>r. ?sum (neg_part f) (neg_part enum) r)"
+ unfolding neg_part_to_pos_part
+ proof (rule pos)
+ show "uminus \<circ> f \<in> borel_measurable M" unfolding comp_def
+ by (rule borel_measurable_uminus_borel_measurable[OF borel])
+
+ show "bij_betw (uminus \<circ> enum) S ((uminus \<circ> f) ` space M)"
+ using bij unfolding bij_betw_def
+ by (auto intro!: comp_inj_on simp: image_compose)
+
+ show "(uminus \<circ> enum) ` (- S) \<subseteq> {0}"
+ using enum_zero by auto
+
+ have minus_image: "\<And>r. (uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M = f -` {enum r} \<inter> space M"
+ by auto
+ show "summable (\<lambda>r. \<bar>(uminus \<circ> enum) r * measure_space.measure M ((uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M)\<bar>)"
+ unfolding minus_image using abs_summable by simp
+ qed
+ ultimately show "integrable f" unfolding integrable_def by auto
+
+ { fix r
+ have "?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r = ?sum f enum r"
+ proof (cases rule: linorder_cases)
+ assume "0 < enum r"
+ hence "pos_part f -` {enum r} \<inter> space M = f -` {enum r} \<inter> space M"
+ unfolding pos_part_def max_def by (auto split: split_if_asm)
+ with `0 < enum r` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq
+ by auto
+ next
+ assume "enum r < 0"
+ hence "neg_part f -` {- enum r} \<inter> space M = f -` {enum r} \<inter> space M"
+ unfolding neg_part_def min_def by (auto split: split_if_asm)
+ with `enum r < 0` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq
+ by auto
+ qed (simp add: neg_part_def pos_part_def) }
+ note sum_diff_eq_sum = this
+
+ have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) - (\<Sum>r. ?sum (neg_part f) (neg_part enum) r)
+ = (\<Sum>r. ?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r)"
+ using neg_part pos_part by (auto intro: suminf_diff)
+ also have "... = (\<Sum>r. ?sum f enum r)" unfolding sum_diff_eq_sum ..
+ finally show "integral f = suminf (?sum f enum)"
+ unfolding integral_def using pos_part neg_part
+ by (auto dest: the_nnfis)
+qed
+
+section "Lebesgue integration on finite space"
+
lemma integral_finite_on_sets:
assumes "f \<in> borel_measurable M" and "finite (space M)" and "a \<in> sets M"
shows "integral (\<lambda>x. f x * indicator_fn a x) =
@@ -1489,51 +1566,12 @@
qed (auto intro!: image_eqI inj_onI)
qed
-lemma borel_measurable_inverse:
- assumes "f \<in> borel_measurable M"
- shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
- unfolding borel_measurable_ge_iff
-proof (safe, rule linorder_cases)
- fix a :: real assume "0 < a"
- { fix w
- from `0 < a` have "a \<le> inverse (f w) \<longleftrightarrow> 0 < f w \<and> f w \<le> 1 / a"
- by (metis inverse_eq_divide inverse_inverse_eq le_imp_inverse_le
- linorder_not_le real_le_refl real_le_trans real_less_def
- xt1(7) zero_less_divide_1_iff) }
- hence "{w \<in> space M. a \<le> inverse (f w)} =
- {w \<in> space M. 0 < f w} \<inter> {w \<in> space M. f w \<le> 1 / a}" by auto
- with Int assms[unfolded borel_measurable_gr_iff]
- assms[unfolded borel_measurable_le_iff]
- show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
-next
- fix a :: real assume "0 = a"
- { fix w have "a \<le> inverse (f w) \<longleftrightarrow> 0 \<le> f w"
- unfolding `0 = a`[symmetric] by auto }
- thus "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M"
- using assms[unfolded borel_measurable_ge_iff] by simp
-next
- fix a :: real assume "a < 0"
- { fix w
- from `a < 0` have "a \<le> inverse (f w) \<longleftrightarrow> f w \<le> 1 / a \<or> 0 \<le> f w"
- apply (cases "0 \<le> f w")
- apply (metis inverse_eq_divide linorder_not_le xt1(8) xt1(9)
- zero_le_divide_1_iff)
- apply (metis inverse_eq_divide inverse_inverse_eq inverse_le_imp_le_neg
- linorder_not_le real_le_refl real_le_trans)
- done }
- hence "{w \<in> space M. a \<le> inverse (f w)} =
- {w \<in> space M. f w \<le> 1 / a} \<union> {w \<in> space M. 0 \<le> f w}" by auto
- with Un assms[unfolded borel_measurable_ge_iff]
- assms[unfolded borel_measurable_le_iff]
- show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
-qed
+section "Radon–Nikodym derivative"
-lemma borel_measurable_divide:
- assumes "f \<in> borel_measurable M"
- and "g \<in> borel_measurable M"
- shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
- unfolding field_divide_inverse
- by (rule borel_measurable_inverse borel_measurable_times_borel_measurable assms)+
+definition
+ "RN_deriv v \<equiv> SOME f. measure_space (M\<lparr>measure := v\<rparr>) \<and>
+ f \<in> borel_measurable M \<and>
+ (\<forall>a \<in> sets M. (integral (\<lambda>x. f x * indicator_fn a x) = v a))"
lemma RN_deriv_finite_singleton:
fixes v :: "'a set \<Rightarrow> real"
@@ -1577,214 +1615,6 @@
using `measure M {x} \<noteq> 0` by (simp add: eq_divide_eq)
qed fact
-lemma countable_space_integral_reduce:
- assumes "positive M (measure M)" and "f \<in> borel_measurable M"
- and "countable (f ` space M)"
- and "\<not> finite (pos_part f ` space M)"
- and "\<not> finite (neg_part f ` space M)"
- and "((\<lambda>r. r * measure m (neg_part f -` {r} \<inter> space m)) o enumerate (neg_part f ` space M)) sums n"
- and "((\<lambda>r. r * measure m (pos_part f -` {r} \<inter> space m)) o enumerate (pos_part f ` space M)) sums p"
- shows "integral f = p - n"
-oops
-
-(*
-val countable_space_integral_reduce = store_thm
- ("countable_space_integral_reduce",
- "\<forall>m f p n. measure_space m \<and>
- positive m \<and>
- f \<in> borel_measurable (space m, sets m) \<and>
- countable (IMAGE f (space m)) \<and>
- ~(FINITE (IMAGE (pos_part f) (space m))) \<and>
- ~(FINITE (IMAGE (neg_part f) (space m))) \<and>
- (\<lambda>r. r *
- measure m (PREIMAGE (neg_part f) {r} \<inter> space m)) o
- enumerate ((IMAGE (neg_part f) (space m))) sums n \<and>
- (\<lambda>r. r *
- measure m (PREIMAGE (pos_part f) {r} \<inter> space m)) o
- enumerate ((IMAGE (pos_part f) (space m))) sums p ==>
- (integral m f = p - n)",
- RW_TAC std_ss [integral_def]
- ++ Suff `((@i. i \<in> nnfis m (pos_part f)) = p) \<and> ((@i. i \<in> nnfis m (neg_part f)) = n)`
- >> RW_TAC std_ss []
- ++ (CONJ_TAC ++ MATCH_MP_TAC SELECT_UNIQUE ++ RW_TAC std_ss [])
- >> (Suff `p \<in> nnfis m (pos_part f)` >> METIS_TAC [nnfis_unique]
- ++ MATCH_MP_TAC nnfis_mon_conv
- ++ `BIJ (enumerate(IMAGE (pos_part f) (space m))) UNIV (IMAGE (pos_part f) (space m))`
- by (`countable (IMAGE (pos_part f) (space m))`
- by (MATCH_MP_TAC COUNTABLE_SUBSET
- ++ Q.EXISTS_TAC `0 INSERT (IMAGE f (space m))`
- ++ RW_TAC std_ss [SUBSET_DEF, IN_IMAGE, pos_part_def, COUNTABLE_INSERT, IN_INSERT]
- ++ METIS_TAC [])
- ++ METIS_TAC [COUNTABLE_ALT])
- ++ FULL_SIMP_TAC std_ss [sums, BIJ_DEF, INJ_DEF, SURJ_DEF, IN_UNIV]
- ++ Q.EXISTS_TAC `(\<lambda>n t. if t \<in> space m \<and> pos_part f t \<in> IMAGE (enumerate (IMAGE (pos_part f) (space m)))
- (pred_set$count n) then pos_part f t else 0)`
- ++ Q.EXISTS_TAC `(\<lambda>n.
- sum (0,n)
- ((\<lambda>r.
- r *
- measure m (PREIMAGE (pos_part f) {r} \<inter> space m)) o
- enumerate (IMAGE (pos_part f) (space m))))`
- ++ ASM_SIMP_TAC std_ss []
- ++ CONJ_TAC
- >> (RW_TAC std_ss [mono_convergent_def]
- >> (RW_TAC real_ss [IN_IMAGE, pred_setTheory.IN_COUNT, pos_part_def] ++ METIS_TAC [LESS_LESS_EQ_TRANS])
- ++ RW_TAC std_ss [SEQ]
- ++ `\<exists>N. enumerate (IMAGE (pos_part f) (space m)) N = (pos_part f) t`
- by METIS_TAC [IN_IMAGE]
- ++ Q.EXISTS_TAC `SUC N`
- ++ RW_TAC real_ss [GREATER_EQ, real_ge, IN_IMAGE, REAL_SUB_LZERO]
- ++ FULL_SIMP_TAC std_ss [pred_setTheory.IN_COUNT]
- ++ METIS_TAC [DECIDE "!n:num. n < SUC n", LESS_LESS_EQ_TRANS, pos_part_def])
- ++ STRIP_TAC ++ MATCH_MP_TAC psfis_nnfis ++ ASM_SIMP_TAC std_ss [GSYM REAL_SUM_IMAGE_COUNT]
- ++ `(\<lambda>t. (if t \<in> space m \<and> pos_part f t \<in> IMAGE (enumerate (IMAGE (pos_part f) (space m))) (pred_set$count n')
- then pos_part f t else 0)) =
- (\<lambda>t. SIGMA (\<lambda>i. (\<lambda>i. enumerate (IMAGE (pos_part f) (space m)) i) i *
- indicator_fn ((\<lambda>i. PREIMAGE (pos_part f) {enumerate (IMAGE (pos_part f) (space m)) i}
- \<inter> (space m)) i) t)
- (pred_set$count n'))`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [IN_IMAGE]
- >> (`pred_set$count n' = x INSERT (pred_set$count n')`
- by (ONCE_REWRITE_TAC [EXTENSION] ++ RW_TAC std_ss [IN_INSERT] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC std_ss [REAL_SUM_IMAGE_THM, pred_setTheory.FINITE_COUNT]
- ++ ONCE_REWRITE_TAC [(REWRITE_RULE [pred_setTheory.FINITE_COUNT] o
- REWRITE_RULE [FINITE_DELETE] o Q.ISPEC `pred_set$count n' DELETE x`) REAL_SUM_IMAGE_IN_IF]
- ++ RW_TAC real_ss [IN_DELETE, indicator_fn_def, IN_INTER, IN_SING, IN_PREIMAGE]
- ++ `(\x'. (if x' \<in> pred_set$count n' \<and> ~(x' = x) then
- enumerate (IMAGE (pos_part f) (space m)) x' *
- (if enumerate (IMAGE (pos_part f) (space m)) x =
- enumerate (IMAGE (pos_part f) (space m)) x'
- then 1 else 0) else 0)) = (\<lambda>x. 0)`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT, FINITE_DELETE])
- ++ FULL_SIMP_TAC real_ss [IN_IMAGE, indicator_fn_def, IN_INTER, IN_PREIMAGE, IN_SING]
- >> RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT, FINITE_DELETE]
- ++ ONCE_REWRITE_TAC [(REWRITE_RULE [pred_setTheory.FINITE_COUNT] o Q.ISPEC `pred_set$count n'`)
- REAL_SUM_IMAGE_IN_IF]
- ++ `(\<lambda>x. (if x \<in> pred_set$count n' then
- (\<lambda>i. enumerate (IMAGE (pos_part f) (space m)) i *
- (if (pos_part f t = enumerate (IMAGE (pos_part f) (space m)) i) \<and>
- t \<in> space m then 1 else 0)) x else 0)) = (\<lambda>x. 0)`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT])
- ++ POP_ORW
- ++ `((\<lambda>r. r * measure m (PREIMAGE (pos_part f) {r} \<inter> space m)) o
- enumerate (IMAGE (pos_part f) (space m))) =
- (\<lambda>i. (\<lambda>i. enumerate (IMAGE (pos_part f) (space m)) i) i *
- measure m ((\<lambda>i.
- PREIMAGE (pos_part f)
- {enumerate (IMAGE (pos_part f) (space m)) i} \<inter>
- space m) i))`
- by (RW_TAC std_ss [FUN_EQ_THM, o_DEF] ++ RW_TAC real_ss [])
- ++ POP_ORW
- ++ MATCH_MP_TAC psfis_intro
- ++ ASM_SIMP_TAC std_ss [pred_setTheory.FINITE_COUNT]
- ++ REVERSE CONJ_TAC
- >> (FULL_SIMP_TAC real_ss [IN_IMAGE, pos_part_def]
- ++ METIS_TAC [REAL_LE_REFL])
- ++ `(pos_part f) \<in> borel_measurable (space m, sets m)`
- by METIS_TAC [pos_part_neg_part_borel_measurable]
- ++ REPEAT STRIP_TAC
- ++ `PREIMAGE (pos_part f) {enumerate (IMAGE (pos_part f) (space m)) i} \<inter> space m =
- {w | w \<in> space m \<and> ((pos_part f) w = (\<lambda>w. enumerate (IMAGE (pos_part f) (space m)) i) w)}`
- by (ONCE_REWRITE_TAC [EXTENSION] ++ RW_TAC std_ss [GSPECIFICATION, IN_INTER, IN_PREIMAGE, IN_SING]
- ++ DECIDE_TAC)
- ++ POP_ORW
- ++ MATCH_MP_TAC borel_measurable_eq_borel_measurable
- ++ METIS_TAC [borel_measurable_const, measure_space_def])
- ++ Suff `n \<in> nnfis m (neg_part f)` >> METIS_TAC [nnfis_unique]
- ++ MATCH_MP_TAC nnfis_mon_conv
- ++ `BIJ (enumerate(IMAGE (neg_part f) (space m))) UNIV (IMAGE (neg_part f) (space m))`
- by (`countable (IMAGE (neg_part f) (space m))`
- by (MATCH_MP_TAC COUNTABLE_SUBSET
- ++ Q.EXISTS_TAC `0 INSERT (IMAGE abs (IMAGE f (space m)))`
- ++ RW_TAC std_ss [SUBSET_DEF, IN_IMAGE, abs, neg_part_def, COUNTABLE_INSERT, IN_INSERT, COUNTABLE_IMAGE]
- ++ METIS_TAC [])
- ++ METIS_TAC [COUNTABLE_ALT])
- ++ FULL_SIMP_TAC std_ss [sums, BIJ_DEF, INJ_DEF, SURJ_DEF, IN_UNIV]
- ++ Q.EXISTS_TAC `(\<lambda>n t. if t \<in> space m \<and> neg_part f t \<in> IMAGE (enumerate (IMAGE (neg_part f) (space m)))
- (pred_set$count n) then neg_part f t else 0)`
- ++ Q.EXISTS_TAC `(\<lambda>n.
- sum (0,n)
- ((\<lambda>r.
- r *
- measure m (PREIMAGE (neg_part f) {r} \<inter> space m)) o
- enumerate (IMAGE (neg_part f) (space m))))`
- ++ ASM_SIMP_TAC std_ss []
- ++ CONJ_TAC
- >> (RW_TAC std_ss [mono_convergent_def]
- >> (RW_TAC real_ss [IN_IMAGE, pred_setTheory.IN_COUNT, neg_part_def] ++ METIS_TAC [LESS_LESS_EQ_TRANS, REAL_LE_NEGTOTAL])
- ++ RW_TAC std_ss [SEQ]
- ++ `\<exists>N. enumerate (IMAGE (neg_part f) (space m)) N = (neg_part f) t`
- by METIS_TAC [IN_IMAGE]
- ++ Q.EXISTS_TAC `SUC N`
- ++ RW_TAC real_ss [GREATER_EQ, real_ge, IN_IMAGE, REAL_SUB_LZERO]
- ++ FULL_SIMP_TAC std_ss [pred_setTheory.IN_COUNT]
- ++ METIS_TAC [DECIDE "!n:num. n < SUC n", LESS_LESS_EQ_TRANS, neg_part_def])
- ++ STRIP_TAC ++ MATCH_MP_TAC psfis_nnfis ++ ASM_SIMP_TAC std_ss [GSYM REAL_SUM_IMAGE_COUNT]
- ++ `(\<lambda>t. (if t \<in> space m \<and> neg_part f t \<in> IMAGE (enumerate (IMAGE (neg_part f) (space m))) (pred_set$count n')
- then neg_part f t else 0)) =
- (\<lambda>t. SIGMA (\<lambda>i. (\<lambda>i. enumerate (IMAGE (neg_part f) (space m)) i) i *
- indicator_fn ((\<lambda>i. PREIMAGE (neg_part f) {enumerate (IMAGE (neg_part f) (space m)) i}
- \<inter> (space m)) i) t)
- (pred_set$count n'))`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [IN_IMAGE]
- >> (`pred_set$count n' = x INSERT (pred_set$count n')`
- by (ONCE_REWRITE_TAC [EXTENSION] ++ RW_TAC std_ss [IN_INSERT] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC std_ss [REAL_SUM_IMAGE_THM, pred_setTheory.FINITE_COUNT]
- ++ ONCE_REWRITE_TAC [(REWRITE_RULE [pred_setTheory.FINITE_COUNT] o
- REWRITE_RULE [FINITE_DELETE] o Q.ISPEC `pred_set$count n' DELETE x`) REAL_SUM_IMAGE_IN_IF]
- ++ RW_TAC real_ss [IN_DELETE, indicator_fn_def, IN_INTER, IN_SING, IN_PREIMAGE]
- ++ `(\x'. (if x' \<in> pred_set$count n' \<and> ~(x' = x) then
- enumerate (IMAGE (neg_part f) (space m)) x' *
- (if enumerate (IMAGE (neg_part f) (space m)) x =
- enumerate (IMAGE (neg_part f) (space m)) x'
- then 1 else 0) else 0)) = (\<lambda>x. 0)`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT, FINITE_DELETE])
- ++ FULL_SIMP_TAC real_ss [IN_IMAGE, indicator_fn_def, IN_INTER, IN_PREIMAGE, IN_SING]
- >> RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT, FINITE_DELETE]
- ++ ONCE_REWRITE_TAC [(REWRITE_RULE [pred_setTheory.FINITE_COUNT] o Q.ISPEC `pred_set$count n'`)
- REAL_SUM_IMAGE_IN_IF]
- ++ `(\<lambda>x. (if x \<in> pred_set$count n' then
- (\<lambda>i. enumerate (IMAGE (neg_part f) (space m)) i *
- (if (neg_part f t = enumerate (IMAGE (neg_part f) (space m)) i) \<and>
- t \<in> space m then 1 else 0)) x else 0)) = (\<lambda>x. 0)`
- by (RW_TAC std_ss [FUN_EQ_THM] ++ RW_TAC real_ss [] ++ METIS_TAC [])
- ++ POP_ORW
- ++ RW_TAC real_ss [REAL_SUM_IMAGE_0, pred_setTheory.FINITE_COUNT])
- ++ POP_ORW
- ++ `((\<lambda>r. r * measure m (PREIMAGE (neg_part f) {r} \<inter> space m)) o
- enumerate (IMAGE (neg_part f) (space m))) =
- (\<lambda>i. (\<lambda>i. enumerate (IMAGE (neg_part f) (space m)) i) i *
- measure m ((\<lambda>i.
- PREIMAGE (neg_part f)
- {enumerate (IMAGE (neg_part f) (space m)) i} \<inter>
- space m) i))`
- by (RW_TAC std_ss [FUN_EQ_THM, o_DEF] ++ RW_TAC real_ss [])
- ++ POP_ORW
- ++ MATCH_MP_TAC psfis_intro
- ++ ASM_SIMP_TAC std_ss [pred_setTheory.FINITE_COUNT]
- ++ REVERSE CONJ_TAC
- >> (FULL_SIMP_TAC real_ss [IN_IMAGE, neg_part_def]
- ++ METIS_TAC [REAL_LE_REFL, REAL_LE_NEGTOTAL])
- ++ `(neg_part f) \<in> borel_measurable (space m, sets m)`
- by METIS_TAC [pos_part_neg_part_borel_measurable]
- ++ REPEAT STRIP_TAC
- ++ `PREIMAGE (neg_part f) {enumerate (IMAGE (neg_part f) (space m)) i} \<inter> space m =
- {w | w \<in> space m \<and> ((neg_part f) w = (\<lambda>w. enumerate (IMAGE (neg_part f) (space m)) i) w)}`
- by (ONCE_REWRITE_TAC [EXTENSION] ++ RW_TAC std_ss [GSPECIFICATION, IN_INTER, IN_PREIMAGE, IN_SING]
- ++ DECIDE_TAC)
- ++ POP_ORW
- ++ MATCH_MP_TAC borel_measurable_eq_borel_measurable
- ++ METIS_TAC [borel_measurable_const, measure_space_def]);
-*)
+end
end
-
-end
\ No newline at end of file
--- a/src/HOL/Probability/Measure.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Probability/Measure.thy Fri Mar 12 16:02:42 2010 +0100
@@ -997,4 +997,25 @@
\<Longrightarrow> measure_space M"
by (rule measure_down [OF sigma_algebra.finite_additivity_sufficient], auto)
+lemma (in measure_space) measure_setsum_split:
+ assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
+ assumes "(\<Union>i \<in> r. b i) = space M"
+ assumes "disjoint_family_on b r"
+ shows "(measure M) a = (\<Sum> i \<in> r. measure M (a \<inter> (b i)))"
+proof -
+ have *: "measure M a = measure M (\<Union>i \<in> r. a \<inter> b i)"
+ using assms by auto
+ show ?thesis unfolding *
+ proof (rule measure_finitely_additive''[symmetric])
+ show "finite r" using `finite r` by auto
+ { fix i assume "i \<in> r"
+ hence "b i \<in> sets M" using br_in_M by auto
+ thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto
+ }
+ show "disjoint_family_on (\<lambda>i. a \<inter> b i) r"
+ using `disjoint_family_on b r`
+ unfolding disjoint_family_on_def by auto
+ qed
+qed
+
end
\ No newline at end of file
--- a/src/HOL/Probability/SeriesPlus.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Probability/SeriesPlus.thy Fri Mar 12 16:02:42 2010 +0100
@@ -1,5 +1,5 @@
theory SeriesPlus
- imports Complex_Main Nat_Int_Bij
+ imports Complex_Main Nat_Bijection
begin
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
@@ -16,7 +16,7 @@
assumes f_nneg: "!!m n. 0 \<le> f(m,n)"
and fsums: "!!m. (\<lambda>n. f (m,n)) sums (g m)"
and sumg: "summable g"
- shows "(f o nat_to_nat2) sums suminf g"
+ shows "(f o prod_decode) sums suminf g"
proof (simp add: sums_def)
{
fix x
@@ -31,18 +31,18 @@
thus "0 \<le> g m" using fsums [of m]
by (auto simp add: sums_iff)
qed
- show "(\<lambda>n. \<Sum>x = 0..<n. f (nat_to_nat2 x)) ----> suminf g"
+ show "(\<lambda>n. \<Sum>x = 0..<n. f (prod_decode x)) ----> suminf g"
proof (rule increasing_LIMSEQ, simp add: f_nneg)
fix n
- def i \<equiv> "Max (Domain (nat_to_nat2 ` {0..<n}))"
- def j \<equiv> "Max (Range (nat_to_nat2 ` {0..<n}))"
- have ij: "nat_to_nat2 ` {0..<n} \<subseteq> ({0..i} \<times> {0..j})"
+ def i \<equiv> "Max (Domain (prod_decode ` {0..<n}))"
+ def j \<equiv> "Max (Range (prod_decode ` {0..<n}))"
+ have ij: "prod_decode ` {0..<n} \<subseteq> ({0..i} \<times> {0..j})"
by (force simp add: i_def j_def
intro: finite_imageI Max_ge finite_Domain finite_Range)
- have "(\<Sum>x = 0..<n. f (nat_to_nat2 x)) = setsum f (nat_to_nat2 ` {0..<n})"
- using setsum_reindex [of nat_to_nat2 "{0..<n}" f]
+ have "(\<Sum>x = 0..<n. f (prod_decode x)) = setsum f (prod_decode ` {0..<n})"
+ using setsum_reindex [of prod_decode "{0..<n}" f]
by (simp add: o_def)
- (metis nat_to_nat2_inj subset_inj_on subset_UNIV nat_to_nat2_inj)
+ (metis inj_prod_decode inj_prod_decode)
also have "... \<le> setsum f ({0..i} \<times> {0..j})"
by (rule setsum_mono2) (auto simp add: ij)
also have "... = setsum (\<lambda>m. setsum (\<lambda>n. f (m,n)) {0..j}) {0..<Suc i}"
@@ -56,10 +56,10 @@
by (auto simp add: sums_iff)
(metis atLeastLessThanSuc_atLeastAtMost f_nneg series_pos_le f_nneg)
qed
- finally have "(\<Sum>x = 0..<n. f (nat_to_nat2 x)) \<le> setsum g {0..<Suc i}" .
+ finally have "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> setsum g {0..<Suc i}" .
also have "... \<le> suminf g"
by (rule series_pos_le [OF sumg]) (simp add: g_nneg)
- finally show "(\<Sum>x = 0..<n. f (nat_to_nat2 x)) \<le> suminf g" .
+ finally show "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> suminf g" .
next
fix e :: real
assume e: "0 < e"
@@ -99,26 +99,25 @@
by auto
have finite_IJ: "finite IJ"
by (simp add: IJ_def)
- def NIJ \<equiv> "Max (nat_to_nat2 -` IJ)"
- have IJsb: "IJ \<subseteq> nat_to_nat2 ` {0..NIJ}"
+ def NIJ \<equiv> "Max (prod_decode -` IJ)"
+ have IJsb: "IJ \<subseteq> prod_decode ` {0..NIJ}"
proof (auto simp add: NIJ_def)
fix i j
assume ij:"(i,j) \<in> IJ"
- hence "(i,j) = nat_to_nat2 (inv nat_to_nat2 (i,j))"
- by (metis nat_to_nat2_surj surj_f_inv_f)
- thus "(i,j) \<in> nat_to_nat2 ` {0..Max (nat_to_nat2 -` IJ)}"
+ hence "(i,j) = prod_decode (prod_encode (i,j))"
+ by simp
+ thus "(i,j) \<in> prod_decode ` {0..Max (prod_decode -` IJ)}"
by (rule image_eqI)
- (simp add: ij finite_vimageI [OF finite_IJ nat_to_nat2_inj]
- nat_to_nat2_surj surj_f_inv_f)
+ (simp add: ij finite_vimageI [OF finite_IJ inj_prod_decode])
qed
- have "setsum f IJ \<le> setsum f (nat_to_nat2 `{0..NIJ})"
+ have "setsum f IJ \<le> setsum f (prod_decode `{0..NIJ})"
by (rule setsum_mono2) (auto simp add: IJsb)
- also have "... = (\<Sum>k = 0..NIJ. f (nat_to_nat2 k))"
- by (simp add: setsum_reindex subset_inj_on [OF nat_to_nat2_inj subset_UNIV])
- also have "... = (\<Sum>k = 0..<Suc NIJ. f (nat_to_nat2 k))"
+ also have "... = (\<Sum>k = 0..NIJ. f (prod_decode k))"
+ by (simp add: setsum_reindex inj_prod_decode)
+ also have "... = (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))"
by (metis atLeast0AtMost atLeast0LessThan lessThan_Suc_atMost)
- finally have "setsum f IJ \<le> (\<Sum>k = 0..<Suc NIJ. f (nat_to_nat2 k))" .
- thus "\<exists>n. suminf g \<le> (\<Sum>x = 0..<n. f (nat_to_nat2 x)) + e" using glessf
+ finally have "setsum f IJ \<le> (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))" .
+ thus "\<exists>n. suminf g \<le> (\<Sum>x = 0..<n. f (prod_decode x)) + e" using glessf
by (metis add_right_mono local.glessf not_leE order_le_less_trans less_asym)
qed
qed
--- a/src/HOL/Rat.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Rat.thy Fri Mar 12 16:02:42 2010 +0100
@@ -1104,11 +1104,11 @@
lemma rat_less_eq_code [code]:
"p \<le> q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d \<le> c * b)"
- by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+ by (cases p, cases q) (simp add: quotient_of_Fract mult.commute)
lemma rat_less_code [code]:
"p < q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d < c * b)"
- by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+ by (cases p, cases q) (simp add: quotient_of_Fract mult.commute)
lemma [code]:
"of_rat p = (let (a, b) = quotient_of p in of_int a / of_int b)"
--- a/src/HOL/SEQ.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/SEQ.thy Fri Mar 12 16:02:42 2010 +0100
@@ -981,6 +981,24 @@
by (blast intro: eq_refl X)
qed
+lemma incseq_SucI:
+ assumes "\<And>n. X n \<le> X (Suc n)"
+ shows "incseq X" unfolding incseq_def
+proof safe
+ fix m n :: nat
+ { fix d m :: nat
+ have "X m \<le> X (m + d)"
+ proof (induct d)
+ case (Suc d)
+ also have "X (m + d) \<le> X (m + Suc d)"
+ using assms by simp
+ finally show ?case .
+ qed simp }
+ note this[of m "n - m"]
+ moreover assume "m \<le> n"
+ ultimately show "X m \<le> X n" by simp
+qed
+
lemma decseq_imp_monoseq: "decseq X \<Longrightarrow> monoseq X"
by (simp add: decseq_def monoseq_def)
--- a/src/HOL/SET_Protocol/Message_SET.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/SET_Protocol/Message_SET.thy Fri Mar 12 16:02:42 2010 +0100
@@ -7,7 +7,7 @@
header{*The Message Theory, Modified for SET*}
theory Message_SET
-imports Main Nat_Int_Bij
+imports Main Nat_Bijection
begin
subsection{*General Lemmas*}
@@ -81,17 +81,16 @@
definition nat_of_agent :: "agent => nat" where
- "nat_of_agent == agent_case (curry nat2_to_nat 0)
- (curry nat2_to_nat 1)
- (curry nat2_to_nat 2)
- (curry nat2_to_nat 3)
- (nat2_to_nat (4,0))"
+ "nat_of_agent == agent_case (curry prod_encode 0)
+ (curry prod_encode 1)
+ (curry prod_encode 2)
+ (curry prod_encode 3)
+ (prod_encode (4,0))"
--{*maps each agent to a unique natural number, for specifications*}
text{*The function is indeed injective*}
lemma inj_nat_of_agent: "inj nat_of_agent"
-by (simp add: nat_of_agent_def inj_on_def curry_def
- nat2_to_nat_inj [THEN inj_eq] split: agent.split)
+by (simp add: nat_of_agent_def inj_on_def curry_def prod_encode_eq split: agent.split)
constdefs
--- a/src/HOL/SET_Protocol/Purchase.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/SET_Protocol/Purchase.thy Fri Mar 12 16:02:42 2010 +0100
@@ -280,9 +280,9 @@
inj_PANSecret: "inj PANSecret"
CardSecret_neq_PANSecret: "CardSecret k \<noteq> PANSecret k'"
--{*No CardSecret equals any PANSecret*}
- apply (rule_tac x="curry nat2_to_nat 0" in exI)
- apply (rule_tac x="curry nat2_to_nat 1" in exI)
- apply (simp add: nat2_to_nat_inj [THEN inj_eq] inj_on_def)
+ apply (rule_tac x="curry prod_encode 0" in exI)
+ apply (rule_tac x="curry prod_encode 1" in exI)
+ apply (simp add: prod_encode_eq inj_on_def)
done
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
--- a/src/HOL/SET_Protocol/ROOT.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/SET_Protocol/ROOT.ML Fri Mar 12 16:02:42 2010 +0100
@@ -5,5 +5,5 @@
Root file for the SET protocol proofs.
*)
-no_document use_thys ["Nat_Int_Bij"];
+no_document use_thys ["Nat_Bijection"];
use_thys ["Cardholder_Registration", "Merchant_Registration", "Purchase"];
--- a/src/HOL/Tools/Nitpick/HISTORY Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/HISTORY Fri Mar 12 16:02:42 2010 +0100
@@ -2,15 +2,17 @@
* Added and implemented "binary_ints" and "bits" options
* Added "std" option and implemented support for nonstandard models
+ * Added and implemented "finitize" option to improve the precision of infinite
+ datatypes based on a monotonicity analysis
* Added support for quotient types
- * Added support for local definitions
- * Disabled "fast_descrs" option by default
+ * Added support for local definitions (for "function" and "termination"
+ proofs)
+ * Added support for term postprocessors
* Optimized "Multiset.multiset" and "FinFun.finfun"
* Improved efficiency of "destroy_constrs" optimization
- * Improved precision of infinite datatypes whose constructors don't appear
- in the formula to falsify based on a monotonicity analysis
* Fixed soundness bugs related to "destroy_constrs" optimization and record
getters
+ * Fixed soundness bug related to higher-order constructors
* Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and "SAT4JLight" to
"MiniSat_JNI", "zChaff_JNI", "BerkMin_Alloy", and "SAT4J_Light"
--- a/src/HOL/Tools/Nitpick/kodkod.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Fri Mar 12 16:02:42 2010 +0100
@@ -157,7 +157,8 @@
type raw_bound = n_ary_index * int list list
datatype outcome =
- NotInstalled |
+ JavaNotInstalled |
+ KodkodiNotInstalled |
Normal of (int * raw_bound list) list * int list * string |
TimedOut of int list |
Interrupted of int list option |
@@ -303,7 +304,8 @@
type raw_bound = n_ary_index * int list list
datatype outcome =
- NotInstalled |
+ JavaNotInstalled |
+ KodkodiNotInstalled |
Normal of (int * raw_bound list) list * int list * string |
TimedOut of int list |
Interrupted of int list option |
@@ -317,6 +319,8 @@
int_expr_func: int_expr -> 'a -> 'a
}
+(** Auxiliary functions on ML representation of Kodkod problems **)
+
(* 'a fold_expr_funcs -> formula -> 'a -> 'a *)
fun fold_formula (F : 'a fold_expr_funcs) formula =
case formula of
@@ -510,18 +514,7 @@
filter (is_relevant_setting o fst) (#settings p1)
= filter (is_relevant_setting o fst) (#settings p2)
-val created_temp_dir = Unsynchronized.ref false
-
-(* bool -> string * string *)
-fun serial_string_and_temporary_dir_for_overlord overlord =
- if overlord then
- ("", getenv "ISABELLE_HOME_USER")
- else
- let
- val dir = getenv "ISABELLE_TMP"
- val _ = if !created_temp_dir then ()
- else (created_temp_dir := true; File.mkdir (Path.explode dir))
- in (serial_string (), dir) end
+(** Serialization of problem **)
(* int -> string *)
fun base_name j =
@@ -893,12 +886,14 @@
map (fn b => (out_assign b; out ";")) expr_assigns;
out "solve "; out_outmost_f formula; out ";\n")
in
- out ("// This file was generated by Isabelle (probably Nitpick)\n" ^
+ out ("// This file was generated by Isabelle (most likely Nitpick)\n" ^
"// " ^ Date.fmt "%Y-%m-%d %H:%M:%S"
(Date.fromTimeLocal (Time.now ())) ^ "\n");
map out_problem problems
end
+(** Parsing of solution **)
+
(* string -> bool *)
fun is_ident_char s =
Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse
@@ -1014,6 +1009,21 @@
is partly due to the JVM and partly due to the ML "bash" function. *)
val fudge_ms = 250
+(** Main Kodkod entry point **)
+
+val created_temp_dir = Unsynchronized.ref false
+
+(* bool -> string * string *)
+fun serial_string_and_temporary_dir_for_overlord overlord =
+ if overlord then
+ ("", getenv "ISABELLE_HOME_USER")
+ else
+ let
+ val dir = getenv "ISABELLE_TMP"
+ val _ = if !created_temp_dir then ()
+ else (created_temp_dir := true; File.mkdir (Path.explode dir))
+ in (serial_string (), dir) end
+
(* bool -> Time.time option -> int -> int -> problem list -> outcome *)
fun solve_any_problem overlord deadline max_threads max_solutions problems =
let
@@ -1083,8 +1093,11 @@
||> apfst (map (apfst reindex)) ||> apsnd (map reindex)
val js = triv_js @ nontriv_js
val first_error =
- File.fold_lines (fn line => fn "" => line | s => s) err_path ""
- handle IO.Io _ => "" | OS.SysErr _ => ""
+ (File.fold_lines (fn line => fn "" => line | s => s) err_path ""
+ handle IO.Io _ => "" | OS.SysErr _ => "")
+ |> perhaps (try (unsuffix "\r"))
+ |> perhaps (try (unsuffix "."))
+ |> perhaps (try (unprefix "Error: "))
in
if null ps then
if code = 2 then
@@ -1092,12 +1105,15 @@
else if code = 0 then
Normal ([], js, first_error)
else if first_error |> Substring.full
+ |> Substring.position "exec: java" |> snd
+ |> Substring.isEmpty |> not then
+ JavaNotInstalled
+ else if first_error |> Substring.full
|> Substring.position "NoClassDefFoundError" |> snd
|> Substring.isEmpty |> not then
- NotInstalled
+ KodkodiNotInstalled
else if first_error <> "" then
- Error (first_error |> perhaps (try (unsuffix "."))
- |> perhaps (try (unprefix "Error: ")), js)
+ Error (first_error, js)
else if io_error then
Error ("I/O error", js)
else
--- a/src/HOL/Tools/Nitpick/minipick.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/minipick.ML Fri Mar 12 16:02:42 2010 +0100
@@ -336,7 +336,8 @@
val max_solutions = 1
in
case solve_any_problem overlord NONE max_threads max_solutions problems of
- NotInstalled => "unknown"
+ JavaNotInstalled => "unknown"
+ | KodkodiNotInstalled => "unknown"
| Normal ([], _, _) => "none"
| Normal _ => "genuine"
| TimedOut _ => "unknown"
--- a/src/HOL/Tools/Nitpick/nitpick.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Fri Mar 12 16:02:42 2010 +0100
@@ -8,6 +8,7 @@
signature NITPICK =
sig
type styp = Nitpick_Util.styp
+ type term_postprocessor = Nitpick_Model.term_postprocessor
type params = {
cards_assigns: (typ option * int list) list,
maxes_assigns: (styp option * int list) list,
@@ -58,6 +59,9 @@
val unregister_frac_type : string -> theory -> theory
val register_codatatype : typ -> string -> styp list -> theory -> theory
val unregister_codatatype : typ -> theory -> theory
+ val register_term_postprocessor :
+ typ -> term_postprocessor -> theory -> theory
+ val unregister_term_postprocessor : typ -> theory -> theory
val pick_nits_in_term :
Proof.state -> params -> bool -> int -> int -> int -> (term * term) list
-> term list -> term -> string * Proof.state
@@ -149,6 +153,9 @@
(length ts downto 1) ts))]
(* unit -> string *)
+fun install_java_message () =
+ "Nitpick requires a Java 1.5 virtual machine called \"java\"."
+(* unit -> string *)
fun install_kodkodi_message () =
"Nitpick requires the external Java program Kodkodi. To install it, download \
\the package from Isabelle's web page and add the \"kodkodi-x.y.z\" \
@@ -326,8 +333,6 @@
val sound_finitizes =
none_true (filter_out (fn (SOME (Type (@{type_name fun}, _)), _) => true
| _ => false) finitizes)
- val genuine_means_genuine =
- got_all_user_axioms andalso none_true wfs andalso sound_finitizes
val standard = forall snd stds
(*
val _ = List.app (priority o string_for_nut ctxt) (nondef_us @ def_us)
@@ -603,7 +608,7 @@
fun show_kodkod_warning "" = ()
| show_kodkod_warning s = print_m (fn () => "Kodkod warning: " ^ s ^ ".")
- (* bool -> KK.raw_bound list -> problem_extension -> bool option *)
+ (* bool -> KK.raw_bound list -> problem_extension -> bool * bool option *)
fun print_and_check_model genuine bounds
({free_names, sel_names, nonsel_names, rel_table, scope, ...}
: problem_extension) =
@@ -614,119 +619,129 @@
show_consts = show_consts}
scope formats frees free_names sel_names nonsel_names rel_table
bounds
- val genuine_means_genuine = genuine_means_genuine andalso codatatypes_ok
+ val genuine_means_genuine =
+ got_all_user_axioms andalso none_true wfs andalso
+ sound_finitizes andalso codatatypes_ok
in
- pprint (Pretty.chunks
- [Pretty.blk (0,
- (pstrs ("Nitpick found a" ^
- (if not genuine then " potential "
- else if genuine_means_genuine then " "
- else " likely genuine ") ^ das_wort_model) @
- (case pretties_for_scope scope verbose of
- [] => []
- | pretties => pstrs " for " @ pretties) @
- [Pretty.str ":\n"])),
- Pretty.indent indent_size reconstructed_model]);
- if genuine then
- (if check_genuine andalso standard then
- (case prove_hol_model scope tac_timeout free_names sel_names
+ (pprint (Pretty.chunks
+ [Pretty.blk (0,
+ (pstrs ("Nitpick found a" ^
+ (if not genuine then " potential "
+ else if genuine_means_genuine then " "
+ else " quasi genuine ") ^ das_wort_model) @
+ (case pretties_for_scope scope verbose of
+ [] => []
+ | pretties => pstrs " for " @ pretties) @
+ [Pretty.str ":\n"])),
+ Pretty.indent indent_size reconstructed_model]);
+ if genuine then
+ (if check_genuine andalso standard then
+ case prove_hol_model scope tac_timeout free_names sel_names
rel_table bounds assms_t of
- SOME true => print ("Confirmation by \"auto\": The above " ^
- das_wort_model ^ " is really genuine.")
+ SOME true =>
+ print ("Confirmation by \"auto\": The above " ^
+ das_wort_model ^ " is really genuine.")
| SOME false =>
if genuine_means_genuine then
error ("A supposedly genuine " ^ das_wort_model ^ " was \
\shown to be spurious by \"auto\".\nThis should never \
\happen.\nPlease send a bug report to blanchet\
\te@in.tum.de.")
- else
- print ("Refutation by \"auto\": The above " ^ das_wort_model ^
- " is spurious.")
- | NONE => print "No confirmation by \"auto\".")
- else
- ();
- if not standard andalso likely_in_struct_induct_step then
- print "The existence of a nonstandard model suggests that the \
- \induction hypothesis is not general enough or perhaps even \
- \wrong. See the \"Inductive Properties\" section of the \
- \Nitpick manual for details (\"isabelle doc nitpick\")."
- else
- ();
- if has_syntactic_sorts orig_t then
- print "Hint: Maybe you forgot a type constraint?"
+ else
+ print ("Refutation by \"auto\": The above " ^
+ das_wort_model ^ " is spurious.")
+ | NONE => print "No confirmation by \"auto\"."
+ else
+ ();
+ if not standard andalso likely_in_struct_induct_step then
+ print "The existence of a nonstandard model suggests that the \
+ \induction hypothesis is not general enough or perhaps \
+ \even wrong. See the \"Inductive Properties\" section of \
+ \the Nitpick manual for details (\"isabelle doc nitpick\")."
+ else
+ ();
+ if has_syntactic_sorts orig_t then
+ print "Hint: Maybe you forgot a type constraint?"
+ else
+ ();
+ if not genuine_means_genuine then
+ if no_poly_user_axioms then
+ let
+ val options =
+ [] |> not got_all_mono_user_axioms
+ ? cons ("user_axioms", "\"true\"")
+ |> not (none_true wfs)
+ ? cons ("wf", "\"smart\" or \"false\"")
+ |> not sound_finitizes
+ ? cons ("finitize", "\"smart\" or \"false\"")
+ |> not codatatypes_ok
+ ? cons ("bisim_depth", "a nonnegative value")
+ val ss =
+ map (fn (name, value) => quote name ^ " set to " ^ value)
+ options
+ in
+ print ("Try again with " ^
+ space_implode " " (serial_commas "and" ss) ^
+ " to confirm that the " ^ das_wort_model ^
+ " is genuine.")
+ end
+ else
+ print ("Nitpick is unable to guarantee the authenticity of \
+ \the " ^ das_wort_model ^ " in the presence of \
+ \polymorphic axioms.")
+ else
+ ();
+ NONE)
+ else
+ if not genuine then
+ (Unsynchronized.inc met_potential;
+ if check_potential then
+ let
+ val status = prove_hol_model scope tac_timeout free_names
+ sel_names rel_table bounds assms_t
+ in
+ (case status of
+ SOME true => print ("Confirmation by \"auto\": The \
+ \above " ^ das_wort_model ^
+ " is genuine.")
+ | SOME false => print ("Refutation by \"auto\": The above " ^
+ das_wort_model ^ " is spurious.")
+ | NONE => print "No confirmation by \"auto\".");
+ status
+ end
+ else
+ NONE)
else
- ();
- if not genuine_means_genuine then
- if no_poly_user_axioms then
- let
- val options =
- [] |> not got_all_mono_user_axioms
- ? cons ("user_axioms", "\"true\"")
- |> not (none_true wfs)
- ? cons ("wf", "\"smart\" or \"false\"")
- |> not sound_finitizes
- ? cons ("finitize", "\"smart\" or \"false\"")
- |> not codatatypes_ok
- ? cons ("bisim_depth", "a nonnegative value")
- val ss =
- map (fn (name, value) => quote name ^ " set to " ^ value)
- options
- in
- print ("Try again with " ^
- space_implode " " (serial_commas "and" ss) ^
- " to confirm that the " ^ das_wort_model ^
- " is genuine.")
- end
- else
- print ("Nitpick is unable to guarantee the authenticity of \
- \the " ^ das_wort_model ^ " in the presence of \
- \polymorphic axioms.")
- else
- ();
- NONE)
- else
- if not genuine then
- (Unsynchronized.inc met_potential;
- if check_potential then
- let
- val status = prove_hol_model scope tac_timeout free_names
- sel_names rel_table bounds assms_t
- in
- (case status of
- SOME true => print ("Confirmation by \"auto\": The above " ^
- das_wort_model ^ " is genuine.")
- | SOME false => print ("Refutation by \"auto\": The above " ^
- das_wort_model ^ " is spurious.")
- | NONE => print "No confirmation by \"auto\".");
- status
- end
- else
- NONE)
- else
- NONE
+ NONE)
+ |> pair genuine_means_genuine
end
- (* int -> int -> int -> bool -> rich_problem list -> int * int * int *)
- fun solve_any_problem max_potential max_genuine donno first_time problems =
+ (* bool * int * int * int -> bool -> rich_problem list
+ -> bool * int * int * int *)
+ fun solve_any_problem (found_really_genuine, max_potential, max_genuine,
+ donno) first_time problems =
let
val max_potential = Int.max (0, max_potential)
val max_genuine = Int.max (0, max_genuine)
- (* bool -> int * KK.raw_bound list -> bool option *)
+ (* bool -> int * KK.raw_bound list -> bool * bool option *)
fun print_and_check genuine (j, bounds) =
print_and_check_model genuine bounds (snd (nth problems j))
val max_solutions = max_potential + max_genuine
|> not need_incremental ? curry Int.min 1
in
if max_solutions <= 0 then
- (0, 0, donno)
+ (found_really_genuine, 0, 0, donno)
else
case KK.solve_any_problem overlord deadline max_threads max_solutions
(map fst problems) of
- KK.NotInstalled =>
+ KK.JavaNotInstalled =>
+ (print_m install_java_message;
+ (found_really_genuine, max_potential, max_genuine, donno + 1))
+ | KK.KodkodiNotInstalled =>
(print_m install_kodkodi_message;
- (max_potential, max_genuine, donno + 1))
+ (found_really_genuine, max_potential, max_genuine, donno + 1))
| KK.Normal ([], unsat_js, s) =>
(update_checked_problems problems unsat_js; show_kodkod_warning s;
- (max_potential, max_genuine, donno))
+ (found_really_genuine, max_potential, max_genuine, donno))
| KK.Normal (sat_ps, unsat_js, s) =>
let
val (lib_ps, con_ps) =
@@ -736,16 +751,19 @@
show_kodkod_warning s;
if null con_ps then
let
- val num_genuine = take max_potential lib_ps
- |> map (print_and_check false)
- |> filter (curry (op =) (SOME true))
- |> length
+ val genuine_codes =
+ lib_ps |> take max_potential
+ |> map (print_and_check false)
+ |> filter (curry (op =) (SOME true) o snd)
+ val found_really_genuine =
+ found_really_genuine orelse exists fst genuine_codes
+ val num_genuine = length genuine_codes
val max_genuine = max_genuine - num_genuine
val max_potential = max_potential
- (length lib_ps - num_genuine)
in
if max_genuine <= 0 then
- (0, 0, donno)
+ (found_really_genuine, 0, 0, donno)
else
let
(* "co_js" is the list of sound problems whose unsound
@@ -765,18 +783,21 @@
|> max_potential <= 0
? filter_out (#unsound o snd)
in
- solve_any_problem max_potential max_genuine donno false
- problems
+ solve_any_problem (found_really_genuine, max_potential,
+ max_genuine, donno) false problems
end
end
else
let
- val _ = take max_genuine con_ps
- |> List.app (ignore o print_and_check true)
- val max_genuine = max_genuine - length con_ps
+ val genuine_codes =
+ con_ps |> take max_genuine
+ |> map (print_and_check true)
+ val max_genuine = max_genuine - length genuine_codes
+ val found_really_genuine =
+ found_really_genuine orelse exists fst genuine_codes
in
if max_genuine <= 0 orelse not first_time then
- (0, max_genuine, donno)
+ (found_really_genuine, 0, max_genuine, donno)
else
let
val bye_js = sort_distinct int_ord
@@ -784,7 +805,10 @@
val problems =
problems |> filter_out_indices bye_js
|> filter_out (#unsound o snd)
- in solve_any_problem 0 max_genuine donno false problems end
+ in
+ solve_any_problem (found_really_genuine, 0, max_genuine,
+ donno) false problems
+ end
end
end
| KK.TimedOut unsat_js =>
@@ -795,11 +819,13 @@
| KK.Error (s, unsat_js) =>
(update_checked_problems problems unsat_js;
print_v (K ("Kodkod error: " ^ s ^ "."));
- (max_potential, max_genuine, donno + 1))
+ (found_really_genuine, max_potential, max_genuine, donno + 1))
end
- (* int -> int -> scope list -> int * int * int -> int * int * int *)
- fun run_batch j n scopes (max_potential, max_genuine, donno) =
+ (* int -> int -> scope list -> bool * int * int * int
+ -> bool * int * int * int *)
+ fun run_batch j n scopes (found_really_genuine, max_potential, max_genuine,
+ donno) =
let
val _ =
if null scopes then
@@ -869,7 +895,8 @@
else
()
in
- solve_any_problem max_potential max_genuine donno true (rev problems)
+ solve_any_problem (found_really_genuine, max_potential, max_genuine,
+ donno) true (rev problems)
end
(* rich_problem list -> scope -> int *)
@@ -894,17 +921,17 @@
in
"Nitpick " ^ did_so_and_so ^
(if is_some (!checked_problems) andalso total > 0 then
- " after checking " ^
- string_of_int (Int.min (total - 1, unsat)) ^ " of " ^
+ " " ^ string_of_int (Int.min (total - 1, unsat)) ^ " of " ^
string_of_int total ^ " scope" ^ plural_s total
else
"") ^ "."
end
- (* int -> int -> scope list -> int * int * int -> KK.outcome *)
- fun run_batches _ _ [] (max_potential, max_genuine, donno) =
+ (* int -> int -> scope list -> bool * int * int * int -> KK.outcome *)
+ fun run_batches _ _ []
+ (found_really_genuine, max_potential, max_genuine, donno) =
if donno > 0 andalso max_genuine > 0 then
- (print_m (fn () => excipit "ran out of some resource"); "unknown")
+ (print_m (fn () => excipit "checked"); "unknown")
else if max_genuine = original_max_genuine then
if max_potential = original_max_potential then
(print_m (fn () =>
@@ -919,10 +946,12 @@
"Nitpick could not find a" ^
(if max_genuine = 1 then " better " ^ das_wort_model ^ "."
else "ny better " ^ das_wort_model ^ "s.")); "potential")
+ else if found_really_genuine then
+ "genuine"
else
- if genuine_means_genuine then "genuine" else "likely_genuine"
+ "quasi_genuine"
| run_batches j n (batch :: batches) z =
- let val (z as (_, max_genuine, _)) = run_batch j n batch z in
+ let val (z as (_, _, max_genuine, _)) = run_batch j n batch z in
run_batches (j + 1) n (if max_genuine > 0 then batches else []) z
end
@@ -942,13 +971,15 @@
val batches = batch_list batch_size (!scopes)
val outcome_code =
- (run_batches 0 (length batches) batches (max_potential, max_genuine, 0)
+ (run_batches 0 (length batches) batches
+ (false, max_potential, max_genuine, 0)
handle Exn.Interrupt => do_interrupted ())
handle TimeLimit.TimeOut =>
- (print_m (fn () => excipit "ran out of time");
+ (print_m (fn () => excipit "ran out of time after checking");
if !met_potential > 0 then "potential" else "unknown")
- | Exn.Interrupt => if auto orelse debug then raise Interrupt
- else error (excipit "was interrupted")
+ | Exn.Interrupt =>
+ if auto orelse debug then raise Interrupt
+ else error (excipit "was interrupted after checking")
val _ = print_v (fn () => "Total time: " ^
signed_string_of_int (Time.toMilliseconds
(Timer.checkRealTimer timer)) ^ " ms.")
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Mar 12 16:02:42 2010 +0100
@@ -52,6 +52,9 @@
val name_sep : string
val numeral_prefix : string
+ val base_prefix : string
+ val step_prefix : string
+ val unrolled_prefix : string
val ubfp_prefix : string
val lbfp_prefix : string
val quot_normal_prefix : string
@@ -59,6 +62,8 @@
val special_prefix : string
val uncurry_prefix : string
val eval_prefix : string
+ val iter_var_prefix : string
+ val strip_first_name_sep : string -> string * string
val original_name : string -> string
val s_conj : term * term -> term
val s_disj : term * term -> term
@@ -72,9 +77,11 @@
val shortest_name : string -> string
val short_name : string -> string
val shorten_names_in_term : term -> term
+ val strict_type_match : theory -> typ * typ -> bool
val type_match : theory -> typ * typ -> bool
val const_match : theory -> styp * styp -> bool
val term_match : theory -> term * term -> bool
+ val frac_from_term_pair : typ -> term -> term -> term
val is_TFree : typ -> bool
val is_higher_order_type : typ -> bool
val is_fun_type : typ -> bool
@@ -167,6 +174,7 @@
val term_under_def : term -> term
val case_const_names :
theory -> (typ option * bool) list -> (string * int) list
+ val unfold_defs_in_term : hol_context -> term -> term
val const_def_table :
Proof.context -> (term * term) list -> term list -> const_table
val const_nondef_table : term list -> const_table
@@ -182,13 +190,13 @@
val optimized_quot_type_axioms :
Proof.context -> (typ option * bool) list -> string * typ list -> term list
val def_of_const : theory -> const_table -> styp -> term option
+ val fixpoint_kind_of_rhs : term -> fixpoint_kind
val fixpoint_kind_of_const :
theory -> const_table -> string * typ -> fixpoint_kind
val is_inductive_pred : hol_context -> styp -> bool
val is_equational_fun : hol_context -> styp -> bool
val is_constr_pattern_lhs : theory -> term -> bool
val is_constr_pattern_formula : theory -> term -> bool
- val unfold_defs_in_term : hol_context -> term -> term
val codatatype_bisim_axioms : hol_context -> typ -> term list
val is_well_founded_inductive_pred : hol_context -> styp -> bool
val unrolled_inductive_pred_const : hol_context -> bool -> styp -> term
@@ -197,13 +205,6 @@
val merge_type_vars_in_terms : term list -> term list
val ground_types_in_type : hol_context -> bool -> typ -> typ list
val ground_types_in_terms : hol_context -> bool -> term list -> typ list
- val format_type : int list -> int list -> typ -> typ
- val format_term_type :
- theory -> const_table -> (term option * int list) list -> term -> typ
- val user_friendly_const :
- hol_context -> string * string -> (term option * int list) list
- -> styp -> term * typ
- val assign_operator_for_const : styp -> string
end;
structure Nitpick_HOL : NITPICK_HOL =
@@ -281,7 +282,8 @@
val uncurry_prefix = nitpick_prefix ^ "unc"
val eval_prefix = nitpick_prefix ^ "eval"
val iter_var_prefix = "i"
-val arg_var_prefix = "x"
+
+(** Constant/type information and term/type manipulation **)
(* int -> string *)
fun sel_prefix_for j = sel_prefix ^ string_of_int j ^ name_sep
@@ -299,7 +301,13 @@
case strip_first_name_sep s of (s1, "") => s1 | (_, s2) => original_name s2
else
s
-val after_name_sep = snd o strip_first_name_sep
+
+(* term * term -> term *)
+fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
+ | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t
+ | s_betapply p = betapply p
+(* term * term list -> term *)
+val s_betapplys = Library.foldl s_betapply
(* term * term -> term *)
fun s_conj (t1, @{const True}) = t1
@@ -361,7 +369,8 @@
(@{const_name finite}, 1),
(@{const_name unknown}, 0),
(@{const_name is_unknown}, 1),
- (@{const_name Tha}, 1),
+ (@{const_name safe_The}, 1),
+ (@{const_name safe_Eps}, 1),
(@{const_name Frac}, 0),
(@{const_name norm_frac}, 0)]
val built_in_nat_consts =
@@ -456,6 +465,15 @@
const_match thy ((shortest_name s1, T1), (shortest_name s2, T2))
| term_match _ (t1, t2) = t1 aconv t2
+(* typ -> term -> term -> term *)
+fun frac_from_term_pair T t1 t2 =
+ case snd (HOLogic.dest_number t1) of
+ 0 => HOLogic.mk_number T 0
+ | n1 => case snd (HOLogic.dest_number t2) of
+ 1 => HOLogic.mk_number T n1
+ | n2 => Const (@{const_name divide}, T --> T --> T)
+ $ HOLogic.mk_number T n1 $ HOLogic.mk_number T n2
+
(* typ -> bool *)
fun is_TFree (TFree _) = true
| is_TFree _ = false
@@ -493,7 +511,8 @@
Type ((if gfp then gfp_iterator_prefix else lfp_iterator_prefix) ^ s,
binder_types T)
(* typ -> styp *)
-fun const_for_iterator_type (Type (s, Ts)) = (after_name_sep s, Ts ---> bool_T)
+fun const_for_iterator_type (Type (s, Ts)) =
+ (strip_first_name_sep s |> snd, Ts ---> bool_T)
| const_for_iterator_type T =
raise TYPE ("Nitpick_HOL.const_for_iterator_type", [T], [])
@@ -528,13 +547,6 @@
fun dest_n_tuple 1 t = [t]
| dest_n_tuple n t = HOLogic.dest_prod t ||> dest_n_tuple (n - 1) |> op ::
-(* int -> typ -> typ list *)
-fun dest_n_tuple_type 1 T = [T]
- | dest_n_tuple_type n (Type (_, [T1, T2])) =
- T1 :: dest_n_tuple_type (n - 1) T2
- | dest_n_tuple_type _ T =
- raise TYPE ("Nitpick_HOL.dest_n_tuple_type", [T], [])
-
type typedef_info =
{rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string,
set_def: thm option, prop_of_Rep: thm, set_name: string,
@@ -1451,106 +1463,57 @@
(Datatype.get_all thy) [] @
map (apsnd length o snd) (#codatatypes (Data.get thy))
-(* Proof.context -> (term * term) list -> term list -> const_table *)
-fun const_def_table ctxt subst ts =
- table_for (map prop_of o Nitpick_Defs.get) ctxt subst
- |> fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
- (map pair_for_prop ts)
-(* term list -> const_table *)
-fun const_nondef_table ts =
- fold (fn t => append (map (fn s => (s, t)) (Term.add_const_names t []))) ts []
- |> AList.group (op =) |> Symtab.make
-(* Proof.context -> (term * term) list -> const_table *)
-val const_simp_table = table_for (map prop_of o Nitpick_Simps.get)
-val const_psimp_table = table_for (map prop_of o Nitpick_Psimps.get)
-(* Proof.context -> (term * term) list -> const_table -> const_table *)
-fun inductive_intro_table ctxt subst def_table =
- table_for (map (unfold_mutually_inductive_preds (ProofContext.theory_of ctxt)
- def_table o prop_of)
- o Nitpick_Intros.get) ctxt subst
-(* theory -> term list Inttab.table *)
-fun ground_theorem_table thy =
- fold ((fn @{const Trueprop} $ t1 =>
- is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1)
- | _ => I) o prop_of o snd) (PureThy.all_thms_of thy) Inttab.empty
+(* theory -> const_table -> string * typ -> fixpoint_kind *)
+fun fixpoint_kind_of_const thy table x =
+ if is_built_in_const thy [(NONE, false)] false x then
+ NoFp
+ else
+ fixpoint_kind_of_rhs (the (def_of_const thy table x))
+ handle Option.Option => NoFp
-val basic_ersatz_table =
- [(@{const_name prod_case}, @{const_name split}),
- (@{const_name card}, @{const_name card'}),
- (@{const_name setsum}, @{const_name setsum'}),
- (@{const_name fold_graph}, @{const_name fold_graph'}),
- (@{const_name wf}, @{const_name wf'}),
- (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
- (@{const_name wfrec}, @{const_name wfrec'})]
-
-(* theory -> (string * string) list *)
-fun ersatz_table thy =
- fold (append o snd) (#frac_types (Data.get thy)) basic_ersatz_table
-
-(* const_table Unsynchronized.ref -> string -> term list -> unit *)
-fun add_simps simp_table s eqs =
- Unsynchronized.change simp_table
- (Symtab.update (s, eqs @ these (Symtab.lookup (!simp_table) s)))
+(* hol_context -> styp -> bool *)
+fun is_real_inductive_pred ({thy, stds, fast_descrs, def_table, intro_table,
+ ...} : hol_context) x =
+ fixpoint_kind_of_const thy def_table x <> NoFp andalso
+ not (null (def_props_for_const thy stds fast_descrs intro_table x))
+fun is_real_equational_fun ({thy, stds, fast_descrs, simp_table, psimp_table,
+ ...} : hol_context) x =
+ exists (fn table => not (null (def_props_for_const thy stds fast_descrs table
+ x)))
+ [!simp_table, psimp_table]
+fun is_inductive_pred hol_ctxt =
+ is_real_inductive_pred hol_ctxt andf (not o is_real_equational_fun hol_ctxt)
+fun is_equational_fun hol_ctxt =
+ (is_real_equational_fun hol_ctxt orf is_real_inductive_pred hol_ctxt orf
+ (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
-(* theory -> styp -> term list *)
-fun inverse_axioms_for_rep_fun thy (x as (_, T)) =
- let val abs_T = domain_type T in
- typedef_info thy (fst (dest_Type abs_T)) |> the
- |> pairf #Abs_inverse #Rep_inverse
- |> pairself (Refute.specialize_type thy x o prop_of o the)
- ||> single |> op ::
- end
-(* theory -> string * typ list -> term list *)
-fun optimized_typedef_axioms thy (abs_z as (abs_s, _)) =
- let val abs_T = Type abs_z in
- if is_univ_typedef thy abs_T then
- []
- else case typedef_info thy abs_s of
- SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
- let
- val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
- val rep_t = Const (Rep_name, abs_T --> rep_T)
- val set_t = Const (set_name, rep_T --> bool_T)
- val set_t' =
- prop_of_Rep |> HOLogic.dest_Trueprop
- |> Refute.specialize_type thy (dest_Const rep_t)
- |> HOLogic.dest_mem |> snd
- in
- [HOLogic.all_const abs_T
- $ Abs (Name.uu, abs_T, set_t $ (rep_t $ Bound 0))]
- |> set_t <> set_t' ? cons (HOLogic.mk_eq (set_t, set_t'))
- |> map HOLogic.mk_Trueprop
- end
- | NONE => []
- end
-(* Proof.context -> string * typ list -> term list *)
-fun optimized_quot_type_axioms ctxt stds abs_z =
- let
- val thy = ProofContext.theory_of ctxt
- val abs_T = Type abs_z
- val rep_T = rep_type_for_quot_type thy abs_T
- val equiv_rel = equiv_relation_for_quot_type thy abs_T
- val a_var = Var (("a", 0), abs_T)
- val x_var = Var (("x", 0), rep_T)
- val y_var = Var (("y", 0), rep_T)
- val x = (@{const_name Quot}, rep_T --> abs_T)
- val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
- val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
- val normal_x = normal_t $ x_var
- val normal_y = normal_t $ y_var
- val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
- in
- [Logic.mk_equals (normal_t $ sel_a_t, sel_a_t),
- Logic.list_implies
- ([@{const Not} $ (is_unknown_t $ normal_x),
- @{const Not} $ (is_unknown_t $ normal_y),
- equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
- Logic.mk_equals (normal_x, normal_y)),
- Logic.list_implies
- ([HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)),
- HOLogic.mk_Trueprop (@{const Not} $ HOLogic.mk_eq (normal_x, x_var))],
- HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
- end
+(* term -> term *)
+fun lhs_of_equation t =
+ case t of
+ Const (@{const_name all}, _) $ Abs (_, _, t1) => lhs_of_equation t1
+ | Const (@{const_name "=="}, _) $ t1 $ _ => SOME t1
+ | @{const "==>"} $ _ $ t2 => lhs_of_equation t2
+ | @{const Trueprop} $ t1 => lhs_of_equation t1
+ | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
+ | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
+ | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
+ | _ => NONE
+(* theory -> term -> bool *)
+fun is_constr_pattern _ (Bound _) = true
+ | is_constr_pattern _ (Var _) = true
+ | is_constr_pattern thy t =
+ case strip_comb t of
+ (Const x, args) =>
+ is_constr_like thy x andalso forall (is_constr_pattern thy) args
+ | _ => false
+fun is_constr_pattern_lhs thy t =
+ forall (is_constr_pattern thy) (snd (strip_comb t))
+fun is_constr_pattern_formula thy t =
+ case lhs_of_equation t of
+ SOME t' => is_constr_pattern_lhs thy t'
+ | NONE => false
+
+(** Constant unfolding **)
(* theory -> (typ option * bool) list -> int * styp -> term *)
fun constr_case_body thy stds (j, (x as (_, T))) =
@@ -1574,7 +1537,6 @@
|> fold_rev (add_constr_case hol_ctxt res_T) (length xs downto 2 ~~ xs')
|> fold_rev (curry absdummy) (func_Ts @ [dataT])
end
-
(* hol_context -> string -> typ -> typ -> term -> term *)
fun optimized_record_get (hol_ctxt as {thy, stds, ...}) s rec_T res_T t =
let val constr_x = hd (datatype_constrs hol_ctxt rec_T) in
@@ -1612,63 +1574,6 @@
end) (index_seq 0 n) Ts
in list_comb (Const constr_x, ts) end
-(* theory -> const_table -> string * typ -> fixpoint_kind *)
-fun fixpoint_kind_of_const thy table x =
- if is_built_in_const thy [(NONE, false)] false x then
- NoFp
- else
- fixpoint_kind_of_rhs (the (def_of_const thy table x))
- handle Option.Option => NoFp
-
-(* hol_context -> styp -> bool *)
-fun is_real_inductive_pred ({thy, stds, fast_descrs, def_table, intro_table,
- ...} : hol_context) x =
- fixpoint_kind_of_const thy def_table x <> NoFp andalso
- not (null (def_props_for_const thy stds fast_descrs intro_table x))
-fun is_real_equational_fun ({thy, stds, fast_descrs, simp_table, psimp_table,
- ...} : hol_context) x =
- exists (fn table => not (null (def_props_for_const thy stds fast_descrs table
- x)))
- [!simp_table, psimp_table]
-fun is_inductive_pred hol_ctxt =
- is_real_inductive_pred hol_ctxt andf (not o is_real_equational_fun hol_ctxt)
-fun is_equational_fun hol_ctxt =
- (is_real_equational_fun hol_ctxt orf is_real_inductive_pred hol_ctxt orf
- (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
-
-(* term * term -> term *)
-fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
- | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t
- | s_betapply p = betapply p
-(* term * term list -> term *)
-val s_betapplys = Library.foldl s_betapply
-
-(* term -> term *)
-fun lhs_of_equation t =
- case t of
- Const (@{const_name all}, _) $ Abs (_, _, t1) => lhs_of_equation t1
- | Const (@{const_name "=="}, _) $ t1 $ _ => SOME t1
- | @{const "==>"} $ _ $ t2 => lhs_of_equation t2
- | @{const Trueprop} $ t1 => lhs_of_equation t1
- | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
- | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
- | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
- | _ => NONE
-(* theory -> term -> bool *)
-fun is_constr_pattern _ (Bound _) = true
- | is_constr_pattern _ (Var _) = true
- | is_constr_pattern thy t =
- case strip_comb t of
- (Const x, args) =>
- is_constr_like thy x andalso forall (is_constr_pattern thy) args
- | _ => false
-fun is_constr_pattern_lhs thy t =
- forall (is_constr_pattern thy) (snd (strip_comb t))
-fun is_constr_pattern_formula thy t =
- case lhs_of_equation t of
- SOME t' => is_constr_pattern_lhs thy t'
- | NONE => false
-
(* Prevents divergence in case of cyclic or infinite definition dependencies. *)
val unfold_max_depth = 255
@@ -1811,6 +1716,109 @@
in s_betapplys (const, map (do_term depth Ts) ts) |> Envir.beta_norm end
in do_term 0 [] end
+(** Axiom extraction/generation **)
+
+(* Proof.context -> (term * term) list -> term list -> const_table *)
+fun const_def_table ctxt subst ts =
+ table_for (map prop_of o Nitpick_Defs.get) ctxt subst
+ |> fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
+ (map pair_for_prop ts)
+(* term list -> const_table *)
+fun const_nondef_table ts =
+ fold (fn t => append (map (fn s => (s, t)) (Term.add_const_names t []))) ts []
+ |> AList.group (op =) |> Symtab.make
+(* Proof.context -> (term * term) list -> const_table *)
+val const_simp_table = table_for (map prop_of o Nitpick_Simps.get)
+val const_psimp_table = table_for (map prop_of o Nitpick_Psimps.get)
+(* Proof.context -> (term * term) list -> const_table -> const_table *)
+fun inductive_intro_table ctxt subst def_table =
+ table_for (map (unfold_mutually_inductive_preds (ProofContext.theory_of ctxt)
+ def_table o prop_of)
+ o Nitpick_Intros.get) ctxt subst
+(* theory -> term list Inttab.table *)
+fun ground_theorem_table thy =
+ fold ((fn @{const Trueprop} $ t1 =>
+ is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1)
+ | _ => I) o prop_of o snd) (PureThy.all_thms_of thy) Inttab.empty
+
+val basic_ersatz_table =
+ [(@{const_name prod_case}, @{const_name split}),
+ (@{const_name card}, @{const_name card'}),
+ (@{const_name setsum}, @{const_name setsum'}),
+ (@{const_name fold_graph}, @{const_name fold_graph'}),
+ (@{const_name wf}, @{const_name wf'}),
+ (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
+ (@{const_name wfrec}, @{const_name wfrec'})]
+
+(* theory -> (string * string) list *)
+fun ersatz_table thy =
+ fold (append o snd) (#frac_types (Data.get thy)) basic_ersatz_table
+
+(* const_table Unsynchronized.ref -> string -> term list -> unit *)
+fun add_simps simp_table s eqs =
+ Unsynchronized.change simp_table
+ (Symtab.update (s, eqs @ these (Symtab.lookup (!simp_table) s)))
+
+(* theory -> styp -> term list *)
+fun inverse_axioms_for_rep_fun thy (x as (_, T)) =
+ let val abs_T = domain_type T in
+ typedef_info thy (fst (dest_Type abs_T)) |> the
+ |> pairf #Abs_inverse #Rep_inverse
+ |> pairself (Refute.specialize_type thy x o prop_of o the)
+ ||> single |> op ::
+ end
+(* theory -> string * typ list -> term list *)
+fun optimized_typedef_axioms thy (abs_z as (abs_s, _)) =
+ let val abs_T = Type abs_z in
+ if is_univ_typedef thy abs_T then
+ []
+ else case typedef_info thy abs_s of
+ SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
+ let
+ val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
+ val rep_t = Const (Rep_name, abs_T --> rep_T)
+ val set_t = Const (set_name, rep_T --> bool_T)
+ val set_t' =
+ prop_of_Rep |> HOLogic.dest_Trueprop
+ |> Refute.specialize_type thy (dest_Const rep_t)
+ |> HOLogic.dest_mem |> snd
+ in
+ [HOLogic.all_const abs_T
+ $ Abs (Name.uu, abs_T, set_t $ (rep_t $ Bound 0))]
+ |> set_t <> set_t' ? cons (HOLogic.mk_eq (set_t, set_t'))
+ |> map HOLogic.mk_Trueprop
+ end
+ | NONE => []
+ end
+(* Proof.context -> string * typ list -> term list *)
+fun optimized_quot_type_axioms ctxt stds abs_z =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val abs_T = Type abs_z
+ val rep_T = rep_type_for_quot_type thy abs_T
+ val equiv_rel = equiv_relation_for_quot_type thy abs_T
+ val a_var = Var (("a", 0), abs_T)
+ val x_var = Var (("x", 0), rep_T)
+ val y_var = Var (("y", 0), rep_T)
+ val x = (@{const_name Quot}, rep_T --> abs_T)
+ val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
+ val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
+ val normal_x = normal_t $ x_var
+ val normal_y = normal_t $ y_var
+ val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
+ in
+ [Logic.mk_equals (normal_t $ sel_a_t, sel_a_t),
+ Logic.list_implies
+ ([@{const Not} $ (is_unknown_t $ normal_x),
+ @{const Not} $ (is_unknown_t $ normal_y),
+ equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
+ Logic.mk_equals (normal_x, normal_y)),
+ Logic.list_implies
+ ([HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)),
+ HOLogic.mk_Trueprop (@{const Not} $ HOLogic.mk_eq (normal_x, x_var))],
+ HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
+ end
+
(* hol_context -> typ -> term list *)
fun codatatype_bisim_axioms (hol_ctxt as {thy, stds, ...}) T =
let
@@ -1821,7 +1829,7 @@
val bisim_max = @{const bisim_iterator_max}
val n_var = Var (("n", 0), iter_T)
val n_var_minus_1 =
- Const (@{const_name Tha}, (iter_T --> bool_T) --> iter_T)
+ Const (@{const_name safe_The}, (iter_T --> bool_T) --> iter_T)
$ Abs ("m", iter_T, HOLogic.eq_const iter_T
$ (suc_const iter_T $ Bound 0) $ n_var)
val x_var = Var (("x", 0), T)
@@ -2143,7 +2151,7 @@
end
fun inductive_pred_axiom hol_ctxt (x as (s, T)) =
if String.isPrefix ubfp_prefix s orelse String.isPrefix lbfp_prefix s then
- let val x' = (after_name_sep s, T) in
+ let val x' = (strip_first_name_sep s |> snd, T) in
raw_inductive_pred_axiom hol_ctxt x' |> subst_atomic [(Const x', Const x)]
end
else
@@ -2165,6 +2173,8 @@
strip_comb t1 |> snd |> forall is_Var
| _ => false
+(** Type preprocessing **)
+
(* term list -> term list *)
fun merge_type_vars_in_terms ts =
let
@@ -2210,198 +2220,4 @@
fun ground_types_in_terms hol_ctxt binarize ts =
fold (fold_types (add_ground_types hol_ctxt binarize)) ts []
-(* theory -> const_table -> styp -> int list *)
-fun const_format thy def_table (x as (s, T)) =
- if String.isPrefix unrolled_prefix s then
- const_format thy def_table (original_name s, range_type T)
- else if String.isPrefix skolem_prefix s then
- let
- val k = unprefix skolem_prefix s
- |> strip_first_name_sep |> fst |> space_explode "@"
- |> hd |> Int.fromString |> the
- in [k, num_binder_types T - k] end
- else if original_name s <> s then
- [num_binder_types T]
- else case def_of_const thy def_table x of
- SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then
- let val k = length (strip_abs_vars t') in
- [k, num_binder_types T - k]
- end
- else
- [num_binder_types T]
- | NONE => [num_binder_types T]
-(* int list -> int list -> int list *)
-fun intersect_formats _ [] = []
- | intersect_formats [] _ = []
- | intersect_formats ks1 ks2 =
- let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in
- intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else []))
- (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @
- [Int.min (k1, k2)]
- end
-
-(* theory -> const_table -> (term option * int list) list -> term -> int list *)
-fun lookup_format thy def_table formats t =
- case AList.lookup (fn (SOME x, SOME y) =>
- (term_match thy) (x, y) | _ => false)
- formats (SOME t) of
- SOME format => format
- | NONE => let val format = the (AList.lookup (op =) formats NONE) in
- case t of
- Const x => intersect_formats format
- (const_format thy def_table x)
- | _ => format
- end
-
-(* int list -> int list -> typ -> typ *)
-fun format_type default_format format T =
- let
- val T = uniterize_unarize_unbox_etc_type T
- val format = format |> filter (curry (op <) 0)
- in
- if forall (curry (op =) 1) format then
- T
- else
- let
- val (binder_Ts, body_T) = strip_type T
- val batched =
- binder_Ts
- |> map (format_type default_format default_format)
- |> rev |> chunk_list_unevenly (rev format)
- |> map (HOLogic.mk_tupleT o rev)
- in List.foldl (op -->) body_T batched end
- end
-(* theory -> const_table -> (term option * int list) list -> term -> typ *)
-fun format_term_type thy def_table formats t =
- format_type (the (AList.lookup (op =) formats NONE))
- (lookup_format thy def_table formats t) (fastype_of t)
-
-(* int list -> int -> int list -> int list *)
-fun repair_special_format js m format =
- m - 1 downto 0 |> chunk_list_unevenly (rev format)
- |> map (rev o filter_out (member (op =) js))
- |> filter_out null |> map length |> rev
-
-(* hol_context -> string * string -> (term option * int list) list
- -> styp -> term * typ *)
-fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...}
- : hol_context) (base_name, step_name) formats =
- let
- val default_format = the (AList.lookup (op =) formats NONE)
- (* styp -> term * typ *)
- fun do_const (x as (s, T)) =
- (if String.isPrefix special_prefix s then
- let
- (* term -> term *)
- val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t')
- val (x' as (_, T'), js, ts) =
- AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T)
- |> the_single
- val max_j = List.last js
- val Ts = List.take (binder_types T', max_j + 1)
- val missing_js = filter_out (member (op =) js) (0 upto max_j)
- val missing_Ts = filter_indices missing_js Ts
- (* int -> indexname *)
- fun nth_missing_var n =
- ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
- val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
- val vars = special_bounds ts @ missing_vars
- val ts' =
- map (fn j =>
- case AList.lookup (op =) (js ~~ ts) j of
- SOME t => do_term t
- | NONE =>
- Var (nth missing_vars
- (find_index (curry (op =) j) missing_js)))
- (0 upto max_j)
- val t = do_const x' |> fst
- val format =
- case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
- | _ => false) formats (SOME t) of
- SOME format =>
- repair_special_format js (num_binder_types T') format
- | NONE =>
- const_format thy def_table x'
- |> repair_special_format js (num_binder_types T')
- |> intersect_formats default_format
- in
- (list_comb (t, ts') |> fold_rev abs_var vars,
- format_type default_format format T)
- end
- else if String.isPrefix uncurry_prefix s then
- let
- val (ss, s') = unprefix uncurry_prefix s
- |> strip_first_name_sep |>> space_explode "@"
- in
- if String.isPrefix step_prefix s' then
- do_const (s', T)
- else
- let
- val k = the (Int.fromString (hd ss))
- val j = the (Int.fromString (List.last ss))
- val (before_Ts, (tuple_T, rest_T)) =
- strip_n_binders j T ||> (strip_n_binders 1 #>> hd)
- val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T
- in do_const (s', T') end
- end
- else if String.isPrefix unrolled_prefix s then
- let val t = Const (original_name s, range_type T) in
- (lambda (Free (iter_var_prefix, nat_T)) t,
- format_type default_format
- (lookup_format thy def_table formats t) T)
- end
- else if String.isPrefix base_prefix s then
- (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T),
- format_type default_format default_format T)
- else if String.isPrefix step_prefix s then
- (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
- format_type default_format default_format T)
- else if String.isPrefix quot_normal_prefix s then
- let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
- (t, format_term_type thy def_table formats t)
- end
- else if String.isPrefix skolem_prefix s then
- let
- val ss = the (AList.lookup (op =) (!skolems) s)
- val (Ts, Ts') = chop (length ss) (binder_types T)
- val frees = map Free (ss ~~ Ts)
- val s' = original_name s
- in
- (fold lambda frees (Const (s', Ts' ---> T)),
- format_type default_format
- (lookup_format thy def_table formats (Const x)) T)
- end
- else if String.isPrefix eval_prefix s then
- let
- val t = nth evals (the (Int.fromString (unprefix eval_prefix s)))
- in (t, format_term_type thy def_table formats t) end
- else if s = @{const_name undefined_fast_The} then
- (Const (nitpick_prefix ^ "The fallback", T),
- format_type default_format
- (lookup_format thy def_table formats
- (Const (@{const_name The}, (T --> bool_T) --> T))) T)
- else if s = @{const_name undefined_fast_Eps} then
- (Const (nitpick_prefix ^ "Eps fallback", T),
- format_type default_format
- (lookup_format thy def_table formats
- (Const (@{const_name Eps}, (T --> bool_T) --> T))) T)
- else
- let val t = Const (original_name s, T) in
- (t, format_term_type thy def_table formats t)
- end)
- |>> map_types uniterize_unarize_unbox_etc_type
- |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
- in do_const end
-
-(* styp -> string *)
-fun assign_operator_for_const (s, T) =
- if String.isPrefix ubfp_prefix s then
- if is_fun_type T then "\<subseteq>" else "\<le>"
- else if String.isPrefix lbfp_prefix s then
- if is_fun_type T then "\<supseteq>" else "\<ge>"
- else if original_name s <> s then
- assign_operator_for_const (after_name_sep s, T)
- else
- "="
-
end;
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Fri Mar 12 16:02:42 2010 +0100
@@ -1388,7 +1388,7 @@
(Func (R1, Formula Neut)) r))
(to_opt R1 u1)
| _ => raise NUT ("Nitpick_Kodkod.to_r (SingletonSet)", [u]))
- | Op1 (Tha, _, R, u1) =>
+ | Op1 (SafeThe, _, R, u1) =>
if is_opt_rep R then
kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom
else
@@ -1653,6 +1653,7 @@
else
kk_comprehension [KK.DeclOne ((1, ~1), to_r discr_u)]
(kk_rel_eq (kk_join (KK.Var (1, ~1)) sel_r) arg_r)
+ |> is_opt_rep (rep_of arg_u) ? to_guard [arg_u] R1
end) sel_us arg_us
in fold1 kk_intersect set_rs end
| BoundRel (x, _, _, _) => KK.Var x
@@ -1723,10 +1724,8 @@
map2 (kk_not oo kk_n_ary_function kk)
(map (unopt_rep o rep_of) func_guard_us) func_guard_rs
in
- if null guard_fs then
- r
- else
- kk_rel_if (fold1 kk_or guard_fs) (empty_rel_for_rep R) r
+ if null guard_fs then r
+ else kk_rel_if (fold1 kk_or guard_fs) (empty_rel_for_rep R) r
end
(* rep -> rep -> KK.rel_expr -> int -> KK.rel_expr *)
and to_project new_R old_R r j0 =
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Fri Mar 12 16:02:42 2010 +0100
@@ -16,11 +16,20 @@
show_skolems: bool,
show_datatypes: bool,
show_consts: bool}
+ type term_postprocessor =
+ Proof.context -> string -> (typ -> term list) -> typ -> term -> term
structure NameTable : TABLE
+ val irrelevant : string
+ val unknown : string
+ val unrep : string
+ val register_term_postprocessor :
+ typ -> term_postprocessor -> theory -> theory
+ val unregister_term_postprocessor : typ -> theory -> theory
val tuple_list_for_name :
nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list
+ val dest_plain_fun : term -> bool * (term list * term list)
val reconstruct_hol_model :
params -> scope -> (term option * int list) list -> styp list -> nut list
-> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list
@@ -47,12 +56,23 @@
show_datatypes: bool,
show_consts: bool}
+type term_postprocessor =
+ Proof.context -> string -> (typ -> term list) -> typ -> term -> term
+
+structure Data = Theory_Data(
+ type T = (typ * term_postprocessor) list
+ val empty = []
+ val extend = I
+ fun merge (ps1, ps2) = AList.merge (op =) (K true) (ps1, ps2))
+
+val irrelevant = "_"
val unknown = "?"
val unrep = "\<dots>"
val maybe_mixfix = "_\<^sup>?"
val base_mixfix = "_\<^bsub>base\<^esub>"
val step_mixfix = "_\<^bsub>step\<^esub>"
val abs_mixfix = "\<guillemotleft>_\<guillemotright>"
+val arg_var_prefix = "x"
val cyclic_co_val_name = "\<omega>"
val cyclic_const_prefix = "\<xi>"
val cyclic_type_name = nitpick_prefix ^ cyclic_const_prefix
@@ -61,6 +81,31 @@
type atom_pool = ((string * int) * int list) list
+(* Proof.context -> ((string * string) * (string * string)) * Proof.context *)
+fun add_wacky_syntax ctxt =
+ let
+ (* term -> string *)
+ val name_of = fst o dest_Const
+ val thy = ProofContext.theory_of ctxt |> Context.reject_draft
+ val (maybe_t, thy) =
+ Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
+ Mixfix (maybe_mixfix, [1000], 1000)) thy
+ val (abs_t, thy) =
+ Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
+ Mixfix (abs_mixfix, [40], 40)) thy
+ val (base_t, thy) =
+ Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
+ Mixfix (base_mixfix, [1000], 1000)) thy
+ val (step_t, thy) =
+ Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
+ Mixfix (step_mixfix, [1000], 1000)) thy
+ in
+ (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
+ ProofContext.transfer_syntax thy ctxt)
+ end
+
+(** Term reconstruction **)
+
(* atom_pool Unsynchronized.ref -> string -> int -> int -> string *)
fun nth_atom_suffix pool s j k =
(case AList.lookup (op =) (!pool) (s, k) of
@@ -105,6 +150,11 @@
| ord => ord)
| _ => Term_Ord.fast_term_ord tp
+(* typ -> term_postprocessor -> theory -> theory *)
+fun register_term_postprocessor T p = Data.map (cons (T, p))
+(* typ -> theory -> theory *)
+fun unregister_term_postprocessor T = Data.map (AList.delete (op =) T)
+
(* nut NameTable.table -> KK.raw_bound list -> nut -> int list list *)
fun tuple_list_for_name rel_table bounds name =
the (AList.lookup (op =) bounds (the_rel rel_table name)) handle NUT _ => [[]]
@@ -165,9 +215,9 @@
(* typ -> typ -> (term * term) list -> term *)
fun aux T1 T2 [] =
Const (if maybe_opt then opt_flag else non_opt_flag, T1 --> T2)
- | aux T1 T2 ((t1, t2) :: ps) =
+ | aux T1 T2 ((t1, t2) :: tps) =
Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
- $ aux T1 T2 ps $ t1 $ t2
+ $ aux T1 T2 tps $ t1 $ t2
in aux T1 T2 o rev end
(* term -> bool *)
fun is_plain_fun (Const (s, _)) = (s = opt_flag orelse s = non_opt_flag)
@@ -178,9 +228,12 @@
val dest_plain_fun =
let
(* term -> bool * (term list * term list) *)
- fun aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
+ fun aux (Abs (_, _, Const (s, _))) = (s <> irrelevant, ([], []))
+ | aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
| aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
- let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end
+ let val (maybe_opt, (ts1, ts2)) = aux t0 in
+ (maybe_opt, (t1 :: ts1, t2 :: ts2))
+ end
| aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
in apsnd (pairself rev) o aux end
@@ -210,22 +263,22 @@
(* typ -> typ -> typ -> typ -> term -> term *)
fun do_curry T1 T1a T1b T2 t =
let
- val (maybe_opt, ps) = dest_plain_fun t
- val ps =
- ps |>> map (break_in_two T1 T1a T1b)
- |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
- |> AList.coalesce (op =)
- |> map (apsnd (make_plain_fun maybe_opt T1b T2))
- in make_plain_fun maybe_opt T1a (T1b --> T2) ps end
+ val (maybe_opt, tsp) = dest_plain_fun t
+ val tps =
+ tsp |>> map (break_in_two T1 T1a T1b)
+ |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
+ |> AList.coalesce (op =)
+ |> map (apsnd (make_plain_fun maybe_opt T1b T2))
+ in make_plain_fun maybe_opt T1a (T1b --> T2) tps end
(* typ -> typ -> term -> term *)
and do_uncurry T1 T2 t =
let
val (maybe_opt, tsp) = dest_plain_fun t
- val ps =
+ val tps =
tsp |> op ~~
|> maps (fn (t1, t2) =>
multi_pair_up T1 t1 (snd (dest_plain_fun t2)))
- in make_plain_fun maybe_opt T1 T2 ps end
+ in make_plain_fun maybe_opt T1 T2 tps end
(* typ -> typ -> typ -> typ -> term -> term *)
and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2')
| do_arrow T1' T2' T1 T2
@@ -271,12 +324,40 @@
| mk_tuple _ (t :: _) = t
| mk_tuple T [] = raise TYPE ("Nitpick_Model.mk_tuple", [T], [])
+(* theory -> typ * typ -> bool *)
+fun varified_type_match thy (candid_T, pat_T) =
+ strict_type_match thy (candid_T, Logic.varifyT pat_T)
+
+(* atom_pool -> (string * string) * (string * string) -> scope -> nut list
+ -> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list -> typ
+ -> term list *)
+fun all_values_of_type pool wacky_names (scope as {card_assigns, ...} : scope)
+ sel_names rel_table bounds card T =
+ let
+ val card = if card = 0 then card_of_type card_assigns T else card
+ (* nat -> term *)
+ fun nth_value_of_type n =
+ let
+ (* bool -> term *)
+ fun term unfold =
+ reconstruct_term unfold pool wacky_names scope sel_names rel_table
+ bounds T T (Atom (card, 0)) [[n]]
+ in
+ case term false of
+ t as Const (s, _) =>
+ if String.isPrefix cyclic_const_prefix s then
+ HOLogic.mk_eq (t, term true)
+ else
+ t
+ | t => t
+ end
+ in index_seq 0 card |> map nth_value_of_type |> sort nice_term_ord end
(* bool -> atom_pool -> (string * string) * (string * string) -> scope
-> nut list -> nut list -> nut list -> nut NameTable.table
-> KK.raw_bound list -> typ -> typ -> rep -> int list list -> term *)
-fun reconstruct_term unfold pool ((maybe_name, abs_name), _)
- ({hol_ctxt as {thy, stds, ...}, binarize, card_assigns, bits, datatypes,
- ofs, ...} : scope) sel_names rel_table bounds =
+and reconstruct_term unfold pool (wacky_names as ((maybe_name, abs_name), _))
+ (scope as {hol_ctxt as {ctxt, thy, stds, ...}, binarize, card_assigns,
+ bits, datatypes, ofs, ...}) sel_names rel_table bounds =
let
val for_auto = (maybe_name = "")
(* int list list -> int *)
@@ -288,6 +369,25 @@
fold (fn j => Integer.add (reasonable_power 2 j |> j = bits ? op ~))
js 0
end
+ (* typ -> term list *)
+ val all_values =
+ all_values_of_type pool wacky_names scope sel_names rel_table bounds 0
+ (* typ -> term -> term *)
+ fun postprocess_term (Type (@{type_name fun}, _)) = I
+ | postprocess_term T =
+ if null (Data.get thy) then
+ I
+ else case AList.lookup (varified_type_match thy) (Data.get thy) T of
+ SOME postproc => postproc ctxt maybe_name all_values T
+ | NONE => I
+ (* typ list -> term -> term *)
+ fun postprocess_subterms Ts (t1 $ t2) =
+ let val t = postprocess_subterms Ts t1 $ postprocess_subterms Ts t2 in
+ postprocess_term (fastype_of1 (Ts, t)) t
+ end
+ | postprocess_subterms Ts (Abs (s, T, t')) =
+ Abs (s, T, postprocess_subterms (T :: Ts) t')
+ | postprocess_subterms Ts t = postprocess_term (fastype_of1 (Ts, t)) t
(* bool -> typ -> typ -> (term * term) list -> term *)
fun make_set maybe_opt T1 T2 tps =
let
@@ -322,16 +422,16 @@
(T1 --> T2) --> T1 --> T2 --> T1 --> T2)
(* (term * term) list -> term *)
fun aux' [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
- | aux' ((t1, t2) :: ps) =
+ | aux' ((t1, t2) :: tps) =
(case t2 of
- Const (@{const_name None}, _) => aux' ps
- | _ => update_const $ aux' ps $ t1 $ t2)
- fun aux ps =
+ Const (@{const_name None}, _) => aux' tps
+ | _ => update_const $ aux' tps $ t1 $ t2)
+ fun aux tps =
if maybe_opt andalso not (is_complete_type datatypes false T1) then
- update_const $ aux' ps $ Const (unrep, T1)
+ update_const $ aux' tps $ Const (unrep, T1)
$ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2'))
else
- aux' ps
+ aux' tps
in aux end
(* typ list -> term -> term *)
fun polish_funs Ts t =
@@ -366,7 +466,11 @@
| Abs (s, T, t') => Abs (s, T, polish_funs (T :: Ts) t')
| Const (s, Type (@{type_name fun}, [T1, T2])) =>
if s = opt_flag orelse s = non_opt_flag then
- Abs ("x", T1, Const (unknown, T2))
+ Abs ("x", T1,
+ Const (if is_complete_type datatypes false T1 then
+ irrelevant
+ else
+ unknown, T2))
else
t
| t => t
@@ -476,23 +580,11 @@
|> dest_n_tuple (length uncur_arg_Ts)
val t =
if constr_s = @{const_name Abs_Frac} then
- let
- val num_T = body_type T
- (* int -> term *)
- val mk_num = HOLogic.mk_number num_T
- in
- case ts of
- [Const (@{const_name Pair}, _) $ t1 $ t2] =>
- (case snd (HOLogic.dest_number t1) of
- 0 => mk_num 0
- | n1 => case HOLogic.dest_number t2 |> snd of
- 1 => mk_num n1
- | n2 => Const (@{const_name divide},
- num_T --> num_T --> num_T)
- $ mk_num n1 $ mk_num n2)
- | _ => raise TERM ("Nitpick_Model.reconstruct_term.\
- \term_for_atom (Abs_Frac)", ts)
- end
+ case ts of
+ [Const (@{const_name Pair}, _) $ t1 $ t2] =>
+ frac_from_term_pair (body_type T) t1 t2
+ | _ => raise TERM ("Nitpick_Model.reconstruct_term.\
+ \term_for_atom (Abs_Frac)", ts)
else if not for_auto andalso
(is_abs_fun thy constr_x orelse
constr_s = @{const_name Quot}) then
@@ -575,7 +667,215 @@
raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
string_of_int (length jss))
- in polish_funs [] o unarize_unbox_etc_term oooo term_for_rep true [] end
+ in
+ postprocess_subterms [] o polish_funs [] o unarize_unbox_etc_term
+ oooo term_for_rep true []
+ end
+
+(** Constant postprocessing **)
+
+(* int -> typ -> typ list *)
+fun dest_n_tuple_type 1 T = [T]
+ | dest_n_tuple_type n (Type (_, [T1, T2])) =
+ T1 :: dest_n_tuple_type (n - 1) T2
+ | dest_n_tuple_type _ T =
+ raise TYPE ("Nitpick_Model.dest_n_tuple_type", [T], [])
+
+(* theory -> const_table -> styp -> int list *)
+fun const_format thy def_table (x as (s, T)) =
+ if String.isPrefix unrolled_prefix s then
+ const_format thy def_table (original_name s, range_type T)
+ else if String.isPrefix skolem_prefix s then
+ let
+ val k = unprefix skolem_prefix s
+ |> strip_first_name_sep |> fst |> space_explode "@"
+ |> hd |> Int.fromString |> the
+ in [k, num_binder_types T - k] end
+ else if original_name s <> s then
+ [num_binder_types T]
+ else case def_of_const thy def_table x of
+ SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then
+ let val k = length (strip_abs_vars t') in
+ [k, num_binder_types T - k]
+ end
+ else
+ [num_binder_types T]
+ | NONE => [num_binder_types T]
+(* int list -> int list -> int list *)
+fun intersect_formats _ [] = []
+ | intersect_formats [] _ = []
+ | intersect_formats ks1 ks2 =
+ let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in
+ intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else []))
+ (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @
+ [Int.min (k1, k2)]
+ end
+
+(* theory -> const_table -> (term option * int list) list -> term -> int list *)
+fun lookup_format thy def_table formats t =
+ case AList.lookup (fn (SOME x, SOME y) =>
+ (term_match thy) (x, y) | _ => false)
+ formats (SOME t) of
+ SOME format => format
+ | NONE => let val format = the (AList.lookup (op =) formats NONE) in
+ case t of
+ Const x => intersect_formats format
+ (const_format thy def_table x)
+ | _ => format
+ end
+
+(* int list -> int list -> typ -> typ *)
+fun format_type default_format format T =
+ let
+ val T = uniterize_unarize_unbox_etc_type T
+ val format = format |> filter (curry (op <) 0)
+ in
+ if forall (curry (op =) 1) format then
+ T
+ else
+ let
+ val (binder_Ts, body_T) = strip_type T
+ val batched =
+ binder_Ts
+ |> map (format_type default_format default_format)
+ |> rev |> chunk_list_unevenly (rev format)
+ |> map (HOLogic.mk_tupleT o rev)
+ in List.foldl (op -->) body_T batched end
+ end
+(* theory -> const_table -> (term option * int list) list -> term -> typ *)
+fun format_term_type thy def_table formats t =
+ format_type (the (AList.lookup (op =) formats NONE))
+ (lookup_format thy def_table formats t) (fastype_of t)
+
+(* int list -> int -> int list -> int list *)
+fun repair_special_format js m format =
+ m - 1 downto 0 |> chunk_list_unevenly (rev format)
+ |> map (rev o filter_out (member (op =) js))
+ |> filter_out null |> map length |> rev
+
+(* hol_context -> string * string -> (term option * int list) list
+ -> styp -> term * typ *)
+fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...}
+ : hol_context) (base_name, step_name) formats =
+ let
+ val default_format = the (AList.lookup (op =) formats NONE)
+ (* styp -> term * typ *)
+ fun do_const (x as (s, T)) =
+ (if String.isPrefix special_prefix s then
+ let
+ (* term -> term *)
+ val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t')
+ val (x' as (_, T'), js, ts) =
+ AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T)
+ |> the_single
+ val max_j = List.last js
+ val Ts = List.take (binder_types T', max_j + 1)
+ val missing_js = filter_out (member (op =) js) (0 upto max_j)
+ val missing_Ts = filter_indices missing_js Ts
+ (* int -> indexname *)
+ fun nth_missing_var n =
+ ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
+ val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
+ val vars = special_bounds ts @ missing_vars
+ val ts' =
+ map (fn j =>
+ case AList.lookup (op =) (js ~~ ts) j of
+ SOME t => do_term t
+ | NONE =>
+ Var (nth missing_vars
+ (find_index (curry (op =) j) missing_js)))
+ (0 upto max_j)
+ val t = do_const x' |> fst
+ val format =
+ case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
+ | _ => false) formats (SOME t) of
+ SOME format =>
+ repair_special_format js (num_binder_types T') format
+ | NONE =>
+ const_format thy def_table x'
+ |> repair_special_format js (num_binder_types T')
+ |> intersect_formats default_format
+ in
+ (list_comb (t, ts') |> fold_rev abs_var vars,
+ format_type default_format format T)
+ end
+ else if String.isPrefix uncurry_prefix s then
+ let
+ val (ss, s') = unprefix uncurry_prefix s
+ |> strip_first_name_sep |>> space_explode "@"
+ in
+ if String.isPrefix step_prefix s' then
+ do_const (s', T)
+ else
+ let
+ val k = the (Int.fromString (hd ss))
+ val j = the (Int.fromString (List.last ss))
+ val (before_Ts, (tuple_T, rest_T)) =
+ strip_n_binders j T ||> (strip_n_binders 1 #>> hd)
+ val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T
+ in do_const (s', T') end
+ end
+ else if String.isPrefix unrolled_prefix s then
+ let val t = Const (original_name s, range_type T) in
+ (lambda (Free (iter_var_prefix, nat_T)) t,
+ format_type default_format
+ (lookup_format thy def_table formats t) T)
+ end
+ else if String.isPrefix base_prefix s then
+ (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T),
+ format_type default_format default_format T)
+ else if String.isPrefix step_prefix s then
+ (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
+ format_type default_format default_format T)
+ else if String.isPrefix quot_normal_prefix s then
+ let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
+ (t, format_term_type thy def_table formats t)
+ end
+ else if String.isPrefix skolem_prefix s then
+ let
+ val ss = the (AList.lookup (op =) (!skolems) s)
+ val (Ts, Ts') = chop (length ss) (binder_types T)
+ val frees = map Free (ss ~~ Ts)
+ val s' = original_name s
+ in
+ (fold lambda frees (Const (s', Ts' ---> T)),
+ format_type default_format
+ (lookup_format thy def_table formats (Const x)) T)
+ end
+ else if String.isPrefix eval_prefix s then
+ let
+ val t = nth evals (the (Int.fromString (unprefix eval_prefix s)))
+ in (t, format_term_type thy def_table formats t) end
+ else if s = @{const_name undefined_fast_The} then
+ (Const (nitpick_prefix ^ "The fallback", T),
+ format_type default_format
+ (lookup_format thy def_table formats
+ (Const (@{const_name The}, (T --> bool_T) --> T))) T)
+ else if s = @{const_name undefined_fast_Eps} then
+ (Const (nitpick_prefix ^ "Eps fallback", T),
+ format_type default_format
+ (lookup_format thy def_table formats
+ (Const (@{const_name Eps}, (T --> bool_T) --> T))) T)
+ else
+ let val t = Const (original_name s, T) in
+ (t, format_term_type thy def_table formats t)
+ end)
+ |>> map_types uniterize_unarize_unbox_etc_type
+ |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
+ in do_const end
+
+(* styp -> string *)
+fun assign_operator_for_const (s, T) =
+ if String.isPrefix ubfp_prefix s then
+ if is_fun_type T then "\<subseteq>" else "\<le>"
+ else if String.isPrefix lbfp_prefix s then
+ if is_fun_type T then "\<supseteq>" else "\<ge>"
+ else if original_name s <> s then
+ assign_operator_for_const (strip_first_name_sep s |> snd, T)
+ else
+ "="
+
+(** Model reconstruction **)
(* atom_pool -> scope -> nut list -> nut NameTable.table -> KK.raw_bound list
-> nut -> term *)
@@ -586,29 +886,6 @@
rel_table bounds T T (rep_of name)
end
-(* Proof.context -> ((string * string) * (string * string)) * Proof.context *)
-fun add_wacky_syntax ctxt =
- let
- (* term -> string *)
- val name_of = fst o dest_Const
- val thy = ProofContext.theory_of ctxt |> Context.reject_draft
- val (maybe_t, thy) =
- Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
- Mixfix (maybe_mixfix, [1000], 1000)) thy
- val (abs_t, thy) =
- Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
- Mixfix (abs_mixfix, [40], 40)) thy
- val (base_t, thy) =
- Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
- Mixfix (base_mixfix, [1000], 1000)) thy
- val (step_t, thy) =
- Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
- Mixfix (step_mixfix, [1000], 1000)) thy
- in
- (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
- ProofContext.transfer_syntax thy ctxt)
- end
-
(* term -> term *)
fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
$ Abs (s, T, Const (@{const_name "op ="}, _)
@@ -683,14 +960,14 @@
t
| t => t
end
- (* nat -> typ -> typ list *)
- fun all_values_of_type card T =
- index_seq 0 card |> map (nth_value_of_type card T) |> sort nice_term_ord
+ (* nat -> typ -> term list *)
+ val all_values =
+ all_values_of_type pool wacky_names scope sel_names rel_table bounds
(* dtype_spec list -> dtype_spec -> bool *)
fun is_codatatype_wellformed (cos : dtype_spec list)
({typ, card, ...} : dtype_spec) =
let
- val ts = all_values_of_type card typ
+ val ts = all_values card typ
val max_depth = Integer.sum (map #card cos)
in
forall (not o bisimilar_values (map #typ cos) max_depth)
@@ -731,7 +1008,7 @@
| _ => []) @
[Pretty.str "=",
Pretty.enum "," "{" "}"
- (map (Syntax.pretty_term ctxt) (all_values_of_type card typ) @
+ (map (Syntax.pretty_term ctxt) (all_values card typ) @
(if fun_from_pair complete false then []
else [Pretty.str unrep]))]))
(* typ -> dtype_spec list *)
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Fri Mar 12 16:02:42 2010 +0100
@@ -264,20 +264,22 @@
$ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3)))
| fin_fun_body _ _ _ = NONE
-(* mdata -> typ -> typ -> mtyp * sign_atom * mtyp *)
-fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) T1 T2 =
+(* mdata -> bool -> typ -> typ -> mtyp * sign_atom * mtyp *)
+fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus
+ T1 T2 =
let
- val M1 = fresh_mtype_for_type mdata T1
- val M2 = fresh_mtype_for_type mdata T2
- val a = if is_fin_fun_supported_type (body_type T2) andalso
- exists_alpha_sub_mtype_fresh M1 then
+ val M1 = fresh_mtype_for_type mdata all_minus T1
+ val M2 = fresh_mtype_for_type mdata all_minus T2
+ val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
+ is_fin_fun_supported_type (body_type T2) then
V (Unsynchronized.inc max_fresh)
else
S Minus
in (M1, a, M2) end
-(* mdata -> typ -> mtyp *)
+(* mdata -> bool -> typ -> mtyp *)
and fresh_mtype_for_type (mdata as {hol_ctxt as {thy, ...}, binarize, alpha_T,
- datatype_mcache, constr_mcache, ...}) =
+ datatype_mcache, constr_mcache, ...})
+ all_minus =
let
(* typ -> mtyp *)
fun do_type T =
@@ -285,7 +287,7 @@
MAlpha
else case T of
Type (@{type_name fun}, [T1, T2]) =>
- MFun (fresh_mfun_for_fun_type mdata T1 T2)
+ MFun (fresh_mfun_for_fun_type mdata false T1 T2)
| Type (@{type_name "*"}, [T1, T2]) => MPair (pairself do_type (T1, T2))
| Type (z as (s, _)) =>
if could_exist_alpha_sub_mtype thy alpha_T T then
@@ -347,14 +349,14 @@
case AList.lookup (op =) (!constr_mcache) x of
SOME M => M
| NONE => if T = alpha_T then
- let val M = fresh_mtype_for_type mdata T in
+ let val M = fresh_mtype_for_type mdata false T in
(Unsynchronized.change constr_mcache (cons (x, M)); M)
end
else
- (fresh_mtype_for_type mdata (body_type T);
+ (fresh_mtype_for_type mdata false (body_type T);
AList.lookup (op =) (!constr_mcache) x |> the)
else
- fresh_mtype_for_type mdata T
+ fresh_mtype_for_type mdata false T
fun mtype_for_sel (mdata as {hol_ctxt, binarize, ...}) (x as (s, _)) =
x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
|> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
@@ -629,7 +631,7 @@
alpha_T, max_fresh, ...}) =
let
(* typ -> mtyp *)
- val mtype_for = fresh_mtype_for_type mdata
+ val mtype_for = fresh_mtype_for_type mdata false
(* mtyp -> mtyp *)
fun plus_set_mtype_for_dom M =
MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
@@ -765,15 +767,6 @@
val ba_set_M = range_type T |> mtype_for_set
in (MFun (ab_set_M, S Minus, ba_set_M), accum) end
| @{const_name trancl} => do_fragile_set_operation T accum
- | @{const_name rtrancl} =>
- (print_g "*** rtrancl"; mtype_unsolvable)
- | @{const_name finite} =>
- if is_finite_type hol_ctxt T then
- let val M1 = mtype_for (domain_type (domain_type T)) in
- (MFun (plus_set_mtype_for_dom M1, S Minus, bool_M), accum)
- end
- else
- (print_g "*** finite"; mtype_unsolvable)
| @{const_name rel_comp} =>
let
val x = Unsynchronized.inc max_fresh
@@ -796,6 +789,10 @@
MFun (plus_set_mtype_for_dom a_M, S Minus,
plus_set_mtype_for_dom b_M)), accum)
end
+ | @{const_name finite} =>
+ let val M1 = mtype_for (domain_type (domain_type T)) in
+ (MFun (plus_set_mtype_for_dom M1, S Minus, bool_M), accum)
+ end
| @{const_name Sigma} =>
let
val x = Unsynchronized.inc max_fresh
@@ -813,14 +810,15 @@
(MFun (a_set_M, S Minus,
MFun (a_to_b_set_M, S Minus, ab_set_M)), accum)
end
- | @{const_name Tha} =>
- let
- val a_M = mtype_for (domain_type (domain_type T))
- val a_set_M = plus_set_mtype_for_dom a_M
- in (MFun (a_set_M, S Minus, a_M), accum) end
| _ =>
- if s = @{const_name minus_class.minus} andalso
- is_set_type (domain_type T) then
+ if s = @{const_name safe_The} orelse
+ s = @{const_name safe_Eps} then
+ let
+ val a_set_M = mtype_for (domain_type T)
+ val a_M = dest_MFun a_set_M |> #1
+ in (MFun (a_set_M, S Minus, a_M), accum) end
+ else if s = @{const_name minus_class.minus} andalso
+ is_set_type (domain_type T) then
let
val set_T = domain_type T
val left_set_M = mtype_for set_T
@@ -842,13 +840,8 @@
(mtype_for_sel mdata x, accum)
else if is_constr thy stds x then
(mtype_for_constr mdata x, accum)
- else if is_built_in_const thy stds fast_descrs x andalso
- s <> @{const_name is_unknown} andalso
- s <> @{const_name unknown} then
- (* the "unknown" part is a hack *)
- case def_of_const thy def_table x of
- SOME t' => do_term t' accum |>> mtype_of_mterm
- | NONE => (print_g ("*** built-in " ^ s); mtype_unsolvable)
+ else if is_built_in_const thy stds fast_descrs x then
+ (fresh_mtype_for_type mdata true T, accum)
else
let val M = mtype_for T in
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
@@ -936,7 +929,7 @@
val M1 = mtype_of_mterm m1
val M2 = mtype_of_mterm m2
val accum = accum ||> add_mtypes_equal M1 M2
- val body_M = fresh_mtype_for_type mdata (nth_range_type 2 T)
+ val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T)
val m = MApp (MApp (MRaw (Const x,
MFun (M1, S Minus, MFun (M2, S Minus, body_M))), m1), m2)
in
@@ -952,7 +945,7 @@
fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
let
(* typ -> mtyp *)
- val mtype_for = fresh_mtype_for_type mdata
+ val mtype_for = fresh_mtype_for_type mdata false
(* term -> accumulator -> mterm * accumulator *)
val do_term = consider_term mdata
(* sign -> term -> accumulator -> mterm * accumulator *)
@@ -1059,7 +1052,7 @@
else
let
(* typ -> mtyp *)
- val mtype_for = fresh_mtype_for_type mdata
+ val mtype_for = fresh_mtype_for_type mdata false
(* term -> accumulator -> mterm * accumulator *)
val do_term = consider_term mdata
(* term -> string -> typ -> term -> accumulator -> mterm * accumulator *)
@@ -1205,8 +1198,26 @@
val T' = type_from_mtype T M
in
case t of
- Const (x as (s, T)) =>
- if member (op =) [@{const_name "=="}, @{const_name "op ="}] s then
+ Const (x as (s, _)) =>
+ if s = @{const_name insert} then
+ case nth_range_type 2 T' of
+ set_T' as Type (@{type_name fin_fun}, [elem_T', _]) =>
+ Abs (Name.uu, elem_T', Abs (Name.uu, set_T',
+ Const (@{const_name If},
+ bool_T --> set_T' --> set_T' --> set_T')
+ $ (Const (@{const_name is_unknown}, elem_T' --> bool_T)
+ $ Bound 1)
+ $ (Const (@{const_name unknown}, set_T'))
+ $ (coerce_term hol_ctxt Ts T' T (Const x)
+ $ Bound 1 $ Bound 0)))
+ | _ => Const (s, T')
+ else if s = @{const_name finite} then
+ case domain_type T' of
+ set_T' as Type (@{type_name fin_fun}, _) =>
+ Abs (Name.uu, set_T', @{const True})
+ | _ => Const (s, T')
+ else if s = @{const_name "=="} orelse
+ s = @{const_name "op ="} then
Const (s, T')
else if is_built_in_const thy stds fast_descrs x then
coerce_term hol_ctxt Ts T' T t
@@ -1228,16 +1239,6 @@
| _ => raise MTERM ("Nitpick_Mono.finitize_funs.term_from_mterm",
[m])
end
- | MAbs (s, T, M, a, m') =>
- let
- val T = type_from_mtype T M
- val t' = term_from_mterm (T :: Ts) m'
- val T' = fastype_of1 (T :: Ts, t')
- in
- Abs (s, T, t')
- |> should_finitize (T --> T') a
- ? construct_value thy stds (fin_fun_constr T T') o single
- end
| MApp (m1, m2) =>
let
val (t1, t2) = pairself (term_from_mterm Ts) (m1, m2)
@@ -1253,6 +1254,16 @@
| _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm",
[T1], [])
in betapply (t1', coerce_term hol_ctxt Ts T2' T2 t2) end
+ | MAbs (s, T, M, a, m') =>
+ let
+ val T = type_from_mtype T M
+ val t' = term_from_mterm (T :: Ts) m'
+ val T' = fastype_of1 (T :: Ts, t')
+ in
+ Abs (s, T, t')
+ |> should_finitize (T --> T') a
+ ? construct_value thy stds (fin_fun_constr T T') o single
+ end
in
Unsynchronized.change constr_cache (map (apsnd (map finitize_constr)));
pairself (map (term_from_mterm [])) msp
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Fri Mar 12 16:02:42 2010 +0100
@@ -39,7 +39,7 @@
Closure |
SingletonSet |
IsUnknown |
- Tha |
+ SafeThe |
First |
Second |
Cast
@@ -158,7 +158,7 @@
Closure |
SingletonSet |
IsUnknown |
- Tha |
+ SafeThe |
First |
Second |
Cast
@@ -232,7 +232,7 @@
| string_for_op1 Closure = "Closure"
| string_for_op1 SingletonSet = "SingletonSet"
| string_for_op1 IsUnknown = "IsUnknown"
- | string_for_op1 Tha = "Tha"
+ | string_for_op1 SafeThe = "SafeThe"
| string_for_op1 First = "First"
| string_for_op1 Second = "Second"
| string_for_op1 Cast = "Cast"
@@ -516,8 +516,15 @@
val t1 = list_comb (t0, ts')
val T1 = fastype_of1 (Ts, t1)
in Op2 (Apply, range_type T1, Any, sub t1, sub t2) end
- (* styp -> term list -> term *)
- fun construct (x as (_, T)) ts =
+ (* op2 -> string -> styp -> term -> nut *)
+ fun do_description_operator oper undef_s (x as (_, T)) t1 =
+ if fast_descrs then
+ Op2 (oper, range_type T, Any, sub t1,
+ sub (Const (undef_s, range_type T)))
+ else
+ do_apply (Const x) [t1]
+ (* styp -> term list -> nut *)
+ fun do_construct (x as (_, T)) ts =
case num_binder_types T - length ts of
0 => Construct (map ((fn (s', T') => ConstName (s', T', Any))
o nth_sel_for_constr x)
@@ -552,18 +559,10 @@
do_quantifier Exist s T t1
| (t0 as Const (@{const_name Ex}, _), [t1]) =>
sub' (t0 $ eta_expand Ts t1 1)
- | (t0 as Const (@{const_name The}, T), [t1]) =>
- if fast_descrs then
- Op2 (The, range_type T, Any, sub t1,
- sub (Const (@{const_name undefined_fast_The}, range_type T)))
- else
- do_apply t0 [t1]
- | (t0 as Const (@{const_name Eps}, T), [t1]) =>
- if fast_descrs then
- Op2 (Eps, range_type T, Any, sub t1,
- sub (Const (@{const_name undefined_fast_Eps}, range_type T)))
- else
- do_apply t0 [t1]
+ | (Const (x as (@{const_name The}, _)), [t1]) =>
+ do_description_operator The @{const_name undefined_fast_The} x t1
+ | (Const (x as (@{const_name Eps}, _)), [t1]) =>
+ do_description_operator Eps @{const_name undefined_fast_Eps} x t1
| (Const (@{const_name "op ="}, T), [t1, t2]) => sub_equals T t1 t2
| (Const (@{const_name "op &"}, _), [t1, t2]) =>
Op2 (And, bool_T, Any, sub' t1, sub' t2)
@@ -605,7 +604,7 @@
Op2 (Image, nth_range_type 2 T, Any, sub t1, sub t2)
| (Const (x as (s as @{const_name Suc}, T)), []) =>
if is_built_in_const thy stds false x then Cst (Suc, T, Any)
- else if is_constr thy stds x then construct x []
+ else if is_constr thy stds x then do_construct x []
else ConstName (s, T, Any)
| (Const (@{const_name finite}, T), [t1]) =>
(if is_finite_type hol_ctxt (domain_type T) then
@@ -616,7 +615,7 @@
| (Const (@{const_name nat}, T), []) => Cst (IntToNat, T, Any)
| (Const (x as (s as @{const_name zero_class.zero}, T)), []) =>
if is_built_in_const thy stds false x then Cst (Num 0, T, Any)
- else if is_constr thy stds x then construct x []
+ else if is_constr thy stds x then do_construct x []
else ConstName (s, T, Any)
| (Const (x as (s as @{const_name one_class.one}, T)), []) =>
if is_built_in_const thy stds false x then Cst (Num 1, T, Any)
@@ -670,8 +669,11 @@
| (Const (@{const_name unknown}, T), []) => Cst (Unknown, T, Any)
| (Const (@{const_name is_unknown}, _), [t1]) =>
Op1 (IsUnknown, bool_T, Any, sub t1)
- | (Const (@{const_name Tha}, Type (@{type_name fun}, [_, T2])), [t1]) =>
- Op1 (Tha, T2, Any, sub t1)
+ | (Const (@{const_name safe_The},
+ Type (@{type_name fun}, [_, T2])), [t1]) =>
+ Op1 (SafeThe, T2, Any, sub t1)
+ | (Const (x as (@{const_name safe_Eps}, _)), [t1]) =>
+ do_description_operator Eps @{const_name undefined_fast_Eps} x t1
| (Const (@{const_name Frac}, T), []) => Cst (Fracs, T, Any)
| (Const (@{const_name norm_frac}, T), []) => Cst (NormFrac, T, Any)
| (Const (@{const_name of_nat}, T as @{typ "nat => int"}), []) =>
@@ -691,7 +693,7 @@
Op2 (Union, T1, Any, sub t1, sub t2)
| (t0 as Const (x as (s, T)), ts) =>
if is_constr thy stds x then
- construct x ts
+ do_construct x ts
else if String.isPrefix numeral_prefix s then
Cst (Num (the (Int.fromString (unprefix numeral_prefix s))), T, Any)
else
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Fri Mar 12 16:02:42 2010 +0100
@@ -33,27 +33,40 @@
binary coding. *)
val binary_int_threshold = 3
+(* bool -> term -> bool *)
+val may_use_binary_ints =
+ let
+ (* bool -> term -> bool *)
+ fun aux def (Const (@{const_name "=="}, _) $ t1 $ t2) =
+ aux def t1 andalso aux false t2
+ | aux def (@{const "==>"} $ t1 $ t2) = aux false t1 andalso aux def t2
+ | aux def (Const (@{const_name "op ="}, _) $ t1 $ t2) =
+ aux def t1 andalso aux false t2
+ | aux def (@{const "op -->"} $ t1 $ t2) = aux false t1 andalso aux def t2
+ | aux def (t1 $ t2) = aux def t1 andalso aux def t2
+ | aux def (t as Const (s, _)) =
+ (not def orelse t <> @{const Suc}) andalso
+ not (member (op =) [@{const_name Abs_Frac}, @{const_name Rep_Frac},
+ @{const_name nat_gcd}, @{const_name nat_lcm},
+ @{const_name Frac}, @{const_name norm_frac}] s)
+ | aux def (Abs (_, _, t')) = aux def t'
+ | aux _ _ = true
+ in aux end
(* term -> bool *)
-fun may_use_binary_ints (t1 $ t2) =
- may_use_binary_ints t1 andalso may_use_binary_ints t2
- | may_use_binary_ints (t as Const (s, _)) =
- t <> @{const Suc} andalso
- not (member (op =) [@{const_name Abs_Frac}, @{const_name Rep_Frac},
- @{const_name nat_gcd}, @{const_name nat_lcm},
- @{const_name Frac}, @{const_name norm_frac}] s)
- | may_use_binary_ints (Abs (_, _, t')) = may_use_binary_ints t'
- | may_use_binary_ints _ = true
-fun should_use_binary_ints (t1 $ t2) =
- should_use_binary_ints t1 orelse should_use_binary_ints t2
- | should_use_binary_ints (Const (s, T)) =
- ((s = @{const_name times} orelse s = @{const_name div}) andalso
- is_integer_type (body_type T)) orelse
- (String.isPrefix numeral_prefix s andalso
- let val n = the (Int.fromString (unprefix numeral_prefix s)) in
- n < ~ binary_int_threshold orelse n > binary_int_threshold
- end)
- | should_use_binary_ints (Abs (_, _, t')) = should_use_binary_ints t'
- | should_use_binary_ints _ = false
+val should_use_binary_ints =
+ let
+ (* term -> bool *)
+ fun aux (t1 $ t2) = aux t1 orelse aux t2
+ | aux (Const (s, T)) =
+ ((s = @{const_name times} orelse s = @{const_name div}) andalso
+ is_integer_type (body_type T)) orelse
+ (String.isPrefix numeral_prefix s andalso
+ let val n = the (Int.fromString (unprefix numeral_prefix s)) in
+ n < ~ binary_int_threshold orelse n > binary_int_threshold
+ end)
+ | aux (Abs (_, _, t')) = aux t'
+ | aux _ = false
+ in aux end
(** Uncurrying **)
@@ -221,7 +234,8 @@
$ do_term new_Ts old_Ts polar t2
| Const (s as @{const_name The}, T) => do_description_operator s T
| Const (s as @{const_name Eps}, T) => do_description_operator s T
- | Const (s as @{const_name Tha}, T) => do_description_operator s T
+ | Const (s as @{const_name safe_The}, T) => do_description_operator s T
+ | Const (s as @{const_name safe_Eps}, T) => do_description_operator s T
| Const (x as (s, T)) =>
Const (s, if s = @{const_name converse} orelse
s = @{const_name trancl} then
@@ -1341,6 +1355,8 @@
let
val Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)
|> filter_out (fn Type (@{type_name fun_box}, _) => true
+ | @{typ signed_bit} => true
+ | @{typ unsigned_bit} => true
| T => is_small_finite_type hol_ctxt T orelse
triple_lookup (type_match thy) monos T
= SOME (SOME false))
@@ -1366,7 +1382,8 @@
is_standard_datatype thy stds nat_T andalso
case binary_ints of
SOME false => false
- | _ => forall may_use_binary_ints (nondef_ts @ def_ts) andalso
+ | _ => forall (may_use_binary_ints false) nondef_ts andalso
+ forall (may_use_binary_ints true) def_ts andalso
(binary_ints = SOME true orelse
exists should_use_binary_ints (nondef_ts @ def_ts))
val box = exists (not_equal (SOME false) o snd) boxes
--- a/src/HOL/Tools/TFL/tfl.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/TFL/tfl.ML Fri Mar 12 16:02:42 2010 +0100
@@ -360,7 +360,7 @@
(*For Isabelle, the lhs of a definition must be a constant.*)
-fun mk_const_def sign (c, Ty, rhs) =
+fun legacy_const_def sign (c, Ty, rhs) =
singleton (Syntax.check_terms (ProofContext.init sign))
(Sign.intern_term sign (Const("==",dummyT) $ Const(c,Ty) $ rhs));
@@ -385,7 +385,7 @@
val wfrec_R_M = map_types poly_tvars
(wfrec $ map_types poly_tvars R)
$ functional
- val def_term = mk_const_def thy (x, Ty, wfrec_R_M)
+ val def_term = legacy_const_def thy (x, Ty, wfrec_R_M)
val ([def], thy') = PureThy.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy
in (thy', def) end;
end;
@@ -540,7 +540,7 @@
val {lhs,rhs} = S.dest_eq proto_def'
val (c,args) = S.strip_comb lhs
val (name,Ty) = dest_atom c
- val defn = mk_const_def thy (name, Ty, S.list_mk_abs (args,rhs))
+ val defn = legacy_const_def thy (name, Ty, S.list_mk_abs (args,rhs))
val ([def0], theory) =
thy
|> PureThy.add_defs false
--- a/src/HOL/Tools/recdef.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/recdef.ML Fri Mar 12 16:02:42 2010 +0100
@@ -184,6 +184,7 @@
fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
let
+ val _ = legacy_feature ("\"recdef\"; prefer \"function\" instead");
val _ = requires_recdef thy;
val name = Sign.intern_const thy raw_name;
--- a/src/HOL/Tools/transfer.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/transfer.ML Fri Mar 12 16:02:42 2010 +0100
@@ -8,10 +8,11 @@
signature TRANSFER =
sig
datatype selection = Direction of term * term | Hints of string list | Prop
- val transfer: Context.generic -> selection -> string list -> thm -> thm
+ val transfer: Context.generic -> selection -> string list -> thm -> thm list
type entry
- val add: entry * entry -> thm -> Context.generic -> Context.generic
- val del: thm -> Context.generic -> Context.generic
+ val add: thm -> bool -> entry -> Context.generic -> Context.generic
+ val del: thm -> entry -> Context.generic -> Context.generic
+ val drop: thm -> Context.generic -> Context.generic
val setup: theory -> theory
end;
@@ -29,14 +30,15 @@
("Transfer: expected theorem of the form " ^ quote (Display.string_of_thm ctxt @{thm transfer_morphismI}));
in direction_of key end;
-type entry = { inj : thm list, emb : thm list, ret : thm list, cong : thm list,
- guess : bool, hints : string list };
+type entry = { inj : thm list, embed : thm list, return : thm list, cong : thm list,
+ hints : string list };
-fun merge_entry ({ inj = inj1, emb = emb1, ret = ret1, cong = cong1, guess = guess1, hints = hints1 } : entry,
- { inj = inj2, emb = emb2, ret = ret2, cong = cong2, guess = guess2, hints = hints2 } : entry) =
- { inj = merge Thm.eq_thm (inj1, inj2), emb = merge Thm.eq_thm (emb1, emb2),
- ret = merge Thm.eq_thm (ret1, ret2), cong = merge Thm.eq_thm (cong1, cong2),
- guess = guess1 andalso guess2, hints = merge (op =) (hints1, hints2) };
+val empty_entry = { inj = [], embed = [], return = [], cong = [], hints = [] };
+fun merge_entry ({ inj = inj1, embed = embed1, return = return1, cong = cong1, hints = hints1 } : entry,
+ { inj = inj2, embed = embed2, return = return2, cong = cong2, hints = hints2 } : entry) =
+ { inj = merge Thm.eq_thm (inj1, inj2), embed = merge Thm.eq_thm (embed1, embed2),
+ return = merge Thm.eq_thm (return1, return2), cong = merge Thm.eq_thm (cong1, cong2),
+ hints = merge (op =) (hints1, hints2) };
structure Data = Generic_Data
(
@@ -49,6 +51,9 @@
(* data lookup *)
+fun transfer_rules_of ({ inj, embed, return, cong, ... } : entry) =
+ (inj, embed, return, cong);
+
fun get_by_direction context (a, D) =
let
val ctxt = Context.proof_of context;
@@ -58,9 +63,9 @@
fun eq_direction ((a, D), thm') =
let
val (a', D') = direction_of thm';
- in a0 aconvc a' andalso D0 aconvc D' end;
- in case AList.lookup eq_direction (Data.get context) (a, D) of
- SOME e => ((a0, D0), e)
+ in a aconvc a' andalso D aconvc D' end;
+ in case AList.lookup eq_direction (Data.get context) (a0, D0) of
+ SOME e => ((a0, D0), transfer_rules_of e)
| NONE => error ("Transfer: no such instance: ("
^ Syntax.string_of_term ctxt a ^ ", " ^ Syntax.string_of_term ctxt D ^ ")")
end;
@@ -68,7 +73,7 @@
fun get_by_hints context hints =
let
val insts = map_filter (fn (k, e) => if exists (member (op =) (#hints e)) hints
- then SOME (direction_of k, e) else NONE) (Data.get context);
+ then SOME (direction_of k, transfer_rules_of e) else NONE) (Data.get context);
val _ = if null insts then error ("Transfer: no such labels: " ^ commas (map quote hints)) else ();
in insts end;
@@ -84,9 +89,9 @@
val _ = if null tys then error "Transfer: unable to guess instance" else ();
val tyss = splits (curry Type.could_unify) tys;
val get_ty = typ_of o ctyp_of_term o fst o direction_of;
- val insts = map_filter (fn tys => get_first (fn (k, ss) =>
+ val insts = map_filter (fn tys => get_first (fn (k, e) =>
if Type.could_unify (hd tys, range_type (get_ty k))
- then SOME (direction_of k, ss)
+ then SOME (direction_of k, transfer_rules_of e)
else NONE) (Data.get context)) tyss;
val _ = if null insts then
error "Transfer: no instances, provide direction or hints explicitly" else ();
@@ -95,36 +100,42 @@
(* applying transfer data *)
-fun transfer_thm inj_only a0 D0 {inj = inj, emb = emb, ret = ret, cong = cg, guess = _, hints = _}
- leave ctxt0 th =
+fun transfer_thm ((raw_a, raw_D), (inj, embed, return, cong)) leave ctxt1 thm =
let
- val ([a, D], ctxt) = apfst (map Drule.dest_term o snd)
- (Variable.import true (map Drule.mk_term [a0, D0]) ctxt0);
- val (aT, bT) =
- let val T = typ_of (ctyp_of_term a)
- in (Term.range_type T, Term.domain_type T) end;
- val ctxt' = (Variable.declare_term (term_of a) o Variable.declare_term (term_of D)
- o Variable.declare_thm th) ctxt;
- val ns = filter (fn i => Type.could_unify (snd i, aT) andalso
- not (member (op =) leave (fst (fst i)))) (Term.add_vars (prop_of th) []);
- val (ins, ctxt'') = Variable.variant_fixes (map (fst o fst) ns) ctxt';
- val certify = Thm.cterm_of (ProofContext.theory_of ctxt'');
- val cns = map (certify o Var) ns;
- val cfis = map (certify o (fn n => Free (n, bT))) ins;
- val cis = map (Thm.capply a) cfis;
- val (hs, ctxt''') = Assumption.add_assumes (map (fn ct =>
- Thm.capply @{cterm "Trueprop"} (Thm.capply D ct)) cfis) ctxt'';
- val th1 = Drule.cterm_instantiate (cns ~~ cis) th;
- val th2 = fold Thm.elim_implies hs (fold_rev implies_intr (map cprop_of hs) th1);
- val simpset = (Simplifier.context ctxt''' HOL_ss)
- addsimps inj addsimps (if inj_only then [] else emb @ ret) addcongs cg;
- val th3 = Simplifier.asm_full_simplify simpset
- (fold_rev implies_intr (map cprop_of hs) th2);
- in hd (Variable.export ctxt''' ctxt0 [th3]) end;
+ (* identify morphism function *)
+ val ([a, D], ctxt2) = ctxt1
+ |> Variable.import true (map Drule.mk_term [raw_a, raw_D])
+ |>> map Drule.dest_term o snd;
+ val transform = Thm.capply @{cterm "Trueprop"} o Thm.capply D;
+ val T = Thm.typ_of (Thm.ctyp_of_term a);
+ val (aT, bT) = (Term.range_type T, Term.domain_type T);
+
+ (* determine variables to transfer *)
+ val ctxt3 = ctxt2
+ |> Variable.declare_thm thm
+ |> Variable.declare_term (term_of a)
+ |> Variable.declare_term (term_of D);
+ val certify = Thm.cterm_of (ProofContext.theory_of ctxt3);
+ val vars = filter (fn ((v, _), T) => Type.could_unify (T, aT) andalso
+ not (member (op =) leave v)) (Term.add_vars (Thm.prop_of thm) []);
+ val c_vars = map (certify o Var) vars;
+ val (vs', ctxt4) = Variable.variant_fixes (map (fst o fst) vars) ctxt3;
+ val c_vars' = map (certify o (fn v => Free (v, bT))) vs';
+ val c_exprs' = map (Thm.capply a) c_vars';
-fun transfer_thm_multiple inj_only insts leave ctxt thm =
- Conjunction.intr_balanced (map
- (fn ((a, D), e) => transfer_thm false a D e leave ctxt thm) insts);
+ (* transfer *)
+ val (hyps, ctxt5) = ctxt4
+ |> Assumption.add_assumes (map transform c_vars');
+ val simpset = Simplifier.context ctxt5 HOL_ss
+ addsimps (inj @ embed @ return) addcongs cong;
+ val thm' = thm
+ |> Drule.cterm_instantiate (c_vars ~~ c_exprs')
+ |> fold_rev Thm.implies_intr (map cprop_of hyps)
+ |> Simplifier.asm_full_simplify simpset
+ in singleton (Variable.export ctxt5 ctxt1) thm' end;
+
+fun transfer_thm_multiple insts leave ctxt thm =
+ map (fn inst => transfer_thm inst leave ctxt thm) insts;
datatype selection = Direction of term * term | Hints of string list | Prop;
@@ -133,49 +144,46 @@
| insts_for context thm Prop = get_by_prop context (Thm.prop_of thm);
fun transfer context selection leave thm =
- transfer_thm_multiple false (insts_for context thm selection) leave (Context.proof_of context) thm;
+ transfer_thm_multiple (insts_for context thm selection) leave (Context.proof_of context) thm;
(* maintaining transfer data *)
-fun merge_update eq m (k, v) [] = [(k, v)]
- | merge_update eq m (k, v) ((k', v') :: al) =
- if eq (k, k') then (k', m (v, v')) :: al else (k', v') :: merge_update eq m (k, v) al;
-
-(*? fun merge_update eq m (k, v) = AList.map_entry eq k (fn v' => m (v, v'));*)
-
-fun merge_entries {inj = inj0, emb = emb0, ret = ret0, cong = cg0, guess = g0, hints = hints0}
- ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
- {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2} : entry) =
+fun extend_entry ctxt (a, D) guess
+ { inj = inj1, embed = embed1, return = return1, cong = cong1, hints = hints1 }
+ { inj = inj2, embed = embed2, return = return2, cong = cong2, hints = hints2 } =
let
- fun h xs0 xs ys = subtract Thm.eq_thm xs0 (merge Thm.eq_thm (xs, ys))
+ fun add_del eq del add = union eq add #> subtract eq del;
+ val guessed = if guess
+ then map (fn thm => transfer_thm
+ ((a, D), (if null inj1 then inj2 else inj1, [], [], cong1)) [] ctxt thm RS sym) embed1
+ else [];
in
- {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
- ret = h ret0 ret1 ret2, cong = h cg0 cg1 cg2, guess = g1 andalso g2,
- hints = subtract (op =) hints0 (union (op =) hints1 hints2) }
+ { inj = union Thm.eq_thm inj1 inj2, embed = union Thm.eq_thm embed1 embed2,
+ return = union Thm.eq_thm guessed (union Thm.eq_thm return1 return2),
+ cong = union Thm.eq_thm cong1 cong2, hints = union (op =) hints1 hints2 }
end;
-fun add (e0 as {inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa},
- ed as {inj = injd, emb = embd, ret = retd, cong = cgd, guess = _, hints = hintsd}) key context =
- context
- |> Data.map (fn al =>
- let
- val ctxt = Context.proof_of context;
- val (a0, D0) = check_morphism_key ctxt key;
- val entry = if g then
- let
- val inj' = if null inja then #inj
- (case AList.lookup Thm.eq_thm al key of SOME e => e
- | NONE => error "Transfer: cannot generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual")
- else inja
- val ret' = merge Thm.eq_thm (reta, map
- (fn th => transfer_thm true a0 D0 {inj = inj', emb = [], ret = [], cong = cga, guess = g,
- hints = hintsa} [] ctxt th RS sym) emba);
- in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
- else e0;
- in merge_update Thm.eq_thm (merge_entries ed) (key, entry) al end);
+fun diminish_entry
+ { inj = inj0, embed = embed0, return = return0, cong = cong0, hints = hints0 }
+ { inj = inj2, embed = embed2, return = return2, cong = cong2, hints = hints2 } =
+ { inj = subtract Thm.eq_thm inj0 inj2, embed = subtract Thm.eq_thm embed0 embed2,
+ return = subtract Thm.eq_thm return0 return2, cong = subtract Thm.eq_thm cong0 cong2,
+ hints = subtract (op =) hints0 hints2 };
-fun del key = Data.map (remove (eq_fst Thm.eq_thm) (key, []));
+fun add key guess entry context =
+ let
+ val ctxt = Context.proof_of context;
+ val a_D = check_morphism_key ctxt key;
+ in
+ context
+ |> Data.map (AList.map_default Thm.eq_thm
+ (key, empty_entry) (extend_entry ctxt a_D guess entry))
+ end;
+
+fun del key entry = Data.map (AList.map_entry Thm.eq_thm key (diminish_entry entry));
+
+fun drop key = Data.map (AList.delete Thm.eq_thm key);
(* syntax *)
@@ -188,22 +196,25 @@
fun keyword k = Scan.lift (Args.$$$ k) >> K ();
fun keyword_colon k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
-val congN = "cong";
+val addN = "add";
+val delN = "del";
+val keyN = "key";
+val modeN = "mode";
+val automaticN = "automatic";
+val manualN = "manual";
val injN = "inj";
val embedN = "embed";
val returnN = "return";
-val addN = "add";
-val delN = "del";
-val modeN = "mode";
-val automaticN = "automatic";
-val manualN = "manual";
+val congN = "cong";
+val labelsN = "labels";
+
+val leavingN = "leaving";
val directionN = "direction";
-val labelsN = "labels";
-val leavingN = "leaving";
-val any_keyword = keyword_colon congN || keyword_colon injN || keyword_colon embedN
- || keyword_colon returnN || keyword_colon directionN || keyword_colon modeN
- || keyword_colon delN || keyword_colon labelsN || keyword_colon leavingN;
+val any_keyword = keyword_colon addN || keyword_colon delN || keyword_colon keyN
+ || keyword_colon modeN || keyword_colon injN || keyword_colon embedN || keyword_colon returnN
+ || keyword_colon congN || keyword_colon labelsN
+ || keyword_colon leavingN || keyword_colon directionN;
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
val names = Scan.repeat (Scan.unless any_keyword (Scan.lift Args.name))
@@ -216,23 +227,30 @@
val cong = (keyword_colon congN |-- thms) -- these (keyword_colon delN |-- thms);
val labels = (keyword_colon labelsN |-- names) -- these (keyword_colon delN |-- names);
-val entry_pair = Scan.optional mode true -- these_pair inj -- these_pair embed
+val entry_pair = these_pair inj -- these_pair embed
-- these_pair return -- these_pair cong -- these_pair labels
- >> (fn (((((g, (inja, injd)), (emba, embd)), (reta, retd)), (cga, cgd)), (hintsa, hintsd)) =>
- ({inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa},
- {inj = injd, emb = embd, ret = retd, cong = cgd, guess = g, hints = hintsd}));
+ >> (fn (((((inja, injd), (embeda, embedd)), (returna, returnd)), (conga, congd)),
+ (hintsa, hintsd)) =>
+ ({ inj = inja, embed = embeda, return = returna, cong = conga, hints = hintsa },
+ { inj = injd, embed = embedd, return = returnd, cong = congd, hints = hintsd }));
val selection = (keyword_colon directionN |-- (Args.term -- Args.term) >> Direction)
|| these names >> (fn hints => if null hints then Prop else Hints hints);
in
-val transfer_attribute = Scan.lift (Args.$$$ delN >> K (Thm.declaration_attribute del))
- || Scan.unless any_keyword (keyword addN) |-- entry_pair
- >> (fn entry_pair => Thm.declaration_attribute (add entry_pair))
+val transfer_attribute = keyword delN >> K (Thm.declaration_attribute drop)
+ || keyword addN |-- Scan.optional mode true -- entry_pair
+ >> (fn (guess, (entry_add, entry_del)) => Thm.declaration_attribute
+ (fn thm => add thm guess entry_add #> del thm entry_del))
+ || keyword_colon keyN |-- Attrib.thm
+ >> (fn key => Thm.declaration_attribute
+ (fn thm => add key false
+ { inj = [], embed = [], return = [thm], cong = [], hints = [] }));
val transferred_attribute = selection -- these (keyword_colon leavingN |-- names)
- >> (fn (selection, leave) => Thm.rule_attribute (fn context => transfer context selection leave));
+ >> (fn (selection, leave) => Thm.rule_attribute (fn context =>
+ Conjunction.intr_balanced o transfer context selection leave));
end;
--- a/src/HOL/Tools/typedef.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Tools/typedef.ML Fri Mar 12 16:02:42 2010 +0100
@@ -130,7 +130,7 @@
in Drule.export_without_context (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
fun typedef_result inhabited =
- Typedecl.typedecl (tname, vs, mx)
+ Typedecl.typedecl_global (tname, vs, mx)
#> snd
#> Sign.add_consts_i
[(Rep_name, newT --> oldT, NoSyn),
--- a/src/HOL/Wellfounded.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/Wellfounded.thy Fri Mar 12 16:02:42 2010 +0100
@@ -8,7 +8,7 @@
header {*Well-founded Recursion*}
theory Wellfounded
-imports Finite_Set Transitive_Closure
+imports Transitive_Closure
uses ("Tools/Function/size.ML")
begin
--- a/src/HOL/ex/Summation.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/ex/Summation.thy Fri Mar 12 16:02:42 2010 +0100
@@ -10,7 +10,7 @@
lemma add_setsum_orient:
"setsum f {k..<j} + setsum f {l..<k} = setsum f {l..<k} + setsum f {k..<j}"
- by (fact plus.commute)
+ by (fact add.commute)
lemma add_setsum_int:
fixes j k l :: int
--- a/src/HOL/ex/Transfer_Ex.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOL/ex/Transfer_Ex.thy Fri Mar 12 16:02:42 2010 +0100
@@ -2,41 +2,46 @@
header {* Various examples for transfer procedure *}
theory Transfer_Ex
-imports Complex_Main
+imports Main
begin
-(* nat to int *)
-
lemma ex1: "(x::nat) + y = y + x"
by auto
-thm ex1 [transferred]
+lemma "(0\<Colon>int) \<le> (y\<Colon>int) \<Longrightarrow> (0\<Colon>int) \<le> (x\<Colon>int) \<Longrightarrow> x + y = y + x"
+ by (fact ex1 [transferred])
lemma ex2: "(a::nat) div b * b + a mod b = a"
by (rule mod_div_equality)
-thm ex2 [transferred]
+lemma "(0\<Colon>int) \<le> (b\<Colon>int) \<Longrightarrow> (0\<Colon>int) \<le> (a\<Colon>int) \<Longrightarrow> a div b * b + a mod b = a"
+ by (fact ex2 [transferred])
lemma ex3: "ALL (x::nat). ALL y. EX z. z >= x + y"
by auto
-thm ex3 [transferred natint]
+lemma "\<forall>x\<ge>0\<Colon>int. \<forall>y\<ge>0\<Colon>int. \<exists>xa\<ge>0\<Colon>int. x + y \<le> xa"
+ by (fact ex3 [transferred nat_int])
lemma ex4: "(x::nat) >= y \<Longrightarrow> (x - y) + y = x"
by auto
-thm ex4 [transferred]
+lemma "(0\<Colon>int) \<le> (x\<Colon>int) \<Longrightarrow> (0\<Colon>int) \<le> (y\<Colon>int) \<Longrightarrow> y \<le> x \<Longrightarrow> tsub x y + y = x"
+ by (fact ex4 [transferred])
-lemma ex5: "(2::nat) * (SUM i <= n. i) = n * (n + 1)"
+lemma ex5: "(2::nat) * \<Sum>{..n} = n * (n + 1)"
by (induct n rule: nat_induct, auto)
-thm ex5 [transferred]
+lemma "(0\<Colon>int) \<le> (n\<Colon>int) \<Longrightarrow> (2\<Colon>int) * \<Sum>{0\<Colon>int..n} = n * (n + (1\<Colon>int))"
+ by (fact ex5 [transferred])
+
+lemma "(0\<Colon>nat) \<le> (n\<Colon>nat) \<Longrightarrow> (2\<Colon>nat) * \<Sum>{0\<Colon>nat..n} = n * (n + (1\<Colon>nat))"
+ by (fact ex5 [transferred, transferred])
theorem ex6: "0 <= (n::int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
by (rule ex5 [transferred])
-thm ex6 [transferred]
-
-thm ex5 [transferred, transferred]
+lemma "(0\<Colon>nat) \<le> (n\<Colon>nat) \<Longrightarrow> (2\<Colon>nat) * \<Sum>{0\<Colon>nat..n} = n * (n + (1\<Colon>nat))"
+ by (fact ex6 [transferred])
end
\ No newline at end of file
--- a/src/HOLCF/IsaMakefile Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOLCF/IsaMakefile Fri Mar 12 16:02:42 2010 +0100
@@ -47,7 +47,6 @@
HOLCF.thy \
Lift.thy \
LowerPD.thy \
- NatIso.thy \
One.thy \
Pcpodef.thy \
Pcpo.thy \
--- a/src/HOLCF/NatIso.thy Tue Mar 09 15:42:23 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,241 +0,0 @@
-(* Title: HOLCF/NatIso.thy
- Author: Brian Huffman
-*)
-
-header {* Isomorphisms of the Natural Numbers *}
-
-theory NatIso
-imports Parity
-begin
-
-subsection {* Isomorphism between naturals and sums of naturals *}
-
-primrec
- sum2nat :: "nat + nat \<Rightarrow> nat"
-where
- "sum2nat (Inl a) = a + a"
-| "sum2nat (Inr b) = Suc (b + b)"
-
-primrec
- nat2sum :: "nat \<Rightarrow> nat + nat"
-where
- "nat2sum 0 = Inl 0"
-| "nat2sum (Suc n) = (case nat2sum n of
- Inl a \<Rightarrow> Inr a | Inr b \<Rightarrow> Inl (Suc b))"
-
-lemma nat2sum_sum2nat [simp]: "nat2sum (sum2nat x) = x"
-by (induct x, rule_tac [!] nat.induct, simp_all)
-
-lemma sum2nat_nat2sum [simp]: "sum2nat (nat2sum n) = n"
-by (induct n, auto split: sum.split)
-
-lemma inj_sum2nat: "inj sum2nat"
-by (rule inj_on_inverseI, rule nat2sum_sum2nat)
-
-lemma sum2nat_eq_iff [simp]: "sum2nat x = sum2nat y \<longleftrightarrow> x = y"
-by (rule inj_sum2nat [THEN inj_eq])
-
-lemma inj_nat2sum: "inj nat2sum"
-by (rule inj_on_inverseI, rule sum2nat_nat2sum)
-
-lemma nat2sum_eq_iff [simp]: "nat2sum x = nat2sum y \<longleftrightarrow> x = y"
-by (rule inj_nat2sum [THEN inj_eq])
-
-declare sum2nat.simps [simp del]
-declare nat2sum.simps [simp del]
-
-
-subsection {* Isomorphism between naturals and pairs of naturals *}
-
-function
- prod2nat :: "nat \<times> nat \<Rightarrow> nat"
-where
- "prod2nat (0, 0) = 0"
-| "prod2nat (0, Suc n) = Suc (prod2nat (n, 0))"
-| "prod2nat (Suc m, n) = Suc (prod2nat (m, Suc n))"
-by pat_completeness auto
-
-termination by (relation
- "inv_image (measure id <*lex*> measure id) (\<lambda>(x, y). (x+y, x))") auto
-
-primrec
- nat2prod :: "nat \<Rightarrow> nat \<times> nat"
-where
- "nat2prod 0 = (0, 0)"
-| "nat2prod (Suc x) =
- (case nat2prod x of (m, n) \<Rightarrow>
- (case n of 0 \<Rightarrow> (0, Suc m) | Suc n \<Rightarrow> (Suc m, n)))"
-
-lemma nat2prod_prod2nat [simp]: "nat2prod (prod2nat x) = x"
-by (induct x, rule prod2nat.induct, simp_all)
-
-lemma prod2nat_nat2prod [simp]: "prod2nat (nat2prod n) = n"
-by (induct n, auto split: prod.split nat.split)
-
-lemma inj_prod2nat: "inj prod2nat"
-by (rule inj_on_inverseI, rule nat2prod_prod2nat)
-
-lemma prod2nat_eq_iff [simp]: "prod2nat x = prod2nat y \<longleftrightarrow> x = y"
-by (rule inj_prod2nat [THEN inj_eq])
-
-lemma inj_nat2prod: "inj nat2prod"
-by (rule inj_on_inverseI, rule prod2nat_nat2prod)
-
-lemma nat2prod_eq_iff [simp]: "nat2prod x = nat2prod y \<longleftrightarrow> x = y"
-by (rule inj_nat2prod [THEN inj_eq])
-
-subsubsection {* Ordering properties *}
-
-lemma fst_snd_le_prod2nat: "fst x \<le> prod2nat x \<and> snd x \<le> prod2nat x"
-by (induct x rule: prod2nat.induct) auto
-
-lemma fst_le_prod2nat: "fst x \<le> prod2nat x"
-by (rule fst_snd_le_prod2nat [THEN conjunct1])
-
-lemma snd_le_prod2nat: "snd x \<le> prod2nat x"
-by (rule fst_snd_le_prod2nat [THEN conjunct2])
-
-lemma le_prod2nat_1: "a \<le> prod2nat (a, b)"
-using fst_le_prod2nat [where x="(a, b)"] by simp
-
-lemma le_prod2nat_2: "b \<le> prod2nat (a, b)"
-using snd_le_prod2nat [where x="(a, b)"] by simp
-
-declare prod2nat.simps [simp del]
-declare nat2prod.simps [simp del]
-
-
-subsection {* Isomorphism between naturals and finite sets of naturals *}
-
-subsubsection {* Preliminaries *}
-
-lemma finite_vimage_Suc_iff: "finite (Suc -` F) \<longleftrightarrow> finite F"
-apply (safe intro!: finite_vimageI inj_Suc)
-apply (rule finite_subset [where B="insert 0 (Suc ` Suc -` F)"])
-apply (rule subsetI, case_tac x, simp, simp)
-apply (rule finite_insert [THEN iffD2])
-apply (erule finite_imageI)
-done
-
-lemma vimage_Suc_insert_0: "Suc -` insert 0 A = Suc -` A"
-by auto
-
-lemma vimage_Suc_insert_Suc:
- "Suc -` insert (Suc n) A = insert n (Suc -` A)"
-by auto
-
-lemma even_nat_Suc_div_2: "even x \<Longrightarrow> Suc x div 2 = x div 2"
-by (simp only: numeral_2_eq_2 even_nat_plus_one_div_two)
-
-lemma div2_even_ext_nat:
- "\<lbrakk>x div 2 = y div 2; even x = even y\<rbrakk> \<Longrightarrow> (x::nat) = y"
-apply (rule mod_div_equality [of x 2, THEN subst])
-apply (rule mod_div_equality [of y 2, THEN subst])
-apply (case_tac "even x")
-apply (simp add: numeral_2_eq_2 even_nat_equiv_def)
-apply (simp add: numeral_2_eq_2 odd_nat_equiv_def)
-done
-
-subsubsection {* From sets to naturals *}
-
-definition
- set2nat :: "nat set \<Rightarrow> nat" where
- "set2nat = setsum (op ^ 2)"
-
-lemma set2nat_empty [simp]: "set2nat {} = 0"
-by (simp add: set2nat_def)
-
-lemma set2nat_insert [simp]:
- "\<lbrakk>finite A; n \<notin> A\<rbrakk> \<Longrightarrow> set2nat (insert n A) = 2^n + set2nat A"
-by (simp add: set2nat_def)
-
-lemma even_set2nat_iff: "finite A \<Longrightarrow> even (set2nat A) = (0 \<notin> A)"
-by (unfold set2nat_def, erule finite_induct, auto)
-
-lemma set2nat_vimage_Suc: "set2nat (Suc -` A) = set2nat A div 2"
-apply (case_tac "finite A")
-apply (erule finite_induct, simp)
-apply (case_tac x)
-apply (simp add: even_nat_Suc_div_2 even_set2nat_iff vimage_Suc_insert_0)
-apply (simp add: finite_vimageI add_commute vimage_Suc_insert_Suc)
-apply (simp add: set2nat_def finite_vimage_Suc_iff)
-done
-
-lemmas set2nat_div_2 = set2nat_vimage_Suc [symmetric]
-
-subsubsection {* From naturals to sets *}
-
-definition
- nat2set :: "nat \<Rightarrow> nat set" where
- "nat2set x = {n. odd (x div 2 ^ n)}"
-
-lemma nat2set_0 [simp]: "0 \<in> nat2set x \<longleftrightarrow> odd x"
-by (simp add: nat2set_def)
-
-lemma nat2set_Suc [simp]:
- "Suc n \<in> nat2set x \<longleftrightarrow> n \<in> nat2set (x div 2)"
-by (simp add: nat2set_def div_mult2_eq)
-
-lemma nat2set_zero [simp]: "nat2set 0 = {}"
-by (simp add: nat2set_def)
-
-lemma nat2set_div_2: "nat2set (x div 2) = Suc -` nat2set x"
-by auto
-
-lemma nat2set_plus_power_2:
- "n \<notin> nat2set z \<Longrightarrow> nat2set (2 ^ n + z) = insert n (nat2set z)"
- apply (induct n arbitrary: z, simp_all)
- apply (rule set_ext, induct_tac x, simp, simp add: even_nat_Suc_div_2)
- apply (rule set_ext, induct_tac x, simp, simp add: add_commute)
-done
-
-lemma finite_nat2set [simp]: "finite (nat2set n)"
-apply (induct n rule: nat_less_induct)
-apply (case_tac "n = 0", simp)
-apply (drule_tac x="n div 2" in spec, simp)
-apply (simp add: nat2set_div_2)
-apply (simp add: finite_vimage_Suc_iff)
-done
-
-subsubsection {* Proof of isomorphism *}
-
-lemma set2nat_nat2set [simp]: "set2nat (nat2set n) = n"
-apply (induct n rule: nat_less_induct)
-apply (case_tac "n = 0", simp)
-apply (drule_tac x="n div 2" in spec, simp)
-apply (simp add: nat2set_div_2 set2nat_vimage_Suc)
-apply (erule div2_even_ext_nat)
-apply (simp add: even_set2nat_iff)
-done
-
-lemma nat2set_set2nat [simp]: "finite A \<Longrightarrow> nat2set (set2nat A) = A"
-apply (erule finite_induct, simp_all)
-apply (simp add: nat2set_plus_power_2)
-done
-
-lemma inj_nat2set: "inj nat2set"
-by (rule inj_on_inverseI, rule set2nat_nat2set)
-
-lemma nat2set_eq_iff [simp]: "nat2set x = nat2set y \<longleftrightarrow> x = y"
-by (rule inj_eq [OF inj_nat2set])
-
-lemma inj_on_set2nat: "inj_on set2nat (Collect finite)"
-by (rule inj_on_inverseI [where g="nat2set"], simp)
-
-lemma set2nat_eq_iff [simp]:
- "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> set2nat A = set2nat B \<longleftrightarrow> A = B"
-by (rule iffI, simp add: inj_onD [OF inj_on_set2nat], simp)
-
-subsubsection {* Ordering properties *}
-
-lemma nat_less_power2: "n < 2^n"
-by (induct n) simp_all
-
-lemma less_set2nat: "\<lbrakk>finite A; x \<in> A\<rbrakk> \<Longrightarrow> x < set2nat A"
-unfolding set2nat_def
-apply (rule order_less_le_trans [where y="setsum (op ^ 2) {x}"])
-apply (simp add: nat_less_power2)
-apply (erule setsum_mono2, simp, simp)
-done
-
-end
\ No newline at end of file
--- a/src/HOLCF/ROOT.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOLCF/ROOT.ML Fri Mar 12 16:02:42 2010 +0100
@@ -4,6 +4,6 @@
HOLCF -- a semantic extension of HOL by the LCF logic.
*)
-no_document use_thys ["Nat_Int_Bij"];
+no_document use_thys ["Nat_Bijection"];
use_thys ["HOLCF"];
--- a/src/HOLCF/Universal.thy Tue Mar 09 15:42:23 2010 +0100
+++ b/src/HOLCF/Universal.thy Fri Mar 12 16:02:42 2010 +0100
@@ -3,7 +3,7 @@
*)
theory Universal
-imports CompactBasis NatIso
+imports CompactBasis Nat_Bijection
begin
subsection {* Basis datatype *}
@@ -13,7 +13,7 @@
definition
node :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis set \<Rightarrow> ubasis"
where
- "node i a S = Suc (prod2nat (i, prod2nat (a, set2nat S)))"
+ "node i a S = Suc (prod_encode (i, prod_encode (a, set_encode S)))"
lemma node_not_0 [simp]: "node i a S \<noteq> 0"
unfolding node_def by simp
@@ -24,30 +24,30 @@
lemma node_inject [simp]:
"\<lbrakk>finite S; finite T\<rbrakk>
\<Longrightarrow> node i a S = node j b T \<longleftrightarrow> i = j \<and> a = b \<and> S = T"
-unfolding node_def by simp
+unfolding node_def by (simp add: prod_encode_eq set_encode_eq)
lemma node_gt0: "i < node i a S"
unfolding node_def less_Suc_eq_le
-by (rule le_prod2nat_1)
+by (rule le_prod_encode_1)
lemma node_gt1: "a < node i a S"
unfolding node_def less_Suc_eq_le
-by (rule order_trans [OF le_prod2nat_1 le_prod2nat_2])
+by (rule order_trans [OF le_prod_encode_1 le_prod_encode_2])
lemma nat_less_power2: "n < 2^n"
by (induct n) simp_all
lemma node_gt2: "\<lbrakk>finite S; b \<in> S\<rbrakk> \<Longrightarrow> b < node i a S"
-unfolding node_def less_Suc_eq_le set2nat_def
-apply (rule order_trans [OF _ le_prod2nat_2])
-apply (rule order_trans [OF _ le_prod2nat_2])
+unfolding node_def less_Suc_eq_le set_encode_def
+apply (rule order_trans [OF _ le_prod_encode_2])
+apply (rule order_trans [OF _ le_prod_encode_2])
apply (rule order_trans [where y="setsum (op ^ 2) {b}"])
apply (simp add: nat_less_power2 [THEN order_less_imp_le])
apply (erule setsum_mono2, simp, simp)
done
-lemma eq_prod2nat_pairI:
- "\<lbrakk>fst (nat2prod x) = a; snd (nat2prod x) = b\<rbrakk> \<Longrightarrow> x = prod2nat (a, b)"
+lemma eq_prod_encode_pairI:
+ "\<lbrakk>fst (prod_decode x) = a; snd (prod_decode x) = b\<rbrakk> \<Longrightarrow> x = prod_encode (a, b)"
by (erule subst, erule subst, simp)
lemma node_cases:
@@ -57,10 +57,10 @@
apply (cases x)
apply (erule 1)
apply (rule 2)
- apply (rule finite_nat2set)
+ apply (rule finite_set_decode)
apply (simp add: node_def)
- apply (rule eq_prod2nat_pairI [OF refl])
- apply (rule eq_prod2nat_pairI [OF refl refl])
+ apply (rule eq_prod_encode_pairI [OF refl])
+ apply (rule eq_prod_encode_pairI [OF refl refl])
done
lemma node_induct:
--- a/src/Provers/Arith/fast_lin_arith.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML Fri Mar 12 16:02:42 2010 +0100
@@ -849,17 +849,20 @@
fun splitasms ctxt (asms : thm list) : splittree =
let val {neqE, ...} = get_data ctxt
- fun elim_neq (asms', []) = Tip (rev asms')
- | elim_neq (asms', asm::asms) =
- (case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) neqE of
+ fun elim_neq [] (asms', []) = Tip (rev asms')
+ | elim_neq [] (asms', asms) = Tip (rev asms' @ asms)
+ | elim_neq (neq :: neqs) (asms', []) = elim_neq neqs ([],rev asms')
+ | elim_neq (neqs as (neq :: _)) (asms', asm::asms) =
+ (case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) [neq] of
SOME spl =>
let val (ct1, ct2) = extract (cprop_of spl)
val thm1 = assume ct1
val thm2 = assume ct2
- in Spl (spl, ct1, elim_neq (asms', asms@[thm1]), ct2, elim_neq (asms', asms@[thm2]))
+ in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]),
+ ct2, elim_neq neqs (asms', asms@[thm2]))
end
- | NONE => elim_neq (asm::asms', asms))
-in elim_neq ([], asms) end;
+ | NONE => elim_neq neqs (asm::asms', asms))
+in elim_neq neqE ([], asms) end;
fun fwdproof ss (Tip asms : splittree) (j::js : injust list) = (mkthm ss asms j, js)
| fwdproof ss (Spl (thm, ct1, tree1, ct2, tree2)) js =
--- a/src/Pure/General/name_space.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/General/name_space.ML Fri Mar 12 16:02:42 2010 +0100
@@ -47,6 +47,7 @@
val transform_binding: naming -> binding -> binding
val full_name: naming -> binding -> string
val declare: bool -> naming -> binding -> T -> string * T
+ val alias: naming -> binding -> string -> T -> T
type 'a table = T * 'a Symtab.table
val define: bool -> naming -> binding * 'a -> 'a table -> string * 'a table
val empty_table: string -> 'a table
@@ -72,8 +73,7 @@
(* datatype entry *)
type entry =
- {externals: xstring list,
- concealed: bool,
+ {concealed: bool,
group: serial option,
theory_name: string,
pos: Position.T,
@@ -96,7 +96,7 @@
Name_Space of
{kind: string,
internals: (string list * string list) Symtab.table, (*visible, hidden*)
- entries: entry Symtab.table};
+ entries: (xstring list * entry) Symtab.table}; (*externals, entry*)
fun make_name_space (kind, internals, entries) =
Name_Space {kind = kind, internals = internals, entries = entries};
@@ -115,8 +115,7 @@
fun the_entry (Name_Space {kind, entries, ...}) name =
(case Symtab.lookup entries name of
NONE => error ("Unknown " ^ kind ^ " " ^ quote name)
- | SOME {concealed, group, theory_name, pos, id, ...} =>
- {concealed = concealed, group = group, theory_name = theory_name, pos = pos, id = id});
+ | SOME (_, entry) => entry);
fun is_concealed space name = #concealed (the_entry space name);
@@ -134,7 +133,7 @@
fun get_accesses (Name_Space {entries, ...}) name =
(case Symtab.lookup entries name of
NONE => [name]
- | SOME {externals, ...} => externals);
+ | SOME (externals, _) => externals);
fun valid_accesses (Name_Space {internals, ...}) name =
Symtab.fold (fn (xname, (names, _)) =>
@@ -212,7 +211,7 @@
then raise Symtab.SAME
else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
val entries' = (entries1, entries2) |> Symtab.join
- (fn name => fn (entry1, entry2) =>
+ (fn name => fn ((_, entry1), (_, entry2)) =>
if #id entry1 = #id entry2 then raise Symtab.SAME
else err_dup kind' (name, entry1) (name, entry2));
in make_name_space (kind', internals', entries') end;
@@ -311,13 +310,13 @@
(* declaration *)
-fun new_entry strict entry =
+fun new_entry strict (name, (externals, entry)) =
map_name_space (fn (kind, internals, entries) =>
let
val entries' =
- (if strict then Symtab.update_new else Symtab.update) entry entries
+ (if strict then Symtab.update_new else Symtab.update) (name, (externals, entry)) entries
handle Symtab.DUP dup =>
- err_dup kind (dup, the (Symtab.lookup entries dup)) entry;
+ err_dup kind (dup, #2 (the (Symtab.lookup entries dup))) (name, entry);
in (kind, internals, entries') end);
fun declare strict naming binding space =
@@ -330,16 +329,35 @@
val _ = name = "" andalso err_bad binding;
val entry =
- {externals = accs',
- concealed = concealed,
+ {concealed = concealed,
group = group,
theory_name = theory_name,
pos = Position.default (Binding.pos_of binding),
id = serial ()};
- val space' = space |> fold (add_name name) accs |> new_entry strict (name, entry);
+ val space' = space
+ |> fold (add_name name) accs
+ |> new_entry strict (name, (accs', entry));
in (name, space') end;
+(* alias *)
+
+fun alias naming binding name space =
+ let
+ val (accs, accs') = accesses naming binding;
+ val space' = space
+ |> fold (add_name name) accs
+ |> map_name_space (fn (kind, internals, entries) =>
+ let
+ val _ = Symtab.defined entries name orelse
+ error ("Undefined " ^ kind ^ " " ^ quote name);
+ val entries' = entries
+ |> Symtab.map_entry name (fn (externals, entry) =>
+ (Library.merge (op =) (externals, accs'), entry))
+ in (kind, internals, entries') end);
+ in space' end;
+
+
(** name spaces coupled with symbol tables **)
--- a/src/Pure/General/sha1_polyml.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/General/sha1_polyml.ML Fri Mar 12 16:02:42 2010 +0100
@@ -15,15 +15,21 @@
let val c = CInterface.fromCchar (CInterface.offset i CInterface.Cchar arr)
in (op ^) (pairself hex_digit (Integer.div_mod (Char.ord c) 16)) end
+val lib_path =
+ ("$ML_HOME/" ^ (if String.isSuffix "cygwin" ml_system then "sha1.dll" else "libsha1.so"))
+ |> Path.explode;
+
fun digest_external str =
let
val digest = CInterface.alloc 20 CInterface.Cchar;
- val _ = CInterface.call3 (CInterface.get_sym "sha1" "sha1_buffer")
- (CInterface.STRING, CInterface.INT, CInterface.POINTER)
- CInterface.POINTER (str, size str, CInterface.address digest);
+ val _ =
+ CInterface.call3 (CInterface.get_sym (File.platform_path lib_path) "sha1_buffer")
+ (CInterface.STRING, CInterface.INT, CInterface.POINTER)
+ CInterface.POINTER (str, size str, CInterface.address digest);
in fold (suffix o hex_string digest) (0 upto 19) "" end;
fun digest str = digest_external str
- handle CInterface.Foreign _ => SHA1.digest str;
+ handle CInterface.Foreign msg =>
+ (warning (msg ^ "\nUsing slow ML implementation of SHA1.digest"); SHA1.digest str);
end;
--- a/src/Pure/Isar/class.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/class.ML Fri Mar 12 16:02:42 2010 +0100
@@ -338,7 +338,7 @@
val subclass = gen_subclass (K I) user_proof;
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
-val subclass_cmd = gen_subclass Sign.read_class user_proof;
+val subclass_cmd = gen_subclass (ProofContext.read_class o ProofContext.init) user_proof;
end; (*local*)
--- a/src/Pure/Isar/class_target.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/class_target.ML Fri Mar 12 16:02:42 2010 +0100
@@ -417,7 +417,8 @@
fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
let
- val all_arities = map (fn raw_tyco => Sign.read_arity thy
+ val ctxt = ProofContext.init thy;
+ val all_arities = map (fn raw_tyco => ProofContext.read_arity ctxt
(raw_tyco, raw_sorts, raw_sort)) raw_tycos;
val tycos = map #1 all_arities;
val (_, sorts, sort) = hd all_arities;
--- a/src/Pure/Isar/constdefs.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/constdefs.ML Fri Mar 12 16:02:42 2010 +0100
@@ -22,6 +22,8 @@
fun gen_constdef prep_vars prep_prop prep_att
structs (raw_decl, ((raw_name, raw_atts), raw_prop)) thy =
let
+ val _ = legacy_feature ("\"constdefs\"; prefer \"definition\" instead");
+
fun err msg ts = error (cat_lines (msg :: map (Syntax.string_of_term_global thy) ts));
val thy_ctxt = ProofContext.init thy;
--- a/src/Pure/Isar/isar_syn.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/isar_syn.ML Fri Mar 12 16:02:42 2010 +0100
@@ -103,16 +103,15 @@
(* types *)
val _ =
- OuterSyntax.command "typedecl" "type declaration" K.thy_decl
- (P.type_args -- P.binding -- P.opt_mixfix >> (fn ((args, a), mx) =>
- Toplevel.theory (Typedecl.typedecl (a, args, mx) #> snd)));
+ OuterSyntax.local_theory "typedecl" "type declaration" K.thy_decl
+ (P.type_args -- P.binding -- P.opt_mixfix
+ >> (fn ((args, a), mx) => Typedecl.typedecl (a, args, mx) #> snd));
val _ =
OuterSyntax.command "types" "declare type abbreviations" K.thy_decl
(Scan.repeat1
(P.type_args -- P.binding -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_mixfix')))
- >> (Toplevel.theory o Sign.add_tyabbrs o
- map (fn ((args, a), (T, mx)) => (a, args, T, mx))));
+ >> (Toplevel.theory o Sign.add_tyabbrs o map (fn ((args, a), (T, mx)) => (a, args, T, mx))));
val _ =
OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"
--- a/src/Pure/Isar/local_defs.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/local_defs.ML Fri Mar 12 16:02:42 2010 +0100
@@ -17,7 +17,7 @@
val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
(term * term) * Proof.context
val export: Proof.context -> Proof.context -> thm -> thm list * thm
- val export_cterm: Proof.context -> Proof.context -> cterm -> cterm * thm
+ val export_cterm: Proof.context -> Proof.context -> cterm -> thm list * cterm
val trans_terms: Proof.context -> thm list -> thm
val trans_props: Proof.context -> thm list -> thm
val print_rules: Proof.context -> unit
@@ -155,8 +155,7 @@
TERM b as b xs == b as
*)
fun export_cterm inner outer ct =
- let val (defs, ct') = export inner outer (Drule.mk_term ct) ||> Drule.dest_term
- in (ct', MetaSimplifier.rewrite true defs ct) end;
+ export inner outer (Drule.mk_term ct) ||> Drule.dest_term;
(* basic transitivity reasoning -- modulo beta-eta *)
--- a/src/Pure/Isar/proof_context.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/proof_context.ML Fri Mar 12 16:02:42 2010 +0100
@@ -27,6 +27,8 @@
val restore_naming: Proof.context -> Proof.context -> Proof.context
val full_name: Proof.context -> binding -> string
val syn_of: Proof.context -> Syntax.syntax
+ val tsig_of: Proof.context -> Type.tsig
+ val default_sort: Proof.context -> indexname -> sort
val consts_of: Proof.context -> Consts.T
val the_const_constraint: Proof.context -> string -> typ
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
@@ -41,6 +43,9 @@
val pretty_term_abbrev: Proof.context -> term -> Pretty.T
val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T
val pretty_fact: Proof.context -> string * thm list -> Pretty.T
+ val read_class: Proof.context -> xstring -> class
+ val read_arity: Proof.context -> xstring * string list * string -> arity
+ val cert_arity: Proof.context -> arity -> arity
val read_typ: Proof.context -> string -> typ
val read_typ_syntax: Proof.context -> string -> typ
val read_typ_abbrev: Proof.context -> string -> typ
@@ -53,6 +58,7 @@
val inferred_param: string -> Proof.context -> typ * Proof.context
val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
val read_type_name: Proof.context -> bool -> string -> typ
+ val read_type_name_proper: Proof.context -> bool -> string -> typ
val read_const_proper: Proof.context -> bool -> string -> term
val read_const: Proof.context -> bool -> string -> term
val allow_dummies: Proof.context -> Proof.context
@@ -122,6 +128,9 @@
Context.generic -> Context.generic
val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->
Context.generic -> Context.generic
+ val class_alias: binding -> class -> Proof.context -> Proof.context
+ val type_alias: binding -> string -> Proof.context -> Proof.context
+ val const_alias: binding -> string -> Proof.context -> Proof.context
val add_const_constraint: string * typ option -> Proof.context -> Proof.context
val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
val revert_abbrev: string -> string -> Proof.context -> Proof.context
@@ -167,16 +176,17 @@
datatype ctxt =
Ctxt of
- {mode: mode, (*inner syntax mode*)
- naming: Name_Space.naming, (*local naming conventions*)
- syntax: Local_Syntax.T, (*local syntax*)
- consts: Consts.T * Consts.T, (*local/global consts*)
- facts: Facts.T, (*local facts*)
+ {mode: mode, (*inner syntax mode*)
+ naming: Name_Space.naming, (*local naming conventions*)
+ syntax: Local_Syntax.T, (*local syntax*)
+ tsigs: Type.tsig * Type.tsig, (*local/global type signature -- local name space only*)
+ consts: Consts.T * Consts.T, (*local/global consts -- local name space/abbrevs only*)
+ facts: Facts.T, (*local facts*)
cases: (string * (Rule_Cases.T * bool)) list}; (*named case contexts*)
-fun make_ctxt (mode, naming, syntax, consts, facts, cases) =
+fun make_ctxt (mode, naming, syntax, tsigs, consts, facts, cases) =
Ctxt {mode = mode, naming = naming, syntax = syntax,
- consts = consts, facts = facts, cases = cases};
+ tsigs = tsigs, consts = consts, facts = facts, cases = cases};
val local_naming = Name_Space.default_naming |> Name_Space.add_path "local";
@@ -185,41 +195,46 @@
type T = ctxt;
fun init thy =
make_ctxt (mode_default, local_naming, Local_Syntax.init thy,
+ (Sign.tsig_of thy, Sign.tsig_of thy),
(Sign.consts_of thy, Sign.consts_of thy), Facts.empty, []);
);
fun rep_context ctxt = ContextData.get ctxt |> (fn Ctxt args => args);
fun map_context f =
- ContextData.map (fn Ctxt {mode, naming, syntax, consts, facts, cases} =>
- make_ctxt (f (mode, naming, syntax, consts, facts, cases)));
+ ContextData.map (fn Ctxt {mode, naming, syntax, tsigs, consts, facts, cases} =>
+ make_ctxt (f (mode, naming, syntax, tsigs, consts, facts, cases)));
-fun set_mode mode = map_context (fn (_, naming, syntax, consts, facts, cases) =>
- (mode, naming, syntax, consts, facts, cases));
+fun set_mode mode = map_context (fn (_, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, syntax, tsigs, consts, facts, cases));
fun map_mode f =
- map_context (fn (Mode {stmt, pattern, schematic, abbrev}, naming, syntax, consts, facts, cases) =>
- (make_mode (f (stmt, pattern, schematic, abbrev)), naming, syntax, consts, facts, cases));
+ map_context (fn (Mode {stmt, pattern, schematic, abbrev}, naming, syntax, tsigs, consts, facts, cases) =>
+ (make_mode (f (stmt, pattern, schematic, abbrev)), naming, syntax, tsigs, consts, facts, cases));
fun map_naming f =
- map_context (fn (mode, naming, syntax, consts, facts, cases) =>
- (mode, f naming, syntax, consts, facts, cases));
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, f naming, syntax, tsigs, consts, facts, cases));
fun map_syntax f =
- map_context (fn (mode, naming, syntax, consts, facts, cases) =>
- (mode, naming, f syntax, consts, facts, cases));
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, f syntax, tsigs, consts, facts, cases));
+
+fun map_tsigs f =
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, syntax, f tsigs, consts, facts, cases));
fun map_consts f =
- map_context (fn (mode, naming, syntax, consts, facts, cases) =>
- (mode, naming, syntax, f consts, facts, cases));
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, syntax, tsigs, f consts, facts, cases));
fun map_facts f =
- map_context (fn (mode, naming, syntax, consts, facts, cases) =>
- (mode, naming, syntax, consts, f facts, cases));
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, syntax, tsigs, consts, f facts, cases));
fun map_cases f =
- map_context (fn (mode, naming, syntax, consts, facts, cases) =>
- (mode, naming, syntax, consts, facts, f cases));
+ map_context (fn (mode, naming, syntax, tsigs, consts, facts, cases) =>
+ (mode, naming, syntax, tsigs, consts, facts, f cases));
val get_mode = #mode o rep_context;
val restore_mode = set_mode o get_mode;
@@ -237,6 +252,9 @@
val set_syntax_mode = map_syntax o Local_Syntax.set_mode;
val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of;
+val tsig_of = #1 o #tsigs o rep_context;
+fun default_sort ctxt = the_default (Type.defaultS (tsig_of ctxt)) o Variable.def_sort ctxt;
+
val consts_of = #1 o #consts o rep_context;
val the_const_constraint = Consts.the_constraint o consts_of;
@@ -246,8 +264,13 @@
(* theory transfer *)
-fun transfer_syntax thy =
- map_syntax (Local_Syntax.rebuild thy) #>
+fun transfer_syntax thy ctxt = ctxt |>
+ map_syntax (Local_Syntax.rebuild thy) |>
+ map_tsigs (fn tsigs as (local_tsig, global_tsig) =>
+ let val thy_tsig = Sign.tsig_of thy in
+ if Type.eq_tsig (thy_tsig, global_tsig) then tsigs
+ else (Type.merge_tsigs (Syntax.pp ctxt) (local_tsig, thy_tsig), thy_tsig)
+ end) |>
map_consts (fn consts as (local_consts, global_consts) =>
let val thy_consts = Sign.consts_of thy in
if Consts.eq_consts (thy_consts, global_consts) then consts
@@ -299,23 +322,49 @@
(** prepare types **)
-(* read_typ *)
+(* classes *)
+
+fun read_class ctxt text =
+ let
+ val tsig = tsig_of ctxt;
+ val (syms, pos) = Syntax.read_token text;
+ val c = Type.cert_class tsig (Type.intern_class tsig (Symbol_Pos.content syms))
+ handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos);
+ val _ = Position.report (Markup.tclass c) pos;
+ in c end;
+
+
+(* type arities *)
+
+local
+
+fun prep_arity prep_tycon prep_sort ctxt (t, Ss, S) =
+ let val arity = (prep_tycon ctxt t, map (prep_sort ctxt) Ss, prep_sort ctxt S)
+ in Type.add_arity (Syntax.pp ctxt) arity (tsig_of ctxt); arity end;
+
+in
+
+val read_arity = prep_arity (Type.intern_type o tsig_of) Syntax.read_sort;
+val cert_arity = prep_arity (K I) (Type.cert_sort o tsig_of);
+
+end;
+
+
+(* types *)
fun read_typ_mode mode ctxt s =
Syntax.read_typ (Type.set_mode mode ctxt) s;
-val read_typ = read_typ_mode Type.mode_default;
+val read_typ = read_typ_mode Type.mode_default;
val read_typ_syntax = read_typ_mode Type.mode_syntax;
val read_typ_abbrev = read_typ_mode Type.mode_abbrev;
-(* cert_typ *)
-
fun cert_typ_mode mode ctxt T =
- Sign.certify_typ_mode mode (theory_of ctxt) T
+ Type.cert_typ_mode mode (tsig_of ctxt) T
handle TYPE (msg, _, _) => error msg;
-val cert_typ = cert_typ_mode Type.mode_default;
+val cert_typ = cert_typ_mode Type.mode_default;
val cert_typ_syntax = cert_typ_mode Type.mode_syntax;
val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev;
@@ -424,16 +473,16 @@
fun read_type_name ctxt strict text =
let
- val thy = theory_of ctxt;
+ val tsig = tsig_of ctxt;
val (c, pos) = token_content text;
in
if Syntax.is_tid c then
(Position.report Markup.tfree pos;
- TFree (c, the_default (Sign.defaultS thy) (Variable.def_sort ctxt (c, ~1))))
+ TFree (c, default_sort ctxt (c, ~1)))
else
let
- val d = Sign.intern_type thy c;
- val decl = Sign.the_type_decl thy d;
+ val d = Type.intern_type tsig c;
+ val decl = Type.the_decl tsig d;
val _ = Position.report (Markup.tycon d) pos;
fun err () = error ("Bad type name: " ^ quote d);
val args =
@@ -444,6 +493,12 @@
in Type (d, replicate args dummyT) end
end;
+fun read_type_name_proper ctxt strict text =
+ (case read_type_name ctxt strict text of
+ T as Type _ => T
+ | T => error ("Not a type constructor: " ^ Syntax.string_of_typ ctxt T));
+
+
fun read_const_proper ctxt strict = prep_const_proper ctxt strict o token_content;
fun read_const ctxt strict text =
@@ -470,8 +525,6 @@
(* local abbreviations *)
-val tsig_of = Sign.tsig_of o ProofContext.theory_of;
-
local
fun certify_consts ctxt = Consts.certify (Syntax.pp ctxt) (tsig_of ctxt)
@@ -553,9 +606,9 @@
(* types *)
-fun get_sort thy def_sort raw_env =
+fun get_sort ctxt def_sort raw_env =
let
- val tsig = Sign.tsig_of thy;
+ val tsig = tsig_of ctxt;
fun eq ((xi, S), (xi', S')) =
Term.eq_ix (xi, xi') andalso Type.eq_sort tsig (S, S');
@@ -571,8 +624,8 @@
| (SOME S, NONE) => S
| (SOME S, SOME S') =>
if Type.eq_sort tsig (S, S') then S'
- else error ("Sort constraint " ^ Syntax.string_of_sort_global thy S ^
- " inconsistent with default " ^ Syntax.string_of_sort_global thy S' ^
+ else error ("Sort constraint " ^ Syntax.string_of_sort ctxt S ^
+ " inconsistent with default " ^ Syntax.string_of_sort ctxt S' ^
" for type variable " ^ quote (Term.string_of_vname' xi)));
in get end;
@@ -594,12 +647,10 @@
in
fun term_context ctxt =
- let val thy = theory_of ctxt in
- {get_sort = get_sort thy (Variable.def_sort ctxt),
- map_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt false a)))
- handle ERROR _ => (false, Consts.intern (consts_of ctxt) a)),
- map_free = intern_skolem ctxt (Variable.def_type ctxt false)}
- end;
+ {get_sort = get_sort ctxt (Variable.def_sort ctxt),
+ map_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt false a)))
+ handle ERROR _ => (false, Consts.intern (consts_of ctxt) a)),
+ map_free = intern_skolem ctxt (Variable.def_type ctxt false)};
fun decode_term ctxt =
let val {get_sort, map_const, map_free} = term_context ctxt
@@ -680,8 +731,7 @@
fun parse_typ ctxt text =
let
- val thy = theory_of ctxt;
- val get_sort = get_sort thy (Variable.def_sort ctxt);
+ val get_sort = get_sort ctxt (Variable.def_sort ctxt);
val (syms, pos) = Syntax.parse_token Markup.typ text;
val T = Syntax.standard_parse_typ ctxt (syn_of ctxt) get_sort (syms, pos)
handle ERROR msg => cat_error msg ("Failed to parse type" ^ Position.str_of pos);
@@ -689,7 +739,6 @@
fun parse_term T ctxt text =
let
- val thy = theory_of ctxt;
val {get_sort, map_const, map_free} = term_context ctxt;
val (T', _) = TypeInfer.paramify_dummies T 0;
@@ -700,33 +749,33 @@
handle ERROR msg => SOME msg;
val t =
Syntax.standard_parse_term (Syntax.pp ctxt) check get_sort map_const map_free
- ctxt (Sign.is_logtype thy) (syn_of ctxt) T' (syms, pos)
- handle ERROR msg => cat_error msg ("Failed to parse " ^ kind ^ Position.str_of pos);
+ ctxt (Type.is_logtype (tsig_of ctxt)) (syn_of ctxt) T' (syms, pos)
+ handle ERROR msg => cat_error msg ("Failed to parse " ^ kind ^ Position.str_of pos);
in t end;
fun unparse_sort ctxt =
- Syntax.standard_unparse_sort {extern_class = Sign.extern_class (theory_of ctxt)}
+ Syntax.standard_unparse_sort {extern_class = Type.extern_class (tsig_of ctxt)}
ctxt (syn_of ctxt);
fun unparse_typ ctxt =
let
- val thy = theory_of ctxt;
- val extern = {extern_class = Sign.extern_class thy, extern_type = Sign.extern_type thy};
+ val tsig = tsig_of ctxt;
+ val extern = {extern_class = Type.extern_class tsig, extern_type = Type.extern_type tsig};
in Syntax.standard_unparse_typ extern ctxt (syn_of ctxt) end;
fun unparse_term ctxt =
let
- val thy = theory_of ctxt;
+ val tsig = tsig_of ctxt;
val syntax = syntax_of ctxt;
val consts = consts_of ctxt;
val extern =
- {extern_class = Sign.extern_class thy,
- extern_type = Sign.extern_type thy,
+ {extern_class = Type.extern_class tsig,
+ extern_type = Type.extern_type tsig,
extern_const = Consts.extern consts};
in
Syntax.standard_unparse_term (Local_Syntax.idents_of syntax) extern ctxt
- (Local_Syntax.syn_of syntax) (not (PureThy.old_appl_syntax thy))
+ (Local_Syntax.syn_of syntax) (not (PureThy.old_appl_syntax (theory_of ctxt)))
end;
in
@@ -967,7 +1016,7 @@
-(** parameters **)
+(** basic logical entities **)
(* variables *)
@@ -999,7 +1048,14 @@
end;
-(* authentic constants *)
+(* authentic syntax *)
+
+val _ = Context.>>
+ (Context.map_theory
+ (Sign.add_advanced_trfuns
+ (Syntax.type_ast_trs
+ {read_class = read_class,
+ read_type = fn ctxt => #1 o dest_Type o read_type_name_proper ctxt false}, [], [], [])));
local
@@ -1070,8 +1126,14 @@
in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end);
in Context.mapping (Sign.notation add mode args') (notation add mode args') end;
+end;
-end;
+
+(* aliases *)
+
+fun class_alias b c ctxt = (map_tsigs o apfst) (Type.class_alias (naming_of ctxt) b c) ctxt;
+fun type_alias b c ctxt = (map_tsigs o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt;
+fun const_alias b c ctxt = (map_consts o apfst) (Consts.alias (naming_of ctxt) b c) ctxt;
(* local constants *)
--- a/src/Pure/Isar/theory_target.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/theory_target.ML Fri Mar 12 16:02:42 2010 +0100
@@ -278,8 +278,11 @@
val thy_ctxt = ProofContext.init thy;
val name' = Thm.def_binding_optional b name;
- val (rhs', rhs_conv) =
- Local_Defs.export_cterm lthy thy_ctxt (Thm.cterm_of thy rhs) |>> Thm.term_of;
+
+ val crhs = Thm.cterm_of thy rhs;
+ val (defs, rhs') = Local_Defs.export_cterm lthy thy_ctxt crhs ||> Thm.term_of;
+ val rhs_conv = MetaSimplifier.rewrite true defs crhs;
+
val xs = Variable.add_fixed (Local_Theory.target_of lthy) rhs' [];
val T = Term.fastype_of rhs;
--- a/src/Pure/Isar/typedecl.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Isar/typedecl.ML Fri Mar 12 16:02:42 2010 +0100
@@ -6,26 +6,51 @@
signature TYPEDECL =
sig
- val typedecl: binding * string list * mixfix -> theory -> typ * theory
+ val typedecl: binding * string list * mixfix -> local_theory -> typ * local_theory
+ val typedecl_global: binding * string list * mixfix -> theory -> typ * theory
end;
structure Typedecl: TYPEDECL =
struct
-fun typedecl (b, vs, mx) thy =
+fun typedecl (b, vs, mx) lthy =
let
- val base_sort = Object_Logic.get_base_sort thy;
- val _ = has_duplicates (op =) vs andalso
- error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b));
- val name = Sign.full_name thy b;
+ fun err msg = error (msg ^ " in type declaration " ^ quote (Binding.str_of b));
+ val _ = has_duplicates (op =) vs andalso err "Duplicate parameters";
+
+ val name = Local_Theory.full_name lthy b;
val n = length vs;
- val T = Type (name, map (fn v => TFree (v, [])) vs);
+ val args = map (fn a => TFree (a, ProofContext.default_sort lthy (a, ~1))) vs;
+ val T = Type (name, args);
+
+ val bad_args =
+ #2 (Term.dest_Type (Logic.type_map (singleton (Variable.polymorphic lthy)) T))
+ |> filter_out Term.is_TVar;
+ val _ = not (null bad_args) andalso
+ err ("Locally fixed type arguments " ^ commas_quote (map (Syntax.string_of_typ lthy) bad_args));
+
+ val base_sort = Object_Logic.get_base_sort (ProofContext.theory_of lthy);
in
- thy
- |> Sign.add_types [(b, n, mx)]
- |> (case base_sort of NONE => I | SOME S => AxClass.axiomatize_arity (name, replicate n S, S))
+ lthy
+ |> Local_Theory.theory
+ (Sign.add_types [(b, n, NoSyn)] #>
+ (case base_sort of
+ NONE => I
+ | SOME S => AxClass.axiomatize_arity (name, replicate n S, S)))
+ |> Local_Theory.checkpoint
+ |> Local_Theory.type_notation true Syntax.mode_default [(T, mx)]
+ |> Local_Theory.type_syntax false (fn phi =>
+ let val b' = Morphism.binding phi b
+ in Context.mapping (Sign.type_alias b' name) (ProofContext.type_alias b' name) end)
+ |> ProofContext.type_alias b name
+ |> Variable.declare_typ T
|> pair T
end;
+fun typedecl_global decl =
+ Theory_Target.init NONE
+ #> typedecl decl
+ #> Local_Theory.exit_result_global Morphism.typ;
+
end;
--- a/src/Pure/ML/ml_antiquote.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/ML/ml_antiquote.ML Fri Mar 12 16:02:42 2010 +0100
@@ -102,8 +102,8 @@
(* type classes *)
-fun class syn = Args.theory -- Scan.lift Args.name_source >> (fn (thy, s) =>
- Sign.read_class thy s
+fun class syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
+ ProofContext.read_class ctxt s
|> syn ? Syntax.mark_class
|> ML_Syntax.print_string);
@@ -119,8 +119,8 @@
fun type_name kind check = Args.context -- Scan.lift (OuterParse.position Args.name_source)
>> (fn (ctxt, (s, pos)) =>
let
- val Type (c, _) = ProofContext.read_type_name ctxt false s;
- val decl = Sign.the_type_decl (ProofContext.theory_of ctxt) c;
+ val Type (c, _) = ProofContext.read_type_name_proper ctxt false s;
+ val decl = Type.the_decl (ProofContext.tsig_of ctxt) c;
val res =
(case try check (c, decl) of
SOME res => res
--- a/src/Pure/ROOT.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/ROOT.ML Fri Mar 12 16:02:42 2010 +0100
@@ -148,7 +148,6 @@
use "assumption.ML";
use "display.ML";
use "goal.ML";
-use "axclass.ML";
(* Isar -- Intelligible Semi-Automated Reasoning *)
@@ -209,6 +208,7 @@
use "Isar/local_theory.ML";
use "Isar/overloading.ML";
use "Isar/locale.ML";
+use "axclass.ML";
use "Isar/class_target.ML";
use "Isar/theory_target.ML";
use "Isar/expression.ML";
--- a/src/Pure/Syntax/syntax.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Syntax/syntax.ML Fri Mar 12 16:02:42 2010 +0100
@@ -29,10 +29,7 @@
val mode_default: mode
val mode_input: mode
val merge_syntaxes: syntax -> syntax -> syntax
- val empty_syntax: syntax
- val basic_syntax:
- {read_class: theory -> xstring -> string,
- read_type: theory -> xstring -> string} -> syntax
+ val basic_syntax: syntax
val basic_nonterms: string list
val print_gram: syntax -> unit
val print_trans: syntax -> unit
@@ -383,9 +380,9 @@
(* basic syntax *)
-fun basic_syntax read =
+val basic_syntax =
empty_syntax
- |> update_syntax mode_default (TypeExt.type_ext read)
+ |> update_syntax mode_default TypeExt.type_ext
|> update_syntax mode_default SynExt.pure_ext;
val basic_nonterms =
--- a/src/Pure/Syntax/type_ext.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/Syntax/type_ext.ML Fri Mar 12 16:02:42 2010 +0100
@@ -15,6 +15,10 @@
val term_of_typ: bool -> typ -> term
val no_brackets: unit -> bool
val no_type_brackets: unit -> bool
+ val type_ast_trs:
+ {read_class: Proof.context -> string -> string,
+ read_type: Proof.context -> string -> string} ->
+ (string * (Proof.context -> Ast.ast list -> Ast.ast)) list
end;
signature TYPE_EXT =
@@ -23,9 +27,7 @@
val term_of_sort: sort -> term
val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
val sortT: typ
- val type_ext:
- {read_class: theory -> string -> string,
- read_type: theory -> string -> string} -> SynExt.syn_ext
+ val type_ext: SynExt.syn_ext
end;
structure TypeExt: TYPE_EXT =
@@ -240,7 +242,7 @@
local open Lexicon SynExt in
-fun type_ext {read_class, read_type} = syn_ext' false (K false)
+val type_ext = syn_ext' false (K false)
[Mfix ("_", tidT --> typeT, "", [], max_pri),
Mfix ("_", tvarT --> typeT, "", [], max_pri),
Mfix ("_", idT --> typeT, "_type_name", [], max_pri),
@@ -267,19 +269,18 @@
Mfix ("'(_')", typeT --> typeT, "", [0], max_pri),
Mfix ("'_", typeT, "\\<^type>dummy", [], max_pri)]
["_type_prop"]
- (map SynExt.mk_trfun
- [("_class_name", class_name_tr o read_class o ProofContext.theory_of),
- ("_classes", classes_tr o read_class o ProofContext.theory_of),
- ("_type_name", type_name_tr o read_type o ProofContext.theory_of),
- ("_tapp", tapp_ast_tr o read_type o ProofContext.theory_of),
- ("_tappl", tappl_ast_tr o read_type o ProofContext.theory_of),
- ("_bracket", K bracket_ast_tr)],
- [],
- [],
- map SynExt.mk_trfun [("\\<^type>fun", K fun_ast_tr')])
+ ([], [], [], map SynExt.mk_trfun [("\\<^type>fun", K fun_ast_tr')])
[]
([], []);
+fun type_ast_trs {read_class, read_type} =
+ [("_class_name", class_name_tr o read_class),
+ ("_classes", classes_tr o read_class),
+ ("_type_name", type_name_tr o read_type),
+ ("_tapp", tapp_ast_tr o read_type),
+ ("_tappl", tappl_ast_tr o read_type),
+ ("_bracket", K bracket_ast_tr)];
+
end;
end;
--- a/src/Pure/assumption.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/assumption.ML Fri Mar 12 16:02:42 2010 +0100
@@ -6,7 +6,7 @@
signature ASSUMPTION =
sig
- type export
+ type export = bool -> cterm list -> (thm -> thm) * (term -> term)
val assume_export: export
val presume_export: export
val assume: cterm -> thm
--- a/src/Pure/axclass.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/axclass.ML Fri Mar 12 16:02:42 2010 +0100
@@ -280,7 +280,7 @@
in (c1, c2) end;
fun read_classrel thy raw_rel =
- cert_classrel thy (pairself (Sign.read_class thy) raw_rel)
+ cert_classrel thy (pairself (ProofContext.read_class (ProofContext.init thy)) raw_rel)
handle TYPE (msg, _, _) => error msg;
@@ -387,7 +387,7 @@
fun prove_arity raw_arity tac thy =
let
val ctxt = ProofContext.init thy;
- val arity = Sign.cert_arity thy raw_arity;
+ val arity = ProofContext.cert_arity ctxt raw_arity;
val names = map (prefix arity_prefix) (Logic.name_arities arity);
val props = Logic.mk_arities arity;
val ths = Goal.prove_multi ctxt [] [] props
@@ -530,7 +530,7 @@
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;
fun ax_arity prep =
- axiomatize prep Logic.mk_arities
+ axiomatize (prep o ProofContext.init) Logic.mk_arities
(map (prefix arity_prefix) o Logic.name_arities) add_arity;
fun class_const c =
@@ -550,11 +550,11 @@
in
val axiomatize_class = ax_class Sign.certify_class cert_classrel;
-val axiomatize_class_cmd = ax_class Sign.read_class read_classrel;
+val axiomatize_class_cmd = ax_class (ProofContext.read_class o ProofContext.init) read_classrel;
val axiomatize_classrel = ax_classrel cert_classrel;
val axiomatize_classrel_cmd = ax_classrel read_classrel;
-val axiomatize_arity = ax_arity Sign.cert_arity;
-val axiomatize_arity_cmd = ax_arity Sign.read_arity;
+val axiomatize_arity = ax_arity ProofContext.cert_arity;
+val axiomatize_arity_cmd = ax_arity ProofContext.read_arity;
end;
--- a/src/Pure/consts.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/consts.ML Fri Mar 12 16:02:42 2010 +0100
@@ -19,6 +19,7 @@
val is_monomorphic: T -> string -> bool (*exception TYPE*)
val the_constraint: T -> string -> typ (*exception TYPE*)
val space_of: T -> Name_Space.T
+ val alias: Name_Space.naming -> binding -> string -> T -> T
val is_concealed: T -> string -> bool
val intern: T -> xstring -> string
val extern: T -> string -> xstring
@@ -122,6 +123,9 @@
fun space_of (Consts {decls = (space, _), ...}) = space;
+fun alias naming binding name = map_consts (fn ((space, decls), constraints, rev_abbrevs) =>
+ ((Name_Space.alias naming binding name space, decls), constraints, rev_abbrevs));
+
val is_concealed = Name_Space.is_concealed o space_of;
val intern = Name_Space.intern o space_of;
--- a/src/Pure/more_thm.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/more_thm.ML Fri Mar 12 16:02:42 2010 +0100
@@ -324,8 +324,7 @@
fun add_axiom (b, prop) thy =
let
- val b' = if Binding.is_empty b
- then Binding.name ("axiom_" ^ serial_string ()) else b;
+ val b' = if Binding.is_empty b then Binding.name ("axiom_" ^ serial_string ()) else b;
val thy' = thy |> Theory.add_axioms_i [(b', prop)];
val axm = unvarify (Thm.axiom thy' (Sign.full_name thy' b'));
in (axm, thy') end;
--- a/src/Pure/sign.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/sign.ML Fri Mar 12 16:02:42 2010 +0100
@@ -47,17 +47,16 @@
val class_space: theory -> Name_Space.T
val type_space: theory -> Name_Space.T
val const_space: theory -> Name_Space.T
+ val class_alias: binding -> class -> theory -> theory
+ val type_alias: binding -> string -> theory -> theory
+ val const_alias: binding -> string -> theory -> theory
val intern_class: theory -> xstring -> string
val extern_class: theory -> string -> xstring
val intern_type: theory -> xstring -> string
val extern_type: theory -> string -> xstring
val intern_const: theory -> xstring -> string
val extern_const: theory -> string -> xstring
- val intern_sort: theory -> sort -> sort
- val extern_sort: theory -> sort -> sort
- val intern_typ: theory -> typ -> typ
val intern_term: theory -> term -> term
- val the_type_decl: theory -> string -> Type.decl
val arity_number: theory -> string -> int
val arity_sorts: theory -> string -> sort -> sort list
val certify_class: theory -> class -> class
@@ -71,9 +70,6 @@
val no_frees: Pretty.pp -> term -> term
val no_vars: Pretty.pp -> term -> term
val cert_def: Proof.context -> term -> (string * typ) * term
- val read_class: theory -> xstring -> class
- val read_arity: theory -> xstring * string list * string -> arity
- val cert_arity: theory -> arity -> arity
val add_defsort: string -> theory -> theory
val add_defsort_i: sort -> theory -> theory
val add_types: (binding * int * mixfix) list -> theory -> theory
@@ -154,7 +150,7 @@
make_sign (Name_Space.default_naming, syn, tsig, consts);
val empty =
- make_sign (Name_Space.default_naming, Syntax.empty_syntax, Type.empty_tsig, Consts.empty);
+ make_sign (Name_Space.default_naming, Syntax.basic_syntax, Type.empty_tsig, Consts.empty);
fun merge pp (sign1, sign2) =
let
@@ -236,10 +232,14 @@
(** intern / extern names **)
-val class_space = #1 o #classes o Type.rep_tsig o tsig_of;
-val type_space = #1 o #types o Type.rep_tsig o tsig_of;
+val class_space = Type.class_space o tsig_of;
+val type_space = Type.type_space o tsig_of;
val const_space = Consts.space_of o consts_of;
+fun class_alias b c thy = map_tsig (Type.class_alias (naming_of thy) b c) thy;
+fun type_alias b c thy = map_tsig (Type.type_alias (naming_of thy) b c) thy;
+fun const_alias b c thy = map_consts (Consts.alias (naming_of thy) b c) thy;
+
val intern_class = Name_Space.intern o class_space;
val extern_class = Name_Space.extern o class_space;
val intern_type = Name_Space.intern o type_space;
@@ -247,9 +247,6 @@
val intern_const = Name_Space.intern o const_space;
val extern_const = Name_Space.extern o const_space;
-val intern_sort = map o intern_class;
-val extern_sort = map o extern_class;
-
local
fun map_typ f g (Type (c, Ts)) = Type (g c, map (map_typ f g) Ts)
@@ -265,7 +262,6 @@
in
-fun intern_typ thy = map_typ (intern_class thy) (intern_type thy);
fun intern_term thy = map_term (intern_class thy) (intern_type thy) (intern_const thy);
end;
@@ -276,7 +272,6 @@
(* certify wrt. type signature *)
-val the_type_decl = Type.the_decl o tsig_of;
val arity_number = Type.arity_number o tsig_of;
fun arity_sorts thy = Type.arity_sorts (Syntax.pp_global thy) (tsig_of thy);
@@ -367,51 +362,6 @@
-(** read and certify entities **) (*exception ERROR*)
-
-(* classes *)
-
-fun read_class thy text =
- let
- val (syms, pos) = Syntax.read_token text;
- val c = certify_class thy (intern_class thy (Symbol_Pos.content syms))
- handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos);
- val _ = Position.report (Markup.tclass c) pos;
- in c end;
-
-
-(* type arities *)
-
-fun prep_arity prep_tycon prep_sort thy (t, Ss, S) =
- let val arity = (prep_tycon thy t, map (prep_sort thy) Ss, prep_sort thy S)
- in Type.add_arity (Syntax.pp_global thy) arity (tsig_of thy); arity end;
-
-val read_arity = prep_arity intern_type Syntax.read_sort_global;
-val cert_arity = prep_arity (K I) certify_sort;
-
-
-(* type syntax entities *)
-
-local
-
-fun read_type thy text =
- let
- val (syms, pos) = Syntax.read_token text;
- val c = intern_type thy (Symbol_Pos.content syms);
- val _ = the_type_decl thy c;
- val _ = Position.report (Markup.tycon c) pos;
- in c end;
-
-in
-
-val _ = Context.>>
- (Context.map_theory
- (map_syn (K (Syntax.basic_syntax {read_class = read_class, read_type = read_type}))));
-
-end;
-
-
-
(** signature extension functions **) (*exception ERROR/TYPE*)
(* add default sort *)
--- a/src/Pure/type.ML Tue Mar 09 15:42:23 2010 +0100
+++ b/src/Pure/type.ML Fri Mar 12 16:02:42 2010 +0100
@@ -13,12 +13,17 @@
Abbreviation of string list * typ * bool |
Nonterminal
type tsig
+ val eq_tsig: tsig * tsig -> bool
val rep_tsig: tsig ->
{classes: Name_Space.T * Sorts.algebra,
default: sort,
types: decl Name_Space.table,
log_types: string list}
val empty_tsig: tsig
+ val class_space: tsig -> Name_Space.T
+ val class_alias: Name_Space.naming -> binding -> string -> tsig -> tsig
+ val intern_class: tsig -> xstring -> string
+ val extern_class: tsig -> string -> xstring
val defaultS: tsig -> sort
val logical_types: tsig -> string list
val eq_sort: tsig -> sort * sort -> bool
@@ -35,6 +40,11 @@
val get_mode: Proof.context -> mode
val set_mode: mode -> Proof.context -> Proof.context
val restore_mode: Proof.context -> Proof.context -> Proof.context
+ val type_space: tsig -> Name_Space.T
+ val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig
+ val intern_type: tsig -> xstring -> string
+ val extern_type: tsig -> string -> xstring
+ val is_logtype: tsig -> string -> bool
val the_decl: tsig -> string -> decl
val cert_typ_mode: mode -> tsig -> typ -> typ
val cert_typ: tsig -> typ -> typ
@@ -103,6 +113,13 @@
types: decl Name_Space.table, (*declared types*)
log_types: string list}; (*logical types sorted by number of arguments*)
+fun eq_tsig
+ (TSig {classes = classes1, default = default1, types = types1, log_types = _},
+ TSig {classes = classes2, default = default2, types = types2, log_types = _}) =
+ pointer_eq (classes1, classes2) andalso
+ default1 = default2 andalso
+ pointer_eq (types1, types2);
+
fun rep_tsig (TSig comps) = comps;
fun make_tsig (classes, default, types, log_types) =
@@ -124,6 +141,14 @@
(* classes and sorts *)
+val class_space = #1 o #classes o rep_tsig;
+
+fun class_alias naming binding name = map_tsig (fn ((space, classes), default, types) =>
+ ((Name_Space.alias naming binding name space, classes), default, types));
+
+val intern_class = Name_Space.intern o class_space;
+val extern_class = Name_Space.extern o class_space;
+
fun defaultS (TSig {default, ...}) = default;
fun logical_types (TSig {log_types, ...}) = log_types;
@@ -158,7 +183,18 @@
fun restore_mode ctxt = set_mode (get_mode ctxt);
-(* lookup types *)
+(* types *)
+
+val type_space = #1 o #types o rep_tsig;
+
+fun type_alias naming binding name = map_tsig (fn (classes, default, (space, types)) =>
+ (classes, default, (Name_Space.alias naming binding name space, types)));
+
+val intern_type = Name_Space.intern o type_space;
+val extern_type = Name_Space.extern o type_space;
+
+val is_logtype = member (op =) o logical_types;
+
fun undecl_type c = "Undeclared type constructor: " ^ quote c;