--- a/Admin/MacOS/App2/Isabelle.app/Contents/Info.plist Sun May 23 10:37:43 2010 +0100
+++ b/Admin/MacOS/App2/Isabelle.app/Contents/Info.plist Sun May 23 10:38:11 2010 +0100
@@ -19,17 +19,17 @@
<key>CFBundlePackageType</key>
<string>APPL</string>
<key>CFBundleShortVersionString</key>
- <string>2009</string>
+ <string>????</string>
<key>CFBundleSignature</key>
<string>????</string>
<key>CFBundleVersion</key>
- <string>2009</string>
+ <string>????</string>
<key>Java</key>
<dict>
<key>JVMVersion</key>
- <string>1.5+</string>
+ <string>1.6</string>
<key>VMOptions</key>
- <string>-Xmx384M</string>
+ <string>-Xms128m -Xmx512m -Xss2m</string>
<key>ClassPath</key>
<string>$JAVAROOT/isabelle-scala.jar</string>
<key>MainClass</key>
@@ -40,6 +40,8 @@
<string>$APP_PACKAGE/Contents/Resources/Isabelle</string>
<key>apple.laf.useScreenMenuBar</key>
<string>true</string>
+ <key>com.apple.mrj.application.apple.menu.about.name</key>
+ <string>Isabelle</string>
</dict>
</dict>
</dict>
--- a/NEWS Sun May 23 10:37:43 2010 +0100
+++ b/NEWS Sun May 23 10:38:11 2010 +0100
@@ -143,6 +143,13 @@
*** HOL ***
+* List membership infix mem operation is only an input abbreviation.
+INCOMPATIBILITY.
+
+* Theory Library/Word.thy has been removed. Use library Word/Word.thy for
+future developements; former Library/Word.thy is still present in the AFP
+entry RSAPPS.
+
* Theorem Int.int_induct renamed to Int.int_of_nat_induct and is no
longer shadowed. INCOMPATIBILITY.
@@ -155,6 +162,41 @@
* Dropped normalizing_semiring etc; use the facts in semiring classes
instead. INCOMPATIBILITY.
+* Dropped several real-specific versions of lemmas about floor and
+ceiling; use the generic lemmas from Archimedean_Field.thy instead.
+INCOMPATIBILITY.
+
+ floor_number_of_eq ~> floor_number_of
+ le_floor_eq_number_of ~> number_of_le_floor
+ le_floor_eq_zero ~> zero_le_floor
+ le_floor_eq_one ~> one_le_floor
+ floor_less_eq_number_of ~> floor_less_number_of
+ floor_less_eq_zero ~> floor_less_zero
+ floor_less_eq_one ~> floor_less_one
+ less_floor_eq_number_of ~> number_of_less_floor
+ less_floor_eq_zero ~> zero_less_floor
+ less_floor_eq_one ~> one_less_floor
+ floor_le_eq_number_of ~> floor_le_number_of
+ floor_le_eq_zero ~> floor_le_zero
+ floor_le_eq_one ~> floor_le_one
+ floor_subtract_number_of ~> floor_diff_number_of
+ floor_subtract_one ~> floor_diff_one
+ ceiling_number_of_eq ~> ceiling_number_of
+ ceiling_le_eq_number_of ~> ceiling_le_number_of
+ ceiling_le_zero_eq ~> ceiling_le_zero
+ ceiling_le_eq_one ~> ceiling_le_one
+ less_ceiling_eq_number_of ~> number_of_less_ceiling
+ less_ceiling_eq_zero ~> zero_less_ceiling
+ less_ceiling_eq_one ~> one_less_ceiling
+ ceiling_less_eq_number_of ~> ceiling_less_number_of
+ ceiling_less_eq_zero ~> ceiling_less_zero
+ ceiling_less_eq_one ~> ceiling_less_one
+ le_ceiling_eq_number_of ~> number_of_le_ceiling
+ le_ceiling_eq_zero ~> zero_le_ceiling
+ le_ceiling_eq_one ~> one_le_ceiling
+ ceiling_subtract_number_of ~> ceiling_diff_number_of
+ ceiling_subtract_one ~> ceiling_diff_one
+
* Theory 'Finite_Set': various folding_XXX locales facilitate the
application of the various fold combinators on finite sets.
@@ -347,6 +389,49 @@
Nat_Int_Bij.bij_nat_to_int_bij ~> bij_int_encode
Nat_Int_Bij.bij_int_to_nat_bij ~> bij_int_decode
+* Sledgehammer:
+ - Renamed ATP commands:
+ atp_info ~> sledgehammer running_atps
+ atp_kill ~> sledgehammer kill_atps
+ atp_messages ~> sledgehammer messages
+ atp_minimize ~> sledgehammer minimize
+ print_atps ~> sledgehammer available_atps
+ INCOMPATIBILITY.
+ - Added user's manual ("isabelle doc sledgehammer").
+ - Added option syntax and "sledgehammer_params" to customize
+ Sledgehammer's behavior. See the manual for details.
+ - Modified the Isar proof reconstruction code so that it produces
+ direct proofs rather than proofs by contradiction. (This feature
+ is still experimental.)
+ - Made Isar proof reconstruction work for SPASS, remote ATPs, and in
+ full-typed mode.
+ - Added support for TPTP syntax for SPASS via the "spass_tptp" ATP.
+
+* Nitpick:
+ - 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 "specification" and "ax_specification"
+ constructs.
+ - 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.
+ - Fixed soundness bugs related to "destroy_constrs" optimization and
+ record getters.
+ - Fixed soundness bug related to higher-order constructors
+ - Improved precision of set constructs.
+ - Added cache to speed up repeated Kodkod invocations on the same
+ problems.
+ - Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and
+ "SAT4JLight" to "MiniSat_JNI", "zChaff_JNI", "BerkMin_Alloy", and
+ "SAT4J_Light". INCOMPATIBILITY.
+ - Removed "skolemize", "uncurry", "sym_break", "flatten_prop",
+ "sharing_depth", and "show_skolems" options. INCOMPATIBILITY.
+
*** HOLCF ***
@@ -391,6 +476,10 @@
foo_unfold ~> foo.unfold
foo_induct ~> foo.induct
+* The "fixrec_simp" attribute has been removed. The "fixrec_simp"
+method and internal fixrec proofs now use the default simpset instead.
+INCOMPATIBILITY.
+
* The "contlub" predicate has been removed. Proof scripts should use
lemma contI2 in place of monocontlub2cont, INCOMPATIBILITY.
@@ -459,6 +548,18 @@
context actually works, but under normal circumstances one needs to
pass the proper local context through the code!
+* Renamed some important ML structures, while keeping the old names
+for some time as aliases within the structure Legacy:
+
+ OuterKeyword ~> Keyword
+ OuterLex ~> Token
+ OuterParse ~> Parse
+ OuterSyntax ~> Outer_Syntax
+ SpecParse ~> Parse_Spec
+
+Note that "open Legacy" simplifies porting of sources, but forgetting
+to remove it again will complicate porting again in the future.
+
*** System ***
--- a/doc-src/Dirs Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/Dirs Sun May 23 10:38:11 2010 +0100
@@ -1,1 +1,1 @@
-Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions Nitpick Main
+Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions Nitpick Main Sledgehammer
--- a/doc-src/Makefile.in Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/Makefile.in Sun May 23 10:38:11 2010 +0100
@@ -48,6 +48,9 @@
isabelle_nitpick.eps:
test -r isabelle_nitpick.eps || ln -s ../gfx/isabelle_nitpick.eps .
+isabelle_sledgehammer.eps:
+ test -r isabelle_sledgehammer.eps || ln -s ../gfx/isabelle_sledgehammer.eps .
+
isabelle.pdf:
test -r isabelle.pdf || ln -s ../gfx/isabelle.pdf .
@@ -64,6 +67,9 @@
isabelle_nitpick.pdf:
test -r isabelle_nitpick.pdf || ln -s ../gfx/isabelle_nitpick.pdf .
+isabelle_sledgehammer.pdf:
+ test -r isabelle_sledgehammer.pdf || ln -s ../gfx/isabelle_sledgehammer.pdf .
+
typedef.ps:
test -r typedef.ps || ln -s ../gfx/typedef.ps .
--- a/doc-src/Nitpick/nitpick.tex Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/Nitpick/nitpick.tex Sun May 23 10:38:11 2010 +0100
@@ -81,7 +81,7 @@
\section{Introduction}
\label{introduction}
-Nitpick \cite{blanchette-nipkow-2009} is a counterexample generator for
+Nitpick \cite{blanchette-nipkow-2010} is a counterexample generator for
Isabelle/HOL \cite{isa-tutorial} that is designed to handle formulas
combining (co)in\-duc\-tive datatypes, (co)in\-duc\-tively defined predicates, and
quantifiers. It builds on Kodkod \cite{torlak-jackson-2007}, a highly optimized
@@ -111,6 +111,11 @@
must find a model for the axioms. If it finds no model, we have an indication
that the axioms might be unsatisfiable.
+You can also invoke Nitpick from the ``Commands'' submenu of the
+``Isabelle'' menu in Proof General or by pressing the Emacs key sequence C-c C-a
+C-n. This is equivalent to entering the \textbf{nitpick} command with no
+arguments in the theory text.
+
Nitpick requires the Kodkodi package for Isabelle as well as a Java 1.5 virtual
machine called \texttt{java}. The examples presented in this manual can be found
in Isabelle's \texttt{src/HOL/Nitpick\_Examples/Manual\_Nits.thy} theory.
@@ -137,6 +142,22 @@
suggesting several textual improvements.
% and Perry James for reporting a typo.
+%\section{Installation}
+%\label{installation}
+%
+%MISSING:
+%
+% * Nitpick is part of Isabelle/HOL
+% * but it relies on an external tool called Kodkodi (Kodkod wrapper)
+% * Two options:
+% * if you use a prebuilt Isabelle package, Kodkodi is automatically there
+% * if you work from sources, the latest Kodkodi can be obtained from ...
+% download it, install it in some directory of your choice (e.g.,
+% $ISABELLE_HOME/contrib/kodkodi), and add the absolute path to Kodkodi
+% in your .isabelle/etc/components file
+%
+% * If you're not sure, just try the example in the next section
+
\section{First Steps}
\label{first-steps}
@@ -184,6 +205,9 @@
\hbox{}\qquad\qquad $Q = \textit{False}$
\postw
+%FIXME: If you get the output:...
+%Then do such-and-such.
+
Nitpick can also be invoked on individual subgoals, as in the example below:
\prew
@@ -2017,17 +2041,17 @@
well-founded. The option can take the following values:
\begin{enum}
-\item[$\bullet$] \textbf{\textit{true}}: Tentatively treat the (co)in\-duc\-tive
+\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 ``quasi
genuine.''
-\item[$\bullet$] \textbf{\textit{false}}: Treat the (co)in\-duc\-tive predicate
+\item[$\bullet$] \textbf{\textit{false}:} Treat the (co)in\-duc\-tive predicate
as if it were not well-founded. The predicate is then unrolled as prescribed by
the \textit{star\_linear\_preds}, \textit{iter}~\qty{const}, and \textit{iter}
options.
-\item[$\bullet$] \textbf{\textit{smart}}: Try to prove that the inductive
+\item[$\bullet$] \textbf{\textit{smart}:} Try to prove that the inductive
predicate is well-founded using Isabelle's \textit{lexicographic\_order} and
\textit{size\_change} tactics. If this succeeds (or the predicate occurs with an
appropriate polarity in the formula to falsify), use an efficient fixed point
@@ -2080,10 +2104,10 @@
counterexamples. The option can take the following values:
\begin{enum}
-\item[$\bullet$] \textbf{\textit{true}}: Box the specified type whenever
+\item[$\bullet$] \textbf{\textit{true}:} Box the specified type whenever
practicable.
-\item[$\bullet$] \textbf{\textit{false}}: Never box the type.
-\item[$\bullet$] \textbf{\textit{smart}}: Box the type only in contexts where it
+\item[$\bullet$] \textbf{\textit{false}:} Never box the type.
+\item[$\bullet$] \textbf{\textit{smart}:} Box the type only in contexts where it
is likely to help. For example, $n$-tuples where $n > 2$ and arguments to
higher-order functions are good candidates for boxing.
\end{enum}
@@ -2107,8 +2131,8 @@
then take the following values:
\begin{enum}
-\item[$\bullet$] \textbf{\textit{false}}: Don't attempt to finitize the type.
-\item[$\bullet$] \textbf{\textit{true}} or \textbf{\textit{smart}}: Finitize the
+\item[$\bullet$] \textbf{\textit{false}:} Don't attempt to finitize the type.
+\item[$\bullet$] \textbf{\textit{true}} or \textbf{\textit{smart}:} Finitize the
type wherever possible.
\end{enum}
@@ -2116,11 +2140,11 @@
datatypes:
\begin{enum}
-\item[$\bullet$] \textbf{\textit{true}}: Finitize the datatype. Since this is
+\item[$\bullet$] \textbf{\textit{true}:} Finitize the datatype. Since this is
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
+\item[$\bullet$] \textbf{\textit{false}:} Don't attempt to finitize the datatype.
+\item[$\bullet$] \textbf{\textit{smart}:} Perform a monotonicity analysis to
detect whether the datatype can be safely finitized before finitizing it.
\end{enum}
@@ -2282,13 +2306,13 @@
Specifies the expected outcome, which must be one of the following:
\begin{enum}
-\item[$\bullet$] \textbf{\textit{genuine}}: Nitpick found a genuine counterexample.
-\item[$\bullet$] \textbf{\textit{quasi\_genuine}}: Nitpick found a ``quasi
+\item[$\bullet$] \textbf{\textit{genuine}:} Nitpick found a genuine counterexample.
+\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.
-\item[$\bullet$] \textbf{\textit{none}}: Nitpick found no counterexample.
-\item[$\bullet$] \textbf{\textit{unknown}}: Nitpick encountered some problem (e.g.,
+\item[$\bullet$] \textbf{\textit{potential}:} Nitpick found a potential counterexample.
+\item[$\bullet$] \textbf{\textit{none}:} Nitpick found no counterexample.
+\item[$\bullet$] \textbf{\textit{unknown}:} Nitpick encountered some problem (e.g.,
Kodkod ran out of memory).
\end{enum}
@@ -2314,11 +2338,9 @@
The supported solvers are listed below:
-\let\foo
-
\begin{enum}
-\item[$\bullet$] \textbf{\textit{MiniSat}}: MiniSat is an efficient solver
+\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.%
@@ -2329,13 +2351,13 @@
\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)
+\item[$\bullet$] \textbf{\textit{MiniSat\_JNI}:} The JNI (Java Native Interface)
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.
-\item[$\bullet$] \textbf{\textit{PicoSAT}}: PicoSAT is an efficient solver
+\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.%
@@ -2344,7 +2366,7 @@
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
+\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.%
\footref{cygwin-paths}
@@ -2352,12 +2374,12 @@
\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
+\item[$\bullet$] \textbf{\textit{zChaff\_JNI}:} The JNI version of zChaff 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}.
-\item[$\bullet$] \textbf{\textit{RSat}}: RSat is an efficient solver written in
+\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.%
\footref{cygwin-paths}
@@ -2365,21 +2387,21 @@
\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
+\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.\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
+\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.%
\footref{cygwin-paths}
-\item[$\bullet$] \textbf{\textit{Jerusat}}: Jerusat 1.3 is an efficient solver
+\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.%
@@ -2387,24 +2409,16 @@
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
+\item[$\bullet$] \textbf{\textit{SAT4J}:} SAT4J is a reasonably efficient solver
written in Java that can be used incrementally. It is bundled with Kodkodi and
requires no further installation or configuration steps. Do not attempt to
install the official SAT4J packages, because their API is incompatible with
Kodkod.
-\item[$\bullet$] \textbf{\textit{SAT4J\_Light}}: Variant of SAT4J that is
+\item[$\bullet$] \textbf{\textit{SAT4J\_Light}:} Variant of SAT4J that is
optimized for small problems. It can also be used incrementally.
-\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.%
-\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
+\item[$\bullet$] \textbf{\textit{smart}:} If \textit{sat\_solver} is set to
\textit{smart}, Nitpick selects the first solver among MiniSat,
PicoSAT, zChaff, RSat, BerkMin, BerkMin\_Alloy, Jerusat, MiniSat\_JNI, and zChaff\_JNI
that is recognized by Isabelle. If none is found, it falls back on SAT4J, which
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Sledgehammer/Makefile Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,36 @@
+#
+# $Id$
+#
+
+## targets
+
+default: dvi
+
+
+## dependencies
+
+include ../Makefile.in
+
+NAME = sledgehammer
+FILES = sledgehammer.tex ../iman.sty ../manual.bib
+
+dvi: $(NAME).dvi
+
+$(NAME).dvi: $(FILES) isabelle_sledgehammer.eps
+ $(LATEX) $(NAME)
+ $(BIBTEX) $(NAME)
+ $(LATEX) $(NAME)
+ $(LATEX) $(NAME)
+ $(SEDINDEX) $(NAME)
+ $(LATEX) $(NAME)
+
+pdf: $(NAME).pdf
+
+$(NAME).pdf: $(FILES) isabelle_sledgehammer.pdf
+ $(PDFLATEX) $(NAME)
+ $(BIBTEX) $(NAME)
+ $(PDFLATEX) $(NAME)
+ $(PDFLATEX) $(NAME)
+ $(SEDINDEX) $(NAME)
+ $(FIXBOOKMARKS) $(NAME).out
+ $(PDFLATEX) $(NAME)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Sledgehammer/sledgehammer.tex Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,483 @@
+\documentclass[a4paper,12pt]{article}
+\usepackage[T1]{fontenc}
+\usepackage{amsmath}
+\usepackage{amssymb}
+\usepackage[english,french]{babel}
+\usepackage{color}
+\usepackage{footmisc}
+\usepackage{graphicx}
+%\usepackage{mathpazo}
+\usepackage{multicol}
+\usepackage{stmaryrd}
+%\usepackage[scaled=.85]{beramono}
+\usepackage{../iman,../pdfsetup}
+
+%\oddsidemargin=4.6mm
+%\evensidemargin=4.6mm
+%\textwidth=150mm
+%\topmargin=4.6mm
+%\headheight=0mm
+%\headsep=0mm
+%\textheight=234mm
+
+\def\Colon{\mathord{:\mkern-1.5mu:}}
+%\def\lbrakk{\mathopen{\lbrack\mkern-3.25mu\lbrack}}
+%\def\rbrakk{\mathclose{\rbrack\mkern-3.255mu\rbrack}}
+\def\lparr{\mathopen{(\mkern-4mu\mid}}
+\def\rparr{\mathclose{\mid\mkern-4mu)}}
+
+\def\unk{{?}}
+\def\undef{(\lambda x.\; \unk)}
+%\def\unr{\textit{others}}
+\def\unr{\ldots}
+\def\Abs#1{\hbox{\rm{\flqq}}{\,#1\,}\hbox{\rm{\frqq}}}
+\def\Q{{\smash{\lower.2ex\hbox{$\scriptstyle?$}}}}
+
+\urlstyle{tt}
+
+\begin{document}
+
+\selectlanguage{english}
+
+\title{\includegraphics[scale=0.5]{isabelle_sledgehammer} \\[4ex]
+Hammering Away \\[\smallskipamount]
+\Large A User's Guide to Sledgehammer for Isabelle/HOL}
+\author{\hbox{} \\
+Jasmin Christian Blanchette \\
+{\normalsize Institut f\"ur Informatik, Technische Universit\"at M\"unchen} \\
+\hbox{}}
+
+\maketitle
+
+\tableofcontents
+
+\setlength{\parskip}{.7em plus .2em minus .1em}
+\setlength{\parindent}{0pt}
+\setlength{\abovedisplayskip}{\parskip}
+\setlength{\abovedisplayshortskip}{.9\parskip}
+\setlength{\belowdisplayskip}{\parskip}
+\setlength{\belowdisplayshortskip}{.9\parskip}
+
+% General-purpose enum environment with correct spacing
+\newenvironment{enum}%
+ {\begin{list}{}{%
+ \setlength{\topsep}{.1\parskip}%
+ \setlength{\partopsep}{.1\parskip}%
+ \setlength{\itemsep}{\parskip}%
+ \advance\itemsep by-\parsep}}
+ {\end{list}}
+
+\def\pre{\begingroup\vskip0pt plus1ex\advance\leftskip by\leftmargin
+\advance\rightskip by\leftmargin}
+\def\post{\vskip0pt plus1ex\endgroup}
+
+\def\prew{\pre\advance\rightskip by-\leftmargin}
+\def\postw{\post}
+
+\section{Introduction}
+\label{introduction}
+
+Sledgehammer is a tool that applies first-order automatic theorem provers (ATPs)
+on the current goal. The supported ATPs are E \cite{schulz-2002}, SPASS
+\cite{weidenbach-et-al-2009}, and Vampire \cite{riazanov-voronkov-2002}, which
+can be run locally or remotely via the SystemOnTPTP web service
+\cite{sutcliffe-2000}.
+
+The problem passed to ATPs consists of the current goal together with a
+heuristic selection of facts (theorems) from the current theory context,
+filtered by relevance. The result of a successful ATP proof search is some
+source text that usually (but not always) reconstructs the proof within
+Isabelle, without requiring the ATPs again. The reconstructed proof relies on
+the general-purpose Metis prover \cite{metis}, which is fully integrated into
+Isabelle/HOL, with explicit inferences going through the kernel. Thus its
+results are correct by construction.
+
+\newbox\boxA
+\setbox\boxA=\hbox{\texttt{nospam}}
+
+Examples of Sledgehammer use can be found in Isabelle's
+\texttt{src/HOL/Metis\_Examples} directory.
+Comments and bug reports concerning Sledgehammer or this manual should be
+directed to
+\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak
+in.\allowbreak tum.\allowbreak de}.
+
+\vskip2.5\smallskipamount
+
+%\textbf{Acknowledgment.} The author would like to thank Mark Summerfield for
+%suggesting several textual improvements.
+
+\section{Installation}
+\label{installation}
+
+Sledgehammer is part of Isabelle, so you don't need to install it. However, it
+relies on third-party automatic theorem provers (ATPs). Currently, E, SPASS, and
+Vampire are supported. All of these are available remotely via SystemOnTPTP
+\cite{sutcliffe-2000}, but if you want better performance you will need to
+install at least E and SPASS locally.
+
+There are three main ways to install E and SPASS:
+
+\begin{enum}
+\item[$\bullet$] If you installed an official Isabelle package with everything
+inside, it should already include properly setup executables for E and SPASS,
+ready to use.
+
+\item[$\bullet$] Otherwise, you can download the Isabelle-aware E and SPASS
+binary packages from Isabelle's download page. Extract the archives, then add a
+line to your \texttt{\char`\~/.isabelle/etc/components} file with the absolute path to
+E or SPASS. For example, if \texttt{\char`\~/.isabelle/etc/components} does not exist
+yet and you extracted SPASS to \texttt{/usr/local/spass-3.7}, create
+the file \texttt{\char`\~/.isabelle/etc/components} with the single line
+
+\prew
+\texttt{/usr/local/spass-3.7}
+\postw
+
+\item[$\bullet$] If you prefer to build E or SPASS yourself, feel free to do so
+and set the environment variable \texttt{E\_HOME} or \texttt{SPASS\_HOME} to the
+directory that contains the \texttt{eproof} or \texttt{SPASS} executable,
+respectively.
+\end{enum}
+
+To check whether E and SPASS are installed, follow the example in
+\S\ref{first-steps}.
+
+\section{First Steps}
+\label{first-steps}
+
+To illustrate Sledgehammer in context, let us start a theory file and
+attempt to prove a simple lemma:
+
+\prew
+\textbf{theory}~\textit{Scratch} \\
+\textbf{imports}~\textit{Main} \\
+\textbf{begin} \\[2\smallskipamount]
+%
+\textbf{lemma} ``$[a] = [b] \,\longleftrightarrow\, a = b$'' \\
+\textbf{sledgehammer}
+\postw
+
+After a few seconds, Sledgehammer produces the following output:
+
+\prew
+\slshape
+Sledgehammer: ATP ``\textit{e}'' for subgoal 1: \\
+$([a] = [b]) = (a = b)$ \\
+Try this command: \textbf{by} (\textit{metis hd.simps}). \\
+To minimize the number of lemmas, try this command: \\
+\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{e}] (\textit{hd.simps}). \\[3\smallskipamount]
+%
+Sledgehammer: ATP ``\textit{spass}'' for subgoal 1: \\
+$([a] = [b]) = (a = b)$ \\
+Try this command: \textbf{by} (\textit{metis insert\_Nil last\_ConsL}). \\
+To minimize the number of lemmas, try this command: \\
+\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{spass}] (\textit{insert\_Nil last\_ConsL}). \\[3\smallskipamount]
+%
+Sledgehammer: ATP ``\textit{remote\_vampire}'' for subgoal 1: \\
+$([a] = [b]) = (a = b)$ \\
+Try this command: \textbf{by} (\textit{metis One\_nat\_def\_raw empty\_replicate} \\
+\phantom{Try this command: \textbf{by} (\textit{metis~}}\textit{insert\_Nil last\_ConsL replicate\_Suc}). \\
+To minimize the number of lemmas, try this command: \\
+\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{remote\_vampire}] \\
+\phantom{\textbf{sledgehammer}~}(\textit{One\_nat\_def\_raw empty\_replicate insert\_Nil} \\
+\phantom{\textbf{sledgehammer}~(}\textit{last\_ConsL replicate\_Suc}).
+\postw
+
+Sledgehammer ran E, SPASS, and the remote version of Vampire in parallel. If E
+and SPASS are not installed (\S\ref{installation}), you will see references to
+their remote American cousins \textit{remote\_e} and \textit{remote\_spass}
+instead of \textit{e} and \textit{spass}.
+
+Based on each ATP proof, Sledgehammer gives a one-liner proof that uses the
+\textit{metis} method. You can click them and insert them into the theory text.
+You can click the ``\textbf{sledgehammer} \textit{minimize}'' command if you
+want to look for a shorter (and faster) proof. But here the proof found by E
+looks perfect, so click it to finish the proof.
+
+You can ask Sledgehammer for an Isar text proof by passing the
+\textit{isar\_proof} option:
+
+\prew
+\textbf{sledgehammer} [\textit{isar\_proof}]
+\postw
+
+When Isar proof construction is successful, it can yield proofs that are more
+readable and also faster than the \textit{metis} one-liners. This feature is
+experimental.
+
+\section{Command Syntax}
+\label{command-syntax}
+
+Sledgehammer can be invoked at any point when there is an open goal by entering
+the \textbf{sledgehammer} command in the theory file. Its general syntax is as
+follows:
+
+\prew
+\textbf{sledgehammer} \textit{subcommand\/$^?$ options\/$^?$ facts\_override\/$^?$ num\/$^?$}
+\postw
+
+For convenience, Sledgehammer is also available in the ``Commands'' submenu of
+the ``Isabelle'' menu in Proof General or by pressing the Emacs key sequence C-c
+C-a C-s. This is equivalent to entering the \textbf{sledgehammer} command with
+no arguments in the theory text.
+
+In the general syntax, the \textit{subcommand} may be any of the following:
+
+\begin{enum}
+\item[$\bullet$] \textbf{\textit{run} (the default):} Runs Sledgehammer on subgoal number
+\textit{num} (1 by default), with the given options and facts.
+
+\item[$\bullet$] \textbf{\textit{minimize}:} Attempts to minimize the provided facts
+(specified in the \textit{facts\_override} argument) to obtain a simpler proof
+involving fewer facts. The options and goal number are as for \textit{run}.
+
+\item[$\bullet$] \textbf{\textit{messages}:} Redisplays recent messages issued by
+Sledgehammer. This allows you to examine results that might have been lost due
+to Sledgehammer's asynchronous nature. The \textit{num} argument specifies a
+limit on the number of messages to display (5 by default).
+
+\item[$\bullet$] \textbf{\textit{available\_atps}:} Prints the list of installed ATPs.
+See \S\ref{installation} and \S\ref{mode-of-operation} for more information on
+how to install ATPs.
+
+\item[$\bullet$] \textbf{\textit{running\_atps}:} Prints information about currently
+running ATPs, including elapsed runtime and remaining time until timeout.
+
+\item[$\bullet$] \textbf{\textit{kill\_atps}:} Terminates all running ATPs.
+
+\item[$\bullet$] \textbf{\textit{refresh\_tptp}:} Refreshes the list of remote
+ATPs available at System\-On\-TPTP \cite{sutcliffe-2000}.
+\end{enum}
+
+Sledgehammer's behavior can be influenced by various \textit{options}, which can
+be specified in brackets after the \textbf{sledgehammer} command. The
+\textit{options} are a list of key--value pairs of the form ``[$k_1 = v_1,
+\ldots, k_n = v_n$]''. For Boolean options, ``= \textit{true}'' is optional. For
+example:
+
+\prew
+\textbf{sledgehammer} [\textit{isar\_proof}, \,\textit{timeout} = 120$\,s$]
+\postw
+
+Default values can be set using \textbf{sledgehammer\_\allowbreak params}:
+
+\prew
+\textbf{sledgehammer\_params} \textit{options}
+\postw
+
+The supported options are described in \S\ref{option-reference}.
+
+The \textit{facts\_override} argument lets you alter the set of facts that go
+through the relevance filter. It may be of the form ``(\textit{facts})'', where
+\textit{facts} is a space-separated list of Isabelle facts (theorems, local
+assumptions, etc.), in which case the relevance filter is bypassed and the given
+facts are used. It may also be of the form (\textit{add}:\ \textit{facts}$_1$),
+(\textit{del}:\ \textit{facts}$_2$), or (\textit{add}:\ \textit{facts}$_1$\
+\textit{del}:\ \textit{facts}$_2$), where the relevance filter is instructed to
+proceed as usual except that it should consider \textit{facts}$_1$
+highly-relevant and \textit{facts}$_2$ fully irrelevant.
+
+\section{Option Reference}
+\label{option-reference}
+
+\def\flushitem#1{\item[]\noindent\kern-\leftmargin \textbf{#1}}
+\def\qty#1{$\left<\textit{#1}\right>$}
+\def\qtybf#1{$\mathbf{\left<\textbf{\textit{#1}}\right>}$}
+\def\optrue#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\quad [\textit{true}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
+\def\opfalse#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\quad [\textit{false}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
+\def\opsmart#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\quad [\textit{smart}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
+\def\opsmartx#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\quad [\textit{smart}]\hfill\\\hbox{}\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
+\def\opnodefault#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]}
+\def\opdefault#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\quad [\textit{#3}]} \nopagebreak\\[\parskip]}
+\def\oparg#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]}
+\def\opargbool#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]}
+\def\opargboolorsmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]}
+
+Sledgehammer's options are categorized as follows:\ mode of operation
+(\S\ref{mode-of-operation}), problem encoding (\S\ref{problem-encoding}), output
+format (\S\ref{output-format}), and timeouts (\S\ref{timeouts}).
+
+The descriptions below refer to the following syntactic quantities:
+
+\begin{enum}
+\item[$\bullet$] \qtybf{string}: A string.
+\item[$\bullet$] \qtybf{bool\/}: \textit{true} or \textit{false}.
+\item[$\bullet$] \qtybf{bool\_or\_smart\/}: \textit{true}, \textit{false}, or \textit{smart}.
+\item[$\bullet$] \qtybf{int\/}: An integer.
+\item[$\bullet$] \qtybf{time}: An integer followed by $\textit{min}$ (minutes), $s$ (seconds), or \textit{ms}
+(milliseconds), or the keyword \textit{none} ($\infty$ years).
+\end{enum}
+
+Default values are indicated in square brackets. Boolean options have a negated
+counterpart (e.g., \textit{debug} vs.\ \textit{no\_debug}). When setting
+Boolean options, ``= \textit{true}'' may be omitted.
+
+\subsection{Mode of Operation}
+\label{mode-of-operation}
+
+\begin{enum}
+%\optrue{blocking}{non\_blocking}
+%Specifies whether the \textbf{sledgehammer} command should operate synchronously.
+%The asynchronous (non-blocking) mode lets the user start proving the putative
+%theorem while Sledgehammer looks for a counterexample, but it can also be more
+%confusing. For technical reasons, automatic runs currently always block.
+
+\opnodefault{atps}{string}
+Specifies the ATPs (automated theorem provers) to use as a space-separated list
+(e.g., ``\textit{e}~\textit{spass}''). The following ATPs are supported:
+
+\begin{enum}
+\item[$\bullet$] \textbf{\textit{e}:} E is an ATP developed by Stephan Schulz
+\cite{schulz-2002}. To use E, set the environment variable
+\texttt{E\_HOME} to the directory that contains the \texttt{eproof} executable,
+or install the prebuilt E package from Isabelle's download page. See
+\S\ref{installation} for details.
+
+\item[$\bullet$] \textbf{\textit{spass}:} SPASS is an ATP developed by Christoph
+Weidenbach et al.\ \cite{weidenbach-et-al-2009}. To use SPASS, set the
+environment variable \texttt{SPASS\_HOME} to the directory that contains the
+\texttt{SPASS} executable, or install the prebuilt SPASS package from Isabelle's
+download page. See \S\ref{installation} for details.
+
+\item[$\bullet$] \textbf{\textit{spass\_tptp}:} Same as the above, except that
+Sledgehammer communicates with SPASS using the TPTP syntax rather than the
+native DFG syntax. This ATP is provided for experimental purposes.
+
+\item[$\bullet$] \textbf{\textit{vampire}:} Vampire is an ATP developed by
+Andrei Voronkov and his colleagues \cite{riazanov-voronkov-2002}. To use
+Vampire, set the environment variable \texttt{VAMPIRE\_HOME} to the directory
+that contains the \texttt{vampire} executable.
+
+\item[$\bullet$] \textbf{\textit{remote\_e}:} The remote version of E executes
+on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}.
+
+\item[$\bullet$] \textbf{\textit{remote\_spass}:} The remote version of SPASS
+executes on Geoff Sutcliffe's Miami servers.
+
+\item[$\bullet$] \textbf{\textit{remote\_vampire}:} The remote version of
+Vampire executes on Geoff Sutcliffe's Miami servers. Version 9 is used.
+
+\end{enum}
+
+By default, Sledgehammer will run E, SPASS, and Vampire in parallel. For E and
+SPASS, it will use any locally installed version if available, falling back
+on the remote versions if necessary. For historical reasons, the default value
+of this option can be overridden using the option ``Sledgehammer: ATPs'' from
+the ``Isabelle'' menu in Proof General.
+
+It is a good idea to run several ATPs in parallel, although it could slow down
+your machine. Tobias Nipkow observed that running E, SPASS, and Vampire together
+for 5 seconds yields the same success rate as running the most effective of
+these (Vampire) for 120 seconds \cite{boehme-nipkow-2010}.
+
+\opnodefault{atp}{string}
+Alias for \textit{atps}.
+
+\opfalse{overlord}{no\_overlord}
+Specifies whether Sledgehammer should put its temporary files in
+\texttt{\$ISA\-BELLE\_\allowbreak HOME\_\allowbreak USER}, which is useful for
+debugging Sledgehammer but also unsafe if several instances of the tool are run
+simultaneously. The files are identified by the prefix \texttt{prob\_}; you may
+safely remove them after Sledgehammer has run.
+
+\nopagebreak
+{\small See also \textit{debug} (\S\ref{output-format}).}
+\end{enum}
+
+\subsection{Problem Encoding}
+\label{problem-encoding}
+
+\begin{enum}
+\opfalse{explicit\_apply}{implicit\_apply}
+Specifies whether function application should be encoded as an explicit
+``apply'' operator. If the option is set to \textit{false}, each function will
+be directly applied to as many arguments as possible. Enabling this option can
+sometimes help discover higher-order proofs that otherwise would not be found.
+
+\opfalse{full\_types}{partial\_types}
+Specifies whether full-type information is exported. Enabling this option can
+prevent the discovery of type-incorrect proofs, but it also tends to slow down
+the ATPs significantly. For historical reasons, the default value of this option
+can be overridden using the option ``Sledgehammer: ATPs'' from the ``Isabelle''
+menu in Proof General.
+
+\opdefault{relevance\_threshold}{int}{50}
+Specifies the threshold above which facts are considered relevant by the
+relevance filter. The option ranges from 0 to 100, where 0 means that all
+theorems are relevant.
+
+\opdefault{relevance\_convergence}{int}{320}
+Specifies the convergence quotient, multiplied by 100, used by the relevance
+filter. This quotient is used by the relevance filter to scale down the
+relevance of facts at each iteration of the filter.
+
+\opsmartx{theory\_relevant}{theory\_irrelevant}
+Specifies whether the theory from which a fact comes should be taken into
+consideration by the relevance filter. If the option is set to \textit{smart},
+it is taken to be \textit{true} for SPASS and \textit{false} for E and Vampire,
+because empirical results suggest that these are the best settings.
+
+\opfalse{defs\_relevant}{defs\_irrelevant}
+Specifies whether the definition of constants occurring in the formula to prove
+should be considered particularly relevant. Enabling this option tends to lead
+to larger problems and typically slows down the ATPs.
+
+\optrue{respect\_no\_atp}{ignore\_no\_atp}
+Specifies whether Sledgehammer should honor the \textit{no\_atp} attributes. The
+\textit{no\_atp} attributes marks theorems that tend to confuse ATPs, typically
+because they can lead to unsound ATP proofs \cite{boehme-nipkow-2010}. It is
+normally a good idea to leave this option enabled, unless you are debugging
+Sledgehammer.
+
+\end{enum}
+
+\subsection{Output Format}
+\label{output-format}
+
+\begin{enum}
+
+\opfalse{verbose}{quiet}
+Specifies whether the \textbf{sledgehammer} command should explain what it does.
+
+\opfalse{debug}{no\_debug}
+Specifies whether Nitpick should display additional debugging information beyond
+what \textit{verbose} already displays. Enabling \textit{debug} also enables
+\textit{verbose} behind the scenes.
+
+\nopagebreak
+{\small See also \textit{overlord} (\S\ref{mode-of-operation}).}
+
+\opfalse{isar\_proof}{no\_isar\_proof}
+Specifies whether Isar proofs should be output in addition to one-liner
+\textit{metis} proofs. Isar proof construction is still experimental and often
+fails; however, they are usually faster and sometimes more robust than
+\textit{metis} proofs.
+
+\opdefault{isar\_shrink\_factor}{int}{1}
+Specifies the granularity of the Isar proof. A value of $n$ indicates that each
+Isar proof step should correspond to a group of up to $n$ consecutive proof
+steps in the ATP proof.
+
+\end{enum}
+
+\subsection{Timeouts}
+\label{timeouts}
+
+\begin{enum}
+\opdefault{timeout}{time}{$\mathbf{60}$ s}
+Specifies the maximum amount of time that the ATPs should spend looking for a
+proof. For historical reasons, the default value of this option can be
+overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle''
+menu in Proof General.
+
+\opdefault{minimize\_timeout}{time}{$\mathbf{5}$\,s}
+Specifies the maximum amount of time that the ATPs should spend looking for a
+proof for \textbf{sledgehammer}~\textit{minimize}.
+\end{enum}
+
+\let\em=\sl
+\bibliography{../manual}{}
+\bibliographystyle{abbrv}
+
+\end{document}
--- a/doc-src/antiquote_setup.ML Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/antiquote_setup.ML Sun May 23 10:38:11 2010 +0100
@@ -166,9 +166,9 @@
in
val _ = entity_antiqs no_check "" "syntax";
-val _ = entity_antiqs (K (is_some o OuterKeyword.command_keyword)) "isacommand" "command";
-val _ = entity_antiqs (K OuterKeyword.is_keyword) "isakeyword" "keyword";
-val _ = entity_antiqs (K OuterKeyword.is_keyword) "isakeyword" "element";
+val _ = entity_antiqs (K (is_some o Keyword.command_keyword)) "isacommand" "command";
+val _ = entity_antiqs (K Keyword.is_keyword) "isakeyword" "keyword";
+val _ = entity_antiqs (K Keyword.is_keyword) "isakeyword" "element";
val _ = entity_antiqs (thy_check Method.intern Method.defined) "" "method";
val _ = entity_antiqs (thy_check Attrib.intern Attrib.defined) "" "attribute";
val _ = entity_antiqs no_check "" "fact";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/gfx/isabelle_sledgehammer.eps Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,6488 @@
+%!PS-Adobe-2.0 EPSF-1.2
+%%Title: isabelle_any
+%%Creator: FreeHand 5.5
+%%CreationDate: 24.09.1998 21:04 Uhr
+%%BoundingBox: 0 0 202 178
+%%FHPathName:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any
+%ALDOriginalFile:MacSystem:Home:Markus:TUM:Isabelle Logo:export:isabelle_any
+%ALDBoundingBox: -153 -386 442 456
+%%FHPageNum:1
+%%DocumentSuppliedResources: procset Altsys_header 4 0
+%%ColorUsage: Color
+%%DocumentProcessColors: Cyan Magenta Yellow Black
+%%DocumentNeededResources: font Symbol
+%%+ font ZapfHumanist601BT-Bold
+%%DocumentFonts: Symbol
+%%+ ZapfHumanist601BT-Bold
+%%DocumentNeededFonts: Symbol
+%%+ ZapfHumanist601BT-Bold
+%%EndComments
+%!PS-AdobeFont-1.0: ZapfHumanist601BT-Bold 003.001
+%%CreationDate: Mon Jun 22 16:09:28 1992
+%%VMusage: 35200 38400
+% Bitstream Type 1 Font Program
+% Copyright 1990-1992 as an unpublished work by Bitstream Inc., Cambridge, MA.
+% All rights reserved.
+% Confidential and proprietary to Bitstream Inc.
+% U.S. GOVERNMENT RESTRICTED RIGHTS
+% This software typeface product is provided with RESTRICTED RIGHTS. Use,
+% duplication or disclosure by the Government is subject to restrictions
+% as set forth in the license agreement and in FAR 52.227-19 (c) (2) (May, 1987),
+% when applicable, or the applicable provisions of the DOD FAR supplement
+% 252.227-7013 subdivision (a) (15) (April, 1988) or subdivision (a) (17)
+% (April, 1988). Contractor/manufacturer is Bitstream Inc.,
+% 215 First Street, Cambridge, MA 02142.
+% Bitstream is a registered trademark of Bitstream Inc.
+11 dict begin
+/FontInfo 9 dict dup begin
+ /version (003.001) readonly def
+ /Notice (Copyright 1990-1992 as an unpublished work by Bitstream Inc. All rights reserved. Confidential.) readonly def
+ /FullName (Zapf Humanist 601 Bold) readonly def
+ /FamilyName (Zapf Humanist 601) readonly def
+ /Weight (Bold) readonly def
+ /ItalicAngle 0 def
+ /isFixedPitch false def
+ /UnderlinePosition -136 def
+ /UnderlineThickness 85 def
+end readonly def
+/FontName /ZapfHumanist601BT-Bold def
+/PaintType 0 def
+/FontType 1 def
+/FontMatrix [0.001 0 0 0.001 0 0] readonly def
+/Encoding StandardEncoding def
+/FontBBox {-167 -275 1170 962} readonly def
+/UniqueID 15530396 def
+currentdict end
+currentfile eexec
+a2951840838a4133839ca9d22e2b99f2b61c767cd675080aacfcb24e19cd
+1336739bb64994c56737090b4cec92c9945ff0745ef7ffc61bb0a9a3b849
+e7e98740e56c0b5af787559cc6956ab31e33cf8553d55c0b0e818ef5ec6b
+f48162eac42e7380ca921dae1c82b38fd6bcf2001abb5d001a56157094cf
+e27d8f4eac9693e88372d20358b47e0c3876558ebf757a1fbc5c1cddf62b
+3c57bf727ef1c4879422c142a084d1c7462ac293e097fabe3a3ecfcd8271
+f259833bac7912707218ec9a3063bf7385e02d8c1058ac06df00b33b8c01
+8768b278010eb4dd58c7ba59321899741cb7215d8a55bee8d3398c887f02
+e1f4869387f89141de693fcb429c7884c22dcdeddcaa62b7f5060249dfab
+cfc351201f7d188b6ed68a228abda4d33b3d269ac09cde172bc045e67449
+c0f25d224efbe8c9f9d2968a01edbfb039123c365ed0db58ad38aabe015b
+8881191dd23092f6d53d5c1cd68ebd038e098d32cb24b433b7d5d89c28ee
+05ea0b6070bb785a2974b5a160ee4cf8b6d8c73445d36720af0530441cd9
+61bc0c367f1af1ec1c5ab7255ddda153c1868aba58cd5b44835535d85326
+5d7fed5ff7118adb5d5b76cc3b72e5ff27e21eb857261b3afb7688fca12d
+1663b0d8bdc1dd47a84b65b47d3e76d3b8fa8b319f17e1bb22b45a7482fd
+f9ad1b6129e09ae47f15cd2447484cd2d64f59ab0f2f876c81e7d87ccdf9
+005aa8fc093d02db51a075d571e925f2d309a1c535a1e59d34215c6cd33e
+3c38997b4956461f376399901a8d0943dca6a335baac93fc8482c0659f04
+329c6f040b35828ea6dd1bd1858f2a9be4ef77731b5b75a1c536c6bc9479
+0821e5d88f6e2981835dbfd65ec254ebcf2cf49c917d121cd3bbb476a12b
+69c15f17d9c17bb15ad1e7d31d2afcf58c8f0ad526d68615a0f1ac3b1d1c
+d3beafeea3cf56c8f8a66367c70df9159f0b1b3157ccfd010045c4718e0e
+625c0891e85790c9b97b85517c74c9d55eaca31a01cddc64898bf0eeadf3
+53391a185e507dcb0a6f52661a56357ac818dfc740a540aadf02f4e7a79d
+8008a77cd30abde337025b01217d6a68c306abe145b7260c3478fa5f366f
+b2d37259ead8a8ec2db2f09ae0eb3a682d27b0d73a60935f80254c24426a
+003a87a29a4370cbf1b2ef1e19ad8466ec725fd5b463d06011a5e0da2258
+ff6c1483c4532bc21f2ed9b99b929b2800ddefc1a98d12ba085adc210bac
+e0274b69e24d16af015a51ca73edf779a7534a887aa780337ad966839881
+edc22ba72038aa1a96a6deba24ad676795da711b92f8cf5f54cb4322ec04
+40ef9e15b11e3005f3ff69376ecb29bb66e8fc1b685f2b05fb162fcb35aa
+d7eb2a8ec39b97ab1ff05ef02f8dbbc12085a5cd252cc4010fab7f63dfd5
+7fa1be86f724d37db5faef17ae8e8e537444e8e9db917b270344183473af
+7f47d5703a160d8ef1e772438620d3336b2fbcf6433612e4b5e64fae0329
+8a3c4010c17d93f13ba66d549c69dd58c7d26ddc90285fed831918563a16
+2a7ac2511e2f18c9eb3df905a9dcba65a31cc1c39f07458abb11b4c60346
+aea19070e126982f1dde336e79be0ecd69a8afbe2493d064d4d2ff38788b
+b3038125961302db9761403c3b8019ec641e107425002205a77ae2ae0f4f
+7550d623dd03f0ec0199f42a9a8b89514e2e21baca9b3c8c96ca48cbf9dc
+ee6d6241d713e014b49e83ad85e62a6b2f70b58e4cc72d41ea6fcbdd3b5c
+890c8af0d24200658773b1628c6cc9aaaabb08865ee4c7ff2c513ad7aa23
+155a321213fa94731683a3e282a0e60aa3c87aade3da231465bdd9097f2c
+89a1af8e5b9649143f2d9482546501ea97e8bea2f5d5eea97d4f19bb6835
+3138d3babb2461c08d537491aaede1f23d734c6f099eb5bef6e2ffaaf138
+e5ab71b8b41599091037e440127a4eaedf208c20c8a2fc62eadab191d1ab
+4d5531f826aa6b9fff2797a7f54673e0a3fae09a93a0dfafb8b11d60dc69
+5acf9b7e1a47c31d0b5a0b85b7b50cddff5ac831651d9c7469c2139c7a89
+7d2f868f36c65156921803eccfdbdd1618595ab6d2a9230ef523a1b5ee51
+f2a0d200fc0e94aff7f546593ff2a3eb865d129895af01b8ab6e4616fe20
+9123b6e2b7e0817adc3cdb78ae8b0b1d75f2986ebd8fb24c4de92ac9e8c3
+6afa520636bcad2e6a03d11c268d97fa578561f6e7523e042c4cc73a0eac
+7a841907450e83d8e7a8de4db5085f6e8b25dc85b59e018380f4b9523a7f
+02cbeec989f0221b7681ec427519062b429dcd8fc2e4f81173519f88e2e4
+3798b90a84060920f6ae789afd6a9182e7fec87cd2d4235d37a65c3f3bcc
+c742c89cbe5f3e2ba6c4f40ebba162e12569d20684cc167685285a087e7a
+0a995fe1939bf25c09553512ba2cf77ef21d2ef8da1c73ba6e5826f4099e
+27d8bc7b3545fc592b75228e70127c15a382774357457cd4586d80dc0bd6
+065aee32acfd5c0523303cece85a3dbf46853b917618c0330146f527c15b
+dbb9f6526964368b2b8593eed1551dad75659565d54c6a0a52da7a8e366f
+dd009ef853491c0fb995e19933cba1dbdc8902721c3ea6017ffdd5851cb8
+3c8bada46075ac121afe13a70e87529e40693157adcc999ed4657e017adf
+f7dbac4bc0d204f204c6f47b769aaf714f9ec1d25226f24d0a1b53e28ac5
+374ab99755852c1431b2486df5fd637e2005a25303345a1c95a15a1189ba
+f6f6883de1ad46d48427b137c2003d210ab2b2f5680f2633939f289d7553
+eb943adf8127f1c3ee7d6453b5566393700ad74ab86eb9a89f8b0380af55
+6b62f51b7dbd0c5dcc9a9fb658944d7ad5845d58dedc2d38200d0ef7cb0f
+76041dc104ef3ab89c1dc2f6a75635d48051c8a7dd9f5e60253a53957ec8
+9d1266566d7ed20d79dfc2807b397d7cf056bdaccdb72528a88aa4987682
+c909b2fe1e35a71c2f29e89a2bf32173967e79610367ce4574ba6a1cc031
+cfb176fc0313f74f91a866ef9954b95b29caf917a6b919586f54d23cb7ce
+23305886ae7760ebd6263df0d3c511ac7afc361df78bc2621f66d3268b99
+078fa59124f0eb9476496c938eb4584e87455dc6f2faa999e938460b31c6
+28021c652acfa12d4556aa4302bbcd043e60389480b796c3fc0b2e51b81e
+c2afa4a34335318a1c5a842dcaa120df414acba2e79ab5cc61c99e98108c
+5cb907a96b30d731131782f9df79aabfc16a20ace8852d047497982e11c8
+26321addf679de8a96a2d18743f6f2c3b2bc397370b10ad273fcfb76b48b
+9dad27cf89ca2c370149cd48ab956e9bbce01cbf8e1f0c661b99cf19b36e
+87b6165dd85ae3f3674525e17d85971093e110520d17f2d6253611e35ec9
+e17000e48e2926454c1e8eb4098e0115b49536f2c33505eb5e574f7a414b
+e176398c5ddf6d846ea5ddf2a5e94c0422e0115c57a8c9e56bf8ba962c82
+698c96bd6138baaca7347e44352cc62df4eeba364954ad921a5a43280980
+264df4a7fb29d005423179f7bd1d98b4280d62ce23c927551f1ffc2b8f17
+0a9c23656c0c91b640cdcfdbd88089ffb28d3ac68bad25dbbed82c083696
+1f9f86a6183cc1061ffdb32279796569d05b31c946955402d0be1fb9f2bf
+304d1ad8e1e357be49e6e2ee67f7f1e7bc699d46a5f7853fe659ba2e1930
+0d3e3ea658b9862701dcab08fdd23bf1d751777f25efbe9e02d12b5612b3
+c3fc2275190346b94ec4024e4ade08e54d75c0b4c48f4956b9182e8ce997
+74b54da4a9318c099d89f1ce3b6803a95f48b9fb8b845372be25e54478e8
+49e4707ea03a36e134efa661e4e6250d89649ae074cfd9d6b9e2071c1099
+3b8a5a5ebc3e1cb228c97565aef7f254e3f90af8a3dd281c83792755719d
+c6a5b3bab4aa6be5afe9624050eee8dfb13b018f4088c932cd48ace38dfe
+b1b4218dba8f7fada6628076acf1b54db0c95d4fb12232f1fa9d1ba848f9
+fe80c65b75d6946c00fe78839142c5302707d9269b24565dbcc551aca060
+b4d0f99b961dd3cc795a982063ac42e9fc81fc98add42744a9f92e74b00d
+637ee4606ea2099b6c763493c3159f8e52a90dafca682010c0e92bc9038a
+10abb066c75c8d97f7ad6fb1a37136e52cf2093c4fa485fe12adad10e4d0
+83b78b023628ddc5326cbf8392516027a9f3de4945f93488e4a1997efd2a
+22c2c136dbac1bdb829e082beac48cdd06dcb17bacf09451c7b636bd49a8
+fc60cb1d2292250bea78e1dd276725ab4c526b66ddabf4e1b2bf0a2571df
+2490df70735f5da321fac74fe4fab54444e53dace9832cff38a70c58104a
+4f0c0545dcf7a1a9ecb54e0e32d6d8195d30b2c98a567741fcf325a4ddeb
+244a1a35676e246ddc835fac13b569f35f22ec93191eca3efbe17ff9a950
+d08f863710b2bbecec969068c498eb2338b68f3fc3f5447449fe4de2b065
+e068ecd9255d369b2bb6c4b0b7423774bed69294758aca2bdb1a8d5bf618
+d3fa09462f7838e8a79b7a53bebe6dacb0a1561eaa074bc6f069f6a06fb2
+c4a5cb13e2172bce9be659a82665da5cded73da84322bb16aa6e19ac1958
+7515cb5d2b65e65e803f76700ce5efd3df7fe4ed431fae0e8e286b1d5056
+a0d18df926b2be7a93c503ab903abd4788680a6201fdc299f2cb5d6a9b6e
+2048109c8d1fb633a54128938594b2cce86a7e0185e7d59e6536584039ec
+9e30ff7be6ddba9fdba82de7415fdc47de84d97afb1aa3ba495bd91dee9d
+f3b21ee1164987dd8510925087cd0919f1085cba6e4dd3c7384d26667f94
+ad7f736a60d8bd76dfaa4b53c23536fc309ff2317c48ee0107ff2ca4d1b3
+f78c5a27b901c931128bdb636094ef0cd543a5b62b7dbe10ed83aed5780c
+a16067a4a7e8b7c5bf8a8e822149bc1b5bcdabe13a7f6aa6eaeff24a42f4
+a58a2b70f545103540169137fda9abb589f734b6776cb50402d6123ce802
+10dce05e3697a98c9411cf60a02c278c91e03d540b936cd00c668960e357
+1aeaf4d94cfb496b259ec0d8fdba9199fb46634ff177bc8d310ea1314eef
+d46c927a981c58e88743ed4e07d80fe841edee812e3053412bf2e710146c
+b25dec8ea70c38bb1f6e4db3c2e8ba521963c1584eeb60ea1e9555058f13
+e98307c13cbd15c26b611f543149b1ddf88dd6296ae703f58daeb67f1b03
+ab5b24c72d5770cb9d8ed242c4faaad1dd940ada00e98ff3a61799d13355
+aba916910aa9a6e5ee8af438d0ba8235346fcd139b9d2cb7db7bd3f298a3
+94ff0aff3b9042f32a004e042c346a5ea35917f229848a9c5a32909b0090
+4aa715640277a6ada99f8b2927fda22899ff1347f42bac73e2bd4bbf3945
+55fd7dd30d5c6dadf5c259fdb2455d8c607de1c5da588e20159f18e4e8da
+b714e532d888a0386c52c2b85964251af003ac9d10c0c8b9b3465e1dde48
+2e74a29e17a7cf6c9a43b5af1646f0b8508f98e9a1279ec3908073d89dcb
+aa093e8dd1004c1ecccce0803095b0069d4be7a1eb14b02efc37d137dfe3
+f0190bc9628069abc257f45d0e050e60c7f5281277937dd986fcd5b94a2b
+845a1a75addd74a142800f18ef408c08a2c2ad16a93298f148c0ae7d2990
+ded147f92f053809a60d4f992a227167aad5b5eb2bbe8a4a77dc77a08b72
+6acb34422e2532eec7e274e4a42d71ee584f5f3df5a3a5b379974ede73ab
+5f1b0a2dbfcc8cfac4747ca26eb6030dc2f85a104b754432977850c093b9
+97ed90af711b544ff682e7b1eac82b2b7b44014b09c17ecf084c994a095d
+9eeef5391305caf093b62ac9916f982a436d521fcf4d75c5b8e4d92266fd
+e99a58aa39d7693ecd1796b8851761d64bbca39a6d5a0b4533ae47123327
+f98d0ad0e8b36625cc3647b55459552906d8a1d5766845ffac101980efcf
+79657e365510be5db557cefef21193ca3cf3dad175ee2e7ae91d174fdc06
+2ff5c51ffe2f021122e96df042019d3a1883e662537ec1b69c11fbb6e750
+0132eabf5803c816153ecbff60ca3b3b39708c26cb1751afb2e65d8e5f4a
+c4397a88fb1f112672fcdd24e4ba545c5b2a7968c17b62f8e2530a8acbff
+cfca82c64b7abcab84e2c4a0a7ced67b15669301fe9ff2c756e70ff7ce33
+497be6acc4ac5617e1f043bd8a87416299a39bf17fc31c02d72d75fdc2a1
+e60669fa4d5e4a49d9afea2f53f4626680e9c0dfca223529efa415c4343a
+b6067aa800c484457ea050eaaa5d3fafeedd0eec72f327e02c6b3912b5a8
+c404de4839c9c4a99da42681cde43316606a34c7d2f02269de1aab776857
+e668f35946af4d618d36d444bdc02b1f63ea25b6260b4fb606ac8575b5c9
+782a5de4037350d5753b1537537ccb008c454eeb264e6cd4727c999e240e
+0ac89e95a896b67d54910a3531345f64198ad394b5ceb52881f1dd9e6beb
+95862dc188d45b3e46aacb5fe40097947dab9bc3c1ee46bfc9b1b3ed6167
+efd0d65ceb043d7b24c1456676e4baa47b1209a315f199bb3a91f4374cd9
+cc0b40d3f09f19f8dd8a46915eee55eeeeb3c7b8f437106ee491ef0f4ff9
+2c5c6f779e0fbe7bd5293964bb645ca362b106abeb773571d9ae83b786a3
+d5a4ea3ea970daadc46cc5e6037f76fd20e0fffc47cf4e7af9522b91f96b
+3297720fd45d9bc2200622ad2ca9445556c8a8202b1991bc63da360d021d
+55be2528e043f803e08da99b91ab9cfc5e65b2655d78206b4aecd445a7b0
+1caa0d06b0a55e8f04b70b60b04a860c8e1ab439f4910051e3f7441b47c7
+8aa3ab8519f181a9e833f3242fa58d02ed76bf0031f71f9def8484ecced8
+b6e41aca56176b6b32a2443d12492c8a0f5ba8a3e227219dfd1dd23fcb48
+fcfd255dbbf3e198874e607399db8d8498e719f00e9ed8bdd96c88817606
+357a0063c23854e64ad4e59ddd5060845b2c4cddd00c40081458f8ee02c7
+303c11747bd104440046bf2d09794fca2c4beb23ed1b66d9ccb9a4dd57ad
+a24943461ecc00704c916bdc621bfddb17913dfb0f3513b65f3ab015786a
+caa51ee9546bc8ddf87e2e104137e35ddf8f8d23724e9a53824169bc7cfa
+99562656e6f1c888d4dbff0b269c5d1e733e5f212d91297610201eb43249
+35e336dd0052738db2d64f3e89429903bb5c1810009cf766e9a06223dd2f
+219b706394a121dc029af55c6ada9052af59682ef7c51e121cf16f0319ac
+0aa9512ef900c548d673fe361da19052808797e958209072e145d46ec8cc
+a89fafd76630eff30ae979973bdf0f8c9e469d8edd3b1c93731c72f976b7
+d81142bc15c376403f967affaa5f482efd57c6f91970729d16db851f0ed3
+ea7d82f409307b5b436886c1beda94a1fa3ab1b60686f6574c844fb2c0b3
+a07174dc4f27b4fed2f8bd4d5436be4b343e5efdf0400d235bd910255341
+a20770804a26f8437e9bce6da8e9f8258a343c7aee291f1510be306ae67a
+ab1d7696453530c02fd153bbe49dbf62baad6146029cbd1656cdb76c952c
+b93edfee76fe33832930be59636bb947e8ad285f20f663cccf484fba97d6
+7446c7b6c6f5857428bb1737d9ae801df75d9cb4d7bd59ef7a4cbadde928
+38f15d232005585d2e40483d2d3e08cc8f398bb43afedb84343c3ba3835d
+0ba82a86dce859cf655f85e63e41365e0dbefcf511b9a27a2b6e66b2ad3a
+c657902842287a317e46ceaa93b5088f09d53a65815b44538af90ad3b06b
+4e5e2dc509f02e30a01e05201c67d4d39582bbe64e20b669f5fd787909a3
+30fc50a95b31426bbb57a4fbf9feacdc31f98bcf50da7e50c2bfc169c6fd
+ccf213cdb878653bcea372968ea6e31fd30dd55434cc91c0af22179ce669
+a05493f195e12432c6173ae2ac3c94fb83f38210014a9f969ea2b44e99f5
+e5a7317e848d429ad62167a4fc5001149676c0c28cdf59b8d1c5a582f516
+3eee855312777fee6dacbf993f5c058f355dbde6552dc960d336eee445dd
+11d53fd21b745d1e5ec317efbbef25e070d0a36797a6081c356ac2328e64
+e5c55fbc81dc75d9c1575548ece74b8307eef485aa8e28859a2e0435c831
+23a600efb323c362fe9f02407a5411c41a69566cd50add324b63ab939980
+b9d7a929ae4887163cfa7acbfc9fabaab8987a1f6906b9881491cd055b94
+485c968479dbb05b34ed0cd6844729a692978c6928c3392e33e8324ded88
+814cacdac8128e1425c0091a13558100d7cdbed5992795d94d39c32f32dc
+621ab6f3b75187a66741f61d6a9c91d791b1cfc3d0e94d4a76302e0c3f2e
+cbdc51f14f3251aa5c8bb989f0e13ee500b7b7f2f1e52ca970ad8a7b4b99
+57e93126254297380d67179deb8ff1e99d5cdf7a35c5bb9fa7b402e63234
+78640344e1f10c378ad23c5cd1aa18e1e0b308db70d3a624a455f8e291a2
+ee102ad10776208c2d546cb76d89ca8103a8b95f8acc2d2bdc9791324915
+6c9e05429091071f0c5b76d82c8d1c8a69d840fd460922cd2090624bc218
+0c9391005926a25042a55e322060807363462e1cdeab309033124ba3a884
+1db13f39edae04ec52cde9dbde64ddda1ad805141d4230ec76bd81fd98d7
+0d90fa1aaa26ea551bf687ddd6cdcf3de5a446b266c68434f07d9c0b382d
+5816c4e22f22cc03ff78064c6dffb12315c6bcbbf5dc510f5aaabf23471a
+234efceeb4aa2f9af9ea787c014c5587ef162fc5b35e8f4c23b168c6e247
+41d33dcc11d2a56d3ba9d8eed6e79aebf9f0faf1a3aeb89d792d69041f0b
+b8fadfc0aa090effc6ae5e2f13cdbf54b5bed69b039eef2627769613b6f1
+aefe9b66747fe8feaf7455796740f411a770d4a1764f0483719584880f45
+430e38d3af184145892a08b2add234a3f3ee4ccfc9f6995c02392adafccd
+722f366d748cfe9373fbf5f878ed47e9d221fd156bb28369df9e7d2b79da
+76120d135ebaf36cff93beb7e313c2b2de7477176fc19609a1b906c995cd
+defef08899265b6b8aefb44da1aadefd1c523dce5ca1b84c0c652b3009fd
+057789892d4d31764f181754b2e0a62c465587585509989a219711a5e4e2
+5b3b340ca8fdd3f04fef204b1b722b2f6c2ccb00c3cf1a94ba9bdfbfeda9
+e2a062c6f1ced3b8aae5dae32ade1fca1001f98d0ad0e8b36625cc3647b5
+5459552906d8a788eb8bc734ccb65fe9582c71df94fd95d22c5323de235c
+28220fb9a2ccb37362174d8cd5922c9e5a87b51d0668555100a33e33750e
+f1f795cbed962494a994be7ce8cf71fc58ff4204551b1615ed27cf088171
+fd000b72462b67935961e7c6c3a05bfd67b9ba094ea2c16fdf486da912e1
+e97bfd1c17934535e551cede20c001b5d2adb2be4cbad7d6ba0bdeae4b1a
+a739f90293e67ecbdeea4d35825e092697fb05b215083e3f3d6be260790e
+2a175fd44eb1c4c16759504827a6eb58a838c4d65fec6eef108495577019
+15740cac164111892e8d1cc447cd208e243a89ab847d8ebf4fb98bff49e7
+a3453facf3b0e8cb67590f390173ddba68324531d2e426aed152e12301d7
+538c1f3c0048a9cc00c009a1a9138460082123209c1e007266fbf236eb72
+21f87d4ca38a0b699e84ca230ffb5095f90a6528bf2a9118f95ac9ab8d2d
+ed9eed9b8b27be894b717469758c8d94fa89acc64f530f432d0e5f16c922
+36d6a63410f099c9e909450fd731d698ef658d8ffc1de14817b850814f68
+1a4a9be5cc7a71c381974c249f0b209bfdc2e97f9540c96f57bb4d283622
+00969b82011315289e6a025b137030a0af3b4b42b00fed7cec49df43c59e
+3b2495a036dd1b17a8e6adae63bfbbd48807c44b5bbf71813355e1b0e58e
+22b6fb88005fc55565be49c17244901b73ef02fc4eb7669be5af22d89c0d
+dff0fc6821d810d13e5821d48d4a71d7e463d5b60bc279d0dcf5f8da3a95
+905b56d6f2be95e6d4243b1048e3b662e62401ffaa3bc3f5f90b0854b8a3
+8c38039f61fcb359b06bbb7d59e3b39a295dccd6db9a8b83a6f64ef8dc94
+a77123dd164cfd1c46f1ee51aa19c3d6e7db92a298d10159f2b5eff2caf9
+dc93a6d267fb65bd900d6adf0c6be598050b6d3a9b3a322ab3c9e880d774
+1f58016ff97e5f606b5dbd72ba99252c669209bb556dd5be84fdd7c1ce92
+8a3b3d3aab8d37e6b740227563bb4d60f6bb04052356e1a48d2079feca44
+7ea17fd06f208426d045dee660d1d6460455f8d20dbc5ae64550bbdf60d7
+27d96cc9afef842a8c8c78ea2257e6c6d0d207c80cfe399e8874c693274e
+d2c2022d303ca50a70624b07434fb85040a76a823f446c7454dab4f9c05f
+10274eb5ba164aa3649d1bc90694316ba5cb3e7df4442e777124cff7ebef
+53df2320a0c441ab61666493cb43da46d5711c21699de85bc74359444da2
+e3e397d4c16234f81531505b621aa242a6698886f82b447104b1f1062f60
+b5c87cea9151bb3c627bfa4532b06fd147c556ed8d61ae30a8719dfb8705
+f8a6c74368381403640cc57026d3790c49e2bbd1c0e48285ec6ba44de678
+e3a1394d659c412f09644b83ee1a333a1f51ad8deb4e6d77b3b226ac2c4f
+fe653411a7976ae7c4a3cb7df309788da6b483f8a7bab4a6990db74362f5
+bc41d545a320389b2599fd726e426ed9fa2916ece67b058f6a269544e517
+128bda38d117f402409d0d8f8c88ed509aa2ba882e0c579b45af4be80770
+22d7269684eaf0f9afc3054316da6611e3fd260d67fb6fe52c9ade5dda24
+a0050a819ed21342aac9d25194778beb3145f56a66980f620998923521ea
+3f957b6ed0c5470734af9f416a16427dd03eff9a0e023452097d4ef936d5
+49a90823cef6de340a1ee02a52851b310cbcf41ae274947a62f9d1d8702a
+669023e3caf967204a340694b45fecbda4bf9552f6bdc62d43b3b2c3d571
+9983c182453e22ee34241ab908e667115f7988174684cd70084aefc55caa
+f5352a88e9dac45d1ea0e032af61fe9a9118a3931b2050fc6db66ab96a39
+74353b597f34dfd9f72150de23285eda5e555a607d198c291965a7233715
+3f4946a57af0b440ff8567b01a6f46c6d32fea5f8bf57d89dccbab7da882
+ee6c9260e89443b1d7db099477492bd0468850df3db668d741123e7ebe3d
+c21748ab4c5cbeb5de33b8963aecafe76bba0c4f6ed8e8263a116ed85e58
+fb71ec4ab0071301be7c7d3afd5fa6ad46c0232807bb7fe129e44bfd16e9
+fd0c8bb5e7cdd86a78b5fb0669093c22eda9151d85b6f58a9c8ead3727c0
+09850bd31a8b4a873d0a506240bb2aeccb8dcb6369532f21d9b967aa8443
+fd6d77cb2d65c4678a5fad188db85940f0a187aa1031dcf5b8e0d0cbfb6d
+b3b96fedec5b249b7a69de9b42dfa605bd622de7a220cce9b66e9f3394d6
+13487dc5e82c1e619079cd057b1e19ac05ebdfd7c8bf01c6c66fab49e0b6
+613df9e42beae2f7b9407a2bff8896d8035cea0fd5c11bc5889cb3d90876
+61766138d2625f42d0244adca65d1bc73989328c0eea0b97c7c766285ab3
+351ce2b183f774488a8806c33178090a3808f0ce5e339b87cf7add933301
+ca486742831ca751f0626864ce13172829a8419af5c78794a0eaa17b5bcd
+fcb684f7d4bb7af15deb432e44dc7dedf56eb8bea08b46f1e8123a49a349
+a7cbccf833a528f5e22d2d463040e09b91e543a2f33077b3e7b9ecc64f14
+306186cdae1fc317a6ced7e9b4d51a10bbbcf2fadff876b4d9082e3f4aef
+dfef230e4232572f4fa33a6e065f6895aa2ea96c5659cb579b023179f0fe
+de7ba64bbd9362a7b2b8c4eaec254915629e81d01c839096339b99bc9e25
+84536955feaa52fa20666f65bafd9b2e69c3e8c15d24fa407e7d881679b1
+789a0e2a695d13553c92c0214c9b7562cd6a9a3d77c8b0c2196cef76dc51
+d855c1dac37f96eae4cc7bf07e17dc7c08333d7af33c8b2965ea1f23446b
+3c96c52b30ea628ad572694d145b58a606f90b278290297aa372cff56b6f
+56f4aad6612eb7c7bd07db4f7d1a70d8044d16d0b5c1605ee02a852ffdb4
+450147b3f9b87d72dc431b34fcdc899462dcc1b6bb6ab1758b6a589e91e5
+8f5196251d00133b43749b7a11fb67a22664c5e38e336dbdeb5509c2d9d6
+2642c07275949df0e2db59314ae0fb34641fc171d3fe1289f919136d853c
+d9048ee9db50c699c49e27a8df199590bbc65b23b55bb387eed0c73f2db5
+1cb091f8c22af83103f214199e371f7de1df23f757817200be30610004df
+81fe8ed6eba79e856fca21a126ca326ad2f313c16e15754663ad6a065e08
+4050ff005fc899d6e233691b918a093b5f1ffda8839ab23ae66b1bb7b953
+0a7f896ec55de6fb9faf1b49656ff2e57488cd7f1c44114c75f9d571461f
+767a6040ffa14e9fb43096f164d60ca530d7cca76d526d1999ac1b52a793
+28651112a65db1f2564ecf90ea6bf2c9ecf515640719c3fb5e36cfc58591
+e227793f39b9d3a9025cb10f324a95c29c488724aa74812366ff0b118fc7
+19f9fd0f202a040be47ec99b46b4dfc3d2a17902a5779c8d52b27231a1bb
+5cd794c838daddc3e6824ca8297ba669a818c239b389400faf17aa04b802
+f763029edb9784dfdc42f223e6496a938e613463bf9bbbd59d63300a9ad7
+4e71865cac4b4e81a5864388c3886e70799c8989188341f7d17cb514cd99
+3b211883f171ec6402cc361885f4f4b110757bb3e52941a94bfaebb2faa0
+3e32eb72e25e31abdde82c2a9015478afa0f434ae3f8b97a4bef598d6eda
+44ffe1915c26ee0e8339d2d45a6a080550f538ded5542c8b96ca2f596979
+8bb6223e460e857516ab5a3323136ee8fc4b0556a7c39d0cf7acb45e48be
+4ae9db325e4750b73289e36a61b301795bdb2ca2a8b933be1c09fd0cd2cb
+8677df171d36ef1519a2269b21e4103b2ee151c513df3e10b2a216d6fb22
+18bf2005fa7e0f0563ad96661a7f55e1b5b991f8ca285651b2683c6a7c9d
+2d1941374989b06f2e9b42a6af60193dc758dd8e9fcfc7c1aa06eab47e81
+bd79660666defac0c6b9e484df9c17a61ce7a61ef73150e8cd406af6da17
+4d9c2392cc420eddda40f975ffbeacad8ce1b4e14bee29ba8552ff03376f
+c034784b38dc1d0ab7bc53943d2545b03d39797af8d58d6dffce56a353d9
+bebc833f04db321ca8642bbb7fcc63ed2349ffa08a33a5d0d78f4fd2c5ea
+4258e4671e362036f1f67fcef9d878ae2c203fd9c05200c59cc98633e65a
+99d912ec51d6f74500d5358b70e799a6817f59adfc43365d7bba1fd6766c
+5c8e76248daf3f01e7a8950fe875d657397797a45e7f99a92887300b6806
+b86db61e03c4c09d6cf507800aeead874a94e6f665746752937214302045
+0b19cfa8db69230517183a03a16e5503882ea1e419c333d3e3b73cef6762
+873ac06bec34c3f736494483442619f5bbadd86f128a5a40b854051893ea
+8d31dd6656777ad4ac2572d17c6fb21385b053495d1270e65d78334a4115
+2787ea89b86f97e72718905a11e9c5664837701a3c1c65ccaf26aebe8dab
+c1207d5da2079c37883d9235708f370203b3b2a8ec3a5bb35fab93dae115
+aef626dc44b67ca56fac18caf1c22e6fbab93564829a75776630b9c42513
+721ca0fbb0b402f4d1db8f701d2b29fa60162feaa8a167eb3113c6f57036
+e8361357913eb24dd38dc6d3bf4c3176a07ffc75cecf8e5940a310f79a8e
+f590844383d631796ade04a91144d073a9413cff34fb454f1fd75cfbe5e6
+525c3bd36ddab80138f6c19aad7417d47df1f1e0fc958fb190a8205b5321
+7c43a4dcb0599be404473d6faebe7240dc402a0e0caa21b56a601b154524
+f44988e5074c71ae8e1948bb2a2ce72fc24cf3b1813cf7408a6b097aff22
+f9d285134d09b7053464259531eb7b270cd5f39f81bbf41a36420f61e5f6
+b429036bbf20e27af1a437becd74c5bbc25ee2519402454fc94d430636e1
+736fe65a643d9b9d21c9a54eac5a8fed51ff60a47b85a0e9423e330e00cf
+220c23e056d20aec2fca3e6bc7a61a8366eb940c9bc99fb90e8704e27655
+20335a983eccc7e20b13745c4b4f30a842f1ba64745718c152697c688c73
+6cffcf5cc8eb5756201560413117a45ad3d264291cd51404f98448d31474
+d47d17d201def12867ba679f0e2605de8f3e8135ed0234890cffa68848f0
+6de427741b34c2ea654251ae8450a152538eb806ace3ecfe86d8c4a137ec
+c98c6d6cbdc191a5f8f5b5972c70b4896960037b6d4c7c63586a52d5eb59
+47af8c192eb980d0801fa670bb1d08740819f9da1dd9e153010bf9580a1d
+0925d8327ea1b88db8d934f40266ddf93e5ea137f267847d826cd7999b33
+c795d0ac05abe2ec1770dd98eea67912f1939118defc9b379e237d6477bc
+91ad08e0046b0836fafa1272b0213dce990c90815f5b30d0eb103ac9539c
+2f7bd2280264cd95b4be84cbc5139a7628ed211905dcb92cbc3180ac9e6b
+b9ecc3cb08608b2395827d5729781dea49d328ba0c1b4cf2cec9f6bbc822
+1f2bbbb9d88f9e7682b9ecc06b9705faa8a90a51678183db1e24cc2c4307
+e16b3c20f08f179ec69df7a8c4261427f5886f9179c493bf2d0ef36640d7
+79925585724aba69df6d1b4f0bd2a356eedfd74a88bea667b102420c2300
+ec420e99b9ce8be1472b617e1255a7f43a0b52f11657f1a4dbb624a24886
+9604fe2062b98f5787d010723e520a4f42a0c3943e151ee627f3d5db90e0
+7747e1a88a53c4784c8d2b042b9c23c9e436d7d88343171161a364cd8961
+37a19582a00d774ef01c7c3fc9e9c7be5074c858d2bacd707a6a4f322027
+137d6ca0421ed9f9c7e7229e867678e5272cfc7156a419e893404ad7dabf
+a5d8b6fd0787cb4fe1a901c34dd931f1b64f0c470ff807005fb66350d0ea
+eb84ebef2c2399cd14a4454ea5004bddd99988b39c4134b92121ec77faee
+55cc716eecc58b594b39c41dcab308efa4458ed71943ec5805dcd0194ddc
+1ba04a5d3d42d07ac62a907ea25cd2a7e77aba470324d41dc1a3fe088388
+787b3312f472cb4f23a414fa5f7c7a0cc5d121d7642b5b3f0cf7ca2173af
+3f878f374938251feb3ce5ddd2d7703fc79a130978ac516daf70ae903799
+28bea3a4296f48725d578d2e8fb0f932e398404fa8a242024bc011c0ae81
+7b92bb104712253a5d89c543a744332069e33ca08bd133211d233ef799f2
+fed6a20a9073021e505def8b79e1279dacc062cfd4dddc2e8e0a7fda5dd6
+bb5a745f99cccb7ec1df532308da3da0f236c74639c280ea649b2f7ec27d
+24221470b642567f3b2e1cd0b3ffa65c5ac986b557aa9b444bf470380435
+abae9b51c6da7ff753810ca7938d8a1c47d2b41fafd236cb5998f3ef365e
+1f700bb257679ba3a82e235a3e97a667a6ad94412839c96dcd49dd86ccbb
+6df8ad01756b311e9fd57ccd2eb2f19f035e214804e2b77769319a5389c2
+35f3ca2a73c616c9ef0984abcba167d7d652b330c68f4f6378aba69628b4
+2d59eaa2a7e4c782f6eb96f6758d17d35650b15cb5de9bf973b3b6f67c1d
+f3285be8322fc2b44359640a3ba5d6d7b96142583a00a9a0ef84fbf14046
+09ad55b2aefe8c5c8f58ed21623bf765f81dbb6cca6d2a51fb7730a14839
+392cad6b47f5e03448350ab36a37d9ff2b9dab69be5196511072b10cc91f
+2e6b5160b2b1bd112e6c02d14063a9bb46977b0d4bc79b921fd942f916c9
+c5708e0d133c8309de2f6ee0b1afc996c889c36de20fbbbfd32878f477cd
+7735c7c3fa59e9c46e654ea20b4381d9f6c6431082e6918d532bcd539284
+af0333a783c9e7fd4fa1e4da5ce8fea2ea4037644a24532d65fa5c1ee982
+89e4b9abaf71a35d308a9b8c337f70babc5fc8dbb0327143707ca5b675c5
+2d3cf09f7a4f667fcda03d8c82d157e661517787ce6bfb35ea772de13c66
+2bd24b74ff9ab0fbcf6635d8e06b54b5b3125d17ae13d175cb7922338ec8
+9d1159fea2110995ce48f7d2b094f06d11d59b3a64a44a83d48c78855e47
+21243e82d9858401b094a236fa0a90d61863931c30d13b9bf33a35ac0d11
+a999f2b4dfba6fc187f8c235a5217d777a5a97112e7db6a8a4b06b07d9c9
+f41820e233c8b58b9e47ac56ad1ddcc0b35dd03976bc776c6ac3692ec0ca
+f8c75ea7825bc84156468ca7b269d890ec9d4a365b0b31d2f6530185d5e0
+2acc3ce14eea55ebb5667067825a8682e135d23c78863d32065ddcf1a755
+e0de6dea7220d1a28416b96db40b1e9f159aeb070c9a9515f301f162b0cf
+e32c6c89287de6e2b40458e3393826189a10af8517ff5a10c41c9d05d999
+aa9305a2ee8e7fe46076bc9c5722ee0a140a144ae383e84a8abe70af5d29
+96a0a896cd499caa0ed7867e7c3aac563763216e7769d12218b584d853ec
+01db93ca22d0c8d6b286b20b6b26d6ef19f2cebe7030ecaa68d069fac7a0
+09d61770b5e8f83024a99142f59d88297cb8d093992c3c6c11b043b151e8
+20df640407d8bc829bfc196bf2901e63c6f16102d03ffb7c54a7a560f5f9
+5cf8379f4a2eccdcb604bd553e6157b4381940d1b3c768dbfbf2618812f5
+7fbe744b3d8ad680dd9223d8bf2412ecbb614d05b485e3b4669d22b417f5
+02cce2d705c208b15fa83b5be77ccfc1c840f385a58ae49fbe6ab4e53912
+473630e0cfecefab95ebc632a2b10a2103bfe801ca0302542080cfb4cf4d
+4c241b1a6c8d28114516e3f1bf39dc02db73e6d9a797279acfd79b02a71b
+ae34860dd0e11b18954129f8dd57c039bb7063a4c92f0f6a1e25f4ae59d6
+6c1cc6b73a79d6a56f7f2a8a64d571caa8a760f4f485d770d000ddf393ba
+784bb27b781c47678dd78ae9b5d5e8b57d163c42c7a55e4aae22061686bf
+aebcede728ff2f65e75955585208c176d100912836b5200a79062d4f09b1
+ba9465b0e937e289160ec543a4cedbbe0cdb5ecfbb4838138ee9e1ac757d
+3c5f04fb6b510b389e2f521759e403bfc8ec6bd79e2d40bdd81901c10dd7
+4620acaac9108940daf03af23f09d3c8b785db562b05e597056406557857
+e96fc8bea53c2c2ccd0ea6572abb0acacfe29e737173d665ab6dc2995f60
+807aaa4073a183aed23c26c67eb137c937999fafc63b66a021125e4ee5c1
+a745ad1fff2bd828dcef392052965ce0e9af7a2c88d730fef69da91083fd
+83d9fe9f73d42a8dbdcaba85b0fa93b210dbf49cdcbf5d4b69e07375fab1
+a39038cc51f66f0b10eebe0cc61f697f7025d9755830b2d65f1ad0db91ef
+ebbfb578053de329935bb28d6ed6c12f748a2f70458990f04d56c35557e3
+8bc5d2e5de7f52bcf00c3bcce091aaa8852d53ac686f8f407baf3f7c8968
+69f3b62f44a5e2291aff9d30d7b5c663658a41add74562dbb0f1062f564a
+9b907846291700151de04c1a55cb945eaa2e7a709218ec56d1becce1c0b7
+dc41d5f016ae8080c3b07311590a0def35337fc3c844c0ccd04926be9fec
+509b1255ef12f368d20601b1ac8c68b0a935f987a21de0f8191604e921ea
+0c04b00dc188fd73499852dbcccd4119ef799472b353be7f7dcc904ddfdb
+920839f3d4a13bb1796f2dc886f31217845f8d7a543aabbc720311fd0e6d
+a31ad3daa06d5e7e6270a34304f35ef170a7abe733428e96b0522fddbb5d
+eb35aacec147067fe066c9ef145246fa3d444d176c274b91fddb8a7bd7ff
+7cc7693c25895bf931eb321dc9d79f662a17691f9bd1662fecbcecf6d1f9
+cd8ddcda56d19811f05fa48bcb492feb355b0ec7c04d6046549c56f7799c
+2cd0d9dade8809de7d510702e525ad9cc82c41b4fb36218e3d72e905c507
+159076a9c0e4a008ccca17bd594c69f5eee656426f865fc1988d677b72ce
+b710b29a0aa8f8337552ae30e93bf7c6e5d013555872dba4737dc5f08c0f
+efd428c66fc8da675373f13f89102688977e18e14dedd7f3b676256b0263
+b66b013617d9a026794b0d6040c23c5506a98530249633a6beec46117c96
+ec036eaf6439e25b8e57754af5ebaaf9b57880ad4fc93f002fb03e9fda21
+df4acb78296b0c49a5a852c134c3b10755177a0dbd6c54ea7a2b9bdac62b
+5d7f3da649df856478e4baf97899e0f891a96536c283f5c81200c51c6ab6
+77285450c7f7e96836b6da5660f6cb76782ddfc64b6fc348ebc3ba4a46f7
+19176296d8c5a31132b3fa7d935a5d777c1dc84d669d564cb4fd689a38ce
+680d0b3b130caea0be43864826d0d154019fd0d865f1c389cd367cb5248e
+24640eb6f66603e50581f6fb5aca6cfec1d6dbf4196da10a5e1ebb14e4ca
+0251c4c8412cc1673d6e7a9666b04b090567efa0b830d2362fd384cb0303
+8a40290597bdaffe429bb89fb66b9dfcfa92f39d92a8baba7266d144ac04
+f069093ebb3fcea961ba4497d3628ad207e0c8c4fac0e5f3f2a663a8d05d
+b6dc33b890ae13d84dce64b495d24cc749b121659373ca31cee09bff2e9e
+e5b62e89d5faa4482a75f341dd172500a54b98fc108a69a3ea94db696513
+d4c7691e0095ed3900cd4489ab008b5460b34ae8dedf3721c60de7086605
+6c391137cf23255c565bf11403bdeecf8bf39ad5e4317a4bb37003b2e7c1
+400c3b8ed7f63719bddf07908dc2decdb0f68e8ef722851c4420303f6de1
+b5efc9b2598732fd1f2cbe45a504bd7fbfdafeade3add7274a1e875aba3c
+4e0abfc6444944b79f95b5009560818f7a0599e5bab4405378fadfe084f1
+653e5a0166714047e8bd4e4cb116596d8089bae9147ec1d62cd94491af75
+a1743d58bafa11b63b447c954a8d7fe11d39d969feac8fa93c614f97807d
+ac62cb7a84a974a0fa555a2e3f0ef662706efcb828ef72e2ea83b29e212d
+f89ffecabcb08dbb7119203c4c5db823bf4e8b698b763fbd4d21e57940d9
+1754959d21f3f649d856ac6615eac692ebcbac555f772eb6ba3cece5ebfb
+cfcc2f3d8dcad7edc697df93aef762cd47cc3ba9e2cdd10940be676efe7a
+a3749170edb47b7562805e3f8bd978b18057c9110ff8d19b466ea238af32
+993e2d3021745b238021f824d887d2e01a7ff12fc6f084b35292f4864579
+406c0f61d0ac7cdf7e4770b424e2ccc22353e6c82bf8ff172973df267ded
+bdaabc2a742beea02e35b9b253f98de9ca131f802deee2905ca1a6dc4608
+19a59b4a4265c723007d0215fc8ac2a91ec5f86cd6aac1e370a297103c3a
+3cff58c7ae201cbaaa8a12c93e95e73974f9abcd678451b1db02ebb2e10c
+c5abfa573a2ea4219fd1851765649318bb556b728d432ec05a86e9894aad
+9cdca63d08642655801bb37f28b6e11b958e8e800c8d521ca4aa045fe9ab
+ac02dc015d18b1901d519181ef60227170a07f3328a6d5fe4c5aedb35fc1
+3dbe86564a9b1dd4c7ec648880360cdd1742ed4ac409450f1d9681cb5e46
+5edd1de2a2c7f8ed63436f98e849504ae71bb872683ae107ad5df3ca0b47
+a5b79513e02d7c540257d465ae4521cb3449d79c931e2ce8c5b0a0a4ac88
+cef7b9e5f92bf721ad51682d6b6f6c14747f78eaac1891fe29aed4eaf177
+e3d2fc655ae889c0c30a3575a76c52e95db2f6a4d8ffee9518391954b92d
+39dae4e97c4022031f8ab390b66ada6dc9ab2de4d1dddf26ac4032981a69
+08f73d34b4849ae28832cddc0dcd116a47d9262b0f93c24fbfdf8a78e6ae
+ae3357f3fb89530854257a9db773a1acf5271fc4ca04a06b46dbe661ca11
+9f45e0080cd129e1a7c23a33f1c48af960761b117d9d91fa5a0ed3e47865
+b774a322f7dddfda2960b91fa7ba20c8f9eb213251299ae328b28ef54b0f
+55fd54f8047c555e4045cbd70964e1c953e471408e4f25fe8ca7009bfe44
+0244b1e30dff518ea7ce5078027baba4e07ecf0ebecb497b4bd88f1ff72e
+b261f6dffec0ed895e237b5608d31ef479e8c9ae9003039a5fe67252ee39
+774e1501100c0fcf154f5c5c81c70539e03118ab91f4ce247f6132d46346
+bbbb126c09d7459c1977e6e367a0c83d14edf7dea081e5f795a7c831fd1b
+325b33674ec9c2b68029a0e600746329ea2e1b9bdd5cb2b140468e53c108
+8e8f2567425443f8146ec37101fa4dfccb0e032fff6cdfd76382463551b1
+ae8ca6cbff0e34a3f75ad400a9573217f8cbb00a6d59ff46e48421e97091
+cb17f53f20ebeb89609ea55ed6ba4101f2f3ceccbc7ade21202439ef91d8
+a9a783c22de7e6601b50c4342e094d0eff223494489fa92150425da1b432
+908423fb3f41e0b115ec1ba592a4f920d15610b9fb33f9912aba67912d05
+1ee00a13282c1909a3a56c4ed06f2f4d1739dc296b7492aad0446f87a416
+c6db4d42b504dec3a6756f3d0845ab2d2e151aa5fde12b31a9c3b5ae1cc9
+d97192bc048f00dead66940004281c4d5a92c20b1f77795cb4f98b8eaa7c
+be16f9b9d4a34a1a53e0a0deadb4fb4b20d9e8064d3412ea8d2ebd259b8f
+2f04bf4bf11a5ab7883c99943d762549c3d5866bb6ed85a0e862eafbcfc7
+03bf4b77cecc0d65bce4df33e0d65456397f231f8cbf66672457cf539817
+6aa5292fae24695009e55904a04588659a3a23fa11989b925705ab45f954
+6f862b0e176fddf75b70d9ef7389f750becbffae25d58a1252cc04a79e13
+fbb6a666fd87cec5562c3e14fd78ad05be28ff3871d6fceff5aa8965bb65
+67ec76d105a6348e915b27767f5010011e80e0e2f9c34742a4eeba369e66
+8faf086a45ac9bcdd76c758db01a78602412a4244c759ece0b963d9ea58b
+0efbf4376bf115288803a54cfcf78584c8af80da2a3324096463e3898285
+57de6c6354444b12a74d5e66053f6907c48522cae9e93bccdb4632131add
+52eb374213888125de71994c31dba481b70b2e4c1f10b865d58ef09fc9dd
+2ca7f69bd2855895256caa5dd6bf7d4d8b341d677c56ca08fd7ba37485b1
+444af8be0dcdb233a512088936ab4d7fc8c03139df396b7408747b142782
+d9406db0dcd31368d2f23ddef61b0da3c0704e9049ccf7f904548c3ca963
+76eadf1ccf77f94c157f5b84f74b0c43466134876a90c5fdc2c53af70c3f
+f5c2d13cb665fed9016454bac1a629361c8ea62f4b2399233e8587db6e75
+a9cde3530f20a68ec155d275a4aa6f63aa5cd115244643b54911c954feca
+d57be2a6c40f1bac38e393969617b066f7d94e8b18dd80fccd0168d4a385
+f2f1489d1dd41b68d47e5ec66ec568333d1f584e3dca90f1367a990630d0
+14355be7dc45378aa111c319838edd441f15e125f928e044640f25ffdcc5
+c116c3f6ce0d4d3195187b22200808366eca9b508ec45e664e562186efec
+a97b22835d384758849605a01973cd9ffc1657b124950c9d9fa3e18b1a20
+7156c4f96f08b87824373c2865845d17a0dda71b1d69f5331c5676d0648b
+ca80a7958a2aa034d7e1e9fafead9248e6e64f9ec327c60ae4f724e1fb95
+8a71e82ac3842768b27b506b5982311557432dc3f270ae6eab23a42fef70
+dd0d407a02cbadeb7b8b74a2523cf46a5f61e52b053c2007f75ae053a96d
+e00646662d027d93f950e516cddff40501c76cd0d7cf76c66b7bcd1998d2
+7a19f52635c8e27511324aabbb641dd524d11d48a946937b7fa0d89a5dbc
+4b582d921811b3fd84c2a432dacb67d684a77ac08845e078e2417c7d9e08
+bd555c5265024aeb55fef4579b46f8c5e79770432c5349d5a65a47ce9338
+e1b599328bb1dff2a838f732852f3debf4bb9b828f9274d03d7cf813b123
+687c5e78a26310d87870bfcb0a76bf32aa20e46f6b2826912e562f503aed
+11e427b7765cd2a68da2ec0609259ff14f57c07963d075e96f8bd2eab9a0
+dc32714dd8905f2627c6d6f33563436bda2d7fa9a976f88947b84c72f454
+bf0b66ca84470375d2ff252b4a2df52ab613d0c8ef0465ff1d809ca82025
+c2122a8f44c56ebfa25690bf6a05675ebb8634ddfd24c3734fe8cb32d6d6
+c69c72a4951cb959175770b4286d383e7a3f158450945c8a2ccf7e54fb19
+aa8d2d98a07f0c55f834f2728d89f82a598269750115a02287c4d415cdaa
+14e1d9e7032684002f90603c0108dd26b40fb569bb21cc63d0da7e9e1873
+9df0a9c85bc340d2b0940860d95571dc244628c59bab449f057e409e58ca
+cc3369f4baa8e53c6765a55620e78341dae06e5cdf2fa5e5ba58634b29ee
+ddfee7f78672e55f18a7debbc30862f278f83f4cc123ab591371f548fbf9
+bd24b3453b9b57051c2e67edff2104f3a05a9f0cb7efd81c1b1b0a2bbe95
+21854902526e5d4fa1b3be270811b972e8726623410cec7911c07f871428
+1caaead97c503714eaadb14ae5923f020093722df1b9d9c055d7d5f95af2
+a9fbc5ab6f6c2bd655f685534d7dc5fbb5ebded6ccdcf369bd83c644dc62
+84c2810495888e9d8f464a42228cdc231d5b561c6b210bc493fc1e7bfd66
+5a6c4055a6a629f571f4f05c15cb2104b4f9d0bd1b1f0ab8252da384eeae
+f5fd5c663ad7a2c29f65a48a30ed8de196f9eb8ea314c6e86989298146a5
+589f76f12664c8d008228b33144679d16ff564453b5e4e9f813191b6c99e
+2680e20a410949ac30691b1428a255b6185b7e3802e8511192e73c376f3d
+eb807ad2727fbb4b27538b3213da0746231b1c1b595a958466155835c537
+e0df4a0ef272d4c3f7f2ef011daed38bc58bb0fd7458e48060db98971bd4
+b24bc7bd0de92573a1c7a80a5fa2b34fbe50271dabeb83aaa4235cb7f63d
+6a6b399360df8b1235e4e9ab59698930044a98d5e083b5f5a5772309b390
+9e1ff2a252734b32fee3940f0e1ba61f54dd1d3f6ff0d57c9ae75a302d14
+b9dd9034279aaca80b6bd05c74bf3d968305a5046910871223a3ef8c77d8
+25d7e6d3d2809e76064c473d1cd7c05666040b6eba647e34588f49fd70a0
+3c937933a2272c938d2fd3aa8149f215bb48f3bb45090bcb9a6ace393a44
+f1a9bda2ad09a5f566b2e8887880afa45a603a63ffe7c188e3eae926a903
+4f1803368e773f42c7391dff1b9ce8599161515c549aca46aebae7db23ec
+8f09db0e0f590aab75e8eb890df354b37cd886bdc230369783a4f22ab51e
+0f623738681b0d3f0099c925b93bbb56411205d63f6c05647b3e460ab354
+1bf98c59f7f6c2ea8f29d8fe08df254d8a16aab686baf6856c4fed3ec96b
+0328738183dbc1eebb2a3d301b0390ed8bd128bd8e7801c89941485c3c86
+22b5f223cb07dca74f0e8643240044e8c376abbd8c82ff98c6dba9b6d244
+5b6cf4189d63c6acd6e45f07485a0fa55eff370da7e71c26469740a68627
+a3c297d2bf215121fb67815b7b9403aecca10d21e59fabcbe38f5ca66e7b
+551b22e28f2d1fd7303d15a42c45bf54b40ef7fc93060ae5164e54f91c55
+20bd303a98d0667a02a900813b260c0343021ac01872fd62cb6abebc7ad3
+a4456805159839ca4a3e35db586221169ded66f852e8974e3815d4d7659f
+6a9bb93585aaf264f06cb6da6a26e51683945224158ea69719b8e4e36eb1
+01333aac974db8f84b051724cf245fe7a4c86582b5dbb9a5d9318180e33b
+8d92c22c44b0d18f8ca34dfa4ee9693c1a26fedece01635fc5eac1fefa81
+32458254ad46dfdfd2be12a1e7f32f3728f286f1d5d4394424a073696b65
+e3c459aee9310752231fa703faf35e11796c4eeef698f4109ca8c46ee322
+5dc2e3e04fa787188e583321f8410b68b9624ff60679d3f25c13e5ea7506
+a3ce8d0bebb99d9a959ad92d8cf909988d9250b310629903d6bfcad4581a
+504b91b2c91889987f36d6fd0be1d0ee5aac00aa0cb48d78a1f7a64a777f
+089573ba79452efcc31c8258fb317369feb0d7ccd48cf13da6d1ccb59a4a
+48ea0b398e590c1169113fed81639e13e96aa268d99cfdb7aee977fbe85f
+f784853a06642b5521ae0a7f610c9739af31ba7a5157ebbbad999e23794a
+d2cf25af987dc85dfa29639957cf28e7f2b7671188045130a6e2785f8d8e
+30e91f0f68c1cc9f2de902952730003e816e4f5703db7a97b4c566f80547
+42fa77be563ef681a4513b9a68b2b0956551c74545cc9883428dfa72fd5c
+4eee93256b26bc86ea34f7427cb0c0cc22c0cc343f739c6c0c46d0923675
+5e04d70587426ef875f8c89ff8492ea23e4e4d763b84a6437a440e69eb70
+65ab6d8cf5f8444a844e6ef3d158b451d121daea2d0e2b423eea24254226
+7eff1b4224c4e80af2a7becac1649e4bbef09f39415e9b1e3750d7ac47a1
+068a4f5ce30840b00574eb4e683e3ec25f6e690feeb0d354568efbc354ba
+813ca1400734a67693af127b0f636d58b83e91548f98e3d87da7fd7cdebf
+f3ecb4b9272d1c83d4980170378d32f1d98b87c440881af9ec052510982a
+0c02ba6743bdc7691a44bae5e044c25304c1a2525cf2c0694494a2e9aa34
+f36af43ab288807ffa4bd418ad51d98c75f2b2f01abfd834d3305682b6b8
+62ef69d05962aac485bb4f560583a5dbb74e967eaf6d299160753ec32249
+bb1d9851d5441cb0c624208e69dc876cd8841a66976b5d7f9c99be68363b
+8112d33d971f2c4f2a1feca88ba1a794ddb725c5e2e2c248082231059aef
+729bb5fee5006ab8809f63e162fc0743c047c7984a9e6333b433fa143d73
+72d4a74fe37314508e04f54dc7a1445e2d6178ec9c041d0cd4fda5cae830
+4b16feb21f3222261c293a8b058dc708405c1a97ff34eee4ca69ff4e1ee2
+a03380d52297574e3aa50c8afb826fc94a14e8caa9ba89d6e92913be9e07
+bf7ae011e6bd142d8952d9c2304735e875d1ddcf82fa9fc0c6449df2acf0
+d5f6cff6d21ef6b2d29022ed79c4226c97f163284f2311cf34d5b0524a1a
+a446645b9d05554f8b49075075f0734b3d1ea31410759c174fcc7305d2c1
+d7128781043cba326251a3375784a506cf32d6a11a4876f85ffa2606fbdf
+27dd16d64b2108d808e33c409dd33f6e0c6079e47e7196016f261e824fba
+b0e4f91a189747053e648ad2d942ece8f582f052668b63a23a2fae4c75a5
+180db7811aac654270ec6e341126e3561429f1d41fe7ba3f1de9f8bbb8d9
+fc5cebdef869376a2e42dcaa578c0807835e58d75c39f91a83d5c1eb86a1
+b0f7aab991f65eef030f212d38d10b1913bff71717c06c78d9a1be136f21
+4be157ba11ba309326c55c23ae8512646751fb82ae200c06bd2e644bed38
+c7cee826cb587ee8ff378b7fdc00ec316bd4a9c24e2c250cb3d64f8ecbb8
+7f4d81626d7f1e4491908bf17c48c84bb1736693eb4d0fe634484cdd590f
+a40ae94d44f348ba683a43004b487f047745fcdfdee2e913328a11a99530
+9bd117e0e5be4fb25d176d59dc2b1842418141190ed9ae1f33e5354cacfd
+a5e4bc186119e1461bcd98517e675276ddf0296d3b3cef617dfa36b4759c
+944fd721e1bf63d45cea90b5817a40d153a2f779e03487cad3c1375425ac
+8cbabf7f754d16cabe45c65f1be4441908e0969d5a5111c931e724537dea
+7cd3fbfec9b2f7d3efa747bf586e9218c3106c49276b89fa28f770fa0644
+fe1f3fe3adf07f59c755a5b39a2ac1d6f23c256a293bf3b31b6b9cf4c622
+b188d6e7401c038657c78bfde9ba09f508f1bbe3ed79793772cfc928c4da
+519f7dbf3ff7074284437d2de8d7b7c78829642d924abacf353119e9088d
+14739935a23667c432806085c3af71ffb7c5fe6b4412b9b1044c1e62ee0a
+a5ce7e0322bc65a8c7d874270d84136526e52d0c7f9f93199c6bb7301216
+a19bebcef3c5633f21d012b448d367157ad928e21f8e471e46982bc46a7f
+df1bf816a86dc62657c4ebf286134b327ce363ab6a66634eaa2a42e99034
+069fe1302febf06959eab8e7304da4d94a83ac1650a02c38c1c4b7e65c43
+e3a6fb0213e57ac49e58721a4f36996069caedefeb48f1a59303459d5873
+f3bedcdb9d00c1cf31130c27b60928f210e1aa5e1c8e04b86d2049f31265
+9198fa646c53afa9058eb8ceb41bda65f415c79ac92af5790b176de1d300
+f1c06b782d584f458dbd07d32c427d894f84215a8e7819e295ee98d976d5
+644f11920ff2f49cb1075c3bb42b9fe4b561362902f11a75669b7e7c4475
+b65f1ae48834cd67816eb63b58cda2f50bc22eeb0cc965569b476bedded1
+2701668f609393659b266bb0e37bb27afc90bca271366e34754383363592
+0f9a3b508aabfe8deef585b07a992460c592a150b325b1e50e4214a2f483
+e9dfc826c54b488493a96eaa37276f5a9666f0a5388fe388263d2c0cf614
+c6cd01571da4389f01fcdbd0ade1c435d64c5921b5bf7dbebd5268100a03
+1e1abb8cbd83873089a9e08cf80276c7e30d2bb40280278c29fa818eb079
+87623b1cfe13e0b01e27be0a8320b69b5afee820f4705202158b7f3059b3
+655bc28a754d088fde23d43d6a9389da8bc1cf3e8ea1a6f4328c196e655e
+42184444d8c0614c7167c91a492c24c8357794c61f5e47cdaf4b38004a5c
+8fceaa8151e929328bce1b8f67b22034f3f75e4d105283337c3d460e7d99
+89920c43f5e1449c74ad6ab5ea029cc6e497ea60068451c4ef2132fb87ae
+049077a156c868b768df4a4c475a532e2a22d999931c64f8bcc18f51d25f
+0f94fbd3e9e6c094f78da062f80c4aa2b86fa572cc469e629deb4ba0c553
+55e8422b562ed2f694d0e8e5540144e30841d7593b255edd4a61dd345d5a
+00e411d2c50d64782a3ebedf945fc31c00d2fe4ca800f5aeeaf12ab399db
+956362e979bd7ef0787188e43835e5389ac444d13204af6bf1875622f175
+09f32015c28729cfa3b3cca90308eefaf260e3fd9df10f3e76786b8bc0eb
+a30e8cd33689aabc55e3ce387cdb89a30573495852a48009cb58a0fd34bd
+da911159ccacc94698ffb94c5f45f15ecc9e82365174cefbe746f95eee44
+7a33b4d823487e203478eeb2d8c4bc7b743427778249c56e48fe17d0a501
+7b693509ddfe1f42bdef97aedcc26ceffa9357dd985cdf2c70bbfc987354
+6f0aa7df227ec42f9ca2482f58809e3f9650444568c54d3520bd0a7301ef
+48bfebef1fc4332b5ca851fd786c1ece136fe9e575b69393b5aec2611903
+fae6e7a5046e2ff350becb8700f209b1131044afd32fed1bc1297b6a2f29
+6ec3b87f170e92aabacc8867360e4dbce9ea29f0c1df981f6cecc8986767
+0ccfb4c9faeaad7ca9029b8ff0129fec4a040f80ead041b3bc8af7526675
+ed9e13204e64d76440a097d77c535d34165bfe9ffcade530abcc75ae224e
+890d5c110004e218bd827a02ac7340e18bf3684c43e664e0a37d5fd4fd1c
+4d4489d25a99d542c16e06685652cfa3567da4eb0cb517be1482939da0cd
+d0ea3519ad1e51bd9dc7b9077375a8cd3b5de9888697e853bacddbbdd1a3
+0e442e1d6f2d652046821813d0cc0e8f16c97cdd32daf239f5b2b65ef620
+46f6e9821b2e2ec539302747795fa746318514d38bdf0d0e490c00e114d5
+03e7fc9a8fb83b14337a5bb4d640b52630f5450bb3bfcf7cecfbb1ef5192
+ae401265450db197bcfa07315ff95a809bc5fb4249e3a728a817f2580ae3
+50d8d6577f79c883ab4a3119d9ab98219aed0d1e826023a66da814396058
+d95e52d9af8bdbcb0454721f27855b686d13bdb473f650c9865f3e04f08d
+b10f5256a3e59bcf16b12a84bb7ef3b370647cdad5929b722a05f5b3669e
+14c232bb82fcb9c1dd8155ff4515f4e83c895cafb86754e896f38e5f3beb
+5d29f1bd99cb8a09c5e50f412f6d8a773b79021ab2c4831aa663c5defc4d
+553616874dd5bd8b75c7a2af7d029aab5a72528fbc4b5ee3d30d523412c9
+60b432434017c4cd68b2062d28f307fc287e11663511d1a6b52143afac0d
+ce0f7ba3f326fb707fb8d2c985dd60090e6664f2344e098a7a1a6448026a
+2ee651e8141cd7786b6543f512e4c31d25dcaf6652b1eb52706300b771cc
+0c49295067befc044ea46341927123ad4b7d094784bda7fa7b568853d0b6
+1e4cc39e1abcc9479f91a2501009ae34ef7d5ff56205cf5288503591cc55
+c48abcc78daa4804549562afc713a4c11152e6e4331619b2e474a25ffb62
+7c46112fa4259f07871f8d6882e9a7ec62d20a86a0c502815d0a8f3f5ce7
+cb4a6a74b6db8e17d54bc919b82c7c729cc05b98855b9d8a0fabd8a9bdfd
+4333f395607631f57c0473be0fb290c4f40a7aa6ac49208570ffa1d0f849
+d4871ebcf9ef6f5106301cf54ff8cc9918d6de74d519fccba58bb1c21543
+f3bca9f43c211b2e5c233ff6dff2c9b56d3f656f6070d13dfd0be04653e4
+98c670770e01c07b731ca0e2eb56e608828fedaf1a31087f2d43cb4c0074
+e576769b0830577c86ad5de48ee216df02d7c4e4ec231afd8e76c608fc9d
+06cc86f38cf4d839e0a0829902f56cf2f86f08b975a6bdd0642d6b4c78e2
+57cf9a4f52646a952f6a220c36c91db7f44c7f44bddf33328ea8cc01827b
+5f2d79e3ee6c514a4f8597a847ef5f32c6400736e6ade28faa7bc6e9c6ba
+e4bbff236fa6dd2b0ed23fc77f92649feba149f82488260b0bea2a4fe1f4
+65d96d8c51719e5e10d4c17d1b67e700aac36b1ed55c93b4b2604e72f51e
+b30fbf5b64c6fcaaef764639ebd789f82ed354712c7f9fcd1df257e14c0e
+8fd59a0eddab684bb1b4176d79b22ad2605bf534e4b8fac2272fbdeaf210
+0424a2c5cc65f8dd5faa13313dd926128ed466046ee94bd3eb41f3ea5505
+5a70603a2ae1981bfae8e77d850fc5a5bf1bacb3df9b7cbce68ce7979fad
+a73c2900526b68236c6d37197b0c521c5b1cf5cbbc89238586eceb99818e
+aa47ca94ff615233575fe83d0d50d734351e0363030a12300f7b20450946
+17bb209c346ac1d35402b617d6260fce04ce8b3231ab5c05af30b0f3ccb3
+3616d3df334c8d963279537563222dfbb705c3e14616ad01927f952e6364
+4c4b7fa44ac97616c1521facd066aa33b2296dc03682eb6a3b9dd8e5bf62
+53f10667ecb07bbd50553f1b211067f5cf098b64b84d94ba9ad8b146dc9e
+8e9be06bc14cfe0945e22fd819856d6996e857c0bb5f292defeb493589f4
+515700753885d61eee1b8c19e6e94fe2302c07933f949d6bf119d207fb04
+dae7bcff7578bf33d77e29611c7cf03b2df12c242827ec4c4e5b5343ca3e
+4f7f38ed337583e30dedd78a082f41d60cbad55d59dbba11af1bd296ed6f
+e31d2e10d3a8b5ea698e656ff97755a47ddd862d23309e2e6ed3e3e111c0
+2c3a713d782fe301dbaff0a4225f932576622d1cbae40d20f46958298d01
+783851c894f2712bfc4736d3802e548a704878e2d139348671fb96d0ddbb
+f56d9349172caef0dfed4b84d867116d91063dcdf9ec401dfe8abb269ee6
+0d646bd12e0752313e2ddc272d9f4aeb9d940987596ab623f9198765cec4
+62f7b6c540c9a70c9a872bd28ea62e056560b61ec51fc68eafe008f20760
+246e06374ae5a6bd2577217700507978811ec29985ab644e474e41e8a105
+295fa67ae05e0739e8c7fbc51104522934942f53e1e1df1ec2a66f0a74b5
+9885cf2c2fad1cab3e2b609f126ac8b7350d5408a7df9ed5c27a10ef6505
+6f0d877cd7bb902977ba93e6e8520d2d018560ec8143876ad0dcb95b173d
+af72c0d413bbb5541f14faa57eedb3ac2430e36911d2f486d9ebf9cb6745
+2ccc763e1e46e7a4b8373e06082176a6c66d045e18f90b4b2ad15802f6ef
+cf2130cdc627601ecc19887784b6de7fb6a193bc3d057ace29f74199acae
+69526ba6f7a2c669593f9d0849f12e37201c32c88384e4548a6718cbb2ab
+714ccc917d93b865ac7d7d4dbd13979843f4f5c1f8b937ef12fcdc9aff50
+f09d2625f4367ee70a98772a273d8919952102aa03297e3cbcd876da5abd
+2ceb162b8fe1d9a22ff694495528c09a8819fbfb6946ab205d4b2424f6d5
+6fa1c704065cb64fb2aa0fdf291fd5e7daa38667e6d8e889be7f4c453da0
+59c492cd25fcf4a03a6995897145273a66cd6ba999138bc8e2aa7d080f9d
+231497ed28a9a27b6b0d4785bfaee46fee71b26d6839f2549a14e7ab7347
+0b6cf368d2d49e74c78d93477828e4582589cb447d795181d3f13dd8ad52
+3c750df8f19b3260c17a6598b406472a7204dd26c5988911ce9884de9a1d
+ce33d834becb1dc80efb07f32d3ed6c2a484c5d53746071576c3f67f25ff
+1558986fe2dc2265b4fff79c07e3f4c6c0ce8319e04c14728ed722cf214f
+65066148bc817753dfdcc0950bf80dc515002e1a92e7d8936e9b3aa9635a
+a6d512c68aebc79a62a6bd17a411bba7684e1f06be9bc3d1aca25d50c8bd
+1d75597194cf87c9ffe04ff28bea91b5b9521fd356ed9e036466137586ee
+f0a8795486438d0d9707cb2854f12963929edac394c562235ca71376d938
+e4e1518668180b857d75318bc22e9f0683749047e7649f9e20b35204b6ee
+60c0d47bebf53179a083f0b4cad5b3327a3faf2cf03753e3e46c05773629
+7e9bb305f603369cbb568350b2b5c6d23a35c551e0ab28b082e321ef4ed0
+e2704d35c75b4750af782160c2f2e9aab0e14e541e95b64ebedd66db2c12
+a8935a60177cab634e20a8871a3a72f4b21c3a34d9dac37176a321c2ce3e
+e828d140c8445117e7fe4738000c30ffae8e2a48bd618cc8813e38fa0f86
+92ca634d1e56010987483aa0f08980d91528df3d370ac724acb238e141ab
+595dcb3da7a769de170edd5763078d1084e2ebefadf8a50a816b50722617
+c9539dbd68d9062b015639708dd900aecf4f15adb36339c05a9aec7403ed
+771f9f28c60e52bda3ba6902e06334036c1dfd66d35ed00e3fc0bebf55da
+416093b5cf512217c47f905ccc91fad879d63dd1380519a02025ddf15d70
+eaa1bd8cb6be67608fbc5c94796bd09ba35933f64c5e72a26db1ae40ef49
+af5e972fa44660588292b67ac670bf046cb1f5a7a0d73ffd6df862744786
+4a56393b0f1b4cfcfa362c74634713093161b29c94a2526b7138aa92fdde
+b37a8c1f30a6b3837d9500b340515f0412e681f5bf36e7869fa157df18e5
+c79df3e6aca924d7b7dd2e0d5b87682d7ea6913b26397ac180fb75fabc1b
+8e156ed542b9d8c83079bccd141c187f90d72694de4f6d08520d11cd454b
+bd3c2e6d259694fda0c8decc724bdd650163b7f6ce1181590c06de4c0dd8
+536aba318cabf54782c919e07c2ffa1034143175d05deddfcd7dce6c86a9
+ec9bf6a4437da474aac2dbce2c91aedc20043f179d5c9120f3dfb1cf6906
+c27f2ec68cd75035c283e1672ea90d953a23a1515c420b81c3270fa06573
+4d003eca1bb71a2dacdab67e44f47c266c2ea1776648b62bc110671e6eca
+4546d3c72c8acd956e10452c32532ed51bf3d0518467fa829efd9c896e8e
+1e5c7ff6da0b51e872e403470affc95f25e1d2b9b59ddb0472705e14fdc8
+fc2af16527188508be10d098372cd7eb7d62a85c8d8dd1d0f55ae3ccd0a6
+5dd6bf776dc187bf4de409d5db3fcc5a6d852848a251f4fb4e01dac5e9b9
+587fa8c46ce03689709008b34dfb3dc105def80a1b515abcbe06e73fdf7e
+7136e40cc922fe9a9da1726747e84427f288d934747b6c587490734906b8
+a91144ac82a57957cffab561714e1ff5148a39499dfc8cc96bf5d87ced17
+825e8f80cd943d9a73945fb8bc51cf1f9cb39c605491c1bb8f1c4139974a
+59471ead310d041b1ca1ecd5e9f92007cd8243cb3fb1ec5256444699a9fc
+ed6cb31eaf0912c16fa480a1cb4a8f4a9cb6a4d9a9903d1e2f674286032b
+489b8a23ac4719fe435a9fa2d79abdbaba740e69d5ed611421b1aefcd06a
+362ddbb7b79aac41e3e90657afc0b87a6e8c57ceef70a628efe19f568634
+50f47b5c6d95870039caa3d07a54e58df064bb5f59dbe9b9a2c7c84d7e0f
+32386309560a0efa2cbfa27f861b208b2df4a062ffe2c59c057296aaf5c2
+0f48ffc9ff0692f8cfbd6fc6ed1f3a14537ba40d7267e6b5f69c997a949b
+26577a9a99db3f53167355c4967dabd522292ddaca3c537bcf303ce76add
+eb99f6664227a94d6a698dd5a5d40008349376067d057e28e55972264502
+e035b1f5e33d7b3aeae016f9be50f2aa09aa138d15d7af3c1ccb805f2d5b
+cd4e9b2b5c288b2af4a25abf0a9093749377c9e8232ba1af17962f85064a
+23b0a13f11acbb471cc700f9f1b588f72cb63d3d1a95a93502ef74ed212a
+c452f1a84619bbdf61a1dc79c0d9ba29c7f19b400f682cf66f7705849314
+f5c8bbf973f2c53bdb060932156bf2c9cd8d36cf6271075500b0e3e6ad49
+958af46a9dc950f4c29f1ab5dc0a85924f7ffef259f778459c80118b1eb1
+ed29208d1145b21b19d62f755de4972c57a09b3decb0a8096ab025fe6b9d
+be49ae35394f0ea40d3693980f97f712b27f0e28d8a549acbf1da63518d0
+374941effacf63ac3de0523cfac0dcaeb690de5836741fe58917c7ecffc1
+95e7b560a3e763aa70fc883751bd60ea0a0f893d8e9fe75a66c67e202c24
+84f66708ae74413c0101fe0b5003be20881345d917203b582a247e6c74a8
+1d0479f317aba7b9dbbc0a92e91c51fbe8775a44c57699acc9da84ad60fb
+9629929d1edabbd70b4ef9887ce4ec2469f154fada42de54240cf3302364
+7c492ba17e6936a4d85e0751df0945463368a803fb40d8ded22abe118250
+86cfff1878abe5b100bc08b991cda6fdfd579332360f0c3374842edce6ed
+e43649d6702f34668a29bf387e647f96d78f33395e8d4b3521cb4fb0956d
+12c924c16eee798cde68e319a358cc3524c753177d976d4e14a2e0cb72a4
+80cd87bfb842060b1266568af298bbec58a717c577be73ad808e004348f1
+6aead32a3d57457376ab57197534d6e469ed24474a83618f3ce21df515a1
+22918f4b62c642de0c8a62315ebe02bcfc529c5b8f7c127085c2d819e29a
+f44be20fa077ee01a8d427bbe3d97a9d2bafd77f17835279bf135900aee5
+9bc49582b18d468bf93e47ce0bdd627775264ebe9e4172839a444f928580
+8c95895b7e23592b2dcd41ee82e966c26aa2143e3057161511796e980998
+1f2e4ef5868b3bf4576e3546e6407e35cdf14654bcefa7557d09407545a2
+38173080b4771ea52054736677a8d9749a2b22b46b24fbff93c55aa2274b
+8c7ddbd751bcaf1df00ccbe1f24a80622aff192fd6db2238db941ec44ae0
+dd73f6b2f80d89bd0aa30c038583deba14913d38a7b61b54522755e251b2
+aeca62033a39ec1143b2b960f9cb87f748428bec3243b8164f07d5ff72eb
+f2ef69347bb933241c2401a96ba5ffa3f9ad060c41f4e6bf7280af65293a
+bbae49d723dbc4be61d7e13f7a5931a697e7f2c6582dff416341ccf5a24e
+9a53686a1e13bbe0bb480c19a4e72a5e477bd29f39dce1a17f63f1e8c696
+d5f8855cefdbf7ce681c7d6ac46798ca9bbdc01f9ad78ce26011ee4b0a55
+786bb41995e509058610650d4858836fcedfe72b42e1d8ba4d607e7ddbbe
+3b0222919c85de3cd428fed182f37f0d38e254378c56358e258f8e336126
+9b1f1acd7f387686e8022326a6bbc1511ed3684e2d2fc9b4e53e83e127e7
+84da13550e593bbad1c87493f27b60240852e7fa24392fbf3f478f411047
+3f00a8fdb6dcb8aae629dc7f055d85341d119f7f6951ae612ffa7df82111
+d1ca48306a57a922cf4c3106f0b5e87efba6815f6de4294c7a0394087067
+677889d22a3fd86b0796200300d2716445078027fe0c0b05c86ac80d2095
+ae874324ee6ea3553bcb92fc1522a6d1524f6fa22b71598fbce784a10b5b
+61e50307ef4409ffb7b38f27800f2185140ed08fc4ab396050b068025a9d
+e4bddcad201e72ed9b41c4ffd4cee743c9c2345b95c5071442defc8ba5fa
+9c63c56e209df41d10d93135a8080f7cccacf67e0b0ddb3e0a31df32b83f
+290b3c536e9949973cdc80aa5c8a4feee20290a95f68e59f54050192de42
+f27464ee374e4d2451ee8708933b970402c90ca3070843a449d7c3146347
+1efa666a60fd5cbf55a47e4a3c5c318fc1af944d58d32690a2c7eeef09b2
+d94721896e1e3e76e44a8efd524ed5d6f5eb9da093d277441546c6828745
+ad71b6c13f653dd631bc6fc55d0eb4648b7bd9c0eddb13222542f2b6e8d8
+b80bfab4365f4199a41ac690979285d917de79359a183e6fc254b63e6408
+6d33e3c029f472f40742a99f92999f302f79994ffd615f1a848194cb56c7
+12146850f5e400303bf5bcd4e5fdccd1fe2edf5352d525cb15d8327f45a2
+6e3ac276dc8780c65724d28dc6bf9c7c985840070c35e32859168890d599
+a884dc2a90194cc2e9cc6a20c6c0ee11b20adf3aff01db48eb8dba7b0c81
+7fc10cf5a66e8171a2823a4cd22f0e80c82011ae56dd895ae2d3ebe84ff3
+d521c31453e0909cb9b1cf0b030eb6b7059ec38038cae12d0e1cc4b5b3bf
+e6c821faac9b8792441e2612aa1ee9318b71f9966d7d3a64abe349be68b1
+744de7b212f6be73a0e1eb2fa30850acc3d9562f989cb2d4fbfbcd5d3ef7
+ba55717da1cabf197b06ee4d8650e968518b6103fbe68fcd5aab70bdd21d
+66f09f96208db67c1b345672486657295a39a7fd689b2c9216c6b46a29dd
+1283bdba295dfa839a45b86c14f553ff903a6f7a962f035ce90c241f7cde
+13bab01d8b94d89abdf5288288a5b32879f0532148c188d42233613b7a1a
+7f68e98e63b44af842b924167da2ab0cab8c470a1696a92a19e190a8e84b
+1d307b824506e72e68377107166c9c6b6dc0eed258e71e2c6c7d3e63d921
+39690865d3f347c95070cd9691a025825421be84bd571802c85e2c83ba53
+841223435a9ced5dead103b470a4c6ae9efcc8b53331c61d0e1e6d3246cd
+aa1b0da347685121196a07e97d21b10ad34e7031d95c1bafa37b4141bf33
+a6be401129dcd64086885f4b5f1b25bce75a4cc8be60af35479509e64044
+d49c8a0c286e4158a5f346ef5fe93a6d4b0a9372233c7434a7a6f9e7ea21
+30c0b4b9f62e3a74cc5d2916ebdaa51a1ef81fceb6cf221e70002a8a3106
+bfbccc2d1809dde18e9607fcaac008fabb72e8c50244507f4013c5a268a3
+6135ead9cc25362c37aa9511589f18d812e6039490f9c599f44e88754ac1
+4f6c1841d570efde27958c7f1b2c68772584e1d12fea252e3a6ec3b051a7
+6faebbf6f5101978e24a9ca927c02065e8e49150a55c64dd30757e8a33d5
+2a788437a9181efb47414dbc22fdeda203d4122137bd045611f68314e12d
+1d6a5ec270c8919562c03e3af7b0e0deceeddbdaf3eab8fb5632e44dc1e8
+d46e2396b0236a46659164e33709415e7b347f7f7b87a9224a189ddf5178
+2cf66c9d385470a51efc88696176f6d3ac3b7b95fa074c981194e22981f5
+1d925f980393b7102f1f836b12855149ef1a20d2949371ddba037b53a389
+7617c257bbdfcd74bc51c2b40f8addfe1b5f8bc45aa4d953c0d1d5f4091c
+6af796af6513c820499969593bfd22f8c6dcde1d2ee2c0ceebb5bd6a1ce4
+5fa61094e932b380cee381f4485e39b4b1797f2a7d8d90bcbf89b9cb1006
+2d50fff083743bf318157caac1c0179c87c03a2857fc002979e7cc97feda
+966b09ceb761d3f55cf07637256c6aa8b8e5cb6aa9739452a330afbe7082
+975ee39fad5e8106e8ee05771157e92d99003533d922ccc37add065b6236
+7613d039741f99edc77c230fe8d1baba720a185186662376b947bbe1a686
+4b42c61ebe1abd40d890751ab8945c629de3b6d2a49809dc693f9e397097
+cf1e568c258081242460af2de0ca44b7ba2734573967b3bdec0e5e64598c
+cbf41e630d821491504f414d9b54a3100dd5105a141cf61bd3ec41b67368
+c8cd366c543754ee800ffee3d19c9cd0d408cc772da10e4d8134964b0a61
+232e2dfbeacd0fdee12792504bb327a2e1fc44127f8577ca51d380a760b3
+740e6be46455cbf3917b90f0dfeadaa25d5d9f66cda43ebf9f75e0191a06
+25ba29666bbe8678822a453d4e876bad4a6b0d4b6cf98feb60339c9eba2a
+dce4ef7faba428422c503d0210dcf8d884ca9f5094aab9f3b1a2238b569f
+444748902907cb0d9d7ca33fccdd0cd29bc68e44f7bca5092be6272bc949
+baae5af92c302bb21f91b6ea8463265680f7c16f45d8ff35392a10eab87e
+296f3af4478032b5b021db8510deb617941130d45c46fb3647d94b162fe2
+2738766fb6d76a06ab6803818b27c5ff4205ba668f95b5ec5ce4ce6da545
+c13ff56f417a4e0b3b8554a1e2a985a167e168adc8c4db28a601a80ab451
+91bf32acfd8d25c39c2f17fb3bca1296d3d160f25b43b4d6b94f20ffe012
+b779339b12860dfc897b366e3d400e756f4f9f4d2c86fb9d94c11ebd1450
+eaf720056e2c39529331bdcb104d113b42c94af2c6a5035750b7ae7fdcba
+b6116d74bc07a11d4357ecf73d99221dad5cba4a7136425c2a3ac0e092fd
+606a4ab722195e3b7fdfb5a5e3ccbb85fc701c42bec43b54e964dff3fa04
+193043eead7681cedae9cce6919949ea60ef5630c4b9263c8f98b4bc74a1
+63ccf3d0a0bc1deff39b800ac90bd734dda7ecdc73169ad77e129887db80
+7a253f8807a422eda8a16c9ee9bb8fc0942634bfe035dac9f7e36d09844e
+39477c043399db4d07b3617da9d6eee76d0fde9201da98b906050748b68d
+8c944ace3c96e90a3c2b63eae27b9152cb7274fa336866d71b65a57f1bc2
+bb1f482a67f3993dcb3ff24abb0223f9a026c81b2b33127a1dad8929dec7
+5d46bdd790eb1addd771c5c3965a2f514d3a128117a44560cc10a729bade
+4e6c86de7c09a39602235c803902e34f5c176b18e127d71a011dd9a3a61e
+ebfaa4a4e2a5651be6f4067e5e09bb4f3514d67c2129e4d3ea9568661138
+1e45af07bd84f883c70577a986416747f3bd8d1bf86d3d7b07e8a350899d
+3c2dae237bd5ece45faba7a0ba30fcda7b7eec9fbeaa5a94620686d1e403
+1cd2512e8d89451c7bd8eb432c8862023d66f3f9fcec0d47598e2df59525
+d673a5ff493d458748cd6341f161a0a3e8996ca5b496508578fe4f653924
+2ae28bf4b7397c02b726fd5f9d8b898938bb668a546be6e42865f4f030d9
+5faa289eb24f7b8e249b224a95a2245605d67417a489626df7417855b8d3
+1c0043cadd2b461d32e1b39ccf409757c37b68f84e752bde6b5bbb847bf1
+57ea3434802def983d6ce5ceb3e9fbc4911b5484e99bb94dc3f383e50672
+0e85a91ed378e352838cf02921ee0ea94be01b5a60f9b1f58fcc1b4f527e
+43725de9b9dadc3ef462fa279bd7138095d4cff2a0563039f71e383430dc
+f628dc9611b2e3db08fb2da1d5383dc1a3c784e1e64541fde1d9d7f42505
+de96d3d0a401099fc2879af0293b0eeb143b78cc221f670c0479bc150047
+0cacb9a282e334e428b527acdfbfc56e6aec8d4d60745c1dc000011b6248
+d9ab4a17dca7cc74e17d33c0641710b02cb1edb0addc6be214b17e9f845b
+2d9c8bf03c19e131e00f91f2a393b5f2ae7c3d4ae9021c4d7891d84d5067
+377ce92836e42eacd7e540824f7ac95360ce116d41d17a50748748971c82
+27f089a22ee0d21940de854f737547b73c7517addd9bdaab425a6c2908f6
+87dd990d6cba4d84308bdd4c4435a6480ecfa1a14daabd4d8e2398178e48
+de28b84f7ce4b61d2e6e64fe043c29a941f6de7621ee6f6d8b506221df05
+db238b8fe4323cb5f259d4d3d9c94d4ae1ca37d6c34345489c0284171346
+e9830e2e3c6c167238a7ffe0989d3eac870cd44102cae139469b9d909b5a
+9c34792f693ac94ecd35d2277080e30a2d24b50391b6f2a3d3b6c81f7ed1
+a7b218903e7fed7a63269e27d793a2e0b40320ebf447c71f36d40dee002d
+7257f43c8add31edf2c571123e46fdb413e007cc89e99b6f98d77ab38bff
+cf140f787e45ffb2c7cc4ddbb59a4e32dfc36e2875f204ac851d757c1236
+12deb31324ea4c201d27fdab46e9f3988ad2bcfb8e9cfa8c487831a9b0c6
+60b20fb66b4c77f52359ac96f3b3d189aa0571c1c53db06ddb10f08882db
+0b1e93e9478d4c75626c5fbdbc6044c4d82684b310ab2af144d12bf36f1a
+c0bf6249d1da9ab319453594cb19d0e93c4e047fb49229c0cce76d0cece4
+2e76fabd2425382afe707db032cf617b046a59a2fc1bb3838d98fd5c8053
+ecb918bc14762e4ca45027623988f434ff4cb08bc9bff5d7de21940e3e03
+1ee042d9c30662aa76f96213fb5a92047af60f320e4660eadd1ec19d0086
+072f2202af5f219725f81882f10d1e065a8035a9946d0ca0e48a5e7dcf61
+0283b834eda01e7d94b3453830daade2aa6c947989b290c02ade0d7b2620
+813ad177ed82813b6a985d5c0a2d42419bda763d409da085936e33c817ae
+68e5467eddc30be172de855a0f7f5c527555b3f4d942401b450f08273b1e
+c5b5352fdb8562a71f276284cf7c27537e628f94bcbffe8d669ea2645752
+60830f1e65e83a2204cec393f6d92d4f61f317471b4b93039d298ca2cc94
+eeada0140823a2bcd1573e732e7b4bde7368f2ecca5961ad547f554ae989
+98d87b7e5d07a85c382bcea1693a697224f41eb8b406bc6a0c3eddfe8b5c
+f25b11c3e4bd91ea7d6274cd6b3ee7b8f18cc3fd502a324c645568dce9e0
+d43caa61f7306fd5488fcfc439d85f8160ebf0ac90fc541f9c74d35d7833
+09309807a639477bb038200738342e50136dc64baa7cc1b879c61f7e1b90
+e1f2bd4f6e54c4dc97b8e4adeb102979203a31fe26a7f58c609915a95abc
+4acc263179423f8ab16b04272d5592fc536f29a45cbcdbe15890f119ca9f
+c7a52eef41dfa5c4fed087eef8e698ba738e300bd58f2a1a10da1198c1f9
+b60e2032f8384a86aa84027df21cb87977528e3bb9bea1e3a6879c56402e
+a29063afc6ac0194f4944433f9a5872cf0a2a741382d7f3c0ca7817d5d7c
+4b8bf53af0f18b1eb54480519cebb61d983157e039b13025e7980eb36f54
+3451bbb84e470ffd0f98eba80c74f238729dd6278294388a2e06de68a719
+47b6d478c85f124d14aaa835620e49b7f5a4f21347302c0f0864f7ebaeec
+d0831c36187cbe9c848736764a31056d2cef27c07cca00033dcddca9a2f3
+b9ebf28e67257b69cd38bc23c711b6a2f6e4dda9bf5a19da275e6a8d683c
+723bfbb95a90a344a6f421f0b67ae84c74652288b0597e4c86c28f73808a
+77455f2948e8df634c2d14f221626b019033f9230c9167982cca9ae6dc37
+aecbcb49fd9fc1dbf2d11bba7187888721bc42a7f47c23e07d2fc5a7a91c
+0dfe255a7f9d17e69af1618502a6b90b1dd748c7eaca1e1ebe8b861b04ff
+e5f628f47eb4e7e65311037d7a5713d7cc3552dc85f452ba74c4f12aecd0
+d72892c940c3325640d62fe3bbbc71361dce6d54766e1fb99dedcb2d19d2
+fa6fa21f9116e03952ebbef659816a62db51a9b5b3916ff818518774ccd6
+79d44100d7236f211f36fa80a4cbafb3db76ba1e7e7f12082b0140eed2cb
+5e793e24501715c6c170ad4f856a4bf16bb10210025156e635264d3cf18b
+1fc1e8cd2fcfdc2ab1a24af9087975bfcf6fb703fb36e288e58d0d2ffc98
+bb4318001d931ad6161dcdf8984e6690e0f6bb07af81bf07445f8f57b355
+6b960d24e7cd152708489e4d953ab6a155a757e002ead97585e6c5333d7e
+5aaab2731f047f3490432e0ebf3d0d628eefa8c1f665b9c86aabb0706639
+5bc372e16378f0d9b439c98e7bf87be73e934995d58e4e70d3ae9a5b54c8
+87a19f2826a772c39d41805c642354d9bec75b065f148f7c1e435dabbeaf
+e4a5744e3f2894a928121ab069bffa3218a106a9dbb83971353a7c7a5616
+d9da66fbb908173f9b07aadcbd4d112cc353e7b70476046ce5a92e86eaff
+4eec40acc840005f51f55c9f5874216851e9cf3fa431d95d3032e779e356
+4bdce33966a3a798b170a06c4cc9f73700224c858c36bbf2d0326c337ce9
+46f69c19a84187fa50afc5b36010f9a7612e3a25e846d49bb907af9505e7
+d8c78748d7dcb501bbb3d6603e829deee3784f2f3ca583d3738d6d2ecfb8
+eaa887103606211a3c1b5cd74a3e0e96fb57da91baebaecd3669661e7b1d
+579ba41928a40a7028acff6cd409e601d23ff66ff2c8acb12e535360d727
+60d2e988d801930e0e9443d60dcb9f378fa75d58d73e6a3b6e5b26407c82
+67d50ad97787f8a9b91765e41552283cb67e43e59bf71cf08b9755c8ce47
+0cf374832c72d1e9702b55bcfc8b5a4e966d5072fb2a72a2108574c58601
+03082ac8c4bba3e7eeb34d6b13181365a0fbd4e0aa25ffded22008d76f67
+d44c3e29741961dbe7cbaae1622a9d2c8bca23056d2a609581d5b5e3d697
+08d7e369b48b08fa69660e0ce3157c24f8d6e59bf2f564ce495d0fca4741
+c3a58ec9f924986399480ee547ad1853288e994940bd1d0a2d2519797bf2
+8f345e1bb9cbf6997dae764e69c64534e7f9dd98f86b5710ff8b500e1c4d
+f509da50c64e213ebdf91978553a5d90908eb554f09b8fc2748c9c405903
+e7bfbf0ea7e84254fb6735f09bf865244238e5fed85336c995bc3a3b9948
+947a6eb95db4cd1b64c0fccf82d247a2202e9e7eef5a550557625a0192bc
+8bcc9e461e52833f6b8729ccd957d5c4b6e07016e864fc02b792c7400ace
+d0a8f43c755f87bba6e5c6e1022416e5454cb34a19865d951f7aea527760
+53658cbf306ead832244f3062c39a0a121a1157a8e47008163c5bfc88197
+be16e9a1ba26a035a16dd38cc28dffb666dd4ba7356c66b7bced9e26e905
+4ce25f6d36607d8f5dda1e21ac96a815bb2989f01130ba1aca9aade554fe
+effdfef5d6b0d2a01aad92f599f6a12e121010ae6acc6f150f19e7305271
+97da761b07530ca19b84b119e5edca1fad18462143b8913d6b3f6864b713
+7a93bb9e1bc29c09d660704e8d8292c61072ebfe35c354a2342b2458a353
+31d043874380d439388e46688a53bcfe01bc190ef1a6b5dec9d40aafe822
+261b28bf3e2d76f3dc4302506ce3387b4aa2a51cd4ba1faa2ed1fd7df664
+6772fe9f83d253451eeb0448b444b8ca80cc7cb653c2d1eaa0de6f2b1c72
+47e6d24ae72e620e200aff83a557a1aa7a0ce0a9cfbbeae03c31d8cbf1d8
+20b53b688ed2ffbd83418d743ee31e3d62216ac7be6c12bc1917548cf670
+d69fd2e78d9f7786ada0ea30a6f6d9fbd1f1406337151ffa1d3d40afbe03
+728fd1aa2fa8a4f075796b9de9586b71218b4356fb52daa01d3c18cb75ae
+d4d33fc809dcb6e3dcf7aee408a0cef21353d76ed480bf522fdfe86e0e0a
+b7d097defcb793057f0ce98ea4989a9b6787b14029a4bf10315a2557149a
+fe9c91e7d825f7518b343fb556f0177a8f6ca08fbda9913d52997511590e
+b9942c9813b4cf4d4aae4919401f2fc11fef0620eb5c40532cdb22d5fad6
+919a3a710de6c40d54993b5386636499c866938e33bc703a99c73adc228d
+95cac73ff4f4a275c04d0d787b62c6a184dacc4024d23f593e7721be232e
+9882fb738160e52ab905f0ce2c76ae6ff2c8bbe118a1acdb3b464178cf01
+94bc6a50df1090e9221be11e49f254b06c3236a31569b947ad041d1c6b55
+bfdec3c18c791ace0fe2a59504eef64a4eec4b5c8dd38b092745e0d5ad29
+276bf02c419c546627672a5764a4904635bff86fd0781d36fbdf13485229
+71f355de2b0ad250052f50ad70f61afc870ac7a816561d3232b73360d4ab
+2727b2fd045f254c782bb3f1f49d94c6d625047071b7e32da5c6d21a86de
+9283fd632074430772bfbd85e0c9ccab1dec16bbc049c3e223bec1b65c8a
+9e98cf58b30a74f74f1a842dc91e30c023498e280ac55edd58f4cc731d81
+e443d9b9efdf5fea63c9f357320e01b8740eedaeef2495cd02eb2f338b3e
+674fb074cc497d7b1937b188da857c2c230e9a931cbc00c85a7a36fa80b4
+56588e1bbabbe4ef429a6aef9bd4eb89c5752421bd049aa13f4dcf9b51ce
+2503e90bc118fac78a25d187353d6f5d496cd6130b337666f49619cea985
+dfbeb7e49c67c1e0f0f8e9ec8ba14624ed0982dcbb69415e4b3c8ddba140
+397eb1fc1ddd36c94c374f018873ba41109e45afa51f0e691157d5958c06
+26fbc0903ae25e47ee372389cf65472a3e4d9769550bdc42c0b72f9a297c
+d5d3c16ec67e06036e740ab664abc9f10b9499269b73ad3678daf4474329
+c2c7252c1f0df1e3b5e8f198dfef8325cb1e7e8057897a3d7fb5bb5858e0
+cfc0c115bbd7362d8e8ee41862af6eeda681cabbb06f72ebd2ae0b0be45b
+a9e1be83f1da30687a655e5d148fcc17d9f53b760810a565f6d2f4cd5da3
+5434116edef756adb4d3df544a1de593be988f2bb8d36c34deaac7d9dc15
+cba49764f1e03aa09fe21fcd7c74e3d6487ebe219569e019f10dd163046b
+c1a3cb2bcbaa8558197cb2c18709a998b4efa8ab8c9a71d2ccf942c17662
+1b88dee6b424165d6ce10ac48375e760983818e0085276b1674dd41042e1
+a01a8de111c903f74834199b3230bd475d92c6226ef74eb1daaec3475a6a
+fcb47644a17c7e390ee3b16bef1c1ca6c55eddc44fbefbdde525921b3047
+0d76817bd8ac724739a8e743eb09cf78e88adad527d4f115b8a32ed4898f
+45bab3eb802b8168aec061e3ecdb026c056fb9efe7e2df48bd516ccb12ce
+00de08ed8be4ee0c41f40f4c8f64483e0ade90a78d6d4fe9203fe0b97c60
+3b2f8882bc15a212453c691c52d00fae8a3a26934ff8acf68d4352eef75a
+0b10d938e55b7333dda2db0296a69e9775bf82b1aa6d684fd9080fc1c11f
+ab4369c7a95a9504063db900a6e345bf6dd99be041230b2e60cc86b8c345
+1d84a9c2cb4ab6d74d63dd43dc26eb6b384f5222796d4083dcc3e1651548
+d9469f09a33b213a33ac52a6a2e23802d8f8a75c01a607940daab0051410
+73a88130bc192f303616adb113c0051b65e12086cb319c0a5323fa7def40
+402f5f87a3b2c2cf0e92789985f6775ac2743e1ffe2d0668291059740d45
+43bae7a2897e5e658592bf5a72966097742e0702deecb0cb12499eab701d
+34ba37a08346217a415e44297a181bbf3744f0a49230ad6f030e11462be9
+afc2ae14e0587bc02311b48b8e2122c28cdf14414f3680fa52dbbb63b17f
+6ebe4a1204f3c5d6150cbf89a8023890383153838d4dde77d4c8b1b78823
+8918c564d3babfe58eeb154307dd1997f5ab7105426e35c279008b2677e4
+695c60f956b348799c04b734338018fc27f7de7ad9d73468fdbc5283bd14
+c066ddad9a3562f16baae15d72d7bfcb409e1c874e9db1a8cde233b282b9
+6e76e9c08d85ddfbd3cce7e64104d0b0e95291bd91f405ff82f41601ee20
+8471e613fbbee67f269e4e954c36d1d18ca9880b7cc2b08fc990978efdc5
+1d157deefedaa765c1e26ee125d4a2514a41a3b95e9151a824532d7d6486
+35ad622718fe71219a697e94c2e64f26424cbb767acdef5cda70e179cd29
+b7e318d1c6d3ad26fd5fdcbf2fc221301cc1f10f5ed86b40a1a6bcc01c90
+eafd65183e75609610637b99fea57885efe76437df02a2ffc21223d039b5
+74955d9a54ff41980eddaa8768c5ad883a0c9150877392b990d63c6805db
+7b8d6ab1358cbedaedb6feadb0ee4fb8f9c1ca03a3e755a74227a8930bb7
+2ea0a00b48fc626fa14d7d48624aedc31c556f44e982f3ccbde7ee735f73
+629ab1b65bcbcf0a3586a920477e8c960219802fcb1bc3a179032b324f8d
+c424899b38275886cb5bc771f26a0880767d49cc23426a40a4b6ff8fe48f
+d747565fc537565f6d7fd08706accc60f5fbcb45bc785f45ee9b0812366f
+ae71b23ec43f3549c8224d78baf18719f05108d5741e681457ead8abc050
+462481771a8dc6cfeb98956e163981a98c59ab44d90e9c3a946c453b5071
+db0c769f7fb5144c7ab0c9ef1a6db1addcde1d4ae1daee1b4035af256a04
+df53926c7a2dcdb94caaf12f986e20929ba4e396f3aa7c93a7abaef1294f
+5f13a0dd3c3aaa8fb38da3e15daa32163b7437af683b4f5e64cb14aebbde
+8c69ed2e8cdbfb213fc8129af29ca2c06c8f85a5038d688d1fa5d1b54ebe
+4dea81a49ce24131f8e6702e7aa4e2cba078d5dd373f894ccb275f49c690
+1dc772e1d2f5fb3fe15dbfffac62c87110162074eb72ae4e5e446bf7e650
+a554178d0d64d3c07f330f0d99e99f2239cb1597f2e5f443854cdb0f5fab
+b28fe62f22e7f3419d017980f325351bb04f8f3c3dc57fee03cc029bd29b
+202308d5a800ed2d500d41ace8e54e2557bf25b627883beb8118d800eb94
+f4253f855168f7fc8a2d29c5fcb76bb90a6c4e345722b8991a854047f46e
+4e97336be85470b6be2b9ba573dbc4967ddcdbfc3b6fc35b0c7f3f2f570c
+55dc3fee6d80bc6f46cc7e4d86a0b86f6fa61d062e213d9e442db63fbf11
+d03165b44572096995ed342893bb672f6bb55ff8fed944667995f0f89a48
+a904c47420f32afd14129c6e2bedffce1f07ea69d550b6909bb5beb4aa08
+b0b44f35e018ba5206fdb4df0228462c1fdbb95a429e53eb27bb1b0490db
+f07202c3608d0f4ce08570e3d6aa3d4581c569b57bd8c1ea0e4ed3fc5497
+e316ecec06e6be582d9170d426f6d22d8c7287b8219945c124941ca8812b
+e97efd9105eb6999edc0665016633b3b48820df736125b7c76c9f3a67d93
+8a2a0a6b743fd42aebc46a0249be459f16811ac9eba7b63bad7c2e88f175
+0eff8da5faaab5659824f9d19b3225aad2ac17c52c523414d3031d08a926
+30abf474fe02a32b44d3b7d9fe0c19aec16ca6d018b71d9d395ffaea0788
+0d4501d7cdf0f7077a2d63303d09083080d67f1f714a1b271dab9fc9866e
+4b0571a171eec8a4e351ba2d02438cd108a33b1106acaad0ccdb051061ea
+7f40543748115f29debfb4be4b42cae8762d62114ec6f8ef68c478a8e05d
+ecfa18b0368428efec9eafb2353f95e3d71e1636b9d9f94a77e692843255
+698576dce13b2b858d2d15ee47cdba3ed08d64b77ab46dd29bba6aac2106
+ab847de378cccdaf35c64e50840248915f4fc110992c493cb1b9cd0b483f
+0f1abf5e9b018210b477fea28234ffbe5e0bbe01338e0842a89f1e00a0ca
+7cdde0b2d7c324d5e17d8d3415ccad703507497ac95360ce660b656e5f66
+72a2f50761f3d02ccdc1d5692d7797699b8e2147cfd4817c81a432ff6a5f
+39cc54927fa146cbed56a55f85f123c0a94b7553a8819b329d9dd122c502
+94e3f6314d5117db89ae7597c4691b6c542979a1ca3d26a8e23d3eb698c7
+1841651e08ec771cfb974d6613f2143872c739b62796bd0a45172530793c
+28d93a65b59f79c245248d2c09428657a35b0c0e367bf7a4a4f0425b3f4b
+485d9f402e164328a4b963f456829a39035c00283d2e4fcb71a42da6d42a
+d46cb751287de34e6519c60bb3f1a6ba91f7bfa21dca96ee712af5681701
+18ece8a0535d9ba1dd4bd835e004a2f38c5ba43c9b30d17045e5649fbbac
+188922e442182d4bdafaefb39e00106a5a7765f3d67850471e3629e526af
+8691f935b57bd38465665204a214fef1006ea37dc0781073ced5fc042781
+93650393c3cadfddedcc5550ed483bb6355f54600e9758e647f9c9711f1b
+e7df05d0e50a698615307c18f6d4886f50188011ba499d03831185915f3f
+77c4b9ce708d78423b110776aaaf90396be0381616d1e9b0c1dcf68b6396
+82399da2a7323bf42ae5347599ef4ae9e5c135522c5ecb87e201853eb899
+db60d24acad17d6b7c2c7ea4dc221f3cb6d6caacd1ac0822ea3242ad9b4d
+d15116c3874e3012fad26074a23b3cc7e25d67ef349811dbc6b87b53377f
+0cf972040a037ecb91e3406a9bac68c9cab9be9a6bb28e93e3275b177cd5
+0b66935cbe8dd3d6a8365625db936b2cfc87d4d6e7322df3dbe6ccda2421
+a5e5372566f626a5e9d8bc66959e443286f8eb4bcfdeb6c49a799f1efa69
+63260d0ea2d51260baba9207fb246da927fc4c89e9c4dd5848fd4ef6f81a
+cd836f5f06ff0fe135cafd7ab512af55a57727dd05a5fe1f7c3c7bbe8ea7
+e6680fcb3bbbee1cf2e2c0bba20185f00e2dc3afd42f22de472cdb3eaa5a
+ddf8c6fb3682eea5548c51ddca25ca615221127b4438ea535ab3089c9ed9
+b971f35245cf831d9461a5da9d57bc4e5606d26535a7414cef6aee2a7b95
+bf2276044818ee0f3b0a16532934b8b745d8137b42ec2b28fae7d55fc02c
+9ccfa4e0055f8a4be96e1e235c01b8b6ad509b832a3e90161e0a449934e7
+4be973c939b31cbc19dad4c58e9be89d242f0ce200548cdd4fa2081ab3f8
+e01f358d5db24b7a50eb2096d833378921f561f132cd7988708ee10cffb6
+2256201801c667e176b1dfaecde9756d725bef093457805e16f550e8a7de
+87ecd46e5b09646b73ee74f890a36867636911e4cda2c46a40e7d57cf297
+9696046614c85b1a47ba55c60544ebd3ad7d750d003bda56dd7eed8c4702
+f8b319aaeef9d3cdc59b3e63ee93c6e1e857af273eb90909ecf36ef4c276
+895c78aa762e5376c5c542f854fba864ebce56e4b0207091139f053c2c08
+3b7ddcd0a9909b52100002bc3f8c47bcb19e7a9cb58b1ac03fee95e81195
+072d3aa7c8079632725f63425a3550a947834d29ac9a26d0774e90248e18
+996731fd9aa53ab62b40ce557d98e874b763d9d629a173f0c7babfc00ae7
+82daef5f00cf3608ebeef403dbbc19e16a1d160b889f4a10359d9eacc19d
+7b5f126b31720dce7fc35ec861dfa56ea23fa18423ff4e8fe6e53fc6ba16
+b95a2b5dec00f614e4f835281ee0b4bf549e7e882689e0b445dd46fc40c9
+090e5575fa2c34b02a51ad0bccf6a7bb83ca3b929285e5e9fd054b72c47b
+733a66c5abda526b18b2e49d0746e067e63b948a45eab2f4221c5b62ae21
+a5d9d7cd8aa9eeb49588891d22c56b14b55ceb6488f02b73ab3b7f6c5555
+b75452594658255e4cd58ac4815f2e1bc3888c6777f62aac2f0a57d416c3
+765c991f0f9a33d888aeb2d527b482c042ee23783a04a73ad13dfc590a52
+f3116f8296cacc7ab29b7d87e7864561a5d0a12bde2d36ee697064f41d1b
+ca6ef2f801caab5295d19bf4c02b10c19f73b44635ba48a0806b967d7dfc
+ce9a4850171a78532cb30020c0d66b3b1e7c75eaa7894904c181a022e8bc
+9b2b8ef1202f3c7d36bcab4742d4a4761bb55b64da0d99685d319f5da8fa
+132be6c0483f50e2657ae8af1e28f969440d6ed43eb00e95fd9e1cd490a4
+8646f6d008598751f7a41b43fbec7770fe591012b6b0c4ae18775ccc7db5
+de0ded2dd53e82c89648d46f0d0cc5d3ac5aa104239608d512a4353b9547
+04fe6eb7e73d718323cf9d748b8ec5da01ec9358267de12cc22b05ef0312
+e4b6ac5dbb6d06d7f2d911f20d527f504d62547aef136834b3695df8044c
+383b6145e824d3931a602f081d9d656f84987a1ef121772f1f5b37a116bb
+d2e77d4ccda01411545d24e15ce595db4cd62ee876b8754df0b85b44e011
+b82d76ce45795e6c2c58be8690b734a8880a074f303a70da4a1b086a6de6
+56c02cc7a4c25258eff18cb0fd868214bb46f972e26509f868d065b3cb14
+1c316898cf22293391bd7051ac3a6927aada952a8fd0658ce63357c07f34
+acbf8c99a5537da0023e901f0eb5547e1b466b7d982c8c539798b76ee2a2
+252437a81a37c3b63f625172d682eeed0b795860b2755f020ef52a138353
+003c61be2052cdd7d73b2cdcd26b127660a7b22fc51a6a2f6034f37e3e46
+c1d7f83f8b28c7c965993abba1d358362833580d9c63fa85d4cb949f97de
+579fb6807b95a58b78f596db50055947dd0d0e597d9687083e9bc0266e86
+90b884b27f4094d8fb82ffdbaac4d580340a9ef8aa242be87e54b601af19
+87a48d267c04e371ae77163ebd0de3f5297b1060442ecdeac38334844e38
+0f294d4be73935fd8a38de7fba6d082c3d9156d7e88f2cfff0459377cbb6
+041f37a7e05010753b98e0b67d5827aa312129bb3c3bd883c12323756406
+d555720da8a0bb30edcfa760c01ecc2ba3b15fecccf5a10e9f358822e0ff
+b64178fce2ea6a1105bfb72df0e4bc499b207ae26b8ea960de48e7ee7010
+b4e671dff795e4cdc5b43e81b1604d224f0616ae311f1208859c502c1a10
+940e7b9cd11be728bd3a0c8005ae23aea32c1b642812198a6f1aed32cb75
+97152b1340dd35ada1b81051e393d38f3740fa9523df6a83b8ca7dbceb33
+6e299b54cd998d4dfef804733c76156585e42b7284cbcc4047ba6b290efc
+aa60953e98cd2b4bc2893857fa6a339f820142a52ccab0df09a2709df550
+f22e5921cbca408e7998cc1cccb8adf6d8f8b71e6685ae59d290fa33f5cd
+664d73e434237424060f634262f04e9a71a977556e93b692ddc3aad26d92
+97dde71e4def64932151ad572af6e681082e9944ddbec6e7a8bdfd534233
+9ca3106ca1ccc80eab14f1655978b137fad8f399df7cbfa2d7d3d9675e0e
+9afec37369a8ede2c93145ab3f42a375926946680c215fa16bf7416fc892
+bacd806cd424b9f85b47802c4336918f7486af2a03bf0d39b10169d35494
+419cb1ab7b8f407897f70c18303e91563b497d70b7181ede6aa0c3efe089
+ca6135b34dd1019b298e3677f8da61f864a67023c31eaa716c40cf3d397f
+9a1209564c9ec759c37028079661d2a56374203c78b023ec61340bce5d96
+e477a4f77e5c0db7c0d1257b4bbbc6f889b17e6eaab045b8adef6f931e4d
+0795583d60a6b7002cf61639c6f930671f3b8ac05a1c4e002f4bfc50d8b2
+3029fc4dce1b602cc3a5533336271bccc226559ffb127e3a562f92f89824
+552b9a70466d5a3c74ae515a222b109d490f26e8fc2d9d72bc8af6d1dcc7
+80463c7af81993bac2ce4aece9d95ab736b1dc73e32d1237bc8ec2b52513
+36dbabb4ecc7ceb5d18b02043281eb9a3bfdf19bc4853c9b1722ef1cdcf4
+fcec534923db2e2653dc48545a9850c0ac2e4594abc9f7d18a0bcf2fadfb
+bf085d465a4d10528312f5d790eb9511ca01061c0d94136b99a043bcf278
+c18223b1e0f1cc062b32b79e28dec2dc59a0aaa4b5f3506923c83e6a87fa
+08a1d941bb644c994491cf7f3b0e2ccf6c8a8ba89376f76dfdb592374f93
+528e78e31e0b18719346b9f1486f652638e3120687774030444674cb0778
+96385c41f6566819652d825dd58f9a4308ff79b45d7828dcbfebc406e40a
+c46e866cb0e3e97d6ce7fcac19a9d0fe39bbde66c5f0cf775eb3b1e6d7e1
+1f67e7edb3d5c4facc85c916bf13322b56a0414ca27d145cb740fa2c37cd
+8c142d9301f1ac3704cf6a8e93973a07fde5a331cf0cbb370c7ba555de61
+18a6cea0ecb2c0e37152390cc57e2e4fb3791ddbc383ee26b6f4006d0d68
+4880888011020f856a9de47f45440f127cf27ccaea7d40a3869d39ec7dec
+ebc06382d294717644b6118354e15544fd4c6d88df9245c9a83b30e6ce09
+e2498dd1df488a019b179cb859889e6ad2838f749e3b038b280ebc8d5c3a
+b03e8f15751214691edf0f86281e612d7ec0773c8a5d2b433266402df62f
+fcc06879ca196aaf1fc73a5f01ac46b44d6cbe7743ae9a862c20445ae2be
+1544f413d010280cc2941900bf3c42ec088cb21b44a915bb810e7666b545
+5324465c5943eedcef0c09128a995f431382e2062f5e39f4338c8eba1bca
+e553cb60bb8f3e5038ac8073398c49f06dc734b18afa7921ea0d455e6e73
+db8ad9f77fb5ba6c28af6b4f18cbe46cf842c82d6c960be1520a5fd929df
+ac7e00ede976fb2be0a07f659079a421fca693de89ce9b8fcb42b0176d9d
+f3ddd58f921e13e216933d27b49d175b423751c451be7618eaab054d3b8c
+23e8dd6fd60182d61e9b5c86b3b764a29a62f913ee7524d8cb33737d7224
+d95dc4bb8c2ad6397604a0ffecc8865adcb540e5da1cd769077838515118
+ebc9f0b988545c1881dd2e7a8fd73e11bd7ae9085fb4d45526b23a346b0f
+e4281ee3d588106db5f7c386c488d8f2f4dd02d4c08e74c1034f987a44e5
+d39fd07538de57a42987ce290fb2f6557e8b5cbcaec168f5780927226415
+1e11e3667d33b36a793aa53e9e2d1102c9eb30cb3ba0ebac953e0227fe4a
+3d3c0eb57e4390c3d35db0c41946e45be2830a1ae33fa25cf2c7c9cb4550
+ce9ff6c6e3d628fc7284daa6241604c90dde6339b7f7e7df3733416cdac8
+e5291357e4983d74d3582a490438a7fdb0af97001a31990b1de68e6adb48
+917daa387e647f9f13312db57310c7dedc2a2ea80800b4f4bbaa99c6b7b2
+7ac8345cb659489307e2565ebfd17774642c9ae5d3c18068dc35170c7d58
+4cf4173f1baf98137fa249c81f3347e1dadd6b1ba0f50c3b64c1eab183a0
+937b0f7278eff101e5267fa6480da7d602844416490c2c2c7eb0d44ac8f4
+75cfd611db5ec268db07c0b3608825c3e12834a2b2efaf5e2723c5199c42
+6011cf22e64e4c0d31d563f321097935ea0c6fcbf5acd3748d90079f6ab8
+687288dc55df29fe7958f566b27b73e2ea30747247f7a2b2add0602c7d64
+d23f52e7c96748e6a54ee8c4629b2aab8882169653f0ba7f05236bf14364
+244720f3259cbed73a318b29e4a9305deb65a2c9dec8a9d0f9a9f6fae541
+83e0f4b9a9a567057a1794945168dc23cec25d1c02ea9242c9fb6d8fc11e
+e8874bd80a5226373ae87cea91853d0625c777ceb1f5a6f3debcf2f75a61
+460c7b4067f568ecd01f62901ade8bf8fbc5db9c6720420496f0cb48a002
+99870773c2e7b12e83987a5d0290d9bbf589ac889bf7d4334a5147187a7f
+71008f216ce917ca4cfba5347078f354897fd87ac48af6a6c62711d2eb3a
+5882bf3b32c0f1bfda976f850c9dcb97170e78c229a27fd5e292d161ece9
+a8c47a223cbdc28e24f79f6429c72b5752a08f917feda941582c36d9acb5
+748c86072858d053170fdbf708971a0bd5a8d8034ec769cb72ea88eb5cd7
+49f35be6ee5e9b5df6021926cae9dac3f5ec2b33680b12e95fd4ecbf28eb
+a0503c10c6f2be6c7c47e9d66a0fae6038441c50e6447892f4aaf0a25ccd
+952c2e8b201bb479099f16fc4903993ac18d4667c84c124685ae7648a826
+6bc1701cc600964fdcc01258a72104a0e5e9996b34c2691a66fa20f48d7c
+2522333dfdabf3785f37dd9b021e8ee29fa10f76f43d5f935996cbf9d98d
+92d0a84ce65613f7c4a5052f4c408bf10679fc28a4a9ff848d9e0c4976bb
+dfdfb78bb934cd72434db596cb49e199f386a0bda69449ce2e11e3a4f53d
+be134c6d7fe452a0927cf6a9a15b2406f8bd354adcde0ce136378baa565f
+b9c51a03b1fbe1e166a1f92af26bd9f072250aaa6596a236ba2d5a200c90
+a760ca050421abc78223b2e8b2eea958ab23084fa1947574e846e48aeb12
+26cebb8b5a92089e9ea771557599e2fff44d75bcf600e76ae7289ba98cf3
+98208c5104562834f568ebd62801b988b0a9fdf132b6564566103b3d2d8e
+6a099b7fbad8a13b8cd7f6729bb6651fc1019e66c4bd6ff27410bd5cdae7
+4010bd68b066bffdb4fd5e3dd9cf7e1a1353f7a4c5157e3ad508f4ca0259
+9761b7cdd6a81b3560b8765be3b0432fe4c25dcb4001b00c7fa62874f681
+ed22127dc3974605a05be8d8fcf9701f859ffce4dc598091891ab7596ac3
+4cd851ecfd2dbbaa2f99dac376f7bb40703fd0700d7499a7c24726bdc9bb
+3b88c6a82e52686c1ee945d8825092bc81848a08722ac5a1d24353f95ec8
+18f3fa487d9600318091b0ae9874b42bb3cb683a2518b18cc1bd86c6e5e8
+3d37c14ef4fe0c77b03a3314995b1e7c1066b98c4375bd1fc5fadee1b024
+7ece4f95a0f59978d543910deb2e5761632c74c508269c4e4b9e315bda02
+975dc771fc30c8164b9df9172a4e571d8ca578cd2aaeaa0dd083e74cdc2e
+d938b984b96d76a64b8c5fd12e63220bbac41e5bcd5ccb6b84bdbf6a02d5
+934ac50c654c0853209a6758bcdf560e53566d78987484bb6672ebe93f22
+dcba14e3acc132a2d9ae837adde04d8b16
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+0000000000000000000000000000000000000000000000000000000000000000
+cleartomark
+%%BeginResource: procset Altsys_header 4 0
+userdict begin /AltsysDict 245 dict def end
+AltsysDict begin
+/bdf{bind def}bind def
+/xdf{exch def}bdf
+/defed{where{pop true}{false}ifelse}bdf
+/ndf{1 index where{pop pop pop}{dup xcheck{bind}if def}ifelse}bdf
+/d{setdash}bdf
+/h{closepath}bdf
+/H{}bdf
+/J{setlinecap}bdf
+/j{setlinejoin}bdf
+/M{setmiterlimit}bdf
+/n{newpath}bdf
+/N{newpath}bdf
+/q{gsave}bdf
+/Q{grestore}bdf
+/w{setlinewidth}bdf
+/sepdef{
+ dup where not
+ {
+AltsysSepDict
+ }
+ if
+ 3 1 roll exch put
+}bdf
+/st{settransfer}bdf
+/colorimage defed /_rci xdf
+/_NXLevel2 defed {
+ _NXLevel2 not {
+/colorimage where {
+userdict eq {
+/_rci false def
+} if
+} if
+ } if
+} if
+/md defed{
+ md type /dicttype eq {
+/colorimage where {
+md eq {
+/_rci false def
+}if
+}if
+/settransfer where {
+md eq {
+/st systemdict /settransfer get def
+}if
+}if
+ }if
+}if
+/setstrokeadjust defed
+{
+ true setstrokeadjust
+ /C{curveto}bdf
+ /L{lineto}bdf
+ /m{moveto}bdf
+}
+{
+ /dr{transform .25 sub round .25 add
+exch .25 sub round .25 add exch itransform}bdf
+ /C{dr curveto}bdf
+ /L{dr lineto}bdf
+ /m{dr moveto}bdf
+ /setstrokeadjust{pop}bdf
+}ifelse
+/rectstroke defed /xt xdf
+xt {/yt save def} if
+/privrectpath {
+ 4 -2 roll m
+ dtransform round exch round exch idtransform
+ 2 copy 0 lt exch 0 lt xor
+ {dup 0 exch rlineto exch 0 rlineto neg 0 exch rlineto}
+ {exch dup 0 rlineto exch 0 exch rlineto neg 0 rlineto}
+ ifelse
+ closepath
+}bdf
+/rectclip{newpath privrectpath clip newpath}def
+/rectfill{gsave newpath privrectpath fill grestore}def
+/rectstroke{gsave newpath privrectpath stroke grestore}def
+xt {yt restore} if
+/_fonthacksave false def
+/currentpacking defed
+{
+ /_bfh {/_fonthacksave currentpacking def false setpacking} bdf
+ /_efh {_fonthacksave setpacking} bdf
+}
+{
+ /_bfh {} bdf
+ /_efh {} bdf
+}ifelse
+/packedarray{array astore readonly}ndf
+/`
+{
+ false setoverprint
+
+
+ /-save0- save def
+ 5 index concat
+ pop
+ storerect left bottom width height rectclip
+ pop
+
+ /dict_count countdictstack def
+ /op_count count 1 sub def
+ userdict begin
+
+ /showpage {} def
+
+ 0 setgray 0 setlinecap 1 setlinewidth
+ 0 setlinejoin 10 setmiterlimit [] 0 setdash newpath
+
+} bdf
+/currentpacking defed{true setpacking}if
+/min{2 copy gt{exch}if pop}bdf
+/max{2 copy lt{exch}if pop}bdf
+/xformfont { currentfont exch makefont setfont } bdf
+/fhnumcolors 1
+ statusdict begin
+/processcolors defed
+{
+pop processcolors
+}
+{
+/deviceinfo defed {
+deviceinfo /Colors known {
+pop deviceinfo /Colors get
+} if
+} if
+} ifelse
+ end
+def
+/printerRes
+ gsave
+ matrix defaultmatrix setmatrix
+ 72 72 dtransform
+ abs exch abs
+ max
+ grestore
+ def
+/graycalcs
+[
+ {Angle Frequency}
+ {GrayAngle GrayFrequency}
+ {0 Width Height matrix defaultmatrix idtransform
+dup mul exch dup mul add sqrt 72 exch div}
+ {0 GrayWidth GrayHeight matrix defaultmatrix idtransform
+dup mul exch dup mul add sqrt 72 exch div}
+] def
+/calcgraysteps {
+ forcemaxsteps
+ {
+maxsteps
+ }
+ {
+/currenthalftone defed
+{currenthalftone /dicttype eq}{false}ifelse
+{
+currenthalftone begin
+HalftoneType 4 le
+{graycalcs HalftoneType 1 sub get exec}
+{
+HalftoneType 5 eq
+{
+Default begin
+{graycalcs HalftoneType 1 sub get exec}
+end
+}
+{0 60}
+ifelse
+}
+ifelse
+end
+}
+{
+currentscreen pop exch
+}
+ifelse
+
+printerRes 300 max exch div exch
+2 copy
+sin mul round dup mul
+3 1 roll
+cos mul round dup mul
+add 1 add
+dup maxsteps gt {pop maxsteps} if
+ }
+ ifelse
+} bdf
+/nextrelease defed {
+ /languagelevel defed not {
+/framebuffer defed {
+0 40 string framebuffer 9 1 roll 8 {pop} repeat
+dup 516 eq exch 520 eq or
+{
+/fhnumcolors 3 def
+/currentscreen {60 0 {pop pop 1}}bdf
+/calcgraysteps {maxsteps} bdf
+}if
+}if
+ }if
+}if
+fhnumcolors 1 ne {
+ /calcgraysteps {maxsteps} bdf
+} if
+/currentpagedevice defed {
+
+
+ currentpagedevice /PreRenderingEnhance known
+ {
+currentpagedevice /PreRenderingEnhance get
+{
+/calcgraysteps
+{
+forcemaxsteps
+{maxsteps}
+{256 maxsteps min}
+ifelse
+} def
+} if
+ } if
+} if
+/gradfrequency 144 def
+printerRes 1000 lt {
+ /gradfrequency 72 def
+} if
+/adjnumsteps {
+
+ dup dtransform abs exch abs max
+
+ printerRes div
+
+ gradfrequency mul
+ round
+ 5 max
+ min
+}bdf
+/goodsep {
+ spots exch get 4 get dup sepname eq exch (_vc_Registration) eq or
+}bdf
+/BeginGradation defed
+{/bb{BeginGradation}bdf}
+{/bb{}bdf}
+ifelse
+/EndGradation defed
+{/eb{EndGradation}bdf}
+{/eb{}bdf}
+ifelse
+/bottom -0 def
+/delta -0 def
+/frac -0 def
+/height -0 def
+/left -0 def
+/numsteps1 -0 def
+/radius -0 def
+/right -0 def
+/top -0 def
+/width -0 def
+/xt -0 def
+/yt -0 def
+/df currentflat def
+/tempstr 1 string def
+/clipflatness currentflat def
+/inverted?
+ 0 currenttransfer exec .5 ge def
+/tc1 [0 0 0 1] def
+/tc2 [0 0 0 1] def
+/storerect{/top xdf /right xdf /bottom xdf /left xdf
+/width right left sub def /height top bottom sub def}bdf
+/concatprocs{
+ systemdict /packedarray known
+ {dup type /packedarraytype eq 2 index type /packedarraytype eq or}{false}ifelse
+ {
+/proc2 exch cvlit def /proc1 exch cvlit def
+proc1 aload pop proc2 aload pop
+proc1 length proc2 length add packedarray cvx
+ }
+ {
+/proc2 exch cvlit def /proc1 exch cvlit def
+/newproc proc1 length proc2 length add array def
+newproc 0 proc1 putinterval newproc proc1 length proc2 putinterval
+newproc cvx
+ }ifelse
+}bdf
+/i{dup 0 eq
+ {pop df dup}
+ {dup} ifelse
+ /clipflatness xdf setflat
+}bdf
+version cvr 38.0 le
+{/setrgbcolor{
+currenttransfer exec 3 1 roll
+currenttransfer exec 3 1 roll
+currenttransfer exec 3 1 roll
+setrgbcolor}bdf}if
+/vms {/vmsv save def} bdf
+/vmr {vmsv restore} bdf
+/vmrs{vmsv restore /vmsv save def}bdf
+/eomode{
+ {/filler /eofill load def /clipper /eoclip load def}
+ {/filler /fill load def /clipper /clip load def}
+ ifelse
+}bdf
+/normtaper{}bdf
+/logtaper{9 mul 1 add log}bdf
+/CD{
+ /NF exch def
+ {
+exch dup
+/FID ne 1 index/UniqueID ne and
+{exch NF 3 1 roll put}
+{pop pop}
+ifelse
+ }forall
+ NF
+}bdf
+/MN{
+ 1 index length
+ /Len exch def
+ dup length Len add
+ string dup
+ Len
+ 4 -1 roll
+ putinterval
+ dup
+ 0
+ 4 -1 roll
+ putinterval
+}bdf
+/RC{4 -1 roll /ourvec xdf 256 string cvs(|______)anchorsearch
+ {1 index MN cvn/NewN exch def cvn
+ findfont dup maxlength dict CD dup/FontName NewN put dup
+ /Encoding ourvec put NewN exch definefont pop}{pop}ifelse}bdf
+/RF{
+ dup
+ FontDirectory exch
+ known
+ {pop 3 -1 roll pop}
+ {RC}
+ ifelse
+}bdf
+/FF{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known
+ {exch pop findfont 3 -1 roll pop}
+ {pop dup findfont dup maxlength dict CD dup dup
+ /Encoding exch /Encoding get 256 array copy 7 -1 roll
+ {3 -1 roll dup 4 -2 roll put}forall put definefont}
+ ifelse}bdf
+/RFJ{
+ dup
+ FontDirectory exch
+ known
+ {pop 3 -1 roll pop
+ FontDirectory /Ryumin-Light-83pv-RKSJ-H known
+ {pop pop /Ryumin-Light-83pv-RKSJ-H dup}if
+ }
+ {RC}
+ ifelse
+}bdf
+/FFJ{dup 256 string cvs(|______)exch MN cvn dup FontDirectory exch known
+ {exch pop findfont 3 -1 roll pop}
+ {pop
+dup FontDirectory exch known not
+ {FontDirectory /Ryumin-Light-83pv-RKSJ-H known
+{pop /Ryumin-Light-83pv-RKSJ-H}if
+ }if
+ dup findfont dup maxlength dict CD dup dup
+ /Encoding exch /Encoding get 256 array copy 7 -1 roll
+ {3 -1 roll dup 4 -2 roll put}forall put definefont}
+ ifelse}bdf
+/fps{
+ currentflat
+ exch
+ dup 0 le{pop 1}if
+ {
+dup setflat 3 index stopped
+{1.3 mul dup 3 index gt{pop setflat pop pop stop}if}
+{exit}
+ifelse
+ }loop
+ pop setflat pop pop
+}bdf
+/fp{100 currentflat fps}bdf
+/clipper{clip}bdf
+/W{/clipper load 100 clipflatness dup setflat fps}bdf
+userdict begin /BDFontDict 29 dict def end
+BDFontDict begin
+/bu{}def
+/bn{}def
+/setTxMode{av 70 ge{pop}if pop}def
+/gm{m}def
+/show{pop}def
+/gr{pop}def
+/fnt{pop pop pop}def
+/fs{pop}def
+/fz{pop}def
+/lin{pop pop}def
+/:M {pop pop} def
+/sf {pop} def
+/S {pop} def
+/@b {pop pop pop pop pop pop pop pop} def
+/_bdsave /save load def
+/_bdrestore /restore load def
+/save { dup /fontsave eq {null} {_bdsave} ifelse } def
+/restore { dup null eq { pop } { _bdrestore } ifelse } def
+/fontsave null def
+end
+/MacVec 256 array def
+MacVec 0 /Helvetica findfont
+/Encoding get 0 128 getinterval putinterval
+MacVec 127 /DEL put MacVec 16#27 /quotesingle put MacVec 16#60 /grave put
+/NUL/SOH/STX/ETX/EOT/ENQ/ACK/BEL/BS/HT/LF/VT/FF/CR/SO/SI
+/DLE/DC1/DC2/DC3/DC4/NAK/SYN/ETB/CAN/EM/SUB/ESC/FS/GS/RS/US
+MacVec 0 32 getinterval astore pop
+/Adieresis/Aring/Ccedilla/Eacute/Ntilde/Odieresis/Udieresis/aacute
+/agrave/acircumflex/adieresis/atilde/aring/ccedilla/eacute/egrave
+/ecircumflex/edieresis/iacute/igrave/icircumflex/idieresis/ntilde/oacute
+/ograve/ocircumflex/odieresis/otilde/uacute/ugrave/ucircumflex/udieresis
+/dagger/degree/cent/sterling/section/bullet/paragraph/germandbls
+/registered/copyright/trademark/acute/dieresis/notequal/AE/Oslash
+/infinity/plusminus/lessequal/greaterequal/yen/mu/partialdiff/summation
+/product/pi/integral/ordfeminine/ordmasculine/Omega/ae/oslash
+/questiondown/exclamdown/logicalnot/radical/florin/approxequal/Delta/guillemotleft
+/guillemotright/ellipsis/nbspace/Agrave/Atilde/Otilde/OE/oe
+/endash/emdash/quotedblleft/quotedblright/quoteleft/quoteright/divide/lozenge
+/ydieresis/Ydieresis/fraction/currency/guilsinglleft/guilsinglright/fi/fl
+/daggerdbl/periodcentered/quotesinglbase/quotedblbase
+/perthousand/Acircumflex/Ecircumflex/Aacute
+/Edieresis/Egrave/Iacute/Icircumflex/Idieresis/Igrave/Oacute/Ocircumflex
+/apple/Ograve/Uacute/Ucircumflex/Ugrave/dotlessi/circumflex/tilde
+/macron/breve/dotaccent/ring/cedilla/hungarumlaut/ogonek/caron
+MacVec 128 128 getinterval astore pop
+end %. AltsysDict
+%%EndResource
+%%EndProlog
+%%BeginSetup
+AltsysDict begin
+_bfh
+%%IncludeResource: font Symbol
+_efh
+0 dict dup begin
+end
+/f0 /Symbol FF def
+_bfh
+%%IncludeResource: font ZapfHumanist601BT-Bold
+_efh
+0 dict dup begin
+end
+/f1 /ZapfHumanist601BT-Bold FF def
+end %. AltsysDict
+%%EndSetup
+AltsysDict begin
+/onlyk4{false}ndf
+/ccmyk{dup 5 -1 roll sub 0 max exch}ndf
+/cmyk2gray{
+ 4 -1 roll 0.3 mul 4 -1 roll 0.59 mul 4 -1 roll 0.11 mul
+ add add add 1 min neg 1 add
+}bdf
+/setcmykcolor{1 exch sub ccmyk ccmyk ccmyk pop setrgbcolor}ndf
+/maxcolor {
+ max max max
+} ndf
+/maxspot {
+ pop
+} ndf
+/setcmykcoloroverprint{4{dup -1 eq{pop 0}if 4 1 roll}repeat setcmykcolor}ndf
+/findcmykcustomcolor{5 packedarray}ndf
+/setcustomcolor{exch aload pop pop 4{4 index mul 4 1 roll}repeat setcmykcolor pop}ndf
+/setseparationgray{setgray}ndf
+/setoverprint{pop}ndf
+/currentoverprint false ndf
+/cmykbufs2gray{
+ 0 1 2 index length 1 sub
+ {
+4 index 1 index get 0.3 mul
+4 index 2 index get 0.59 mul
+4 index 3 index get 0.11 mul
+4 index 4 index get
+add add add cvi 255 min
+255 exch sub
+2 index 3 1 roll put
+ }for
+ 4 1 roll pop pop pop
+}bdf
+/colorimage{
+ pop pop
+ [
+5 -1 roll/exec cvx
+6 -1 roll/exec cvx
+7 -1 roll/exec cvx
+8 -1 roll/exec cvx
+/cmykbufs2gray cvx
+ ]cvx
+ image
+}
+%. version 47.1 on Linotronic of Postscript defines colorimage incorrectly (rgb model only)
+version cvr 47.1 le
+statusdict /product get (Lino) anchorsearch{pop pop true}{pop false}ifelse
+and{userdict begin bdf end}{ndf}ifelse
+fhnumcolors 1 ne {/yt save def} if
+/customcolorimage{
+ aload pop
+ (_vc_Registration) eq
+ {
+pop pop pop pop separationimage
+ }
+ {
+/ik xdf /iy xdf /im xdf /ic xdf
+ic im iy ik cmyk2gray /xt xdf
+currenttransfer
+{dup 1.0 exch sub xt mul add}concatprocs
+st
+image
+ }
+ ifelse
+}ndf
+fhnumcolors 1 ne {yt restore} if
+fhnumcolors 3 ne {/yt save def} if
+/customcolorimage{
+ aload pop
+ (_vc_Registration) eq
+ {
+pop pop pop pop separationimage
+ }
+ {
+/ik xdf /iy xdf /im xdf /ic xdf
+1.0 dup ic ik add min sub
+1.0 dup im ik add min sub
+1.0 dup iy ik add min sub
+/ic xdf /iy xdf /im xdf
+currentcolortransfer
+4 1 roll
+{dup 1.0 exch sub ic mul add}concatprocs 4 1 roll
+{dup 1.0 exch sub iy mul add}concatprocs 4 1 roll
+{dup 1.0 exch sub im mul add}concatprocs 4 1 roll
+setcolortransfer
+{/dummy xdf dummy}concatprocs{dummy}{dummy}true 3 colorimage
+ }
+ ifelse
+}ndf
+fhnumcolors 3 ne {yt restore} if
+fhnumcolors 4 ne {/yt save def} if
+/customcolorimage{
+ aload pop
+ (_vc_Registration) eq
+ {
+pop pop pop pop separationimage
+ }
+ {
+/ik xdf /iy xdf /im xdf /ic xdf
+currentcolortransfer
+{1.0 exch sub ik mul ik sub 1 add}concatprocs 4 1 roll
+{1.0 exch sub iy mul iy sub 1 add}concatprocs 4 1 roll
+{1.0 exch sub im mul im sub 1 add}concatprocs 4 1 roll
+{1.0 exch sub ic mul ic sub 1 add}concatprocs 4 1 roll
+setcolortransfer
+{/dummy xdf dummy}concatprocs{dummy}{dummy}{dummy}
+true 4 colorimage
+ }
+ ifelse
+}ndf
+fhnumcolors 4 ne {yt restore} if
+/separationimage{image}ndf
+/newcmykcustomcolor{6 packedarray}ndf
+/inkoverprint false ndf
+/setinkoverprint{pop}ndf
+/setspotcolor {
+ spots exch get
+ dup 4 get (_vc_Registration) eq
+ {pop 1 exch sub setseparationgray}
+ {0 5 getinterval exch setcustomcolor}
+ ifelse
+}ndf
+/currentcolortransfer{currenttransfer dup dup dup}ndf
+/setcolortransfer{st pop pop pop}ndf
+/fas{}ndf
+/sas{}ndf
+/fhsetspreadsize{pop}ndf
+/filler{fill}bdf
+/F{gsave {filler}fp grestore}bdf
+/f{closepath F}bdf
+/S{gsave {stroke}fp grestore}bdf
+/s{closepath S}bdf
+/bc4 [0 0 0 0] def
+/_lfp4 {
+ /iosv inkoverprint def
+ /cosv currentoverprint def
+ /yt xdf
+ /xt xdf
+ /ang xdf
+ storerect
+ /taperfcn xdf
+ /k2 xdf /y2 xdf /m2 xdf /c2 xdf
+ /k1 xdf /y1 xdf /m1 xdf /c1 xdf
+ c1 c2 sub abs
+ m1 m2 sub abs
+ y1 y2 sub abs
+ k1 k2 sub abs
+ maxcolor
+ calcgraysteps mul abs round
+ height abs adjnumsteps
+ dup 2 lt {pop 1} if
+ 1 sub /numsteps1 xdf
+ currentflat mark
+ currentflat clipflatness
+ /delta top bottom sub numsteps1 1 add div def
+ /right right left sub def
+ /botsv top delta sub def
+ {
+{
+W
+xt yt translate
+ang rotate
+xt neg yt neg translate
+dup setflat
+/bottom botsv def
+0 1 numsteps1
+{
+numsteps1 dup 0 eq {pop 0.5 } { div } ifelse
+taperfcn /frac xdf
+bc4 0 c2 c1 sub frac mul c1 add put
+bc4 1 m2 m1 sub frac mul m1 add put
+bc4 2 y2 y1 sub frac mul y1 add put
+bc4 3 k2 k1 sub frac mul k1 add put
+bc4 vc
+1 index setflat
+{
+mark {newpath left bottom right delta rectfill}stopped
+{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if}
+{cleartomark exit}ifelse
+}loop
+/bottom bottom delta sub def
+}for
+}
+gsave stopped grestore
+{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if}
+{exit}ifelse
+ }loop
+ cleartomark setflat
+ iosv setinkoverprint
+ cosv setoverprint
+}bdf
+/bcs [0 0] def
+/_lfs4 {
+ /iosv inkoverprint def
+ /cosv currentoverprint def
+ /yt xdf
+ /xt xdf
+ /ang xdf
+ storerect
+ /taperfcn xdf
+ /tint2 xdf
+ /tint1 xdf
+ bcs exch 1 exch put
+ tint1 tint2 sub abs
+ bcs 1 get maxspot
+ calcgraysteps mul abs round
+ height abs adjnumsteps
+ dup 2 lt {pop 2} if
+ 1 sub /numsteps1 xdf
+ currentflat mark
+ currentflat clipflatness
+ /delta top bottom sub numsteps1 1 add div def
+ /right right left sub def
+ /botsv top delta sub def
+ {
+{
+W
+xt yt translate
+ang rotate
+xt neg yt neg translate
+dup setflat
+/bottom botsv def
+0 1 numsteps1
+{
+numsteps1 div taperfcn /frac xdf
+bcs 0
+1.0 tint2 tint1 sub frac mul tint1 add sub
+put bcs vc
+1 index setflat
+{
+mark {newpath left bottom right delta rectfill}stopped
+{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if}
+{cleartomark exit}ifelse
+}loop
+/bottom bottom delta sub def
+}for
+}
+gsave stopped grestore
+{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if}
+{exit}ifelse
+ }loop
+ cleartomark setflat
+ iosv setinkoverprint
+ cosv setoverprint
+}bdf
+/_rfs4 {
+ /iosv inkoverprint def
+ /cosv currentoverprint def
+ /tint2 xdf
+ /tint1 xdf
+ bcs exch 1 exch put
+ /radius xdf
+ /yt xdf
+ /xt xdf
+ tint1 tint2 sub abs
+ bcs 1 get maxspot
+ calcgraysteps mul abs round
+ radius abs adjnumsteps
+ dup 2 lt {pop 2} if
+ 1 sub /numsteps1 xdf
+ radius numsteps1 div 2 div /halfstep xdf
+ currentflat mark
+ currentflat clipflatness
+ {
+{
+dup setflat
+W
+0 1 numsteps1
+{
+dup /radindex xdf
+numsteps1 div /frac xdf
+bcs 0
+tint2 tint1 sub frac mul tint1 add
+put bcs vc
+1 index setflat
+{
+newpath mark xt yt radius 1 frac sub mul halfstep add 0 360
+{ arc
+radindex numsteps1 ne
+{
+xt yt
+radindex 1 add numsteps1
+div 1 exch sub
+radius mul halfstep add
+dup xt add yt moveto
+360 0 arcn
+} if
+fill
+}stopped
+{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if}
+{cleartomark exit}ifelse
+}loop
+}for
+}
+gsave stopped grestore
+{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if}
+{exit}ifelse
+ }loop
+ cleartomark setflat
+ iosv setinkoverprint
+ cosv setoverprint
+}bdf
+/_rfp4 {
+ /iosv inkoverprint def
+ /cosv currentoverprint def
+ /k2 xdf /y2 xdf /m2 xdf /c2 xdf
+ /k1 xdf /y1 xdf /m1 xdf /c1 xdf
+ /radius xdf
+ /yt xdf
+ /xt xdf
+ c1 c2 sub abs
+ m1 m2 sub abs
+ y1 y2 sub abs
+ k1 k2 sub abs
+ maxcolor
+ calcgraysteps mul abs round
+ radius abs adjnumsteps
+ dup 2 lt {pop 1} if
+ 1 sub /numsteps1 xdf
+ radius numsteps1 dup 0 eq {pop} {div} ifelse
+ 2 div /halfstep xdf
+ currentflat mark
+ currentflat clipflatness
+ {
+{
+dup setflat
+W
+0 1 numsteps1
+{
+dup /radindex xdf
+numsteps1 dup 0 eq {pop 0.5 } { div } ifelse
+/frac xdf
+bc4 0 c2 c1 sub frac mul c1 add put
+bc4 1 m2 m1 sub frac mul m1 add put
+bc4 2 y2 y1 sub frac mul y1 add put
+bc4 3 k2 k1 sub frac mul k1 add put
+bc4 vc
+1 index setflat
+{
+newpath mark xt yt radius 1 frac sub mul halfstep add 0 360
+{ arc
+radindex numsteps1 ne
+{
+xt yt
+radindex 1 add
+numsteps1 dup 0 eq {pop} {div} ifelse
+1 exch sub
+radius mul halfstep add
+dup xt add yt moveto
+360 0 arcn
+} if
+fill
+}stopped
+{cleartomark exch 1.3 mul dup setflat exch 2 copy gt{stop}if}
+{cleartomark exit}ifelse
+}loop
+}for
+}
+gsave stopped grestore
+{exch pop 2 index exch 1.3 mul dup 100 gt{cleartomark setflat stop}if}
+{exit}ifelse
+ }loop
+ cleartomark setflat
+ iosv setinkoverprint
+ cosv setoverprint
+}bdf
+/lfp4{_lfp4}ndf
+/lfs4{_lfs4}ndf
+/rfs4{_rfs4}ndf
+/rfp4{_rfp4}ndf
+/cvc [0 0 0 1] def
+/vc{
+ AltsysDict /cvc 2 index put
+ aload length 4 eq
+ {setcmykcolor}
+ {setspotcolor}
+ ifelse
+}bdf
+/origmtx matrix currentmatrix def
+/ImMatrix matrix currentmatrix def
+0 setseparationgray
+/imgr {1692 1570.1102 2287.2756 2412 } def
+/bleed 0 def
+/clpr {1692 1570.1102 2287.2756 2412 } def
+/xs 1 def
+/ys 1 def
+/botx 0 def
+/overlap 0 def
+/wdist 18 def
+0 2 mul fhsetspreadsize
+0 0 ne {/df 0 def /clipflatness 0 def} if
+/maxsteps 256 def
+/forcemaxsteps false def
+vms
+-1845 -1956 translate
+/currentpacking defed{false setpacking}if
+/spots[
+1 0 0 0 (Process Cyan) false newcmykcustomcolor
+0 1 0 0 (Process Magenta) false newcmykcustomcolor
+0 0 1 0 (Process Yellow) false newcmykcustomcolor
+0 0 0 1 (Process Black) false newcmykcustomcolor
+]def
+/textopf false def
+/curtextmtx{}def
+/otw .25 def
+/msf{dup/curtextmtx xdf makefont setfont}bdf
+/makesetfont/msf load def
+/curtextheight{.707104 .707104 curtextmtx dtransform
+ dup mul exch dup mul add sqrt}bdf
+/ta2{
+tempstr 2 index gsave exec grestore
+cwidth cheight rmoveto
+4 index eq{5 index 5 index rmoveto}if
+2 index 2 index rmoveto
+}bdf
+/ta{exch systemdict/cshow known
+{{/cheight xdf/cwidth xdf tempstr 0 2 index put ta2}exch cshow}
+{{tempstr 0 2 index put tempstr stringwidth/cheight xdf/cwidth xdf ta2}forall}
+ifelse 6{pop}repeat}bdf
+/sts{/textopf currentoverprint def vc setoverprint
+/ts{awidthshow}def exec textopf setoverprint}bdf
+/stol{/xt currentlinewidth def
+ setlinewidth vc newpath
+ /ts{{false charpath stroke}ta}def exec
+ xt setlinewidth}bdf
+
+/strk{/textopf currentoverprint def vc setoverprint
+ /ts{{false charpath stroke}ta}def exec
+ textopf setoverprint
+ }bdf
+n
+[] 0 d
+3.863708 M
+1 w
+0 j
+0 J
+false setoverprint
+0 i
+false eomode
+[0 0 0 1] vc
+vms
+%white border -- disabled
+%1845.2293 2127.8588 m
+%2045.9437 2127.8588 L
+%2045.9437 1956.1412 L
+%1845.2293 1956.1412 L
+%1845.2293 2127.8588 L
+%0.1417 w
+%2 J
+%2 M
+%[0 0 0 0] vc
+%s
+n
+1950.8 2097.2 m
+1958.8 2092.5 1967.3 2089 1975.5 2084.9 C
+1976.7 2083.5 1976.1 2081.5 1976.7 2079.9 C
+1979.6 2081.1 1981.6 2086.8 1985.3 2084 C
+1993.4 2079.3 2001.8 2075.8 2010 2071.7 C
+2010.5 2071.5 2010.5 2071.1 2010.8 2070.8 C
+2011.2 2064.3 2010.9 2057.5 2011 2050.8 C
+2015.8 2046.9 2022.2 2046.2 2026.6 2041.7 C
+2026.5 2032.5 2026.8 2022.9 2026.4 2014.1 C
+2020.4 2008.3 2015 2002.4 2008.8 1997.1 C
+2003.8 1996.8 2000.7 2001.2 1996.1 2002.1 C
+1995.2 1996.4 1996.9 1990.5 1995.6 1984.8 C
+1989.9 1979 1984.5 1973.9 1978.8 1967.8 C
+1977.7 1968.6 1976 1967.6 1974.5 1968.3 C
+1967.4 1972.5 1960.1 1976.1 1952.7 1979.3 C
+1946.8 1976.3 1943.4 1970.7 1938.5 1966.1 C
+1933.9 1966.5 1929.4 1968.8 1925.1 1970.7 C
+1917.2 1978.2 1906 1977.9 1897.2 1983.4 C
+1893.2 1985.6 1889.4 1988.6 1885 1990.1 C
+1884.6 1990.6 1883.9 1991 1883.8 1991.6 C
+1883.7 2000.4 1884 2009.9 1883.6 2018.9 C
+1887.7 2024 1893.2 2028.8 1898 2033.8 C
+1899.1 2035.5 1900.9 2036.8 1902.5 2037.9 C
+1903.9 2037.3 1905.2 2036.6 1906.4 2035.5 C
+1906.3 2039.7 1906.5 2044.6 1906.1 2048.9 C
+1906.3 2049.6 1906.7 2050.2 1907.1 2050.8 C
+1913.4 2056 1918.5 2062.7 1924.8 2068.1 C
+1926.6 2067.9 1928 2066.9 1929.4 2066 C
+1930.2 2071 1927.7 2077.1 1930.6 2081.6 C
+1936.6 2086.9 1941.5 2092.9 1947.9 2097.9 C
+1949 2098.1 1949.9 2097.5 1950.8 2097.2 C
+[0 0 0 0.18] vc
+f
+0.4 w
+S
+n
+1975.2 2084.7 m
+1976.6 2083.4 1975.7 2081.1 1976 2079.4 C
+1979.3 2079.5 1980.9 2086.2 1984.8 2084 C
+1992.9 2078.9 2001.7 2075.6 2010 2071.2 C
+2011 2064.6 2010.2 2057.3 2010.8 2050.6 C
+2015.4 2046.9 2021.1 2045.9 2025.9 2042.4 C
+2026.5 2033.2 2026.8 2022.9 2025.6 2013.9 C
+2020.5 2008.1 2014.5 2003.1 2009.3 1997.6 C
+2004.1 1996.7 2000.7 2001.6 1995.9 2002.6 C
+1995.2 1996.7 1996.3 1990.2 1994.9 1984.6 C
+1989.8 1978.7 1983.6 1973.7 1978.4 1968 C
+1977.3 1969.3 1976 1967.6 1974.8 1968.5 C
+1967.7 1972.7 1960.4 1976.3 1952.9 1979.6 C
+1946.5 1976.9 1943.1 1970.5 1937.8 1966.1 C
+1928.3 1968.2 1920.6 1974.8 1911.6 1978.4 C
+1901.9 1979.7 1893.9 1986.6 1885 1990.6 C
+1884.3 1991 1884.3 1991.7 1884 1992.3 C
+1884.5 2001 1884.2 2011 1884.3 2019.9 C
+1890.9 2025.3 1895.9 2031.9 1902.3 2037.4 C
+1904.2 2037.9 1905.6 2034.2 1906.8 2035.7 C
+1907.4 2040.9 1905.7 2046.1 1907.3 2050.8 C
+1913.6 2056.2 1919.2 2062.6 1925.1 2067.9 C
+1926.9 2067.8 1928 2066.3 1929.6 2065.7 C
+1929.9 2070.5 1929.2 2076 1930.1 2080.8 C
+1936.5 2086.1 1941.6 2092.8 1948.4 2097.6 C
+1957.3 2093.3 1966.2 2088.8 1975.2 2084.7 C
+[0 0 0 0] vc
+f
+S
+n
+1954.8 2093.8 m
+1961.6 2090.5 1968.2 2087 1975 2084 C
+1975 2082.8 1975.6 2080.9 1974.8 2080.6 C
+1974.3 2075.2 1974.6 2069.6 1974.5 2064 C
+1977.5 2059.7 1984.5 2060 1988.9 2056.4 C
+1989.5 2055.5 1990.5 2055.3 1990.8 2054.4 C
+1991.1 2045.7 1991.4 2036.1 1990.6 2027.8 C
+1990.7 2026.6 1992 2027.3 1992.8 2027.1 C
+1997 2032.4 2002.6 2037.8 2007.6 2042.2 C
+2008.7 2042.3 2007.8 2040.6 2007.4 2040 C
+2002.3 2035.6 1997.5 2030 1992.8 2025.2 C
+1991.6 2024.7 1990.8 2024.9 1990.1 2025.4 C
+1989.4 2024.9 1988.1 2025.2 1987.2 2024.4 C
+1987.1 2025.8 1988.3 2026.5 1989.4 2026.8 C
+1989.4 2026.6 1989.3 2026.2 1989.6 2026.1 C
+1989.9 2026.2 1989.9 2026.6 1989.9 2026.8 C
+1989.8 2026.6 1990 2026.5 1990.1 2026.4 C
+1990.2 2027 1991.1 2028.3 1990.1 2028 C
+1989.9 2037.9 1990.5 2044.1 1989.6 2054.2 C
+1985.9 2058 1979.7 2057.4 1976 2061.2 C
+1974.5 2061.6 1975.2 2059.9 1974.5 2059.5 C
+1973.9 2058 1975.6 2057.8 1975 2056.6 C
+1974.5 2057.1 1974.6 2055.3 1973.6 2055.9 C
+1971.9 2059.3 1974.7 2062.1 1973.1 2065.5 C
+1973.1 2071.2 1972.9 2077 1973.3 2082.5 C
+1967.7 2085.6 1962 2088 1956.3 2090.7 C
+1953.9 2092.4 1951 2093 1948.6 2094.8 C
+1943.7 2089.9 1937.9 2084.3 1933 2079.6 C
+1931.3 2076.1 1933.2 2071.3 1932.3 2067.2 C
+1931.3 2062.9 1933.3 2060.6 1932 2057.6 C
+1932.7 2056.5 1930.9 2053.3 1933.2 2051.8 C
+1936.8 2050.1 1940.1 2046.9 1944 2046.8 C
+1946.3 2049.7 1949.3 2051.9 1952 2054.4 C
+1954.5 2054.2 1956.4 2052.3 1958.7 2051.3 C
+1960.8 2050 1963.2 2049 1965.6 2048.4 C
+1968.3 2050.8 1970.7 2054.3 1973.6 2055.4 C
+1973 2052.2 1969.7 2050.4 1967.6 2048.2 C
+1967.1 2046.7 1968.8 2046.6 1969.5 2045.8 C
+1972.8 2043.3 1980.6 2043.4 1979.3 2038.4 C
+1979.4 2038.6 1979.2 2038.7 1979.1 2038.8 C
+1978.7 2038.6 1978.9 2038.1 1978.8 2037.6 C
+1978.9 2037.9 1978.7 2038 1978.6 2038.1 C
+1978.2 2032.7 1978.4 2027.1 1978.4 2021.6 C
+1979.3 2021.1 1980 2020.2 1981.5 2020.1 C
+1983.5 2020.5 1984 2021.8 1985.1 2023.5 C
+1985.7 2024 1987.4 2023.7 1986 2022.8 C
+1984.7 2021.7 1983.3 2020.8 1983.9 2018.7 C
+1987.2 2015.9 1993 2015.4 1994.9 2011.5 C
+1992.2 2004.9 1999.3 2005.2 2002.1 2002.4 C
+2005.9 2002.7 2004.8 1997.4 2009.1 1999 C
+2011 1999.3 2010 2002.9 2012.7 2002.4 C
+2010.2 2000.7 2009.4 1996.1 2005.5 1998.5 C
+2002.1 2000.3 1999 2002.5 1995.4 2003.8 C
+1995.2 2003.6 1994.9 2003.3 1994.7 2003.1 C
+1994.3 1997 1995.6 1991.1 1994.4 1985.3 C
+1994.3 1986 1993.8 1985 1994 1985.6 C
+1993.8 1995.4 1994.4 2001.6 1993.5 2011.7 C
+1989.7 2015.5 1983.6 2014.9 1979.8 2018.7 C
+1978.3 2019.1 1979.1 2017.4 1978.4 2017 C
+1977.8 2015.5 1979.4 2015.3 1978.8 2014.1 C
+1978.4 2014.6 1978.5 2012.8 1977.4 2013.4 C
+1975.8 2016.8 1978.5 2019.6 1976.9 2023 C
+1977 2028.7 1976.7 2034.5 1977.2 2040 C
+1971.6 2043.1 1965.8 2045.6 1960.1 2048.2 C
+1957.7 2049.9 1954.8 2050.5 1952.4 2052.3 C
+1947.6 2047.4 1941.8 2041.8 1936.8 2037.2 C
+1935.2 2033.6 1937.1 2028.8 1936.1 2024.7 C
+1935.1 2020.4 1937.1 2018.1 1935.9 2015.1 C
+1936.5 2014.1 1934.7 2010.8 1937.1 2009.3 C
+1944.4 2004.8 1952 2000.9 1959.9 1997.8 C
+1963.9 1997 1963.9 2001.9 1966.8 2003.3 C
+1970.3 2006.9 1973.7 2009.9 1976.9 2012.9 C
+1977.9 2013 1977.1 2011.4 1976.7 2010.8 C
+1971.6 2006.3 1966.8 2000.7 1962 1995.9 C
+1960 1995.2 1960.1 1996.6 1958.2 1995.6 C
+1957 1997 1955.1 1998.8 1953.2 1998 C
+1951.7 1994.5 1954.1 1993.4 1952.9 1991.1 C
+1952.1 1990.5 1953.3 1990.2 1953.2 1989.6 C
+1954.2 1986.8 1950.9 1981.4 1954.4 1981.2 C
+1954.7 1981.6 1954.7 1981.7 1955.1 1982 C
+1961.9 1979.1 1967.6 1975 1974.3 1971.6 C
+1974.7 1969.8 1976.7 1969.5 1978.4 1969.7 C
+1980.3 1970 1979.3 1973.6 1982 1973.1 C
+1975.8 1962.2 1968 1975.8 1960.8 1976.7 C
+1956.9 1977.4 1953.3 1982.4 1949.1 1978.8 C
+1946 1975.8 1941.2 1971 1939.5 1969.2 C
+1938.5 1968.6 1938.9 1967.4 1937.8 1966.8 C
+1928.7 1969.4 1920.6 1974.5 1912.4 1979.1 C
+1904 1980 1896.6 1985 1889.3 1989.4 C
+1887.9 1990.4 1885.1 1990.3 1885 1992.5 C
+1885.4 2000.6 1885.2 2012.9 1885.2 2019.9 C
+1886.1 2022 1889.7 2019.5 1888.4 2022.8 C
+1889 2023.3 1889.8 2024.4 1890.3 2024 C
+1891.2 2023.5 1891.8 2028.2 1893.4 2026.6 C
+1894.2 2026.3 1893.9 2027.3 1894.4 2027.6 C
+1893.4 2027.6 1894.7 2028.3 1894.1 2028.5 C
+1894.4 2029.6 1896 2030 1896 2029.2 C
+1896.2 2029 1896.3 2029 1896.5 2029.2 C
+1896.8 2029.8 1897.3 2030 1897 2030.7 C
+1896.5 2030.7 1896.9 2031.5 1897.2 2031.6 C
+1898.3 2034 1899.5 2030.6 1899.6 2033.3 C
+1898.5 2033 1899.6 2034.4 1900.1 2034.8 C
+1901.3 2035.8 1903.2 2034.6 1902.5 2036.7 C
+1904.4 2036.9 1906.1 2032.2 1907.6 2035.5 C
+1907.5 2040.1 1907.7 2044.9 1907.3 2049.4 C
+1908 2050.2 1908.3 2051.4 1909.5 2051.6 C
+1910.1 2051.1 1911.6 2051.1 1911.4 2052.3 C
+1909.7 2052.8 1912.4 2054 1912.6 2054.7 C
+1913.4 2055.2 1913 2053.7 1913.6 2054.4 C
+1913.6 2054.5 1913.6 2055.3 1913.6 2054.7 C
+1913.7 2054.4 1913.9 2054.4 1914 2054.7 C
+1914 2054.9 1914.1 2055.3 1913.8 2055.4 C
+1913.7 2056 1915.2 2057.6 1916 2057.6 C
+1915.9 2057.3 1916.1 2057.2 1916.2 2057.1 C
+1917 2056.8 1916.7 2057.7 1917.2 2058 C
+1917 2058.3 1916.7 2058.3 1916.4 2058.3 C
+1917.1 2059 1917.3 2060.1 1918.4 2060.4 C
+1918.1 2059.2 1919.1 2060.6 1919.1 2059.5 C
+1919 2060.6 1920.6 2060.1 1919.8 2061.2 C
+1919.6 2061.2 1919.3 2061.2 1919.1 2061.2 C
+1919.6 2061.9 1921.4 2064.2 1921.5 2062.6 C
+1922.4 2062.1 1921.6 2063.9 1922.2 2064.3 C
+1922.9 2067.3 1926.1 2064.3 1925.6 2067.2 C
+1927.2 2066.8 1928.4 2064.6 1930.1 2065.2 C
+1931.8 2067.8 1931 2071.8 1930.8 2074.8 C
+1930.6 2076.4 1930.1 2078.6 1930.6 2080.4 C
+1936.6 2085.4 1941.8 2091.6 1948.1 2096.9 C
+1950.7 2096.7 1952.6 2094.8 1954.8 2093.8 C
+[0 0.33 0.33 0.99] vc
+f
+S
+n
+1989.4 2080.6 m
+1996.1 2077.3 2002.7 2073.8 2009.6 2070.8 C
+2009.6 2069.6 2010.2 2067.7 2009.3 2067.4 C
+2008.9 2062 2009.1 2056.4 2009.1 2050.8 C
+2012.3 2046.6 2019 2046.6 2023.5 2043.2 C
+2024 2042.3 2025.1 2042.1 2025.4 2041.2 C
+2025.3 2032.7 2025.6 2023.1 2025.2 2014.6 C
+2025 2015.3 2024.6 2014.2 2024.7 2014.8 C
+2024.5 2024.7 2025.1 2030.9 2024.2 2041 C
+2020.4 2044.8 2014.3 2044.2 2010.5 2048 C
+2009 2048.4 2009.8 2046.7 2009.1 2046.3 C
+2008.5 2044.8 2010.2 2044.6 2009.6 2043.4 C
+2009.1 2043.9 2009.2 2042.1 2008.1 2042.7 C
+2006.5 2046.1 2009.3 2048.9 2007.6 2052.3 C
+2007.7 2058 2007.5 2063.8 2007.9 2069.3 C
+2002.3 2072.4 1996.5 2074.8 1990.8 2077.5 C
+1988.4 2079.2 1985.6 2079.8 1983.2 2081.6 C
+1980.5 2079 1977.9 2076.5 1975.5 2074.1 C
+1975.5 2075.1 1975.5 2076.2 1975.5 2077.2 C
+1977.8 2079.3 1980.3 2081.6 1982.7 2083.7 C
+1985.3 2083.5 1987.1 2081.6 1989.4 2080.6 C
+f
+S
+n
+1930.1 2079.9 m
+1931.1 2075.6 1929.2 2071.1 1930.8 2067.2 C
+1930.3 2066.3 1930.1 2064.6 1928.7 2065.5 C
+1927.7 2066.4 1926.5 2067 1925.3 2067.4 C
+1924.5 2066.9 1925.6 2065.7 1924.4 2066 C
+1924.2 2067.2 1923.6 2065.5 1923.2 2065.7 C
+1922.3 2063.6 1917.8 2062.1 1919.6 2060.4 C
+1919.3 2060.5 1919.2 2060.3 1919.1 2060.2 C
+1919.7 2060.9 1918.2 2061 1917.6 2060.2 C
+1917 2059.6 1916.1 2058.8 1916.4 2058 C
+1915.5 2058 1917.4 2057.1 1915.7 2057.8 C
+1914.8 2057.1 1913.4 2056.2 1913.3 2054.9 C
+1913.1 2055.4 1911.3 2054.3 1910.9 2053.2 C
+1910.7 2052.9 1910.2 2052.5 1910.7 2052.3 C
+1911.1 2052.5 1910.9 2052 1910.9 2051.8 C
+1910.5 2051.2 1909.9 2052.6 1909.2 2051.8 C
+1908.2 2051.4 1907.8 2050.2 1907.1 2049.4 C
+1907.5 2044.8 1907.3 2040 1907.3 2035.2 C
+1905.3 2033 1902.8 2039.3 1902.3 2035.7 C
+1899.6 2036 1898.4 2032.5 1896.3 2030.7 C
+1895.7 2030.1 1897.5 2030 1896.3 2029.7 C
+1896.3 2030.6 1895 2029.7 1894.4 2029.2 C
+1892.9 2028.1 1894.2 2027.4 1893.6 2027.1 C
+1892.1 2027.9 1891.7 2025.6 1890.8 2024.9 C
+1891.1 2024.6 1889.1 2024.3 1888.4 2023 C
+1887.5 2022.6 1888.2 2021.9 1888.1 2021.3 C
+1886.7 2022 1885.2 2020.4 1884.8 2019.2 C
+1884.8 2010 1884.6 2000.2 1885 1991.8 C
+1886.9 1989.6 1889.9 1989.3 1892.2 1987.5 C
+1898.3 1982.7 1905.6 1980.1 1912.8 1978.6 C
+1921 1974.2 1928.8 1968.9 1937.8 1966.6 C
+1939.8 1968.3 1938.8 1968.3 1940.4 1970 C
+1945.4 1972.5 1947.6 1981.5 1954.6 1979.3 C
+1952.3 1981 1950.4 1978.4 1948.6 1977.9 C
+1945.1 1973.9 1941.1 1970.6 1938 1966.6 C
+1928.4 1968.5 1920.6 1974.8 1911.9 1978.8 C
+1907.1 1979.2 1902.6 1981.7 1898.2 1983.6 C
+1893.9 1986 1889.9 1989 1885.5 1990.8 C
+1884.9 1991.2 1884.8 1991.8 1884.5 1992.3 C
+1884.9 2001.3 1884.7 2011.1 1884.8 2019.6 C
+1890.6 2025 1896.5 2031.2 1902.3 2036.9 C
+1904.6 2037.6 1905 2033 1907.3 2035.5 C
+1907.2 2040.2 1907 2044.8 1907.1 2049.6 C
+1913.6 2055.3 1918.4 2061.5 1925.1 2067.4 C
+1927.3 2068.2 1929.6 2062.5 1930.6 2066.9 C
+1929.7 2070.7 1930.3 2076 1930.1 2080.1 C
+1935.6 2085.7 1941.9 2090.7 1947.2 2096.7 C
+1942.2 2091.1 1935.5 2085.2 1930.1 2079.9 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1930.8 2061.9 m
+1930.3 2057.8 1931.8 2053.4 1931.1 2050.4 C
+1931.3 2050.3 1931.7 2050.5 1931.6 2050.1 C
+1933 2051.1 1934.4 2049.5 1935.9 2048.7 C
+1937 2046.5 1939.5 2047.1 1941.2 2045.1 C
+1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C
+1934 2039.7 1934.5 2038.1 1933.7 2037.6 C
+1934 2033.3 1933.1 2027.9 1934.4 2024.4 C
+1934.3 2023.8 1933.9 2022.8 1933 2022.8 C
+1931.6 2023.1 1930.5 2024.4 1929.2 2024.9 C
+1928.4 2024.5 1929.8 2023.5 1928.7 2023.5 C
+1927.7 2024.1 1926.2 2022.6 1925.6 2021.6 C
+1926.9 2021.6 1924.8 2020.6 1925.6 2020.4 C
+1924.7 2021.7 1923.9 2019.6 1923.2 2019.2 C
+1923.3 2018.3 1923.8 2018.1 1923.2 2018 C
+1922.9 2017.8 1922.9 2017.5 1922.9 2017.2 C
+1922.8 2018.3 1921.3 2017.3 1920.3 2018 C
+1916.6 2019.7 1913 2022.1 1910 2024.7 C
+1910 2032.9 1910 2041.2 1910 2049.4 C
+1915.4 2055.2 1920 2058.7 1925.3 2064.8 C
+1927.2 2064 1929 2061.4 1930.8 2061.9 C
+[0 0 0 0] vc
+f
+S
+n
+1907.6 2030.4 m
+1907.5 2027.1 1906.4 2021.7 1908.5 2019.9 C
+1908.8 2020.1 1908.9 2019 1909.2 2019.6 C
+1910 2019.6 1912 2019.2 1913.1 2018.2 C
+1913.7 2016.5 1920.2 2015.7 1917.4 2012.7 C
+1918.2 2011.2 1917 2013.8 1917.2 2012 C
+1916.9 2012.3 1916 2012.4 1915.2 2012 C
+1912.5 2010.5 1916.6 2008.8 1913.6 2009.6 C
+1912.6 2009.2 1911.1 2009 1910.9 2007.6 C
+1911 1999.2 1911.8 1989.8 1911.2 1982.2 C
+1910.1 1981.1 1908.8 1982.2 1907.6 1982.2 C
+1900.8 1986.5 1893.2 1988.8 1887.2 1994.2 C
+1887.2 2002.4 1887.2 2010.7 1887.2 2018.9 C
+1892.6 2024.7 1897.2 2028.2 1902.5 2034.3 C
+1904.3 2033.3 1906.2 2032.1 1907.6 2030.4 C
+f
+S
+n
+1910.7 2025.4 m
+1912.7 2022.4 1916.7 2020.8 1919.8 2018.9 C
+1920.2 2018.7 1920.6 2018.6 1921 2018.4 C
+1925 2020 1927.4 2028.5 1932 2024.2 C
+1932.3 2025 1932.5 2023.7 1932.8 2024.4 C
+1932.8 2028 1932.8 2031.5 1932.8 2035 C
+1931.9 2033.9 1932.5 2036.3 1932.3 2036.9 C
+1933.2 2036.4 1932.5 2038.5 1933 2038.4 C
+1933.1 2040.5 1935.6 2042.2 1936.6 2043.2 C
+1936.2 2042.4 1935.1 2040.8 1933.7 2040.3 C
+1932.2 2034.4 1933.8 2029.8 1933 2023.2 C
+1931.1 2024.9 1928.4 2026.4 1926.5 2023.5 C
+1925.1 2021.6 1923 2019.8 1921.5 2018.2 C
+1917.8 2018.9 1915.2 2022.5 1911.6 2023.5 C
+1910.8 2023.8 1911.2 2024.7 1910.4 2025.2 C
+1910.9 2031.8 1910.6 2039.1 1910.7 2045.6 C
+1910.1 2048 1910.7 2045.9 1911.2 2044.8 C
+1910.6 2038.5 1911.2 2031.8 1910.7 2025.4 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1910.7 2048.9 m
+1910.3 2047.4 1911.3 2046.5 1911.6 2045.3 C
+1912.9 2045.3 1913.9 2047.1 1915.2 2045.8 C
+1915.2 2044.9 1916.6 2043.3 1917.2 2042.9 C
+1918.7 2042.9 1919.4 2044.4 1920.5 2043.2 C
+1921.2 2042.2 1921.4 2040.9 1922.4 2040.3 C
+1924.5 2040.3 1925.7 2040.9 1926.8 2039.6 C
+1927.1 2037.9 1926.8 2038.1 1927.7 2037.6 C
+1929 2037.5 1930.4 2037 1931.6 2037.2 C
+1932.3 2038.2 1933.1 2038.7 1932.8 2040.3 C
+1935 2041.8 1935.9 2043.8 1938.5 2044.8 C
+1938.6 2045 1938.3 2045.5 1938.8 2045.3 C
+1939.1 2042.9 1935.4 2044.2 1935.4 2042.2 C
+1932.1 2040.8 1932.8 2037.2 1932 2034.8 C
+1932.3 2034 1932.7 2035.4 1932.5 2034.8 C
+1931.3 2031.8 1935.5 2020.1 1928.9 2025.9 C
+1924.6 2024.7 1922.6 2014.5 1917.4 2020.4 C
+1915.5 2022.8 1912 2022.6 1910.9 2025.4 C
+1911.5 2031.9 1910.9 2038.8 1911.4 2045.3 C
+1911.1 2046.5 1910 2047.4 1910.4 2048.9 C
+1915.1 2054.4 1920.4 2058.3 1925.1 2063.8 C
+1920.8 2058.6 1914.9 2054.3 1910.7 2048.9 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1934.7 2031.9 m
+1934.6 2030.7 1934.9 2029.5 1934.4 2028.5 C
+1934 2029.5 1934.3 2031.2 1934.2 2032.6 C
+1933.8 2031.7 1934.9 2031.6 1934.7 2031.9 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+vmrs
+1934.7 2019.4 m
+1934.1 2015.3 1935.6 2010.9 1934.9 2007.9 C
+1935.1 2007.8 1935.6 2008.1 1935.4 2007.6 C
+1936.8 2008.6 1938.2 2007 1939.7 2006.2 C
+1940.1 2004.3 1942.7 2005 1943.6 2003.8 C
+1945.1 2000.3 1954 2000.8 1950 1996.6 C
+1952.1 1993.3 1948.2 1989.2 1951.2 1985.6 C
+1953 1981.4 1948.4 1982.3 1947.9 1979.8 C
+1945.4 1979.6 1945.1 1975.5 1942.4 1975 C
+1942.4 1972.3 1938 1973.6 1938.5 1970.4 C
+1937.4 1969 1935.6 1970.1 1934.2 1970.2 C
+1927.5 1974.5 1919.8 1976.8 1913.8 1982.2 C
+1913.8 1990.4 1913.8 1998.7 1913.8 2006.9 C
+1919.3 2012.7 1923.8 2016.2 1929.2 2022.3 C
+1931.1 2021.6 1932.8 2018.9 1934.7 2019.4 C
+[0 0 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+2024.2 2038.1 m
+2024.1 2029.3 2024.4 2021.7 2024.7 2014.4 C
+2024.4 2013.6 2020.6 2013.4 2021.3 2011.2 C
+2020.5 2010.3 2018.4 2010.6 2018.9 2008.6 C
+2019 2008.8 2018.8 2009 2018.7 2009.1 C
+2018.2 2006.7 2015.2 2007.9 2015.3 2005.5 C
+2014.7 2004.8 2012.4 2005.1 2013.2 2003.6 C
+2012.3 2004.2 2012.8 2002.4 2012.7 2002.6 C
+2009.4 2003.3 2011.2 1998.6 2008.4 1999.2 C
+2007 1999.1 2006.1 1999.4 2005.7 2000.4 C
+2006.9 1998.5 2007.7 2000.5 2009.3 2000.2 C
+2009.2 2003.7 2012.4 2002.1 2012.9 2005.2 C
+2015.9 2005.6 2015.2 2008.6 2017.7 2008.8 C
+2018.4 2009.6 2018.3 2011.4 2019.6 2011 C
+2021.1 2011.7 2021.4 2014.8 2023.7 2015.1 C
+2023.7 2023.5 2023.9 2031.6 2023.5 2040.5 C
+2021.8 2041.7 2020.7 2043.6 2018.4 2043.9 C
+2020.8 2042.7 2025.5 2041.8 2024.2 2038.1 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+2023.5 2040 m
+2023.5 2031.1 2023.5 2023.4 2023.5 2015.1 C
+2020.2 2015 2021.8 2010.3 2018.4 2011 C
+2018.6 2007.5 2014.7 2009.3 2014.8 2006.4 C
+2011.8 2006.3 2012.2 2002.3 2009.8 2002.4 C
+2009.7 2001.5 2009.2 2000.1 2008.4 2000.2 C
+2008.7 2000.9 2009.7 2001.2 2009.3 2002.4 C
+2008.4 2004.2 2007.5 2003.1 2007.9 2005.5 C
+2007.9 2010.8 2007.7 2018.7 2008.1 2023.2 C
+2009 2024.3 2007.3 2023.4 2007.9 2024 C
+2007.7 2024.6 2007.3 2026.3 2008.6 2027.1 C
+2009.7 2026.8 2010 2027.6 2010.5 2028 C
+2010.5 2028.2 2010.5 2029.1 2010.5 2028.5 C
+2011.5 2028 2010.5 2030 2011.5 2030 C
+2014.2 2029.7 2012.9 2032.2 2014.8 2032.6 C
+2015.1 2033.6 2015.3 2033 2016 2033.3 C
+2017 2033.9 2016.6 2035.4 2017.2 2036.2 C
+2018.7 2036.4 2019.2 2039 2021.3 2038.4 C
+2021.6 2035.4 2019.7 2029.5 2021.1 2027.3 C
+2020.9 2023.5 2021.5 2018.5 2020.6 2016 C
+2020.9 2013.9 2021.5 2015.4 2022.3 2014.4 C
+2022.2 2015.1 2023.3 2014.8 2023.2 2015.6 C
+2022.7 2019.8 2023.3 2024.3 2022.8 2028.5 C
+2022.3 2028.2 2022.6 2027.6 2022.5 2027.1 C
+2022.5 2027.8 2022.5 2029.2 2022.5 2029.2 C
+2022.6 2029.2 2022.7 2029.1 2022.8 2029 C
+2023.9 2032.8 2022.6 2037 2023 2040.8 C
+2022.3 2041.2 2021.6 2041.5 2021.1 2042.2 C
+2022 2041.2 2022.9 2041.4 2023.5 2040 C
+[0 1 1 0.23] vc
+f
+S
+n
+2009.1 1997.8 m
+2003.8 1997.7 2000.1 2002.4 1995.4 2003.1 C
+1995 1999.5 1995.2 1995 1995.2 1992 C
+1995.2 1995.8 1995 1999.7 1995.4 2003.3 C
+2000.3 2002.2 2003.8 1997.9 2009.1 1997.8 C
+2012.3 2001.2 2015.6 2004.8 2018.7 2008.1 C
+2021.6 2011.2 2027.5 2013.9 2025.9 2019.9 C
+2026.1 2017.9 2025.6 2016.2 2025.4 2014.4 C
+2020.2 2008.4 2014 2003.6 2009.1 1997.8 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+2009.3 1997.8 m
+2008.7 1997.4 2007.9 1997.6 2007.2 1997.6 C
+2007.9 1997.6 2008.9 1997.4 2009.6 1997.8 C
+2014.7 2003.6 2020.8 2008.8 2025.9 2014.8 C
+2025.8 2017.7 2026.1 2014.8 2025.6 2014.1 C
+2020.4 2008.8 2014.8 2003.3 2009.3 1997.8 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+2009.6 1997.6 m
+2009 1997.1 2008.1 1997.4 2007.4 1997.3 C
+2008.1 1997.4 2009 1997.1 2009.6 1997.6 C
+2014.8 2003.7 2021.1 2008.3 2025.9 2014.4 C
+2021.1 2008.3 2014.7 2003.5 2009.6 1997.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+2021.8 2011.5 m
+2021.9 2012.2 2022.3 2013.5 2023.7 2013.6 C
+2023.4 2012.7 2022.8 2011.8 2021.8 2011.5 C
+[0 0.33 0.33 0.99] vc
+f
+S
+n
+2021.1 2042 m
+2022.1 2041.1 2020.9 2040.2 2020.6 2039.6 C
+2018.4 2039.5 2018.1 2036.9 2016.3 2036.4 C
+2015.8 2035.5 2015.3 2033.8 2014.8 2033.6 C
+2012.4 2033.8 2013 2030.4 2010.5 2030.2 C
+2009.6 2028.9 2009.6 2028.3 2008.4 2028 C
+2006.9 2026.7 2007.5 2024.3 2006 2023.2 C
+2006.6 2023.2 2005.7 2023.3 2005.7 2023 C
+2006.4 2022.5 2006.3 2021.1 2006.7 2020.6 C
+2006.6 2015 2006.9 2009 2006.4 2003.8 C
+2006.9 2002.5 2007.6 2001.1 2006.9 2000.7 C
+2004.6 2003.6 2003 2002.9 2000.2 2004.3 C
+1999.3 2005.8 1997.9 2006.3 1996.1 2006.7 C
+1995.7 2008.9 1996 2011.1 1995.9 2012.9 C
+1993.4 2015.1 1990.5 2016.2 1987.7 2017.7 C
+1987.1 2019.3 1991.1 2019.4 1990.4 2021.3 C
+1990.5 2021.5 1991.9 2022.3 1992 2023 C
+1994.8 2024.4 1996.2 2027.5 1998.5 2030 C
+2002.4 2033 2005.2 2037.2 2008.8 2041 C
+2010.2 2041.3 2011.6 2042 2011 2043.9 C
+2011.2 2044.8 2010.1 2045.3 2010.5 2046.3 C
+2013.8 2044.8 2017.5 2043.4 2021.1 2042 C
+[0 0.5 0.5 0.2] vc
+f
+S
+n
+2019.4 2008.8 m
+2018.9 2009.2 2019.3 2009.9 2019.6 2010.3 C
+2022.2 2011.5 2020.3 2009.1 2019.4 2008.8 C
+[0 0.33 0.33 0.99] vc
+f
+S
+n
+2018 2007.4 m
+2015.7 2006.7 2015.3 2003.6 2012.9 2002.8 C
+2013.5 2003.7 2013.5 2005.1 2015.6 2005.2 C
+2016.4 2006.1 2015.7 2007.7 2018 2007.4 C
+f
+S
+n
+vmrs
+1993.5 2008.8 m
+1993.4 2000 1993.7 1992.5 1994 1985.1 C
+1993.7 1984.3 1989.9 1984.1 1990.6 1982 C
+1989.8 1981.1 1987.7 1981.4 1988.2 1979.3 C
+1988.3 1979.6 1988.1 1979.7 1988 1979.8 C
+1987.5 1977.5 1984.5 1978.6 1984.6 1976.2 C
+1983.9 1975.5 1981.7 1975.8 1982.4 1974.3 C
+1981.6 1974.9 1982.1 1973.1 1982 1973.3 C
+1979 1973.7 1980 1968.8 1976.9 1969.7 C
+1975.9 1969.8 1975.3 1970.3 1975 1971.2 C
+1976.2 1969.2 1977 1971.2 1978.6 1970.9 C
+1978.5 1974.4 1981.7 1972.8 1982.2 1976 C
+1985.2 1976.3 1984.5 1979.3 1987 1979.6 C
+1987.7 1980.3 1987.5 1982.1 1988.9 1981.7 C
+1990.4 1982.4 1990.7 1985.5 1993 1985.8 C
+1992.9 1994.3 1993.2 2002.3 1992.8 2011.2 C
+1991.1 2012.4 1990 2014.4 1987.7 2014.6 C
+1990.1 2013.4 1994.7 2012.6 1993.5 2008.8 C
+[0 0.87 0.91 0.83] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1992.8 2010.8 m
+1992.8 2001.8 1992.8 1994.1 1992.8 1985.8 C
+1989.5 1985.7 1991.1 1981.1 1987.7 1981.7 C
+1987.9 1978.2 1983.9 1980 1984.1 1977.2 C
+1981.1 1977 1981.5 1973 1979.1 1973.1 C
+1979 1972.2 1978.5 1970.9 1977.6 1970.9 C
+1977.9 1971.6 1979 1971.9 1978.6 1973.1 C
+1977.6 1974.9 1976.8 1973.9 1977.2 1976.2 C
+1977.2 1981.5 1977 1989.4 1977.4 1994 C
+1978.3 1995 1976.6 1994.1 1977.2 1994.7 C
+1977 1995.3 1976.6 1997 1977.9 1997.8 C
+1979 1997.5 1979.3 1998.3 1979.8 1998.8 C
+1979.8 1998.9 1979.8 1999.8 1979.8 1999.2 C
+1980.8 1998.7 1979.7 2000.7 1980.8 2000.7 C
+1983.5 2000.4 1982.1 2003 1984.1 2003.3 C
+1984.4 2004.3 1984.5 2003.7 1985.3 2004 C
+1986.3 2004.6 1985.9 2006.1 1986.5 2006.9 C
+1988 2007.1 1988.4 2009.7 1990.6 2009.1 C
+1990.9 2006.1 1989 2000.2 1990.4 1998 C
+1990.2 1994.3 1990.8 1989.2 1989.9 1986.8 C
+1990.2 1984.7 1990.8 1986.2 1991.6 1985.1 C
+1991.5 1985.9 1992.6 1985.5 1992.5 1986.3 C
+1992 1990.5 1992.6 1995 1992 1999.2 C
+1991.6 1998.9 1991.9 1998.3 1991.8 1997.8 C
+1991.8 1998.5 1991.8 2000 1991.8 2000 C
+1991.9 1999.9 1992 1999.8 1992 1999.7 C
+1993.2 2003.5 1991.9 2007.7 1992.3 2011.5 C
+1991.6 2012 1990.9 2012.2 1990.4 2012.9 C
+1991.3 2011.9 1992.2 2012.1 1992.8 2010.8 C
+[0 1 1 0.23] vc
+f
+S
+n
+1978.4 1968.5 m
+1977 1969.2 1975.8 1968.2 1974.5 1969 C
+1968.3 1973 1961.6 1976 1955.1 1979.1 C
+1962 1975.9 1968.8 1972.5 1975.5 1968.8 C
+1976.5 1968.8 1977.6 1968.8 1978.6 1968.8 C
+1981.7 1972.1 1984.8 1975.7 1988 1978.8 C
+1990.9 1981.9 1996.8 1984.6 1995.2 1990.6 C
+1995.3 1988.6 1994.9 1986.9 1994.7 1985.1 C
+1989.5 1979.1 1983.3 1974.3 1978.4 1968.5 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1978.4 1968.3 m
+1977.9 1968.7 1977.1 1968.5 1976.4 1968.5 C
+1977.3 1968.8 1978.1 1967.9 1978.8 1968.5 C
+1984 1974.3 1990.1 1979.5 1995.2 1985.6 C
+1995.1 1988.4 1995.3 1985.6 1994.9 1984.8 C
+1989.5 1979.4 1983.9 1973.8 1978.4 1968.3 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1978.6 1968 m
+1977.9 1968 1977.4 1968.6 1978.4 1968 C
+1983.9 1973.9 1990.1 1979.1 1995.2 1985.1 C
+1990.2 1979 1983.8 1974.1 1978.6 1968 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1991.1 1982.2 m
+1991.2 1982.9 1991.6 1984.2 1993 1984.4 C
+1992.6 1983.5 1992.1 1982.5 1991.1 1982.2 C
+[0 0.33 0.33 0.99] vc
+f
+S
+n
+1990.4 2012.7 m
+1991.4 2011.8 1990.2 2010.9 1989.9 2010.3 C
+1987.7 2010.2 1987.4 2007.6 1985.6 2007.2 C
+1985.1 2006.2 1984.6 2004.5 1984.1 2004.3 C
+1981.7 2004.5 1982.3 2001.2 1979.8 2000.9 C
+1978.8 1999.6 1978.8 1999.1 1977.6 1998.8 C
+1976.1 1997.4 1976.7 1995 1975.2 1994 C
+1975.8 1994 1975 1994 1975 1993.7 C
+1975.7 1993.2 1975.6 1991.8 1976 1991.3 C
+1975.9 1985.7 1976.1 1979.7 1975.7 1974.5 C
+1976.2 1973.3 1976.9 1971.8 1976.2 1971.4 C
+1973.9 1974.3 1972.2 1973.6 1969.5 1975 C
+1967.9 1977.5 1963.8 1977.1 1961.8 1980 C
+1959 1980 1957.6 1983 1954.8 1982.9 C
+1953.8 1984.2 1954.8 1985.7 1955.1 1987.2 C
+1956.2 1989.5 1959.7 1990.1 1959.9 1991.8 C
+1965.9 1998 1971.8 2005.2 1978.1 2011.7 C
+1979.5 2012 1980.9 2012.7 1980.3 2014.6 C
+1980.5 2015.6 1979.4 2016 1979.8 2017 C
+1983 2015.6 1986.8 2014.1 1990.4 2012.7 C
+[0 0.5 0.5 0.2] vc
+f
+S
+n
+1988.7 1979.6 m
+1988.2 1979.9 1988.6 1980.6 1988.9 1981 C
+1991.4 1982.2 1989.6 1979.9 1988.7 1979.6 C
+[0 0.33 0.33 0.99] vc
+f
+S
+n
+1987.2 1978.1 m
+1985 1977.5 1984.6 1974.3 1982.2 1973.6 C
+1982.7 1974.5 1982.8 1975.8 1984.8 1976 C
+1985.7 1976.9 1985 1978.4 1987.2 1978.1 C
+f
+S
+n
+1975.5 2084 m
+1975.5 2082 1975.3 2080 1975.7 2078.2 C
+1978.8 2079 1980.9 2085.5 1984.8 2083.5 C
+1993 2078.7 2001.6 2075 2010 2070.8 C
+2010.1 2064 2009.9 2057.2 2010.3 2050.6 C
+2014.8 2046.2 2020.9 2045.7 2025.6 2042 C
+2026.1 2035.1 2025.8 2028 2025.9 2021.1 C
+2025.8 2027.8 2026.1 2034.6 2025.6 2041.2 C
+2022.2 2044.9 2017.6 2046.8 2012.9 2048 C
+2012.5 2049.5 2010.4 2049.4 2009.8 2051.1 C
+2009.9 2057.6 2009.6 2064.2 2010 2070.5 C
+2001.2 2075.4 1992 2079.1 1983.2 2084 C
+1980.3 2082.3 1977.8 2079.2 1975.2 2077.5 C
+1974.9 2079.9 1977.2 2084.6 1973.3 2085.2 C
+1964.7 2088.6 1956.8 2093.7 1948.1 2097.2 C
+1949 2097.3 1949.6 2096.9 1950.3 2096.7 C
+1958.4 2091.9 1967.1 2088.2 1975.5 2084 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+vmrs
+1948.6 2094.5 m
+1950.2 2093.7 1951.8 2092.9 1953.4 2092.1 C
+1951.8 2092.9 1950.2 2093.7 1948.6 2094.5 C
+[0 0.87 0.91 0.83] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1971.6 2082.3 m
+1971.6 2081.9 1970.7 2081.1 1970.9 2081.3 C
+1970.7 2081.6 1970.6 2081.6 1970.4 2081.3 C
+1970.8 2080.1 1968.7 2081.7 1968.3 2080.8 C
+1966.6 2080.9 1966.7 2078 1964.2 2078.2 C
+1964.8 2075 1960.1 2075.8 1960.1 2072.9 C
+1958 2072.3 1957.5 2069.3 1955.3 2069.3 C
+1953.9 2070.9 1948.8 2067.8 1950 2072 C
+1949 2074 1943.2 2070.6 1944 2074.8 C
+1942.2 2076.6 1937.6 2073.9 1938 2078.2 C
+1936.7 2078.6 1935 2078.6 1933.7 2078.2 C
+1933.5 2080 1936.8 2080.7 1937.3 2082.8 C
+1939.9 2083.5 1940.6 2086.4 1942.6 2088 C
+1945.2 2089.2 1946 2091.3 1948.4 2093.6 C
+1956 2089.5 1963.9 2086.1 1971.6 2082.3 C
+[0 0.01 1 0] vc
+f
+S
+n
+1958.2 2089.7 m
+1956.4 2090 1955.6 2091.3 1953.9 2091.9 C
+1955.6 2091.9 1956.5 2089.7 1958.2 2089.7 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1929.9 2080.4 m
+1929.5 2077.3 1929.7 2073.9 1929.6 2070.8 C
+1929.8 2074.1 1929.2 2077.8 1930.1 2080.8 C
+1935.8 2085.9 1941.4 2091.3 1946.9 2096.9 C
+1941.2 2091 1935.7 2086 1929.9 2080.4 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1930.1 2080.4 m
+1935.8 2086 1941.5 2090.7 1946.9 2096.7 C
+1941.5 2090.9 1935.7 2085.8 1930.1 2080.4 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1940.9 2087.1 m
+1941.7 2088 1944.8 2090.6 1943.6 2089.2 C
+1942.5 2089 1941.6 2087.7 1940.9 2087.1 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1972.8 2082.8 m
+1973 2075.3 1972.4 2066.9 1973.3 2059.5 C
+1972.5 2058.9 1972.8 2057.3 1973.1 2056.4 C
+1974.8 2055.2 1973.4 2055.5 1972.4 2055.4 C
+1970.1 2053.2 1967.9 2050.9 1965.6 2048.7 C
+1960.9 2049.9 1956.9 2052.7 1952.4 2054.7 C
+1949.3 2052.5 1946.3 2049.5 1943.6 2046.8 C
+1939.9 2047.7 1936.8 2050.1 1933.5 2051.8 C
+1930.9 2054.9 1933.5 2056.2 1932.3 2059.7 C
+1933.2 2059.7 1932.2 2060.5 1932.5 2060.2 C
+1933.2 2062.5 1931.6 2064.6 1932.5 2067.4 C
+1932.9 2069.7 1932.7 2072.2 1932.8 2074.6 C
+1933.6 2070.6 1932.2 2066.3 1933 2062.6 C
+1934.4 2058.2 1929.8 2053.5 1935.2 2051.1 C
+1937.7 2049.7 1940.2 2048 1942.8 2046.8 C
+1945.9 2049.2 1948.8 2052 1951.7 2054.7 C
+1952.7 2054.7 1953.6 2054.6 1954.4 2054.2 C
+1958.1 2052.5 1961.7 2049.3 1965.9 2049.2 C
+1968.2 2052.8 1975.2 2055 1972.6 2060.9 C
+1973.3 2062.4 1972.2 2065.2 1972.6 2067.6 C
+1972.7 2072.6 1972.4 2077.7 1972.8 2082.5 C
+1968.1 2084.9 1963.5 2087.5 1958.7 2089.5 C
+1963.5 2087.4 1968.2 2085 1972.8 2082.8 C
+f
+S
+n
+1935.2 2081.1 m
+1936.8 2083.4 1938.6 2084.6 1940.4 2086.6 C
+1938.8 2084.4 1936.7 2083.4 1935.2 2081.1 C
+f
+S
+n
+1983.2 2081.3 m
+1984.8 2080.5 1986.3 2079.7 1988 2078.9 C
+1986.3 2079.7 1984.8 2080.5 1983.2 2081.3 C
+f
+S
+n
+2006.2 2069.1 m
+2006.2 2068.7 2005.2 2067.9 2005.5 2068.1 C
+2005.3 2068.4 2005.2 2068.4 2005 2068.1 C
+2005.4 2066.9 2003.3 2068.5 2002.8 2067.6 C
+2001.2 2067.7 2001.2 2064.8 1998.8 2065 C
+1999.4 2061.8 1994.7 2062.6 1994.7 2059.7 C
+1992.4 2059.5 1992.4 2055.8 1990.1 2056.8 C
+1985.9 2059.5 1981.1 2061 1976.9 2063.8 C
+1977.2 2067.6 1974.9 2074.2 1978.8 2075.8 C
+1979.6 2077.8 1981.7 2078.4 1982.9 2080.4 C
+1990.6 2076.3 1998.5 2072.9 2006.2 2069.1 C
+[0 0.01 1 0] vc
+f
+S
+n
+vmrs
+1992.8 2076.5 m
+1991 2076.8 1990.2 2078.1 1988.4 2078.7 C
+1990.2 2078.7 1991 2076.5 1992.8 2076.5 C
+[0 0.87 0.91 0.83] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1975.5 2073.4 m
+1976.1 2069.7 1973.9 2064.6 1977.4 2062.4 C
+1973.9 2064.5 1976.1 2069.9 1975.5 2073.6 C
+1976 2074.8 1979.3 2077.4 1978.1 2076 C
+1977 2075.7 1975.8 2074.5 1975.5 2073.4 C
+f
+S
+n
+2007.4 2069.6 m
+2007.6 2062.1 2007 2053.7 2007.9 2046.3 C
+2007.1 2045.7 2007.3 2044.1 2007.6 2043.2 C
+2009.4 2042 2007.9 2042.3 2006.9 2042.2 C
+2002.2 2037.4 1996.7 2032.4 1992.5 2027.3 C
+1992 2027.3 1991.6 2027.3 1991.1 2027.3 C
+1991.4 2035.6 1991.4 2045.6 1991.1 2054.4 C
+1990.5 2055.5 1988.4 2056.6 1990.6 2055.4 C
+1991.6 2055.4 1991.6 2054.1 1991.6 2053.2 C
+1990.8 2044.7 1991.9 2035.4 1991.6 2027.6 C
+1991.8 2027.6 1992 2027.6 1992.3 2027.6 C
+1997 2032.8 2002.5 2037.7 2007.2 2042.9 C
+2007.3 2044.8 2006.7 2047.4 2007.6 2048.4 C
+2006.9 2055.1 2007.1 2062.5 2007.4 2069.3 C
+2002.7 2071.7 1998.1 2074.3 1993.2 2076.3 C
+1998 2074.2 2002.7 2071.8 2007.4 2069.6 C
+f
+S
+n
+2006.7 2069.1 m
+2006.3 2068.6 2005.9 2067.7 2005.7 2066.9 C
+2005.7 2059.7 2005.9 2051.4 2005.5 2045.1 C
+2004.9 2045.3 2004.7 2044.5 2004.3 2045.3 C
+2005.1 2045.3 2004.2 2045.8 2004.8 2046 C
+2004.8 2052.2 2004.8 2059.2 2004.8 2064.5 C
+2005.7 2065.7 2005.1 2065.7 2005 2066.7 C
+2003.8 2067 2002.7 2067.2 2001.9 2066.4 C
+2001.3 2064.6 1998 2063.1 1998 2061.9 C
+1996.1 2062.3 1996.6 2058.3 1994.2 2058.8 C
+1992.6 2057.7 1992.7 2054.8 1989.9 2056.6 C
+1985.6 2059.3 1980.9 2060.8 1976.7 2063.6 C
+1976 2066.9 1976 2071.2 1976.7 2074.6 C
+1977.6 2070.8 1973.1 2062.1 1980.5 2061.2 C
+1984.3 2060.3 1987.5 2058.2 1990.8 2056.4 C
+1991.7 2056.8 1992.9 2057.2 1993.5 2059.2 C
+1994.3 2058.6 1994.4 2060.6 1994.7 2059.2 C
+1995.3 2062.7 1999.2 2061.4 1998.8 2064.8 C
+2001.8 2065.4 2002.5 2068.4 2005.2 2067.4 C
+2004.9 2067.9 2006 2068 2006.4 2069.1 C
+2001.8 2071.1 1997.4 2073.9 1992.8 2075.8 C
+1997.5 2073.8 2002 2071.2 2006.7 2069.1 C
+[0 0.2 1 0] vc
+f
+S
+n
+1988.7 2056.6 m
+1985.1 2058.7 1981.1 2060.1 1977.6 2061.9 C
+1981.3 2060.5 1985.6 2058.1 1988.7 2056.6 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1977.9 2059.5 m
+1975.7 2064.5 1973.7 2054.7 1975.2 2060.9 C
+1976 2060.6 1977.6 2059.7 1977.9 2059.5 C
+f
+S
+n
+1989.6 2051.3 m
+1990.1 2042.3 1989.8 2036.6 1989.9 2028 C
+1989.8 2027 1990.8 2028.3 1990.1 2027.3 C
+1988.9 2026.7 1986.7 2026.9 1986.8 2024.7 C
+1987.4 2023 1985.9 2024.6 1985.1 2023.7 C
+1984.1 2021.4 1982.5 2020.5 1980.3 2020.6 C
+1979.9 2020.8 1979.5 2021.1 1979.3 2021.6 C
+1979.7 2025.8 1978.4 2033 1979.6 2038.1 C
+1983.7 2042.9 1968.8 2044.6 1978.8 2042.7 C
+1979.3 2042.3 1979.6 2041.9 1980 2041.5 C
+1980 2034.8 1980 2027 1980 2021.6 C
+1981.3 2020.5 1981.7 2021.5 1982.9 2021.8 C
+1983.6 2024.7 1986.1 2023.8 1986.8 2026.4 C
+1987.1 2027.7 1988.6 2027.1 1989.2 2028.3 C
+1989.1 2036.7 1989.3 2044.8 1988.9 2053.7 C
+1987.2 2054.9 1986.2 2056.8 1983.9 2057.1 C
+1986.3 2055.9 1990.9 2055 1989.6 2051.3 C
+f
+S
+n
+1971.6 2078.9 m
+1971.4 2070.5 1972.1 2062.2 1971.6 2055.9 C
+1969.9 2053.7 1967.6 2051.7 1965.6 2049.6 C
+1961.4 2050.4 1957.6 2053.6 1953.4 2055.2 C
+1949.8 2055.6 1948.2 2051.2 1945.5 2049.6 C
+1945.1 2048.8 1944.5 2047.9 1943.6 2047.5 C
+1940.1 2047.8 1937.3 2051 1934 2052.3 C
+1933.7 2052.6 1933.7 2053 1933.2 2053.2 C
+1933.7 2060.8 1933.4 2067.2 1933.5 2074.6 C
+1933.8 2068.1 1934 2060.9 1933.2 2054 C
+1935.3 2050.9 1939.3 2049.6 1942.4 2047.5 C
+1942.8 2047.5 1943.4 2047.4 1943.8 2047.7 C
+1947.1 2050.2 1950.3 2057.9 1955.3 2054.4 C
+1955.4 2054.4 1955.5 2054.3 1955.6 2054.2 C
+1955.9 2057.6 1956.1 2061.8 1955.3 2064.8 C
+1955.4 2064.3 1955.1 2063.8 1955.6 2063.6 C
+1956 2066.6 1955.3 2068.7 1958.7 2069.8 C
+1959.2 2071.7 1961.4 2071.7 1962 2074.1 C
+1964.4 2074.2 1964 2077.7 1967.3 2078.4 C
+1967 2079.7 1968.1 2079.9 1969 2080.1 C
+1971.1 2079.9 1970 2079.2 1970.4 2078 C
+1969.5 2077.2 1970.3 2075.9 1969.7 2075.1 C
+1970.1 2069.8 1970.1 2063.6 1969.7 2058.8 C
+1969.2 2058.5 1970 2058.1 1970.2 2057.8 C
+1970.4 2058.3 1971.2 2057.7 1971.4 2058.3 C
+1971.5 2065.3 1971.2 2073.6 1971.6 2081.1 C
+1974.1 2081.4 1969.8 2084.3 1972.4 2082.5 C
+1971.9 2081.4 1971.6 2080.2 1971.6 2078.9 C
+[0 0.4 1 0] vc
+f
+S
+n
+1952.4 2052 m
+1954.1 2051.3 1955.6 2050.4 1957.2 2049.6 C
+1955.6 2050.4 1954.1 2051.3 1952.4 2052 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1975.5 2039.8 m
+1975.5 2039.4 1974.5 2038.7 1974.8 2038.8 C
+1974.6 2039.1 1974.5 2039.1 1974.3 2038.8 C
+1974.6 2037.6 1972.5 2039.3 1972.1 2038.4 C
+1970.4 2038.4 1970.5 2035.5 1968 2035.7 C
+1968.6 2032.5 1964 2033.3 1964 2030.4 C
+1961.9 2029.8 1961.4 2026.8 1959.2 2026.8 C
+1957.7 2028.5 1952.6 2025.3 1953.9 2029.5 C
+1952.9 2031.5 1947 2028.2 1947.9 2032.4 C
+1946 2034.2 1941.5 2031.5 1941.9 2035.7 C
+1940.6 2036.1 1938.9 2036.1 1937.6 2035.7 C
+1937.3 2037.5 1940.7 2038.2 1941.2 2040.3 C
+1943.7 2041.1 1944.4 2043.9 1946.4 2045.6 C
+1949.1 2046.7 1949.9 2048.8 1952.2 2051.1 C
+1959.9 2047.1 1967.7 2043.6 1975.5 2039.8 C
+[0 0.01 1 0] vc
+f
+S
+n
+vmrs
+1962 2047.2 m
+1960.2 2047.5 1959.5 2048.9 1957.7 2049.4 C
+1959.5 2049.5 1960.3 2047.2 1962 2047.2 C
+[0 0.87 0.91 0.83] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+2012.4 2046.3 m
+2010.3 2051.3 2008.3 2041.5 2009.8 2047.7 C
+2010.5 2047.4 2012.2 2046.5 2012.4 2046.3 C
+f
+S
+n
+1944.8 2044.6 m
+1945.5 2045.6 1948.6 2048.1 1947.4 2046.8 C
+1946.3 2046.5 1945.5 2045.2 1944.8 2044.6 C
+f
+S
+n
+1987.2 2054.9 m
+1983.7 2057.3 1979.6 2058 1976 2060.2 C
+1974.7 2058.2 1977.2 2055.8 1974.3 2054.9 C
+1973.1 2052 1970.4 2050.2 1968 2048 C
+1968 2047.7 1968 2047.4 1968.3 2047.2 C
+1969.5 2046.1 1983 2040.8 1972.4 2044.8 C
+1971.2 2046.6 1967.9 2046 1968 2048.2 C
+1970.5 2050.7 1973.8 2052.6 1974.3 2055.6 C
+1975.1 2055 1975.7 2056.7 1975.7 2057.1 C
+1975.7 2058.2 1974.8 2059.3 1975.5 2060.4 C
+1979.3 2058.2 1983.9 2057.7 1987.2 2054.9 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1967.8 2047.5 m
+1968.5 2047 1969.1 2046.5 1969.7 2046 C
+1969.1 2046.5 1968.5 2047 1967.8 2047.5 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1976.7 2040.3 m
+1976.9 2032.8 1976.3 2024.4 1977.2 2017 C
+1976.4 2016.5 1976.6 2014.8 1976.9 2013.9 C
+1978.7 2012.7 1977.2 2013 1976.2 2012.9 C
+1971.5 2008.1 1965.9 2003.1 1961.8 1998 C
+1960.9 1998 1960.1 1998 1959.2 1998 C
+1951.5 2001.1 1944.3 2005.5 1937.1 2009.6 C
+1935 2012.9 1937 2013.6 1936.1 2017.2 C
+1937.1 2017.2 1936 2018 1936.4 2017.7 C
+1937 2020.1 1935.5 2022.1 1936.4 2024.9 C
+1936.8 2027.2 1936.5 2029.7 1936.6 2032.1 C
+1937.4 2028.2 1936 2023.8 1936.8 2020.1 C
+1938.3 2015.7 1933.6 2011 1939 2008.6 C
+1945.9 2004.5 1953.1 2000.3 1960.6 1998.3 C
+1960.9 1998.3 1961.3 1998.3 1961.6 1998.3 C
+1966.2 2003.5 1971.8 2008.4 1976.4 2013.6 C
+1976.6 2015.5 1976 2018.1 1976.9 2019.2 C
+1976.1 2025.8 1976.4 2033.2 1976.7 2040 C
+1971.9 2042.4 1967.4 2045 1962.5 2047 C
+1967.3 2044.9 1972 2042.6 1976.7 2040.3 C
+f
+S
+n
+1939 2038.6 m
+1940.6 2040.9 1942.5 2042.1 1944.3 2044.1 C
+1942.7 2041.9 1940.6 2040.9 1939 2038.6 C
+f
+S
+n
+2006.2 2065.7 m
+2006 2057.3 2006.7 2049 2006.2 2042.7 C
+2002.1 2038.4 1997.7 2033.4 1993 2030 C
+1992.9 2029.3 1992.5 2028.6 1992 2028.3 C
+1992.1 2036.6 1991.9 2046.2 1992.3 2054.9 C
+1990.8 2056.2 1989 2056.7 1987.5 2058 C
+1988.7 2057.7 1990.7 2054.4 1993 2056.4 C
+1993.4 2058.8 1996 2058.2 1996.6 2060.9 C
+1999 2061 1998.5 2064.5 2001.9 2065.2 C
+2001.5 2066.5 2002.7 2066.7 2003.6 2066.9 C
+2005.7 2066.7 2004.6 2066 2005 2064.8 C
+2004 2064 2004.8 2062.7 2004.3 2061.9 C
+2004.6 2056.6 2004.6 2050.4 2004.3 2045.6 C
+2003.7 2045.3 2004.6 2044.9 2004.8 2044.6 C
+2005 2045.1 2005.7 2044.5 2006 2045.1 C
+2006 2052.1 2005.8 2060.4 2006.2 2067.9 C
+2008.7 2068.2 2004.4 2071.1 2006.9 2069.3 C
+2006.4 2068.2 2006.2 2067 2006.2 2065.7 C
+[0 0.4 1 0] vc
+f
+S
+n
+2021.8 2041.7 m
+2018.3 2044.1 2014.1 2044.8 2010.5 2047 C
+2009.3 2045 2011.7 2042.6 2008.8 2041.7 C
+2004.3 2035.1 1997.6 2030.9 1993 2024.4 C
+1992.1 2024 1991.5 2024.3 1990.8 2024 C
+1993.2 2023.9 1995.3 2027.1 1996.8 2029 C
+2000.4 2032.6 2004.9 2036.9 2008.4 2040.8 C
+2008.2 2043.1 2011.4 2042.8 2009.8 2045.8 C
+2009.8 2046.3 2009.7 2046.9 2010 2047.2 C
+2013.8 2045 2018.5 2044.5 2021.8 2041.7 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+2001.6 2034 m
+2000.7 2033.1 1999.9 2032.3 1999 2031.4 C
+1999.9 2032.3 2000.7 2033.1 2001.6 2034 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+vmrs
+1989.4 2024.4 m
+1989.5 2025.4 1988.6 2024.3 1988.9 2024.7 C
+1990.5 2025.8 1990.7 2024.2 1992.8 2024.9 C
+1993.8 2025.9 1995 2027.1 1995.9 2028 C
+1994.3 2026 1991.9 2023.4 1989.4 2024.4 C
+[0 0.87 0.91 0.83] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1984.8 2019.9 m
+1984.6 2018.6 1986.3 2017.2 1987.7 2016.8 C
+1987.2 2017.5 1982.9 2017.9 1984.4 2020.6 C
+1984.1 2019.9 1984.9 2020 1984.8 2019.9 C
+f
+S
+n
+1981.7 2017 m
+1979.6 2022 1977.6 2012.3 1979.1 2018.4 C
+1979.8 2018.1 1981.5 2017.2 1981.7 2017 C
+f
+S
+n
+1884.3 2019.2 m
+1884.7 2010.5 1884.5 2000.6 1884.5 1991.8 C
+1886.6 1989.3 1889.9 1988.9 1892.4 1987 C
+1890.8 1988.7 1886 1989.1 1884.3 1992.3 C
+1884.7 2001 1884.5 2011.3 1884.5 2019.9 C
+1891 2025.1 1895.7 2031.5 1902 2036.9 C
+1896.1 2031 1890 2024.9 1884.3 2019.2 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1884 2019.4 m
+1884.5 2010.6 1884.2 2000.4 1884.3 1991.8 C
+1884.8 1990.4 1887.8 1989 1884.8 1990.8 C
+1884.3 1991.3 1884.3 1992 1884 1992.5 C
+1884.5 2001.2 1884.2 2011.1 1884.3 2019.9 C
+1887.9 2023.1 1891.1 2026.4 1894.4 2030 C
+1891.7 2026.1 1887.1 2022.9 1884 2019.4 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1885 2011.7 m
+1885 2006.9 1885 2001.9 1885 1997.1 C
+1885 2001.9 1885 2006.9 1885 2011.7 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1975.5 2036.4 m
+1975.2 2028 1976 2019.7 1975.5 2013.4 C
+1971.1 2008.5 1965.6 2003.6 1961.6 1999 C
+1958.8 1998 1956 2000 1953.6 2001.2 C
+1948.2 2004.7 1941.9 2006.5 1937.1 2010.8 C
+1937.5 2018.3 1937.3 2024.7 1937.3 2032.1 C
+1937.6 2025.6 1937.9 2018.4 1937.1 2011.5 C
+1937.3 2011 1937.6 2010.5 1937.8 2010 C
+1944.6 2005.7 1951.9 2002.3 1959.2 1999 C
+1960.1 1998.5 1960.1 1999.8 1960.4 2000.4 C
+1959.7 2006.9 1959.7 2014.2 1959.4 2021.1 C
+1959 2021.1 1959.2 2021.9 1959.2 2022.3 C
+1959.2 2021.9 1959 2021.3 1959.4 2021.1 C
+1959.8 2024.1 1959.2 2026.2 1962.5 2027.3 C
+1963 2029.2 1965.3 2029.2 1965.9 2031.6 C
+1968.3 2031.8 1967.8 2035.2 1971.2 2036 C
+1970.8 2037.2 1971.9 2037.5 1972.8 2037.6 C
+1974.9 2037.4 1973.9 2036.7 1974.3 2035.5 C
+1973.3 2034.7 1974.1 2033.4 1973.6 2032.6 C
+1973.9 2027.3 1973.9 2021.1 1973.6 2016.3 C
+1973 2016 1973.9 2015.6 1974 2015.3 C
+1974.3 2015.9 1975 2015.3 1975.2 2015.8 C
+1975.3 2022.8 1975.1 2031.2 1975.5 2038.6 C
+1977.9 2039 1973.7 2041.8 1976.2 2040 C
+1975.7 2039 1975.5 2037.8 1975.5 2036.4 C
+[0 0.4 1 0] vc
+f
+S
+n
+1991.1 2012.4 m
+1987.5 2014.8 1983.4 2015.6 1979.8 2017.7 C
+1978.5 2015.7 1981 2013.3 1978.1 2012.4 C
+1973.6 2005.8 1966.8 2001.6 1962.3 1995.2 C
+1961.4 1994.7 1960.8 1995 1960.1 1994.7 C
+1962.5 1994.6 1964.6 1997.8 1966.1 1999.7 C
+1969.7 2003.3 1974.2 2007.6 1977.6 2011.5 C
+1977.5 2013.8 1980.6 2013.5 1979.1 2016.5 C
+1979.1 2017 1979 2017.6 1979.3 2018 C
+1983.1 2015.7 1987.8 2015.2 1991.1 2012.4 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1970.9 2004.8 m
+1970 2003.9 1969.2 2003 1968.3 2002.1 C
+1969.2 2003 1970 2003.9 1970.9 2004.8 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1887.9 1994.9 m
+1888.5 1992.3 1891.4 1992.2 1893.2 1990.8 C
+1898.4 1987.5 1904 1984.8 1909.5 1982.2 C
+1909.7 1982.7 1910.3 1982.1 1910.4 1982.7 C
+1909.5 1990.5 1910.1 1996.4 1910 2004.5 C
+1909.1 2003.4 1909.7 2005.8 1909.5 2006.4 C
+1910.4 2006 1909.7 2008 1910.2 2007.9 C
+1911.3 2010.6 1912.5 2012.6 1915.7 2013.4 C
+1915.8 2013.7 1915.5 2014.4 1916 2014.4 C
+1916.3 2015 1915.4 2016 1915.2 2016 C
+1916.1 2015.5 1916.5 2014.5 1916 2013.6 C
+1913.4 2013.3 1913.1 2010.5 1910.9 2009.8 C
+1910.7 2008.8 1910.4 2007.9 1910.2 2006.9 C
+1911.1 1998.8 1909.4 1990.7 1910.7 1982.4 C
+1910 1982.1 1908.9 1982.1 1908.3 1982.4 C
+1901.9 1986.1 1895 1988.7 1888.8 1993 C
+1888 1993.4 1888.4 1994.3 1887.6 1994.7 C
+1888.1 2001.3 1887.8 2008.6 1887.9 2015.1 C
+1887.3 2017.5 1887.9 2015.4 1888.4 2014.4 C
+1887.8 2008 1888.4 2001.3 1887.9 1994.9 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+vmrs
+1887.9 2018.4 m
+1887.5 2016.9 1888.5 2016 1888.8 2014.8 C
+1890.1 2014.8 1891.1 2016.6 1892.4 2015.3 C
+1892.4 2014.4 1893.8 2012.9 1894.4 2012.4 C
+1895.9 2012.4 1896.6 2013.9 1897.7 2012.7 C
+1898.4 2011.7 1898.6 2010.4 1899.6 2009.8 C
+1901.7 2009.9 1902.9 2010.4 1904 2009.1 C
+1904.3 2007.4 1904 2007.6 1904.9 2007.2 C
+1906.2 2007 1907.6 2006.5 1908.8 2006.7 C
+1910.6 2008.2 1909.8 2011.5 1912.6 2012 C
+1912.4 2013 1913.8 2012.7 1914 2013.2 C
+1911.5 2011.1 1909.1 2007.9 1909.2 2004.3 C
+1909.5 2003.5 1909.9 2004.9 1909.7 2004.3 C
+1909.9 1996.2 1909.3 1990.5 1910.2 1982.7 C
+1909.5 1982.6 1909.5 1982.6 1908.8 1982.7 C
+1903.1 1985.7 1897 1987.9 1891.7 1992 C
+1890.5 1993 1888.2 1992.9 1888.1 1994.9 C
+1888.7 2001.4 1888.1 2008.4 1888.6 2014.8 C
+1888.3 2016 1887.2 2016.9 1887.6 2018.4 C
+1892.3 2023.9 1897.6 2027.9 1902.3 2033.3 C
+1898 2028.2 1892.1 2023.8 1887.9 2018.4 C
+[0.4 0.4 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1910.9 1995.2 m
+1910.4 1999.8 1911 2003.3 1910.9 2008.1 C
+1910.9 2003.8 1910.9 1999.2 1910.9 1995.2 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1911.2 2004.3 m
+1911.2 2001.9 1911.2 1999.7 1911.2 1997.3 C
+1911.2 1999.7 1911.2 2001.9 1911.2 2004.3 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1958.7 1995.2 m
+1959 1995.6 1956.2 1995 1956.5 1996.8 C
+1955.8 1997.6 1954.2 1998.5 1953.6 1997.3 C
+1953.6 1990.8 1954.9 1989.6 1953.4 1983.9 C
+1953.4 1983.3 1953.3 1982.1 1954.4 1982 C
+1955.5 1982.6 1956.5 1981.3 1957.5 1981 C
+1956.3 1981.8 1954.7 1982.6 1953.9 1981.5 C
+1951.4 1983 1954.7 1988.8 1952.9 1990.6 C
+1953.8 1990.6 1953.2 1992.7 1953.4 1993.7 C
+1953.8 1994.5 1952.3 1996.1 1953.2 1997.8 C
+1956.3 1999.4 1957.5 1994 1959.9 1995.6 C
+1962 1994.4 1963.7 1997.7 1965.2 1998.8 C
+1963.5 1996.7 1961.2 1994.1 1958.7 1995.2 C
+f
+S
+n
+1945 2000.7 m
+1945.4 1998.7 1945.4 1997.9 1945 1995.9 C
+1944.5 1995.3 1944.2 1992.6 1945.7 1993.2 C
+1946 1992.2 1948.7 1992.5 1948.4 1990.6 C
+1947.5 1990.3 1948.1 1988.7 1947.9 1988.2 C
+1948.9 1987.8 1950.5 1986.8 1950.5 1984.6 C
+1951.5 1980.9 1946.7 1983 1947.2 1979.8 C
+1944.5 1979.9 1945.2 1976.6 1943.1 1976.7 C
+1941.8 1975.7 1942.1 1972.7 1939.2 1973.8 C
+1938.2 1974.6 1939.3 1971.6 1938.3 1970.9 C
+1938.8 1969.2 1933.4 1970.3 1937.3 1970 C
+1939.4 1971.2 1937.2 1973 1937.6 1974.3 C
+1937.2 1976.3 1937.1 1981.2 1937.8 1984.1 C
+1938.8 1982.3 1937.9 1976.6 1938.5 1973.1 C
+1938.9 1975 1938.5 1976.4 1939.7 1977.2 C
+1939.5 1983.5 1938.9 1991.3 1940.2 1997.3 C
+1939.4 1999.1 1938.6 1997.1 1937.8 1997.1 C
+1937.4 1996.7 1937.6 1996.1 1937.6 1995.6 C
+1936.5 1998.5 1940.1 1998.4 1940.9 2000.7 C
+1942.1 2000.4 1943.2 2001.3 1943.1 2002.4 C
+1943.6 2003.1 1941.1 2004.6 1942.8 2003.8 C
+1943.9 2002.5 1942.6 2000.6 1945 2000.7 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1914.5 2006.4 m
+1914.1 2004.9 1915.2 2004 1915.5 2002.8 C
+1916.7 2002.8 1917.8 2004.6 1919.1 2003.3 C
+1919 2002.4 1920.4 2000.9 1921 2000.4 C
+1922.5 2000.4 1923.2 2001.9 1924.4 2000.7 C
+1925 1999.7 1925.3 1998.4 1926.3 1997.8 C
+1928.4 1997.9 1929.5 1998.4 1930.6 1997.1 C
+1930.9 1995.4 1930.7 1995.6 1931.6 1995.2 C
+1932.8 1995 1934.3 1994.5 1935.4 1994.7 C
+1936.1 1995.8 1936.9 1996.2 1936.6 1997.8 C
+1938.9 1999.4 1939.7 2001.3 1942.4 2002.4 C
+1942.4 2002.5 1942.2 2003 1942.6 2002.8 C
+1942.9 2000.4 1939.2 2001.8 1939.2 1999.7 C
+1936.2 1998.6 1937 1995.3 1935.9 1993.5 C
+1937.1 1986.5 1935.2 1977.9 1937.6 1971.2 C
+1937.6 1970.3 1936.6 1971 1936.4 1970.4 C
+1930.2 1973.4 1924 1976 1918.4 1980 C
+1917.2 1981 1914.9 1980.9 1914.8 1982.9 C
+1915.3 1989.4 1914.7 1996.4 1915.2 2002.8 C
+1914.9 2004 1913.9 2004.9 1914.3 2006.4 C
+1919 2011.9 1924.2 2015.9 1928.9 2021.3 C
+1924.6 2016.2 1918.7 2011.8 1914.5 2006.4 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1914.5 1982.9 m
+1915.1 1980.3 1918 1980.2 1919.8 1978.8 C
+1925 1975.5 1930.6 1972.8 1936.1 1970.2 C
+1939.4 1970.6 1936.1 1974.2 1936.6 1976.4 C
+1936.5 1981.9 1936.8 1987.5 1936.4 1992.8 C
+1935.9 1992.8 1936.2 1993.5 1936.1 1994 C
+1937.1 1993.6 1936.2 1995.9 1936.8 1995.9 C
+1937 1998 1939.5 1999.7 1940.4 2000.7 C
+1940.1 1998.6 1935 1997.2 1937.6 1993.7 C
+1938.3 1985.7 1935.9 1976.8 1937.8 1970.7 C
+1936.9 1969.8 1935.4 1970.3 1934.4 1970.7 C
+1928.3 1974.4 1921.4 1976.7 1915.5 1981 C
+1914.6 1981.4 1915.1 1982.3 1914.3 1982.7 C
+1914.7 1989.3 1914.5 1996.6 1914.5 2003.1 C
+1913.9 2005.5 1914.5 2003.4 1915 2002.4 C
+1914.5 1996 1915.1 1989.3 1914.5 1982.9 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1939.2 1994.9 m
+1939.3 1995 1939.4 1995.1 1939.5 1995.2 C
+1939.1 1989 1939.3 1981.6 1939 1976.7 C
+1938.6 1976.3 1938.6 1974.6 1938.5 1973.3 C
+1938.7 1976.1 1938.1 1979.4 1939 1981.7 C
+1937.3 1986 1937.7 1991.6 1938 1996.4 C
+1937.3 1994.3 1939.6 1996.2 1939.2 1994.9 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1938.3 1988.4 m
+1938.5 1990.5 1937.9 1994.1 1938.8 1994.7 C
+1937.9 1992.6 1939 1990.6 1938.3 1988.4 C
+[0 0.87 0.91 0.83] vc
+f
+S
+n
+1938.8 1985.8 m
+1938.5 1985.9 1938.4 1985.7 1938.3 1985.6 C
+1938.4 1986.2 1938 1989.5 1938.8 1987.2 C
+1938.8 1986.8 1938.8 1986.3 1938.8 1985.8 C
+f
+S
+n
+vmrs
+1972.8 2062.1 m
+1971.9 2061 1972.5 2059.4 1972.4 2058 C
+1972.2 2063.8 1971.9 2073.7 1972.4 2081.3 C
+1972.5 2074.9 1971.9 2067.9 1972.8 2062.1 C
+[0 1 1 0.36] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1940.2 2071.7 m
+1941.3 2072 1943.1 2072.3 1944 2071.5 C
+1943.6 2069.9 1945.2 2069.1 1946 2068.8 C
+1950 2071.1 1948.7 2065.9 1951.7 2066.2 C
+1953.5 2063.9 1956.9 2069.4 1955.6 2063.8 C
+1955.5 2064.2 1955.7 2064.8 1955.3 2065 C
+1954.3 2063.7 1956.2 2063.6 1955.6 2062.1 C
+1954.5 2060 1958.3 2050.3 1952.2 2055.6 C
+1949.1 2053.8 1946 2051 1943.8 2048 C
+1940.3 2048 1937.5 2051.3 1934.2 2052.5 C
+1933.1 2054.6 1934.4 2057.3 1934 2060 C
+1934 2065.1 1934 2069.7 1934 2074.6 C
+1934.4 2069 1934.1 2061.5 1934.2 2054.9 C
+1934.6 2054.5 1935.3 2054.7 1935.9 2054.7 C
+1937 2055.3 1935.9 2056.1 1935.9 2056.8 C
+1936.5 2063 1935.6 2070.5 1935.9 2074.6 C
+1936.7 2074.4 1937.3 2075.2 1938 2074.6 C
+1937.9 2073.6 1939.1 2072.1 1940.2 2071.7 C
+[0 0.2 1 0] vc
+f
+S
+n
+1933.2 2074.1 m
+1933.2 2071.5 1933.2 2069 1933.2 2066.4 C
+1933.2 2069 1933.2 2071.5 1933.2 2074.1 C
+[0 1 1 0.36] vc
+f
+S
+n
+2007.4 2048.9 m
+2006.5 2047.8 2007.1 2046.2 2006.9 2044.8 C
+2006.7 2050.6 2006.5 2060.5 2006.9 2068.1 C
+2007.1 2061.7 2006.5 2054.7 2007.4 2048.9 C
+f
+S
+n
+1927.2 2062.4 m
+1925.8 2060.1 1928.1 2058.2 1927 2056.4 C
+1927.3 2055.5 1926.5 2053.5 1926.8 2051.8 C
+1926.8 2052.8 1926 2052.5 1925.3 2052.5 C
+1924.1 2052.8 1925 2050.5 1924.4 2050.1 C
+1925.3 2050.2 1925.4 2048.8 1926.3 2049.4 C
+1926.5 2052.3 1928.4 2047.2 1928.4 2051.1 C
+1928.9 2050.5 1929 2051.4 1928.9 2051.8 C
+1928.9 2052 1928.9 2052.3 1928.9 2052.5 C
+1929.4 2051.4 1928.9 2049 1930.1 2048.2 C
+1928.9 2047.1 1930.5 2047.1 1930.4 2046.5 C
+1931.9 2046.2 1933.1 2046.1 1934.7 2046.5 C
+1934.6 2046.9 1935.2 2047.9 1934.4 2048.4 C
+1936.9 2048.1 1933.6 2043.8 1935.9 2043.9 C
+1935.7 2043.9 1934.8 2041.3 1933.2 2041.7 C
+1932.5 2041.6 1932.4 2039.6 1932.3 2041 C
+1930.8 2042.6 1929 2040.6 1927.7 2042 C
+1927.5 2041.4 1927.1 2040.9 1927.2 2040.3 C
+1927.8 2040.6 1927.4 2039.1 1928.2 2038.6 C
+1929.4 2038 1930.5 2038.8 1931.3 2037.9 C
+1931.7 2039 1932.5 2038.6 1931.8 2037.6 C
+1930.9 2037 1928.7 2037.8 1928.2 2037.9 C
+1926.7 2037.8 1928 2039 1927 2038.8 C
+1927.4 2040.4 1925.6 2040.8 1925.1 2041 C
+1924.3 2040.4 1923.2 2040.5 1922.2 2040.5 C
+1921.4 2041.7 1921 2043.9 1919.3 2043.9 C
+1918.8 2043.4 1917.2 2043.3 1916.4 2043.4 C
+1915.9 2044.4 1915.7 2046 1914.3 2046.5 C
+1913.1 2046.6 1912 2044.5 1911.4 2046.3 C
+1912.8 2046.5 1913.8 2047.4 1915.7 2047 C
+1916.9 2047.7 1915.6 2048.8 1916 2049.4 C
+1915.4 2049.3 1913.9 2050.3 1913.3 2051.1 C
+1913.9 2054.1 1916 2050.2 1916.7 2053 C
+1916.9 2053.8 1915.5 2054.1 1916.7 2054.4 C
+1917 2054.7 1920.2 2054.3 1919.3 2056.6 C
+1918.8 2056.1 1920.2 2058.6 1920.3 2057.6 C
+1921.2 2057.9 1922.1 2057.5 1922.4 2059 C
+1922.3 2059.1 1922.2 2059.3 1922 2059.2 C
+1922.1 2059.7 1922.4 2060.3 1922.9 2060.7 C
+1923.2 2060.1 1923.8 2060.4 1924.6 2060.7 C
+1925.9 2062.6 1923.2 2062 1925.6 2063.6 C
+1926.1 2063.1 1927.3 2062.5 1927.2 2062.4 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1933.2 2063.3 m
+1933.2 2060.7 1933.2 2058.2 1933.2 2055.6 C
+1933.2 2058.2 1933.2 2060.7 1933.2 2063.3 C
+[0 1 1 0.36] vc
+f
+S
+n
+1965.2 2049.2 m
+1967.1 2050.1 1969.9 2053.7 1972.1 2056.4 C
+1970.5 2054 1967.6 2051.3 1965.2 2049.2 C
+f
+S
+n
+1991.8 2034.8 m
+1991.7 2041.5 1992 2048.5 1991.6 2055.2 C
+1990.5 2056.4 1991.9 2054.9 1991.8 2054.4 C
+1991.8 2047.9 1991.8 2041.3 1991.8 2034.8 C
+f
+S
+n
+1988.9 2053.2 m
+1988.9 2044.3 1988.9 2036.6 1988.9 2028.3 C
+1985.7 2028.2 1987.2 2023.5 1983.9 2024.2 C
+1983.9 2022.4 1982 2021.6 1981 2021.3 C
+1980.6 2021.1 1980.6 2021.7 1980.3 2021.6 C
+1980.3 2027 1980.3 2034.8 1980.3 2041.5 C
+1979.3 2043.2 1977.6 2043 1976.2 2043.6 C
+1977.1 2043.8 1978.5 2043.2 1978.8 2044.1 C
+1978.5 2045.3 1979.9 2045.3 1980.3 2045.8 C
+1980.5 2046.8 1980.7 2046.2 1981.5 2046.5 C
+1982.4 2047.1 1982 2048.6 1982.7 2049.4 C
+1984.2 2049.6 1984.6 2052.2 1986.8 2051.6 C
+1987.1 2048.6 1985.1 2042.7 1986.5 2040.5 C
+1986.3 2036.7 1986.9 2031.7 1986 2029.2 C
+1986.3 2027.1 1986.9 2028.6 1987.7 2027.6 C
+1987.7 2028.3 1988.7 2028 1988.7 2028.8 C
+1988.1 2033 1988.7 2037.5 1988.2 2041.7 C
+1987.8 2041.4 1988 2040.8 1988 2040.3 C
+1988 2041 1988 2042.4 1988 2042.4 C
+1988 2042.4 1988.1 2042.3 1988.2 2042.2 C
+1989.3 2046 1988 2050.2 1988.4 2054 C
+1987.8 2054.4 1987.1 2054.7 1986.5 2055.4 C
+1987.4 2054.4 1988.4 2054.6 1988.9 2053.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1950.8 2054.4 m
+1949.7 2053.4 1948.7 2052.3 1947.6 2051.3 C
+1948.7 2052.3 1949.7 2053.4 1950.8 2054.4 C
+[0 1 1 0.36] vc
+f
+S
+n
+vmrs
+2006.7 2043.2 m
+2004.5 2040.8 2002.4 2038.4 2000.2 2036 C
+2002.4 2038.4 2004.5 2040.8 2006.7 2043.2 C
+[0 1 1 0.36] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1976.7 2019.6 m
+1975.8 2018.6 1976.4 2016.9 1976.2 2015.6 C
+1976 2021.3 1975.8 2031.2 1976.2 2038.8 C
+1976.4 2032.4 1975.8 2025.5 1976.7 2019.6 C
+f
+S
+n
+1988.4 2053.5 m
+1988.6 2049.2 1988.1 2042.8 1988 2040 C
+1988.4 2040.4 1988.1 2041 1988.2 2041.5 C
+1988.3 2037.2 1988 2032.7 1988.4 2028.5 C
+1987.6 2027.1 1987.2 2028.6 1986.8 2028 C
+1985.9 2028.5 1986.5 2029.7 1986.3 2030.4 C
+1986.9 2029.8 1986.6 2031 1987 2031.2 C
+1987.4 2039.6 1985 2043 1987.2 2050.4 C
+1987.2 2051.6 1985.9 2052.3 1984.6 2051.3 C
+1981.9 2049.7 1982.9 2047 1980.3 2046.5 C
+1980.3 2045.2 1978.1 2046.2 1978.6 2043.9 C
+1975.6 2043.3 1979.3 2045.6 1979.6 2046.5 C
+1980.8 2046.6 1981.5 2048.5 1982.2 2049.9 C
+1983.7 2050.8 1984.8 2052.8 1986.5 2053 C
+1986.7 2053.5 1987.5 2054.1 1987 2054.7 C
+1987.4 2053.9 1988.3 2054.3 1988.4 2053.5 C
+[0 1 1 0.23] vc
+f
+S
+n
+1988 2038.1 m
+1988 2036.7 1988 2035.4 1988 2034 C
+1988 2035.4 1988 2036.7 1988 2038.1 C
+[0 1 1 0.36] vc
+f
+S
+n
+1999.7 2035.7 m
+1997.6 2033.5 1995.4 2031.2 1993.2 2029 C
+1995.4 2031.2 1997.6 2033.5 1999.7 2035.7 C
+f
+S
+n
+1944 2029.2 m
+1945.2 2029.5 1946.9 2029.8 1947.9 2029 C
+1947.4 2027.4 1949 2026.7 1949.8 2026.4 C
+1953.9 2028.6 1952.6 2023.4 1955.6 2023.7 C
+1957.4 2021.4 1960.7 2027 1959.4 2021.3 C
+1959.3 2021.7 1959.6 2022.3 1959.2 2022.5 C
+1958.1 2021.2 1960.1 2021.1 1959.4 2019.6 C
+1959.1 2012.7 1959.9 2005.1 1959.6 1999.2 C
+1955.3 2000.1 1951.3 2003.1 1947.2 2005 C
+1943.9 2006 1941.2 2008.7 1938 2010 C
+1936.9 2012.1 1938.2 2014.8 1937.8 2017.5 C
+1937.8 2022.6 1937.8 2027.3 1937.8 2032.1 C
+1938.2 2026.5 1938 2019 1938 2012.4 C
+1938.5 2012 1939.2 2012.3 1939.7 2012.2 C
+1940.8 2012.8 1939.7 2013.6 1939.7 2014.4 C
+1940.4 2020.5 1939.4 2028 1939.7 2032.1 C
+1940.6 2031.9 1941.2 2032.7 1941.9 2032.1 C
+1941.7 2031.2 1943 2029.7 1944 2029.2 C
+[0 0.2 1 0] vc
+f
+S
+n
+1937.1 2031.6 m
+1937.1 2029.1 1937.1 2026.5 1937.1 2024 C
+1937.1 2026.5 1937.1 2029.1 1937.1 2031.6 C
+[0 1 1 0.36] vc
+f
+S
+n
+1991.8 2028 m
+1992.5 2027.8 1993.2 2029.9 1994 2030.2 C
+1992.9 2029.6 1993.1 2028.1 1991.8 2028 C
+[0 1 1 0.23] vc
+f
+S
+n
+1991.8 2027.8 m
+1992.4 2027.6 1992.6 2028.3 1993 2028.5 C
+1992.6 2028.2 1992.2 2027.6 1991.6 2027.8 C
+1991.6 2028.5 1991.6 2029.1 1991.6 2029.7 C
+1991.6 2029.1 1991.4 2028.3 1991.8 2027.8 C
+[0 1 1 0.36] vc
+f
+S
+n
+1985.8 2025.4 m
+1985.3 2025.2 1984.8 2024.7 1984.1 2024.9 C
+1983.3 2025.3 1983.6 2027.3 1983.9 2027.6 C
+1985 2028 1986.9 2026.9 1985.8 2025.4 C
+[0 1 1 0.23] vc
+f
+S
+n
+vmrs
+1993.5 2024.4 m
+1992.4 2023.7 1991.3 2022.9 1990.1 2023.2 C
+1990.7 2023.7 1989.8 2023.8 1989.4 2023.7 C
+1989.1 2023.7 1988.6 2023.9 1988.4 2023.5 C
+1988.5 2023.2 1988.3 2022.7 1988.7 2022.5 C
+1989 2022.6 1988.9 2023 1988.9 2023.2 C
+1989.1 2022.8 1990.4 2022.3 1990.6 2021.3 C
+1990.4 2021.8 1990 2021.3 1990.1 2021.1 C
+1990.1 2020.9 1990.1 2020.1 1990.1 2020.6 C
+1989.9 2021.1 1989.5 2020.6 1989.6 2020.4 C
+1989.6 2019.8 1988.7 2019.6 1988.2 2019.2 C
+1987.5 2018.7 1987.7 2020.2 1987 2019.4 C
+1987.5 2020.4 1986 2021.1 1987.5 2021.8 C
+1986.8 2023.1 1986.6 2021.1 1986 2021.1 C
+1986.1 2020.1 1985.9 2019 1986.3 2018.2 C
+1986.7 2018.4 1986.5 2019 1986.5 2019.4 C
+1986.5 2018.7 1986.4 2017.8 1987.2 2017.7 C
+1986.5 2017.2 1985.5 2019.3 1985.3 2020.4 C
+1986.2 2022 1987.3 2023.5 1989.2 2024.2 C
+1990.8 2024.3 1991.6 2022.9 1993.2 2024.4 C
+1993.8 2025.4 1995 2026.6 1995.9 2027.1 C
+1995 2026.5 1994.1 2025.5 1993.5 2024.4 C
+[0 1 1 0.36] vc
+f
+0.4 w
+2 J
+2 M
+[0 0.5 0.5 0.2] vc
+S
+n
+2023 2040.3 m
+2023.2 2036 2022.7 2029.6 2022.5 2026.8 C
+2022.9 2027.2 2022.7 2027.8 2022.8 2028.3 C
+2022.8 2024 2022.6 2019.5 2023 2015.3 C
+2022.2 2013.9 2021.7 2015.4 2021.3 2014.8 C
+2020.4 2015.3 2021 2016.5 2020.8 2017.2 C
+2021.4 2016.6 2021.1 2017.8 2021.6 2018 C
+2022 2026.4 2019.6 2029.8 2021.8 2037.2 C
+2021.7 2038.4 2020.5 2039.1 2019.2 2038.1 C
+2016.5 2036.5 2017.5 2033.8 2014.8 2033.3 C
+2014.9 2032 2012.6 2033 2013.2 2030.7 C
+2011.9 2030.8 2011.2 2030.1 2010.8 2029.2 C
+2010.8 2029.1 2010.8 2028.2 2010.8 2028.8 C
+2010 2028.8 2010.4 2026.5 2008.6 2027.3 C
+2007.9 2026.6 2007.3 2025.9 2007.9 2027.1 C
+2009.7 2028 2010 2030.1 2012.2 2030.9 C
+2012.9 2032.1 2013.7 2033.6 2015.1 2033.6 C
+2015.7 2035.1 2016.9 2036.7 2018.4 2038.4 C
+2019.8 2039.3 2022 2039.4 2021.6 2041.5 C
+2021.9 2040.7 2022.9 2041.1 2023 2040.3 C
+[0 1 1 0.23] vc
+f
+S
+n
+2022.5 2024.9 m
+2022.5 2023.5 2022.5 2022.2 2022.5 2020.8 C
+2022.5 2022.2 2022.5 2023.5 2022.5 2024.9 C
+[0 1 1 0.36] vc
+f
+S
+n
+1983.2 2022.8 m
+1982.4 2022.5 1982.1 2021.6 1981.2 2022.3 C
+1981.1 2022.9 1980.5 2024 1981 2024.2 C
+1981.8 2024.6 1982.9 2024.4 1983.2 2022.8 C
+[0 1 1 0.23] vc
+f
+S
+n
+1931.1 2019.9 m
+1929.6 2017.7 1932 2015.7 1930.8 2013.9 C
+1931.1 2013 1930.3 2011 1930.6 2009.3 C
+1930.6 2010.3 1929.8 2010 1929.2 2010 C
+1928 2010.3 1928.8 2008.1 1928.2 2007.6 C
+1929.1 2007.8 1929.3 2006.3 1930.1 2006.9 C
+1930.3 2009.8 1932.2 2004.8 1932.3 2008.6 C
+1932.7 2008 1932.8 2009 1932.8 2009.3 C
+1932.8 2009.6 1932.8 2009.8 1932.8 2010 C
+1933.2 2009 1932.7 2006.6 1934 2005.7 C
+1932.7 2004.6 1934.3 2004.6 1934.2 2004 C
+1935.8 2003.7 1937 2003.6 1938.5 2004 C
+1938.5 2004.5 1939.1 2005.4 1938.3 2006 C
+1940.7 2005.7 1937.4 2001.3 1939.7 2001.4 C
+1939.5 2001.4 1938.6 1998.8 1937.1 1999.2 C
+1936.3 1999.1 1936.2 1997.1 1936.1 1998.5 C
+1934.7 2000.1 1932.9 1998.2 1931.6 1999.5 C
+1931.3 1998.9 1930.9 1998.5 1931.1 1997.8 C
+1931.6 1998.2 1931.3 1996.6 1932 1996.1 C
+1933.2 1995.5 1934.3 1996.4 1935.2 1995.4 C
+1935.5 1996.5 1936.3 1996.1 1935.6 1995.2 C
+1934.7 1994.5 1932.5 1995.3 1932 1995.4 C
+1930.5 1995.3 1931.9 1996.5 1930.8 1996.4 C
+1931.2 1997.9 1929.5 1998.3 1928.9 1998.5 C
+1928.1 1997.9 1927.1 1998 1926 1998 C
+1925.3 1999.2 1924.8 2001.4 1923.2 2001.4 C
+1922.6 2000.9 1921 2000.9 1920.3 2000.9 C
+1919.7 2001.9 1919.6 2003.5 1918.1 2004 C
+1916.9 2004.1 1915.8 2002 1915.2 2003.8 C
+1916.7 2004 1917.6 2004.9 1919.6 2004.5 C
+1920.7 2005.2 1919.4 2006.3 1919.8 2006.9 C
+1919.2 2006.9 1917.7 2007.8 1917.2 2008.6 C
+1917.8 2011.6 1919.8 2007.8 1920.5 2010.5 C
+1920.8 2011.3 1919.3 2011.6 1920.5 2012 C
+1920.8 2012.3 1924 2011.8 1923.2 2014.1 C
+1922.6 2013.6 1924.1 2016.1 1924.1 2015.1 C
+1925.1 2015.4 1925.9 2015 1926.3 2016.5 C
+1926.2 2016.6 1926 2016.8 1925.8 2016.8 C
+1925.9 2017.2 1926.2 2017.8 1926.8 2018.2 C
+1927.1 2017.6 1927.7 2018 1928.4 2018.2 C
+1929.7 2020.1 1927.1 2019.5 1929.4 2021.1 C
+1929.9 2020.7 1931.1 2020 1931.1 2019.9 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1937.1 2020.8 m
+1937.1 2018.3 1937.1 2015.7 1937.1 2013.2 C
+1937.1 2015.7 1937.1 2018.3 1937.1 2020.8 C
+[0 1 1 0.36] vc
+f
+S
+n
+2020.4 2012.2 m
+2019.8 2012 2019.3 2011.5 2018.7 2011.7 C
+2017.9 2012.1 2018.1 2014.1 2018.4 2014.4 C
+2019.6 2014.8 2021.4 2013.7 2020.4 2012.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1976 2013.9 m
+1973.8 2011.5 1971.6 2009.1 1969.5 2006.7 C
+1971.6 2009.1 1973.8 2011.5 1976 2013.9 C
+[0 1 1 0.36] vc
+f
+S
+n
+1995.4 2012.7 m
+1996.1 2010.3 1993.8 2006.2 1997.3 2005.7 C
+1998.9 2005.4 2000 2003.7 2001.4 2003.1 C
+2003.9 2003.1 2005.3 2001.3 2006.9 1999.7 C
+2004.5 2003.5 2000 2002.2 1997.6 2005.7 C
+1996.5 2005.9 1994.8 2006.1 1995.2 2007.6 C
+1995.7 2009.4 1995.2 2011.6 1994.7 2012.9 C
+1992 2015.8 1987.8 2015.7 1985.3 2018.7 C
+1988.3 2016.3 1992.3 2015.3 1995.4 2012.7 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1995.6 2012.4 m
+1995.6 2011.2 1995.6 2010 1995.6 2008.8 C
+1995.6 2010 1995.6 2011.2 1995.6 2012.4 C
+[0 1 1 0.36] vc
+f
+S
+n
+vmrs
+2017.7 2009.6 m
+2016.9 2009.3 2016.7 2008.4 2015.8 2009.1 C
+2014.2 2010.6 2016 2010.6 2016.5 2011.5 C
+2017.2 2010.9 2018.1 2010.8 2017.7 2009.6 C
+[0 1 1 0.23] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+2014.4 2006.4 m
+2013.5 2006.8 2012.1 2005.6 2012 2006.7 C
+2013 2007.3 2011.9 2009.2 2012.9 2008.4 C
+2014.2 2008.3 2014.6 2007.8 2014.4 2006.4 C
+f
+S
+n
+1969 2006.4 m
+1966.5 2003.8 1964 2001.2 1961.6 1998.5 C
+1964 2001.2 1966.5 2003.8 1969 2006.4 C
+[0 1 1 0.36] vc
+f
+S
+n
+2012 2005.2 m
+2012.2 2004.2 2011.4 2003.3 2010.3 2003.3 C
+2009 2003.6 2010 2004.7 2009.6 2004.8 C
+2009.3 2005.7 2011.4 2006.7 2012 2005.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1962.8 1995.2 m
+1961.7 1994.4 1960.6 1993.7 1959.4 1994 C
+1959.5 1994.9 1957.5 1994.1 1956.8 1994.7 C
+1955.9 1995.5 1956.7 1997 1955.1 1997.3 C
+1956.9 1996.7 1956.8 1994 1959.2 1994.7 C
+1961.1 1991 1968.9 2003.2 1962.8 1995.2 C
+[0 1 1 0.36] vc
+f
+S
+n
+1954.6 1995.6 m
+1955.9 1994.7 1955.1 1989.8 1955.3 1988 C
+1954.5 1988.3 1954.9 1986.6 1954.4 1986 C
+1955.7 1989.2 1953.9 1991.1 1954.8 1994.2 C
+1954.5 1995.9 1953.5 1995.3 1953.9 1997.3 C
+1955.3 1998.3 1953.2 1995.5 1954.6 1995.6 C
+f
+S
+n
+1992.3 2011 m
+1992.5 2006.7 1992 2000.3 1991.8 1997.6 C
+1992.2 1997.9 1992 1998.5 1992 1999 C
+1992.1 1994.7 1991.9 1990.2 1992.3 1986 C
+1991.4 1984.6 1991 1986.1 1990.6 1985.6 C
+1989.7 1986 1990.3 1987.2 1990.1 1988 C
+1990.7 1987.4 1990.4 1988.5 1990.8 1988.7 C
+1991.3 1997.1 1988.9 2000.6 1991.1 2007.9 C
+1991 2009.1 1989.8 2009.9 1988.4 2008.8 C
+1985.7 2007.2 1986.8 2004.5 1984.1 2004 C
+1984.2 2002.7 1981.9 2003.7 1982.4 2001.4 C
+1981.2 2001.5 1980.5 2000.8 1980 2000 C
+1980 1999.8 1980 1998.9 1980 1999.5 C
+1979.3 1999.5 1979.7 1997.2 1977.9 1998 C
+1977.2 1997.3 1976.6 1996.7 1977.2 1997.8 C
+1979 1998.7 1979.3 2000.8 1981.5 2001.6 C
+1982.2 2002.8 1983 2004.3 1984.4 2004.3 C
+1985 2005.8 1986.2 2007.5 1987.7 2009.1 C
+1989 2010 1991.3 2010.2 1990.8 2012.2 C
+1991.2 2011.4 1992.2 2011.8 1992.3 2011 C
+[0 1 1 0.23] vc
+f
+S
+n
+1991.8 1995.6 m
+1991.8 1994.3 1991.8 1992.9 1991.8 1991.6 C
+1991.8 1992.9 1991.8 1994.3 1991.8 1995.6 C
+[0 1 1 0.36] vc
+f
+S
+n
+1959.2 1994.2 m
+1958.8 1993.3 1960.7 1993.9 1961.1 1993.7 C
+1961.5 1993.9 1961.2 1994.4 1961.8 1994.2 C
+1960.9 1994 1960.8 1992.9 1959.9 1992.5 C
+1959.6 1993.5 1958.3 1993.5 1958.2 1994.2 C
+1958.1 1994.1 1958 1994 1958 1994 C
+1957.2 1994.9 1958 1993.4 1956.8 1993 C
+1955.6 1992.5 1956 1991 1956.3 1989.9 C
+1956.5 1989.8 1956.6 1990 1956.8 1990.1 C
+1957.1 1989 1956 1989.1 1955.8 1988.2 C
+1955.1 1990.4 1956.2 1995 1954.8 1995.9 C
+1954.1 1995.5 1954.5 1996.5 1954.4 1997.1 C
+1955 1996.8 1954.8 1997.4 1955.6 1996.8 C
+1956 1996 1956.3 1993.2 1958.7 1994.2 C
+1958.9 1994.2 1959.7 1994.2 1959.2 1994.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1958.2 1994 m
+1958.4 1993.5 1959.7 1993.1 1959.9 1992 C
+1959.7 1992.5 1959.3 1992 1959.4 1991.8 C
+1959.4 1991.6 1959.4 1990.8 1959.4 1991.3 C
+1959.2 1991.8 1958.8 1991.3 1958.9 1991.1 C
+1958.9 1990.5 1958 1990.3 1957.5 1989.9 C
+1956.8 1989.5 1956.9 1991 1956.3 1990.1 C
+1956.7 1991 1955.4 1992.1 1956.5 1992.3 C
+1956.8 1993.5 1958.3 1992.9 1957.2 1994 C
+1957.8 1994.3 1958.1 1992.4 1958.2 1994 C
+[0 0.5 0.5 0.2] vc
+f
+S
+n
+vmrs
+1954.4 1982.7 m
+1956.1 1982.7 1954.1 1982.5 1953.9 1982.9 C
+1953.9 1983.7 1953.7 1984.7 1954.1 1985.3 C
+1954.4 1984.2 1953.6 1983.6 1954.4 1982.7 C
+[0 1 1 0.36] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1989.6 1982.9 m
+1989.1 1982.7 1988.6 1982.3 1988 1982.4 C
+1987.2 1982.8 1987.4 1984.8 1987.7 1985.1 C
+1988.9 1985.6 1990.7 1984.4 1989.6 1982.9 C
+[0 1 1 0.23] vc
+f
+S
+n
+1987 1980.3 m
+1986.2 1980 1986 1979.1 1985.1 1979.8 C
+1983.5 1981.4 1985.3 1981.4 1985.8 1982.2 C
+1986.5 1981.7 1987.4 1981.5 1987 1980.3 C
+f
+S
+n
+1983.6 1977.2 m
+1982.7 1977.5 1981.4 1976.3 1981.2 1977.4 C
+1982.3 1978 1981.2 1979.9 1982.2 1979.1 C
+1983.5 1979 1983.9 1978.5 1983.6 1977.2 C
+f
+S
+n
+1981.2 1976 m
+1981.5 1974.9 1980.6 1974 1979.6 1974 C
+1978.3 1974.3 1979.3 1975.4 1978.8 1975.5 C
+1978.6 1976.4 1980.7 1977.4 1981.2 1976 C
+f
+S
+n
+1972.1 2082.3 m
+1971.8 2081.8 1971.3 2080.9 1971.2 2080.1 C
+1971.1 2072.9 1971.3 2064.6 1970.9 2058.3 C
+1970.3 2058.5 1970.1 2057.7 1969.7 2058.5 C
+1970.6 2058.5 1969.7 2059 1970.2 2059.2 C
+1970.2 2065.4 1970.2 2072.4 1970.2 2077.7 C
+1971.1 2078.9 1970.6 2078.9 1970.4 2079.9 C
+1969.2 2080.2 1968.2 2080.4 1967.3 2079.6 C
+1966.8 2077.8 1963.4 2076.3 1963.5 2075.1 C
+1961.5 2075.5 1962 2071.5 1959.6 2072 C
+1959.2 2070 1956.5 2069.3 1955.8 2067.6 C
+1956 2068.4 1955.3 2069.7 1956.5 2069.8 C
+1958.6 2068.9 1958.1 2073.5 1960.1 2072.4 C
+1960.7 2075.9 1964.7 2074.6 1964.2 2078 C
+1967.2 2078.6 1967.9 2081.6 1970.7 2080.6 C
+1970.3 2081.1 1971.5 2081.2 1971.9 2082.3 C
+1967.2 2084.3 1962.9 2087.1 1958.2 2089 C
+1962.9 2087 1967.4 2084.4 1972.1 2082.3 C
+[0 0.2 1 0] vc
+f
+S
+n
+1971.9 2080.1 m
+1971.9 2075.1 1971.9 2070 1971.9 2065 C
+1971.9 2070 1971.9 2075.1 1971.9 2080.1 C
+[0 1 1 0.23] vc
+f
+S
+n
+2010.8 2050.6 m
+2013.2 2049 2010.5 2050.1 2010.5 2051.3 C
+2010.5 2057.7 2010.5 2064.1 2010.5 2070.5 C
+2008.7 2072.4 2006 2073.3 2003.6 2074.4 C
+2016.4 2073.7 2008 2058.4 2010.8 2050.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+2006.4 2066.9 m
+2006.4 2061.9 2006.4 2056.8 2006.4 2051.8 C
+2006.4 2056.8 2006.4 2061.9 2006.4 2066.9 C
+[0 1 1 0.23] vc
+f
+S
+n
+1971.9 2060.7 m
+1972.2 2060.3 1971.4 2068.2 1972.4 2061.9 C
+1971.8 2061.6 1972.4 2060.9 1971.9 2060.7 C
+f
+S
+n
+vmrs
+1986.5 2055.2 m
+1987.5 2054.3 1986.3 2053.4 1986 2052.8 C
+1983.8 2052.7 1983.6 2050.1 1981.7 2049.6 C
+1981.2 2048.7 1980.8 2047 1980.3 2046.8 C
+1978.5 2047 1978 2044.6 1976.7 2043.9 C
+1974 2044.4 1972 2046.6 1969.2 2047 C
+1969 2047.2 1968.8 2047.5 1968.5 2047.7 C
+1970.6 2049.6 1973.1 2051.3 1974.3 2054.2 C
+1975.7 2054.5 1977 2055.2 1976.4 2057.1 C
+1976.7 2058 1975.5 2058.5 1976 2059.5 C
+1979.2 2058 1983 2056.6 1986.5 2055.2 C
+[0 0.5 0.5 0.2] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1970.2 2054.2 m
+1971.5 2055.3 1972.5 2056.8 1972.1 2058.3 C
+1972.8 2056.5 1971.6 2055.6 1970.2 2054.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1992 2052.5 m
+1992 2053.4 1992.2 2054.4 1991.8 2055.2 C
+1992.2 2054.4 1992 2053.4 1992 2052.5 C
+f
+S
+n
+1957.2 2053 m
+1958.1 2052.6 1959 2052.2 1959.9 2051.8 C
+1959 2052.2 1958.1 2052.6 1957.2 2053 C
+f
+S
+n
+2006.4 2047.5 m
+2006.8 2047.1 2006 2055 2006.9 2048.7 C
+2006.4 2048.4 2007 2047.7 2006.4 2047.5 C
+f
+S
+n
+2004.8 2041 m
+2006.1 2042.1 2007.1 2043.6 2006.7 2045.1 C
+2007.3 2043.3 2006.2 2042.4 2004.8 2041 C
+f
+S
+n
+1976 2039.8 m
+1975.6 2039.3 1975.2 2038.4 1975 2037.6 C
+1974.9 2030.4 1975.2 2022.1 1974.8 2015.8 C
+1974.2 2016 1974 2015.3 1973.6 2016 C
+1974.4 2016 1973.5 2016.5 1974 2016.8 C
+1974 2022.9 1974 2030 1974 2035.2 C
+1974.9 2036.4 1974.4 2036.4 1974.3 2037.4 C
+1973.1 2037.7 1972 2037.9 1971.2 2037.2 C
+1970.6 2035.3 1967.3 2033.9 1967.3 2032.6 C
+1965.3 2033 1965.9 2029.1 1963.5 2029.5 C
+1963 2027.6 1960.4 2026.8 1959.6 2025.2 C
+1959.8 2025.9 1959.2 2027.2 1960.4 2027.3 C
+1962.5 2026.4 1961.9 2031 1964 2030 C
+1964.6 2033.4 1968.5 2032.1 1968 2035.5 C
+1971 2036.1 1971.8 2039.1 1974.5 2038.1 C
+1974.2 2038.7 1975.3 2038.7 1975.7 2039.8 C
+1971 2041.8 1966.7 2044.6 1962 2046.5 C
+1966.8 2044.5 1971.3 2041.9 1976 2039.8 C
+[0 0.2 1 0] vc
+f
+S
+n
+1975.7 2037.6 m
+1975.7 2032.6 1975.7 2027.6 1975.7 2022.5 C
+1975.7 2027.6 1975.7 2032.6 1975.7 2037.6 C
+[0 1 1 0.23] vc
+f
+S
+n
+1992 2035.5 m
+1992 2034.2 1992 2032.9 1992 2031.6 C
+1992 2032.9 1992 2034.2 1992 2035.5 C
+f
+S
+n
+2015.3 2036 m
+2015.4 2034.1 2013.3 2034 2012.9 2033.3 C
+2011.5 2031 2009.3 2029.4 2007.4 2028 C
+2006.9 2027.1 2006.6 2023.8 2005 2024.9 C
+2004 2024.9 2002.9 2024.9 2001.9 2024.9 C
+2001.4 2026.5 2001 2028.4 2003.8 2028.3 C
+2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C
+2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C
+2013 2035.5 2015.3 2037.4 2015.3 2036 C
+[0 0 0 0] vc
+f
+S
+n
+vmrs
+2009.1 2030.4 m
+2009.1 2029 2007.5 2029.4 2006.9 2028.3 C
+2007.2 2027.1 2006.5 2025.5 2005.7 2024.7 C
+2004.6 2025.1 2003.1 2024.9 2001.9 2024.9 C
+2001.8 2026.2 2000.9 2027 2002.4 2028 C
+2004.5 2027.3 2004.9 2029.4 2006.9 2029 C
+2007 2030.2 2007.6 2030.7 2008.4 2031.4 C
+2008.8 2031.5 2009.1 2031.1 2009.1 2030.4 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+2003.8 2029.5 m
+2003 2029.4 2001.9 2029.1 2002.4 2030.4 C
+2003.1 2031.3 2005.2 2030.3 2003.8 2029.5 C
+[0 1 1 0.23] vc
+f
+S
+n
+1999.2 2025.2 m
+1999.1 2025.6 1998 2025.7 1998.8 2026.6 C
+2000.9 2028.5 1999.5 2023.4 1999.2 2025.2 C
+f
+S
+n
+2007.6 2024.2 m
+2007.6 2022.9 2008.4 2024.2 2007.6 2022.8 C
+2007.6 2017.5 2007.8 2009.1 2007.4 2003.8 C
+2007.9 2003.7 2008.7 2002.8 2009.1 2002.1 C
+2009.6 2000.8 2008.3 2000.8 2007.9 2000.2 C
+2004.9 2000 2008.9 2001.3 2007.2 2002.1 C
+2006.7 2007.7 2007 2015.1 2006.9 2021.1 C
+2006.7 2022.1 2005.4 2022.8 2006.2 2023.5 C
+2006.6 2023.1 2008 2025.9 2007.6 2024.2 C
+f
+S
+n
+1989.9 2023.5 m
+1989.5 2022.6 1991.4 2023.2 1991.8 2023 C
+1992.2 2023.2 1991.9 2023.7 1992.5 2023.5 C
+1991.6 2023.2 1991.6 2022.2 1990.6 2021.8 C
+1990.4 2022.8 1989 2022.8 1988.9 2023.5 C
+1988.5 2023 1988.7 2022.6 1988.7 2023.5 C
+1989.1 2023.5 1990.2 2023.5 1989.9 2023.5 C
+f
+[0 0.5 0.5 0.2] vc
+S
+n
+2003.3 2023.5 m
+2003.1 2023.3 2003.1 2023.2 2003.3 2023 C
+2003.7 2023.1 2003.9 2022.9 2003.8 2022.5 C
+2003.4 2022.2 2001.2 2022.3 2002.4 2023 C
+2002.6 2022.9 2002.7 2023.1 2002.8 2023.2 C
+2000.7 2023.7 2003.9 2023.4 2003.3 2023.5 C
+[0 1 1 0.23] vc
+f
+S
+n
+1986.8 2019.4 m
+1987.8 2019.8 1987.5 2018.6 1987.2 2018 C
+1986.2 2017.8 1987.3 2020.5 1986.3 2019.2 C
+1986.3 2017.7 1986.3 2020.6 1986.3 2021.3 C
+1988.5 2023.1 1985.6 2020.3 1986.8 2019.4 C
+f
+S
+n
+1975.7 2018.2 m
+1976.1 2017.8 1975.2 2025.7 1976.2 2019.4 C
+1975.7 2019.2 1976.3 2018.4 1975.7 2018.2 C
+f
+S
+n
+1974 2011.7 m
+1975.4 2012.8 1976.4 2014.3 1976 2015.8 C
+1976.6 2014 1975.5 2013.1 1974 2011.7 C
+f
+S
+n
+1984.6 2006.7 m
+1984.7 2004.8 1982.6 2004.8 1982.2 2004 C
+1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C
+1976.1 1997.8 1975.8 1994.5 1974.3 1995.6 C
+1973.3 1995.6 1972.2 1995.6 1971.2 1995.6 C
+1970.7 1997.2 1970.3 1999.1 1973.1 1999 C
+1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C
+1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C
+1982.3 2006.2 1984.5 2008.1 1984.6 2006.7 C
+[0 0 0 0] vc
+f
+S
+n
+vmrs
+1978.4 2001.2 m
+1978.4 1999.7 1976.8 2000.1 1976.2 1999 C
+1976.5 1997.8 1975.8 1996.2 1975 1995.4 C
+1973.9 1995.8 1972.4 1995.6 1971.2 1995.6 C
+1971 1997 1970.2 1997.7 1971.6 1998.8 C
+1973.8 1998 1974.2 2000.1 1976.2 1999.7 C
+1976.3 2000.9 1976.9 2001.4 1977.6 2002.1 C
+1978.1 2002.2 1978.4 2001.8 1978.4 2001.2 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1973.1 2000.2 m
+1972.3 2000.1 1971.2 1999.8 1971.6 2001.2 C
+1972.4 2002 1974.5 2001 1973.1 2000.2 C
+[0 1 1 0.23] vc
+f
+S
+n
+1960.8 1998.5 m
+1961.6 1998.2 1962.6 2000.3 1963.2 2000.9 C
+1962.3 2000.1 1962.2 1998.7 1960.8 1998.5 C
+f
+S
+n
+1968.5 1995.9 m
+1968.4 1996.4 1967.3 1996.4 1968 1997.3 C
+1970.1 1999.2 1968.8 1994.1 1968.5 1995.9 C
+f
+S
+n
+1976.9 1994.9 m
+1976.9 1993.7 1977.6 1994.9 1976.9 1993.5 C
+1976.9 1988.2 1977.1 1979.8 1976.7 1974.5 C
+1977.2 1974.5 1978 1973.5 1978.4 1972.8 C
+1978.8 1971.5 1977.6 1971.5 1977.2 1970.9 C
+1974.2 1970.7 1978.2 1972 1976.4 1972.8 C
+1976 1978.4 1976.3 1985.8 1976.2 1991.8 C
+1976 1992.8 1974.6 1993.5 1975.5 1994.2 C
+1975.9 1993.8 1977.3 1996.6 1976.9 1994.9 C
+f
+S
+n
+1972.6 1994.2 m
+1972.4 1994 1972.4 1993.9 1972.6 1993.7 C
+1973 1993.8 1973.1 1993.7 1973.1 1993.2 C
+1972.7 1992.9 1970.5 1993.1 1971.6 1993.7 C
+1971.9 1993.7 1972 1993.8 1972.1 1994 C
+1970 1994.4 1973.1 1994.1 1972.6 1994.2 C
+f
+S
+n
+1948.1 2093.8 m
+1947 2092.7 1945.9 2091.6 1944.8 2090.4 C
+1945.9 2091.6 1947 2092.7 1948.1 2093.8 C
+[0 0.4 1 0] vc
+f
+S
+n
+1953.4 2091.4 m
+1954.8 2090.7 1956.3 2090 1957.7 2089.2 C
+1956.3 2090 1954.8 2090.7 1953.4 2091.4 C
+[0 0.2 1 0] vc
+f
+S
+n
+1954.1 2091.4 m
+1956.6 2089.6 1957.2 2089.6 1954.1 2091.4 C
+[0 0.4 1 0] vc
+f
+S
+n
+1962.3 2087.3 m
+1963.7 2086.6 1965.2 2085.9 1966.6 2085.2 C
+1965.2 2085.9 1963.7 2086.6 1962.3 2087.3 C
+f
+S
+n
+vmrs
+1967.1 2084.9 m
+1968.3 2084.4 1969.7 2083.8 1970.9 2083.2 C
+1969.7 2083.8 1968.3 2084.4 1967.1 2084.9 C
+[0 0.4 1 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1982.7 2080.6 m
+1981.5 2079.5 1980.5 2078.4 1979.3 2077.2 C
+1980.5 2078.4 1981.5 2079.5 1982.7 2080.6 C
+f
+S
+n
+1988 2078.2 m
+1989.4 2077.5 1990.8 2076.8 1992.3 2076 C
+1990.8 2076.8 1989.4 2077.5 1988 2078.2 C
+[0 0.2 1 0] vc
+f
+S
+n
+1988.7 2078.2 m
+1991.1 2076.4 1991.8 2076.4 1988.7 2078.2 C
+[0 0.4 1 0] vc
+f
+S
+n
+1976.2 2063.8 m
+1978.6 2062.2 1976 2063.3 1976 2064.5 C
+1976.1 2067.8 1975.5 2071.4 1976.4 2074.4 C
+1975.7 2071.1 1975.9 2067.2 1976.2 2063.8 C
+f
+S
+n
+1996.8 2074.1 m
+1998.3 2073.4 1999.7 2072.7 2001.2 2072 C
+1999.7 2072.7 1998.3 2073.4 1996.8 2074.1 C
+f
+S
+n
+2001.6 2071.7 m
+2002.9 2071.2 2004.2 2070.6 2005.5 2070 C
+2004.2 2070.6 2002.9 2071.2 2001.6 2071.7 C
+f
+S
+n
+1981.5 2060.7 m
+1980.2 2061.2 1978.9 2061.5 1977.9 2062.6 C
+1978.9 2061.5 1980.2 2061.2 1981.5 2060.7 C
+f
+S
+n
+1982 2060.4 m
+1982.7 2060.1 1983.6 2059.8 1984.4 2059.5 C
+1983.6 2059.8 1982.7 2060.1 1982 2060.4 C
+f
+S
+n
+1952 2051.3 m
+1950.8 2050.2 1949.7 2049.1 1948.6 2048 C
+1949.7 2049.1 1950.8 2050.2 1952 2051.3 C
+f
+S
+n
+vmrs
+1977.4 2047.7 m
+1975.8 2047.8 1974.8 2046.1 1974.5 2045.3 C
+1974.9 2044.4 1976 2044.5 1976.7 2044.8 C
+1977.9 2045 1977 2048.4 1979.3 2047.5 C
+1979.9 2047.5 1980.8 2048.6 1979.8 2049.2 C
+1978.2 2050.4 1980.8 2049.5 1980.3 2049.4 C
+1981.4 2049.8 1980.3 2048.4 1980.3 2048 C
+1979.8 2047.5 1979 2046.6 1978.4 2046.5 C
+1977.3 2045.9 1977.2 2043.3 1975.2 2044.6 C
+1974.7 2045.3 1973.6 2045 1973.3 2045.8 C
+1975 2046.3 1975.8 2049.8 1978.1 2049.4 C
+1978.4 2050.9 1978.7 2048.5 1977.9 2049.2 C
+1977.7 2048.7 1977.2 2047.8 1977.4 2047.7 C
+[0 0.5 0.5 0.2] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1957.2 2048.9 m
+1958.7 2048.2 1960.1 2047.5 1961.6 2046.8 C
+1960.1 2047.5 1958.7 2048.2 1957.2 2048.9 C
+[0 0.2 1 0] vc
+f
+S
+n
+1958 2048.9 m
+1960.4 2047.1 1961.1 2047.1 1958 2048.9 C
+[0 0.4 1 0] vc
+f
+S
+n
+1966.1 2044.8 m
+1967.6 2044.1 1969 2043.4 1970.4 2042.7 C
+1969 2043.4 1967.6 2044.1 1966.1 2044.8 C
+f
+S
+n
+1970.9 2042.4 m
+1972.2 2041.9 1973.5 2041.3 1974.8 2040.8 C
+1973.5 2041.3 1972.2 2041.9 1970.9 2042.4 C
+f
+S
+n
+2012 2034.5 m
+2010.4 2034.6 2009.3 2032.9 2009.1 2032.1 C
+2009.4 2031 2010.3 2031.3 2011.2 2031.6 C
+2012.5 2031.8 2011.6 2035.2 2013.9 2034.3 C
+2014.4 2034.3 2015.4 2035.4 2014.4 2036 C
+2012.7 2037.2 2015.3 2036.3 2014.8 2036.2 C
+2015.9 2036.6 2014.8 2035.2 2014.8 2034.8 C
+2014.4 2034.3 2013.6 2033.4 2012.9 2033.3 C
+2011.5 2031 2009.3 2029.4 2007.4 2028 C
+2007.5 2026.5 2007.3 2027.9 2007.2 2028.3 C
+2007.9 2028.8 2008.7 2029.1 2009.3 2030 C
+2009.6 2030.7 2009 2031.9 2008.4 2031.6 C
+2006.7 2031 2007.7 2028 2005 2028.8 C
+2004.8 2028.6 2004.3 2028.2 2003.8 2028.3 C
+2006.6 2030.4 2008.9 2033.7 2011.2 2036.2 C
+2011.8 2036.4 2012.9 2035.8 2012.9 2036.7 C
+2012.7 2036.1 2011.8 2035 2012 2034.5 C
+[0 0.5 0.5 0.2] vc
+f
+S
+n
+1981.2 2005.2 m
+1979.7 2005.3 1978.6 2003.6 1978.4 2002.8 C
+1978.7 2001.8 1979.6 2002.1 1980.5 2002.4 C
+1981.8 2002.5 1980.9 2005.9 1983.2 2005 C
+1983.7 2005.1 1984.7 2006.1 1983.6 2006.7 C
+1982 2007.9 1984.6 2007 1984.1 2006.9 C
+1985.2 2007.3 1984.1 2006 1984.1 2005.5 C
+1983.6 2005 1982.9 2004.1 1982.2 2004 C
+1980.8 2001.7 1978.6 2000.1 1976.7 1998.8 C
+1976.7 1997.2 1976.6 1998.6 1976.4 1999 C
+1977.2 1999.5 1978 1999.8 1978.6 2000.7 C
+1978.8 2001.5 1978.3 2002.7 1977.6 2002.4 C
+1976 2001.8 1977 1998.7 1974.3 1999.5 C
+1974.1 1999.3 1973.6 1998.9 1973.1 1999 C
+1975.8 2001.2 1978.2 2004.4 1980.5 2006.9 C
+1981.1 2007.1 1982.1 2006.5 1982.2 2007.4 C
+1982 2006.8 1981.1 2005.7 1981.2 2005.2 C
+f
+S
+n
+1966.8 1976.4 m
+1969.4 1973 1974.4 1974.6 1976.2 1970.4 C
+1972.7 1974 1968 1975.1 1964 1977.4 C
+1960.9 1979.9 1957.1 1981.8 1953.9 1982.7 C
+1958.4 1981.1 1962.6 1978.8 1966.8 1976.4 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1948.4 2093.8 m
+1949.8 2093.1 1951.2 2092.5 1952.7 2091.9 C
+1951.2 2092.5 1949.8 2093.1 1948.4 2093.8 C
+[0 0.2 1 0] vc
+f
+S
+n
+1948.1 2093.6 m
+1947.3 2092.8 1946.5 2091.9 1945.7 2091.2 C
+1946.5 2091.9 1947.3 2092.8 1948.1 2093.6 C
+f
+S
+n
+vmrs
+1942.1 2087.8 m
+1943.5 2088.4 1944.3 2089.5 1945.2 2090.7 C
+1944.8 2089.3 1943.3 2088.3 1942.1 2087.8 C
+[0 0.2 1 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1933.5 2078.4 m
+1933.5 2078 1933.2 2079 1933.7 2079.4 C
+1935 2080.4 1936.2 2081.3 1937.1 2082.8 C
+1936.7 2080.7 1933.7 2080.7 1933.5 2078.4 C
+f
+S
+n
+1982.9 2080.6 m
+1984.4 2079.9 1985.8 2079.3 1987.2 2078.7 C
+1985.8 2079.3 1984.4 2079.9 1982.9 2080.6 C
+f
+S
+n
+1982.7 2080.4 m
+1981.9 2079.6 1981.1 2078.7 1980.3 2078 C
+1981.1 2078.7 1981.9 2079.6 1982.7 2080.4 C
+f
+S
+n
+1977.4 2075.1 m
+1977.9 2075.3 1979.1 2076.4 1979.8 2077.5 C
+1979 2076.8 1978.7 2075.1 1977.4 2075.1 C
+f
+S
+n
+1952.2 2051.3 m
+1953.6 2050.7 1955.1 2050.1 1956.5 2049.4 C
+1955.1 2050.1 1953.6 2050.7 1952.2 2051.3 C
+f
+S
+n
+1952 2051.1 m
+1951.2 2050.3 1950.3 2049.5 1949.6 2048.7 C
+1950.3 2049.5 1951.2 2050.3 1952 2051.1 C
+f
+S
+n
+1946 2045.3 m
+1947.3 2045.9 1948.1 2047 1949.1 2048.2 C
+1948.6 2046.8 1947.1 2045.8 1946 2045.3 C
+f
+S
+n
+1937.3 2036 m
+1937.4 2035.5 1937 2036.5 1937.6 2036.9 C
+1938.8 2037.9 1940.1 2038.8 1940.9 2040.3 C
+1940.6 2038.2 1937.6 2038.2 1937.3 2036 C
+f
+S
+n
+1935.2 2073.2 m
+1936.4 2069.9 1935.8 2061.8 1935.6 2056.4 C
+1935.8 2055.9 1936.3 2055.7 1936.1 2055.2 C
+1935.7 2054.7 1935 2055 1934.4 2054.9 C
+1934.4 2061.5 1934.4 2068.7 1934.4 2074.6 C
+1935.7 2075.1 1936 2073.7 1935.2 2073.2 C
+[0 0.01 1 0] vc
+f
+S
+n
+vmrs
+1939 2030.7 m
+1940.3 2027.4 1939.7 2019.3 1939.5 2013.9 C
+1939.7 2013.5 1940.1 2013.2 1940 2012.7 C
+1939.5 2012.3 1938.8 2012.5 1938.3 2012.4 C
+1938.3 2019 1938.3 2026.2 1938.3 2032.1 C
+1939.5 2032.7 1939.8 2031.2 1939 2030.7 C
+[0 0.01 1 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1975.2 2077.2 m
+1975.3 2077.3 1975.4 2077.4 1975.5 2077.5 C
+1974.7 2073.2 1974.9 2067.5 1975.2 2063.6 C
+1975.4 2064 1974.6 2063.9 1974.8 2064.3 C
+1974.9 2069.9 1974.3 2076.5 1975.2 2081.1 C
+1974.9 2079.9 1974.9 2078.4 1975.2 2077.2 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1930.8 2067.4 m
+1931.5 2070.1 1929.6 2072.1 1930.6 2074.6 C
+1931 2072.6 1930.8 2069.8 1930.8 2067.4 C
+f
+S
+n
+2010 2050.1 m
+2009.8 2050.5 2009.5 2050.9 2009.3 2051.1 C
+2009.5 2056.7 2008.9 2063.3 2009.8 2067.9 C
+2009.5 2062.1 2009.3 2054.7 2010 2050.1 C
+f
+S
+n
+1930.1 2060.9 m
+1929.3 2057.1 1930.7 2054.8 1929.9 2051.3 C
+1930.2 2050.2 1931.1 2049.6 1931.8 2049.2 C
+1931.4 2049.6 1930.4 2049.5 1930.1 2050.1 C
+1928.4 2054.8 1933.4 2063.5 1925.3 2064.3 C
+1927.2 2063.9 1928.5 2062.1 1930.1 2060.9 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1929.6 2061.2 m
+1929.6 2057.6 1929.6 2054.1 1929.6 2050.6 C
+1930 2049.9 1930.5 2049.4 1931.1 2049.2 C
+1930 2048.6 1930.5 2050.2 1929.4 2049.6 C
+1928 2054.4 1932.8 2063 1925.3 2064 C
+1926.9 2063.3 1928.3 2062.4 1929.6 2061.2 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1930.8 2061.6 m
+1930.5 2058 1931.6 2054 1930.8 2051.3 C
+1930.3 2054.5 1930.9 2058.5 1930.4 2061.9 C
+1930.5 2061.2 1931 2062.2 1930.8 2061.6 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1941.2 2045.1 m
+1939.7 2042.6 1937.3 2041.2 1935.4 2039.3 C
+1934.2 2040 1933.7 2036.4 1934 2039.3 C
+1934.9 2040.1 1936.1 2039.9 1936.8 2040.8 C
+1935.3 2044.2 1942.3 2041.7 1939.5 2046 C
+1937.1 2048.5 1940.5 2045.6 1941.2 2045.1 C
+f
+S
+n
+1910 2045.8 m
+1910 2039.4 1910 2033 1910 2026.6 C
+1910 2033 1910 2039.4 1910 2045.8 C
+f
+S
+n
+1978.8 2022.3 m
+1979.1 2021.7 1979.4 2020.4 1978.6 2021.6 C
+1978.6 2026.9 1978.6 2033 1978.6 2037.6 C
+1979.2 2037 1979.1 2038.2 1979.1 2038.6 C
+1978.7 2033.6 1978.9 2026.8 1978.8 2022.3 C
+f
+S
+n
+vmrs
+2026.1 2041.2 m
+2026.1 2034.8 2026.1 2028.3 2026.1 2021.8 C
+2026.1 2028.5 2026.3 2035.4 2025.9 2042 C
+2024.4 2042.9 2022.9 2044.1 2021.3 2044.8 C
+2023.1 2044 2025.1 2042.8 2026.1 2041.2 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+2026.4 2021.8 m
+2026.3 2028.5 2026.5 2035.4 2026.1 2042 C
+2025.6 2042.8 2024.7 2042.7 2024.2 2043.4 C
+2024.7 2042.7 2025.5 2042.7 2026.1 2042.2 C
+2026.5 2035.5 2026.3 2027.9 2026.4 2021.8 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+2025.6 2038.4 m
+2025.6 2033 2025.6 2027.6 2025.6 2022.3 C
+2025.6 2027.6 2025.6 2033 2025.6 2038.4 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1934 2023.5 m
+1934 2024.7 1933.8 2026 1934.2 2027.1 C
+1934 2025.5 1934.7 2024.6 1934 2023.5 C
+f
+S
+n
+1928.2 2023.5 m
+1928 2024.6 1927.4 2023.1 1926.8 2023.2 C
+1926.2 2021 1921.4 2019.3 1923.2 2018 C
+1922.7 2016.5 1923.2 2019.3 1922.2 2018.2 C
+1924.4 2020.4 1926.2 2023.3 1928.9 2024.9 C
+1927.9 2024.2 1929.8 2023.5 1928.2 2023.5 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1934 2019.2 m
+1932 2019.6 1930.8 2022.6 1928.7 2021.8 C
+1924.5 2016.5 1918.2 2011.8 1914 2006.7 C
+1914 2005.7 1914 2004.6 1914 2003.6 C
+1913.6 2004.3 1913.9 2005.8 1913.8 2006.9 C
+1919 2012.4 1924.1 2016.5 1929.2 2022.3 C
+1931 2021.7 1932.2 2019.8 1934 2019.2 C
+f
+S
+n
+1928.7 2024.9 m
+1926.3 2022.7 1924.1 2020.4 1921.7 2018.2 C
+1924.1 2020.4 1926.3 2022.7 1928.7 2024.9 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1914.3 2006.7 m
+1918.7 2011.8 1924.5 2016.4 1928.9 2021.6 C
+1924.2 2016.1 1919 2012.1 1914.3 2006.7 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1924.8 2020.8 m
+1921.2 2016.9 1925.6 2022.5 1926 2021.1 C
+1924.2 2021 1926.7 2019.6 1924.8 2020.8 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1934 2018.4 m
+1933.2 2014.7 1934.5 2012.3 1933.7 2008.8 C
+1934 2007.8 1935 2007.2 1935.6 2006.7 C
+1935.3 2007.1 1934.3 2007 1934 2007.6 C
+1932.2 2012.3 1937.2 2021 1929.2 2021.8 C
+1931.1 2021.4 1932.3 2019.6 1934 2018.4 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+vmrs
+1933.5 2018.7 m
+1933.5 2015.1 1933.5 2011.7 1933.5 2008.1 C
+1933.8 2007.4 1934.3 2006.9 1934.9 2006.7 C
+1933.8 2006.1 1934.3 2007.7 1933.2 2007.2 C
+1931.9 2012 1936.7 2020.5 1929.2 2021.6 C
+1930.7 2020.8 1932.2 2019.9 1933.5 2018.7 C
+[0.4 0.4 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1934.7 2019.2 m
+1934.3 2015.6 1935.4 2011.5 1934.7 2008.8 C
+1934.1 2012 1934.7 2016 1934.2 2019.4 C
+1934.4 2018.7 1934.8 2019.8 1934.7 2019.2 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1917.6 2013.6 m
+1917.8 2011.1 1916.8 2014.2 1917.2 2012.2 C
+1916.3 2012.9 1914.8 2011.8 1914.3 2010.8 C
+1914.2 2010.5 1914.4 2010.4 1914.5 2010.3 C
+1913.9 2008.8 1913.9 2011.9 1914.3 2012 C
+1916.3 2012 1917.6 2013.6 1916.7 2015.6 C
+1913.7 2017.4 1919.6 2014.8 1917.6 2013.6 C
+f
+S
+n
+1887.2 2015.3 m
+1887.2 2008.9 1887.2 2002.5 1887.2 1996.1 C
+1887.2 2002.5 1887.2 2008.9 1887.2 2015.3 C
+f
+S
+n
+1916.7 2014.4 m
+1917 2012.1 1913 2013 1913.8 2010.8 C
+1912.1 2009.8 1910.9 2009.4 1910.7 2007.9 C
+1910.4 2010.6 1913.4 2010.4 1914 2012.4 C
+1914.9 2012.8 1916.6 2012.9 1916.4 2014.4 C
+1916.9 2015.1 1914.5 2016.6 1916.2 2015.8 C
+1916.4 2015.3 1916.7 2015 1916.7 2014.4 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1914 2009.3 m
+1912.8 2010.9 1909.6 2005.3 1911.9 2009.8 C
+1912.3 2009.6 1913.6 2010.2 1914 2009.3 C
+[0.92 0.92 0 0.67] vc
+f
+S
+n
+1951.2 1998.8 m
+1949 1996.4 1951.5 1994 1950.3 1991.8 C
+1949.1 1989.1 1954 1982.7 1948.8 1981.2 C
+1949.2 1981.5 1951 1982.4 1950.8 1983.6 C
+1951.9 1988.6 1947.1 1986.5 1948.1 1990.4 C
+1948.5 1990.3 1948.7 1990.7 1948.6 1991.1 C
+1949 1992.5 1947.3 1991.9 1948.1 1992.5 C
+1947.1 1992.7 1945.7 1993.5 1945.2 1994.7 C
+1944.5 1996.8 1947.7 2000.5 1943.8 2001.4 C
+1943.4 2002 1943.7 2004 1942.4 2004.5 C
+1945.2 2002.2 1948.9 2000.9 1951.2 1998.8 C
+f
+S
+n
+1994.9 1993 m
+1995.1 1996.5 1994.5 2000.3 1995.4 2003.6 C
+1994.5 2000.3 1995.1 1996.5 1994.9 1993 C
+f
+S
+n
+1913.8 2003.3 m
+1913.8 1996.9 1913.8 1990.5 1913.8 1984.1 C
+1913.8 1990.5 1913.8 1996.9 1913.8 2003.3 C
+f
+S
+n
+1941.9 1998 m
+1940.5 1997.3 1940.7 1999.4 1940.7 2000 C
+1942.8 2001.3 1942.6 1998.8 1941.9 1998 C
+[0 0 0 0] vc
+f
+S
+n
+vmrs
+1942.1 1999.2 m
+1942.2 1998.9 1941.8 1998.8 1941.6 1998.5 C
+1940.4 1998 1940.7 1999.7 1940.7 2000 C
+1941.6 2000.3 1942.6 2000.4 1942.1 1999.2 C
+[0.92 0.92 0 0.67] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1940 1997.1 m
+1939.8 1996 1939.7 1995.9 1939.2 1995.2 C
+1939.1 1995.3 1938.5 1997.9 1937.8 1996.4 C
+1938 1997.3 1939.4 1998.6 1940 1997.1 C
+f
+S
+n
+1911.2 1995.9 m
+1911.2 1991.6 1911.3 1987.2 1911.4 1982.9 C
+1911.3 1987.2 1911.2 1991.6 1911.2 1995.9 C
+f
+S
+n
+1947.2 1979.1 m
+1945.1 1978.8 1944.6 1975.7 1942.4 1975 C
+1940.5 1972.6 1942.2 1973.7 1942.4 1975.7 C
+1945.8 1975.5 1944.2 1979.8 1947.6 1979.6 C
+1948.3 1982.3 1948.5 1980 1947.2 1979.1 C
+f
+S
+n
+1939.5 1973.3 m
+1940.1 1972.6 1939.8 1974.2 1940.2 1973.1 C
+1939.1 1972.8 1938.8 1968.5 1935.9 1969.7 C
+1937.4 1969.2 1938.5 1970.6 1939 1971.4 C
+1939.2 1972.7 1938.6 1973.9 1939.5 1973.3 C
+f
+S
+n
+1975.2 2073.2 m
+1975.2 2070.2 1975.2 2067.2 1975.2 2064.3 C
+1975.2 2067.2 1975.2 2070.2 1975.2 2073.2 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1929.9 2065.7 m
+1928.1 2065.6 1926 2068.8 1924.1 2066.9 C
+1918.1 2060.9 1912.9 2055.7 1907.1 2049.9 C
+1906.7 2047.1 1906.9 2043.9 1906.8 2041 C
+1906.8 2043.9 1906.8 2046.8 1906.8 2049.6 C
+1913.2 2055.5 1918.7 2061.9 1925.1 2067.6 C
+1927.1 2067.9 1928.6 2064.4 1930.1 2066.2 C
+1929.7 2070.3 1929.9 2074.7 1929.9 2078.9 C
+1929.6 2074.4 1930.5 2070.1 1929.9 2065.7 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1930.1 2061.6 m
+1928.1 2062.1 1927 2065.1 1924.8 2064.3 C
+1920.7 2058.9 1914.4 2054.3 1910.2 2049.2 C
+1910.2 2048.1 1910.2 2047.1 1910.2 2046 C
+1909.8 2046.8 1910 2048.3 1910 2049.4 C
+1915.1 2054.9 1920.3 2059 1925.3 2064.8 C
+1927.1 2064.2 1928.4 2062.3 1930.1 2061.6 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1932 2049.9 m
+1932.3 2050.3 1932 2050.4 1932.8 2050.4 C
+1932 2050.4 1932.2 2049.2 1931.3 2049.6 C
+1931.4 2050.5 1930.3 2050.4 1930.4 2051.3 C
+1931.1 2051.1 1930.7 2049.4 1932 2049.9 C
+f
+S
+n
+1938.3 2046 m
+1936.3 2046.8 1935.2 2047.2 1934.2 2048.9 C
+1935.3 2047.7 1936.8 2046.2 1938.3 2046 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+vmrs
+1938.3 2047 m
+1937.9 2046.9 1936.6 2047.1 1936.1 2048 C
+1936.5 2047.5 1937.3 2046.7 1938.3 2047 C
+[0.18 0.18 0 0.78] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1910.2 2043.2 m
+1910.1 2037.5 1910 2031.8 1910 2026.1 C
+1910 2031.8 1910.1 2037.5 1910.2 2043.2 C
+f
+S
+n
+1933.5 2032.1 m
+1933.7 2035.2 1932.8 2035.8 1933.7 2038.6 C
+1933.3 2036.6 1934.6 2018 1933.5 2032.1 C
+f
+S
+n
+1907.3 2021.8 m
+1906.6 2025.9 1909.4 2032.6 1903.2 2034 C
+1902.8 2034.1 1902.4 2033.9 1902 2033.8 C
+1897.9 2028.5 1891.6 2023.8 1887.4 2018.7 C
+1887.4 2017.7 1887.4 2016.6 1887.4 2015.6 C
+1887 2016.3 1887.2 2017.8 1887.2 2018.9 C
+1892.3 2024.4 1897.5 2028.5 1902.5 2034.3 C
+1904.3 2033.6 1905.7 2032 1907.3 2030.9 C
+1907.3 2027.9 1907.3 2024.9 1907.3 2021.8 C
+f
+S
+n
+1933.7 2023.2 m
+1932 2021.7 1931.1 2024.9 1929.4 2024.9 C
+1931.2 2024.7 1932.4 2021.5 1933.7 2023.2 C
+f
+S
+n
+1989.2 2024.4 m
+1987.4 2023.7 1985.8 2022.2 1985.1 2020.4 C
+1984.6 2020.1 1986 2018.9 1985.1 2019.2 C
+1985.6 2020.8 1984.1 2019.4 1984.6 2021.1 C
+1986.3 2022.3 1988.1 2025.3 1989.2 2024.4 C
+f
+S
+n
+1904.4 2031.9 m
+1903 2029.7 1905.3 2027.7 1904.2 2025.9 C
+1904.5 2025 1903.7 2023 1904 2021.3 C
+1904 2022.3 1903.2 2022 1902.5 2022 C
+1901.3 2022.3 1902.2 2020.1 1901.6 2019.6 C
+1902.5 2019.8 1902.6 2018.3 1903.5 2018.9 C
+1903.7 2021.8 1905.6 2016.8 1905.6 2020.6 C
+1905.9 2020 1906.3 2020.8 1906.1 2021.1 C
+1905.8 2022.7 1906.7 2020.4 1906.4 2019.9 C
+1906.4 2018.5 1908.2 2017.8 1906.8 2016.5 C
+1906.9 2015.7 1907.7 2017.1 1907.1 2016.3 C
+1908.5 2015.8 1910.3 2015.1 1911.6 2016 C
+1912.2 2016.2 1911.9 2018 1911.6 2018 C
+1914.5 2017.1 1910.4 2013.6 1913.3 2013.4 C
+1912.4 2011.3 1910.5 2011.8 1909.5 2010 C
+1910 2010.5 1909 2010.8 1908.8 2011.2 C
+1907.5 2009.9 1906.1 2011.7 1904.9 2011.5 C
+1904.7 2010.9 1904.3 2010.5 1904.4 2009.8 C
+1905 2010.2 1904.6 2008.6 1905.4 2008.1 C
+1906.6 2007.5 1907.7 2008.4 1908.5 2007.4 C
+1908.9 2008.5 1909.7 2008.1 1909 2007.2 C
+1908.1 2006.5 1905.9 2007.3 1905.4 2007.4 C
+1903.9 2007.3 1905.2 2008.5 1904.2 2008.4 C
+1904.6 2009.9 1902.8 2010.3 1902.3 2010.5 C
+1901.5 2009.9 1900.4 2010 1899.4 2010 C
+1898.6 2011.2 1898.2 2013.4 1896.5 2013.4 C
+1896 2012.9 1894.4 2012.9 1893.6 2012.9 C
+1893.1 2013.9 1892.9 2015.5 1891.5 2016 C
+1890.3 2016.1 1889.2 2014 1888.6 2015.8 C
+1890 2016 1891 2016.9 1892.9 2016.5 C
+1894.1 2017.2 1892.8 2018.3 1893.2 2018.9 C
+1892.6 2018.9 1891.1 2019.8 1890.5 2020.6 C
+1891.1 2023.6 1893.2 2019.8 1893.9 2022.5 C
+1894.1 2023.3 1892.7 2023.6 1893.9 2024 C
+1894.2 2024.3 1897.4 2023.8 1896.5 2026.1 C
+1896 2025.6 1897.4 2028.1 1897.5 2027.1 C
+1898.4 2027.4 1899.3 2027 1899.6 2028.5 C
+1899.5 2028.6 1899.4 2028.8 1899.2 2028.8 C
+1899.3 2029.2 1899.6 2029.8 1900.1 2030.2 C
+1900.4 2029.6 1901 2030 1901.8 2030.2 C
+1903.1 2032.1 1900.4 2031.5 1902.8 2033.1 C
+1903.3 2032.7 1904.5 2032 1904.4 2031.9 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1909.2 2019.4 m
+1908.8 2020.3 1910.2 2019.8 1909.2 2019.2 C
+1908.3 2019.3 1907.6 2020.2 1907.6 2021.3 C
+1908.5 2021 1907.6 2019 1909.2 2019.4 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1915.5 2015.6 m
+1913.5 2016.3 1912.4 2016.8 1911.4 2018.4 C
+1912.5 2017.2 1914 2015.7 1915.5 2015.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1915.5 2016.5 m
+1915.1 2016.4 1913.8 2016.6 1913.3 2017.5 C
+1913.7 2017 1914.5 2016.2 1915.5 2016.5 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+vmrs
+1887.4 2012.7 m
+1887.3 2007 1887.2 2001.3 1887.2 1995.6 C
+1887.2 2001.3 1887.3 2007 1887.4 2012.7 C
+[0.18 0.18 0 0.78] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1935.9 2007.4 m
+1936.2 2007.8 1935.8 2007.9 1936.6 2007.9 C
+1935.9 2007.9 1936.1 2006.7 1935.2 2007.2 C
+1935.2 2008.1 1934.1 2007.9 1934.2 2008.8 C
+1935 2008.7 1934.6 2006.9 1935.9 2007.4 C
+f
+S
+n
+1942.1 2003.6 m
+1940.1 2004.3 1939.1 2004.8 1938 2006.4 C
+1939.1 2005.2 1940.6 2003.7 1942.1 2003.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1942.1 2004.5 m
+1941.8 2004.4 1940.4 2004.6 1940 2005.5 C
+1940.4 2005 1941.2 2004.2 1942.1 2004.5 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+1914 2000.7 m
+1914 1995 1913.9 1989.3 1913.8 1983.6 C
+1913.9 1989.3 1914 1995 1914 2000.7 C
+f
+S
+n
+1941.6 1998.3 m
+1943.4 2001.9 1942.4 1996 1940.9 1998.3 C
+1941.2 1998.3 1941.4 1998.3 1941.6 1998.3 C
+f
+S
+n
+1954.8 1989.9 m
+1953.9 1989.6 1954.7 1991.6 1953.9 1991.1 C
+1954.5 1993.1 1953.6 1998 1954.6 1993.2 C
+1954 1992.2 1954.7 1990.7 1954.8 1989.9 C
+f
+S
+n
+1947.6 1992.5 m
+1946.2 1993.5 1944.9 1993 1944.8 1994.7 C
+1945.5 1994 1947 1992.2 1947.6 1992.5 C
+f
+S
+n
+1910.7 1982.2 m
+1910.3 1981.8 1909.7 1982 1909.2 1982 C
+1909.7 1982 1910.3 1981.9 1910.7 1982.2 C
+1911 1987.1 1910 1992.6 1910.7 1997.3 C
+1910.7 1992.3 1910.7 1987.2 1910.7 1982.2 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1910.9 1992.8 m
+1910.9 1991.3 1910.9 1989.7 1910.9 1988.2 C
+1910.9 1989.7 1910.9 1991.3 1910.9 1992.8 C
+[0.18 0.18 0 0.78] vc
+f
+S
+n
+vmrs
+1953.6 1983.6 m
+1954.1 1985.3 1953.2 1988.6 1954.8 1989.4 C
+1954.1 1987.9 1954.4 1985.4 1953.6 1983.6 C
+[0.18 0.18 0 0.78] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1910.7 1982 m
+1911.6 1982.9 1911 1984.4 1911.2 1985.6 C
+1911 1984.4 1911.6 1982.9 1910.7 1982 C
+f
+S
+n
+1947.2 1979.6 m
+1947.5 1980.6 1948.3 1980.6 1947.4 1979.6 C
+1946.2 1979.4 1945.7 1978.8 1947.2 1979.6 C
+f
+S
+n
+1930.4 2061.4 m
+1930.4 2058 1930.4 2053.5 1930.4 2051.1 C
+1930.7 2054.6 1929.8 2057.4 1930.1 2061.2 C
+1929.5 2061.9 1929.7 2061.2 1930.4 2061.4 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1939.5 2044.8 m
+1940 2041.5 1935.2 2044.3 1936.4 2040.8 C
+1934.9 2040.9 1934.1 2039.7 1933.5 2038.6 C
+1933.3 2035.4 1933.2 2040 1934 2040.3 C
+1936.2 2040.6 1936.3 2043.6 1938.5 2043.4 C
+1939.7 2044.2 1939.4 2045.6 1938.3 2046.5 C
+1939.1 2046.6 1939.6 2045.6 1939.5 2044.8 C
+f
+S
+n
+1910.4 2045.3 m
+1910.4 2039.5 1910.4 2033.6 1910.4 2027.8 C
+1910.4 2033.6 1910.4 2039.5 1910.4 2045.3 C
+f
+S
+n
+1906.8 2030.9 m
+1907.6 2026.8 1905 2020.8 1909 2018.7 C
+1906.5 2018.9 1906.8 2022.4 1906.8 2024.7 C
+1906.4 2028.2 1907.9 2032 1903 2033.8 C
+1902.2 2034 1903.8 2033.4 1904.2 2033.1 C
+1905.1 2032.4 1905.9 2031.5 1906.8 2030.9 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1907.1 2030.7 m
+1907.1 2028.8 1907.1 2027 1907.1 2025.2 C
+1907.1 2027 1907.1 2028.8 1907.1 2030.7 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1932 2023.2 m
+1932.2 2023.6 1931.7 2023.7 1931.6 2024 C
+1932 2023.7 1932.3 2022.8 1933 2023 C
+1933.9 2024.3 1933.3 2026.2 1933.5 2027.8 C
+1933.5 2026.4 1934.9 2022.2 1932 2023.2 C
+f
+S
+n
+2026.1 2021.6 m
+2026.1 2020.8 2026.1 2019.9 2026.1 2019.2 C
+2026.1 2019.9 2026.1 2020.8 2026.1 2021.6 C
+f
+S
+n
+vmrs
+1934.2 2018.9 m
+1934.2 2015.5 1934.2 2011 1934.2 2008.6 C
+1934.5 2012.1 1933.7 2014.9 1934 2018.7 C
+1933.4 2019.5 1933.5 2018.7 1934.2 2018.9 C
+[0.65 0.65 0 0.42] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1887.6 2014.8 m
+1887.6 2009 1887.6 2003.1 1887.6 1997.3 C
+1887.6 2003.1 1887.6 2009 1887.6 2014.8 C
+f
+S
+n
+1914.3 2002.8 m
+1914.3 1997 1914.3 1991.1 1914.3 1985.3 C
+1914.3 1991.1 1914.3 1997 1914.3 2002.8 C
+f
+S
+n
+1995.4 1992.3 m
+1995.4 1991.5 1995.4 1990.7 1995.4 1989.9 C
+1995.4 1990.7 1995.4 1991.5 1995.4 1992.3 C
+f
+S
+n
+1896 1988.4 m
+1896.9 1988 1897.8 1987.7 1898.7 1987.2 C
+1897.8 1987.7 1896.9 1988 1896 1988.4 C
+f
+S
+n
+1899.4 1986.8 m
+1900.4 1986.3 1901.3 1985.8 1902.3 1985.3 C
+1901.3 1985.8 1900.4 1986.3 1899.4 1986.8 C
+f
+S
+n
+1902.8 1985.1 m
+1905.2 1984 1905.2 1984 1902.8 1985.1 C
+f
+S
+n
+1949.1 1983.4 m
+1950.2 1984.4 1947.8 1984.6 1949.3 1985.1 C
+1949.5 1984.4 1949.6 1984.1 1949.1 1983.4 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1906.1 1983.4 m
+1908.6 1982 1908.6 1982 1906.1 1983.4 C
+[0.65 0.65 0 0.42] vc
+f
+S
+n
+1922.7 1976.4 m
+1923.6 1976 1924.4 1975.7 1925.3 1975.2 C
+1924.4 1975.7 1923.6 1976 1922.7 1976.4 C
+f
+S
+n
+vmrs
+1926 1974.8 m
+1927 1974.3 1928 1973.8 1928.9 1973.3 C
+1928 1973.8 1927 1974.3 1926 1974.8 C
+[0.65 0.65 0 0.42] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1929.4 1973.1 m
+1931.9 1972 1931.9 1972 1929.4 1973.1 C
+f
+S
+n
+1932.8 1971.4 m
+1935.3 1970 1935.3 1970 1932.8 1971.4 C
+f
+S
+n
+1949.6 2097.2 m
+1951.1 2096.4 1952.6 2095.5 1954.1 2094.8 C
+1952.6 2095.5 1951.1 2096.4 1949.6 2097.2 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1955.1 2094.3 m
+1956.7 2093.5 1958.3 2092.7 1959.9 2091.9 C
+1958.3 2092.7 1956.7 2093.5 1955.1 2094.3 C
+f
+S
+n
+1960.4 2091.6 m
+1961.3 2091.2 1962.1 2090.9 1963 2090.4 C
+1962.1 2090.9 1961.3 2091.2 1960.4 2091.6 C
+f
+S
+n
+1963.5 2090.2 m
+1964.4 2089.7 1965.2 2089.2 1966.1 2088.8 C
+1965.2 2089.2 1964.4 2089.7 1963.5 2090.2 C
+f
+S
+n
+1966.6 2088.5 m
+1969.5 2087.1 1972.4 2085.8 1975.2 2084.4 C
+1972.4 2085.8 1969.5 2087.1 1966.6 2088.5 C
+f
+S
+n
+1965.2 2086.1 m
+1965.9 2085.7 1966.8 2085.3 1967.6 2084.9 C
+1966.8 2085.3 1965.9 2085.7 1965.2 2086.1 C
+f
+S
+n
+1968.3 2084.7 m
+1969.2 2084.3 1970 2083.9 1970.9 2083.5 C
+1970 2083.9 1969.2 2084.3 1968.3 2084.7 C
+f
+S
+n
+vmrs
+1984.1 2084 m
+1985.6 2083.2 1987.2 2082.3 1988.7 2081.6 C
+1987.2 2082.3 1985.6 2083.2 1984.1 2084 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1976 2078.7 m
+1978.1 2080.1 1980 2082 1982 2083.7 C
+1980 2081.9 1977.9 2080.3 1976 2078.2 C
+1975.5 2079.9 1975.8 2081.9 1975.7 2083.7 C
+1975.8 2082 1975.5 2080.2 1976 2078.7 C
+f
+S
+n
+1989.6 2081.1 m
+1991.3 2080.3 1992.8 2079.5 1994.4 2078.7 C
+1992.8 2079.5 1991.3 2080.3 1989.6 2081.1 C
+f
+S
+n
+1933.2 2074.6 m
+1932.4 2076.2 1932.8 2077.5 1933 2078.7 C
+1933 2077.6 1932.9 2074.8 1933.2 2074.6 C
+f
+S
+n
+1994.9 2078.4 m
+1995.8 2078 1996.7 2077.7 1997.6 2077.2 C
+1996.7 2077.7 1995.8 2078 1994.9 2078.4 C
+f
+S
+n
+1998 2077 m
+1998.9 2076.5 1999.8 2076 2000.7 2075.6 C
+1999.8 2076 1998.9 2076.5 1998 2077 C
+f
+S
+n
+2001.2 2075.3 m
+2004 2073.9 2006.9 2072.6 2009.8 2071.2 C
+2006.9 2072.6 2004 2073.9 2001.2 2075.3 C
+f
+S
+n
+1980.5 2060.7 m
+1979.9 2060.7 1976.7 2062.8 1975.7 2064.5 C
+1975.7 2067.5 1975.7 2070.5 1975.7 2073.4 C
+1976.3 2068.7 1973.9 2061.6 1980.5 2060.7 C
+f
+S
+n
+1999.7 2072.9 m
+2000.5 2072.5 2001.3 2072.1 2002.1 2071.7 C
+2001.3 2072.1 2000.5 2072.5 1999.7 2072.9 C
+f
+S
+n
+2002.8 2071.5 m
+2003.7 2071.1 2004.6 2070.7 2005.5 2070.3 C
+2004.6 2070.7 2003.7 2071.1 2002.8 2071.5 C
+f
+S
+n
+vmrs
+2015.1 2047.5 m
+2014.4 2047.5 2011.2 2049.6 2010.3 2051.3 C
+2010.3 2057.7 2010.3 2064.1 2010.3 2070.5 C
+2010.3 2063.9 2010.1 2057.1 2010.5 2050.6 C
+2012 2049.3 2013.5 2048.3 2015.1 2047.5 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1910.4 2049.2 m
+1914.8 2054.3 1920.7 2058.9 1925.1 2064 C
+1920.4 2058.6 1915.1 2054.6 1910.4 2049.2 C
+f
+S
+n
+1988.2 2057.3 m
+1989.1 2056.8 1989.9 2056.2 1990.8 2055.6 C
+1989.9 2056.2 1989.1 2056.8 1988.2 2057.3 C
+f
+S
+n
+1991.6 2051.3 m
+1991.6 2046.3 1991.6 2041.2 1991.6 2036.2 C
+1991.6 2041.2 1991.6 2046.3 1991.6 2051.3 C
+f
+S
+n
+1935.6 2047.5 m
+1932.9 2051.7 1939.7 2043.8 1935.6 2047.5 C
+f
+S
+n
+1938.8 2043.9 m
+1938.1 2043.3 1938.2 2043.7 1937.3 2043.4 C
+1938.7 2043 1938.2 2044.9 1939 2045.3 C
+1938.2 2045.3 1938.7 2046.6 1937.8 2046.5 C
+1939.1 2046.2 1939.1 2044.5 1938.8 2043.9 C
+f
+S
+n
+1972.4 2045.6 m
+1973.4 2045 1974.5 2044.4 1975.5 2043.9 C
+1974.5 2044.4 1973.4 2045 1972.4 2045.6 C
+f
+S
+n
+1969 2043.6 m
+1969.8 2043.2 1970.6 2042.9 1971.4 2042.4 C
+1970.6 2042.9 1969.8 2043.2 1969 2043.6 C
+f
+S
+n
+1972.1 2042.2 m
+1973 2041.8 1973.9 2041.4 1974.8 2041 C
+1973.9 2041.4 1973 2041.8 1972.1 2042.2 C
+f
+S
+n
+1906.6 2035 m
+1905 2034.7 1904.8 2036.6 1903.5 2036.9 C
+1904.9 2037 1905.8 2033.4 1907.1 2035.7 C
+1907.1 2037.2 1907.1 2038.6 1907.1 2040 C
+1906.9 2038.4 1907.5 2036.4 1906.6 2035 C
+f
+S
+n
+vmrs
+1937.1 2032.1 m
+1936.2 2033.7 1936.6 2035 1936.8 2036.2 C
+1936.8 2035.1 1936.8 2032.4 1937.1 2032.1 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1887.6 2018.7 m
+1892 2023.8 1897.9 2028.4 1902.3 2033.6 C
+1897.6 2028.1 1892.3 2024.1 1887.6 2018.7 C
+f
+S
+n
+1999.7 2031.4 m
+1998.7 2030.3 1997.6 2029.2 1996.6 2028 C
+1997.6 2029.2 1998.7 2030.3 1999.7 2031.4 C
+f
+S
+n
+1912.8 2017 m
+1910.6 2021.1 1913.6 2015.3 1914.5 2016 C
+1914 2016.3 1913.4 2016.7 1912.8 2017 C
+f
+S
+n
+1939.5 2005 m
+1936.7 2009.2 1943.6 2001.3 1939.5 2005 C
+f
+S
+n
+1942.6 2001.4 m
+1941.9 2000.8 1942 2001.2 1941.2 2000.9 C
+1942.5 2000.6 1942.1 2002.4 1942.8 2002.8 C
+1942 2002.8 1942.5 2004.1 1941.6 2004 C
+1943 2003.7 1942.9 2002.1 1942.6 2001.4 C
+f
+S
+n
+2006.2 2000.7 m
+2005.4 2001.5 2004 2002.8 2004 2002.8 C
+2004.5 2002.4 2005.5 2001.4 2006.2 2000.7 C
+f
+S
+n
+1998.5 2001.6 m
+1997.7 2002 1996.8 2002.4 1995.9 2002.6 C
+1995.5 1999.3 1995.7 1995.7 1995.6 1992.3 C
+1995.6 1995.7 1995.6 1999.2 1995.6 2002.6 C
+1996.6 2002.4 1997.7 2002.2 1998.5 2001.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1996.1 2002.8 m
+1995.9 2002.8 1995.8 2002.8 1995.6 2002.8 C
+1995.2 1999.5 1995.5 1995.9 1995.4 1992.5 C
+1995.4 1995.9 1995.4 1999.4 1995.4 2002.8 C
+1996.4 2003.1 1998.2 2001.6 1996.1 2002.8 C
+[0.07 0.06 0 0.58] vc
+f
+S
+n
+1969 2002.1 m
+1968 2001 1966.9 1999.9 1965.9 1998.8 C
+1966.9 1999.9 1968 2001 1969 2002.1 C
+f
+S
+n
+vmrs
+2000 2001.2 m
+2002.1 2000 2004.1 1998.9 2006.2 1997.8 C
+2004.1 1998.9 2002.1 2000 2000 2001.2 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1895.8 1984.8 m
+1898.3 1983.6 1900.8 1982.3 1903.2 1981 C
+1900.8 1982.3 1898.3 1983.6 1895.8 1984.8 C
+f
+S
+n
+1905.2 1980.3 m
+1906.4 1979.9 1907.6 1979.5 1908.8 1979.1 C
+1907.6 1979.5 1906.4 1979.9 1905.2 1980.3 C
+f
+S
+n
+1964.7 1977.4 m
+1963.8 1977.5 1962.5 1980.2 1960.8 1980 C
+1962.5 1980.2 1963.3 1978 1964.7 1977.4 C
+f
+S
+n
+1952 1979.6 m
+1955.2 1979.2 1955.2 1979.2 1952 1979.6 C
+f
+S
+n
+1937.8 1966.4 m
+1941.2 1969.5 1946.1 1976.4 1951.5 1979.3 C
+1946.1 1976.7 1942.8 1970.4 1937.8 1966.4 C
+f
+S
+n
+1911.9 1978.6 m
+1914.3 1977.4 1916.7 1976.2 1919.1 1975 C
+1916.7 1976.2 1914.3 1977.4 1911.9 1978.6 C
+f
+S
+n
+1975.5 1971.4 m
+1974.6 1972.2 1973.3 1973.6 1973.3 1973.6 C
+1973.7 1973.1 1974.8 1972.1 1975.5 1971.4 C
+f
+S
+n
+1922.4 1972.8 m
+1924.9 1971.6 1927.4 1970.3 1929.9 1969 C
+1927.4 1970.3 1924.9 1971.6 1922.4 1972.8 C
+f
+S
+n
+1969.2 1971.9 m
+1971.1 1970.9 1972.9 1969.8 1974.8 1968.8 C
+1972.9 1969.8 1971.1 1970.9 1969.2 1971.9 C
+f
+S
+n
+vmrs
+1931.8 1968.3 m
+1933 1967.9 1934.2 1967.5 1935.4 1967.1 C
+1934.2 1967.5 1933 1967.9 1931.8 1968.3 C
+[0.07 0.06 0 0.58] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1940.7 2072.4 m
+1941.5 2072.4 1942.3 2072.3 1943.1 2072.2 C
+1942.3 2072.3 1941.5 2072.4 1940.7 2072.4 C
+[0 0 0 0.18] vc
+f
+S
+n
+1948.6 2069.3 m
+1947 2069.5 1945.7 2068.9 1944.8 2069.8 C
+1945.9 2068.5 1948.4 2070.2 1948.6 2069.3 C
+f
+S
+n
+1954.6 2066.4 m
+1954.7 2067.9 1955.6 2067.3 1955.6 2068.8 C
+1955.4 2067.8 1956 2066.6 1954.6 2066.4 C
+f
+S
+n
+1929.2 2061.2 m
+1927.8 2062.1 1926.3 2064.1 1924.8 2063.3 C
+1926.3 2064.6 1928 2062 1929.2 2061.2 C
+f
+S
+n
+1924.4 2067.4 m
+1918.5 2061.6 1912.7 2055.9 1906.8 2050.1 C
+1912.7 2055.9 1918.5 2061.6 1924.4 2067.4 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1924.6 2062.8 m
+1923.9 2062.1 1923.2 2061.2 1922.4 2060.4 C
+1923.2 2061.2 1923.9 2062.1 1924.6 2062.8 C
+[0 0 0 0.18] vc
+f
+S
+n
+1919.3 2057.3 m
+1917.5 2055.6 1915.7 2053.8 1913.8 2052 C
+1915.7 2053.8 1917.5 2055.6 1919.3 2057.3 C
+f
+S
+n
+1929.2 2055.2 m
+1929.2 2054.2 1929.2 2053.2 1929.2 2052.3 C
+1929.2 2053.2 1929.2 2054.2 1929.2 2055.2 C
+f
+S
+n
+1926.3 2049.6 m
+1925.4 2049 1925.4 2050.5 1924.4 2050.4 C
+1925.3 2051.3 1924.5 2051.9 1925.6 2052.5 C
+1926.9 2052.6 1926 2050.6 1926.3 2049.6 C
+f
+S
+n
+vmrs
+1911.2 2046.8 m
+1910.1 2048.9 1911.9 2050.1 1913.1 2051.3 C
+1912.1 2049.9 1910.6 2048.8 1911.2 2046.8 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1934 2048.7 m
+1932.6 2048.7 1930.1 2047.7 1929.6 2049.4 C
+1930.9 2048.6 1933.3 2049 1934 2048.7 C
+f
+S
+n
+1980 2048.4 m
+1979.5 2046.8 1976.3 2047.9 1977.2 2045.6 C
+1976.8 2045.1 1976.1 2044.7 1975.2 2044.8 C
+1973.7 2046 1976.3 2046.4 1976.7 2047.5 C
+1977.8 2047.2 1978.2 2050 1979.6 2049.2 C
+1980 2049 1979.6 2048.6 1980 2048.4 C
+f
+S
+n
+1938.3 2045.6 m
+1938.2 2044.4 1936.8 2043.8 1935.9 2043.4 C
+1936.4 2044.4 1939.1 2044.3 1937.6 2045.8 C
+1937 2046.1 1935.9 2046.1 1935.9 2046.8 C
+1936.7 2046.3 1937.8 2046.2 1938.3 2045.6 C
+f
+S
+n
+1932.5 2040 m
+1932.8 2038.1 1932 2038.9 1932.3 2040.3 C
+1933.1 2040.3 1932.7 2041.7 1933.7 2041.5 C
+1933.1 2041 1932.9 2040.5 1932.5 2040 C
+f
+S
+n
+2014.6 2035.2 m
+2014.1 2033.6 2010.9 2034.7 2011.7 2032.4 C
+2011.3 2031.9 2009.4 2030.7 2009.3 2032.1 C
+2009.5 2033.7 2012.9 2033.8 2012.4 2035.7 C
+2013 2036.4 2014.2 2036.5 2014.6 2035.2 C
+f
+S
+n
+1906.4 2030.7 m
+1905 2031.6 1903.5 2033.6 1902 2032.8 C
+1903.4 2034 1905.6 2031.4 1906.4 2030.7 C
+f
+S
+n
+1901.8 2037.2 m
+1899.5 2034.8 1897.2 2032.5 1894.8 2030.2 C
+1897.2 2032.5 1899.5 2034.8 1901.8 2037.2 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1901.8 2032.4 m
+1901.1 2031.6 1900.4 2030.7 1899.6 2030 C
+1900.4 2030.7 1901.1 2031.6 1901.8 2032.4 C
+[0 0 0 0.18] vc
+f
+S
+n
+1944.5 2030 m
+1945.3 2029.9 1946.1 2029.8 1946.9 2029.7 C
+1946.1 2029.8 1945.3 2029.9 1944.5 2030 C
+f
+S
+n
+vmrs
+1997.8 2027.8 m
+1997.7 2027.9 1997.6 2028.1 1997.3 2028 C
+1997.4 2029.1 1998.5 2029.5 1999.2 2030 C
+2000.1 2029.5 1998.9 2028 1997.8 2027.8 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1906.4 2029.2 m
+1906.4 2026.6 1906.4 2024 1906.4 2021.3 C
+1906.4 2024 1906.4 2026.6 1906.4 2029.2 C
+f
+S
+n
+2006.2 2025.9 m
+2006 2025.9 2005.8 2025.8 2005.7 2025.6 C
+2005.7 2025.5 2005.7 2025.3 2005.7 2025.2 C
+2004.6 2025.8 2002.7 2024.7 2001.9 2026.1 C
+2001.9 2027.9 2007.8 2029.2 2006.2 2025.9 C
+[0 0 0 0] vc
+f
+S
+n
+1952.4 2026.8 m
+1950.9 2027 1949.6 2026.4 1948.6 2027.3 C
+1949.7 2026.1 1952.2 2027.7 1952.4 2026.8 C
+[0 0 0 0.18] vc
+f
+S
+n
+1896.5 2026.8 m
+1894.7 2025.1 1892.9 2023.3 1891 2021.6 C
+1892.9 2023.3 1894.7 2025.1 1896.5 2026.8 C
+f
+S
+n
+1958.4 2024 m
+1958.5 2025.5 1959.4 2024.8 1959.4 2026.4 C
+1959.3 2025.3 1959.8 2024.1 1958.4 2024 C
+f
+S
+n
+1903.5 2019.2 m
+1902.6 2018.6 1902.6 2020 1901.6 2019.9 C
+1902.5 2020.8 1901.7 2021.4 1902.8 2022 C
+1904.1 2022.2 1903.2 2020.1 1903.5 2019.2 C
+f
+S
+n
+1933 2018.7 m
+1931.7 2019.6 1930.1 2021.6 1928.7 2020.8 C
+1930.1 2022.1 1931.8 2019.5 1933 2018.7 C
+f
+S
+n
+1888.4 2016.3 m
+1887.3 2018.4 1889.1 2019.6 1890.3 2020.8 C
+1889.3 2019.5 1887.8 2018.3 1888.4 2016.3 C
+f
+S
+n
+1928.4 2020.4 m
+1927.7 2019.6 1927 2018.7 1926.3 2018 C
+1927 2018.7 1927.7 2019.6 1928.4 2020.4 C
+f
+S
+n
+vmrs
+1911.2 2018.2 m
+1909.8 2018.3 1907.3 2017.2 1906.8 2018.9 C
+1908.1 2018.1 1910.5 2018.6 1911.2 2018.2 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1915.5 2015.1 m
+1915.4 2013.9 1914 2013.3 1913.1 2012.9 C
+1913.6 2013.9 1916.3 2013.8 1914.8 2015.3 C
+1914.2 2015.6 1913.1 2015.6 1913.1 2016.3 C
+1913.9 2015.9 1915 2015.7 1915.5 2015.1 C
+f
+S
+n
+1923.2 2014.8 m
+1921.3 2013.1 1919.5 2011.3 1917.6 2009.6 C
+1919.5 2011.3 1921.3 2013.1 1923.2 2014.8 C
+f
+S
+n
+1933 2012.7 m
+1933 2011.7 1933 2010.8 1933 2009.8 C
+1933 2010.8 1933 2011.7 1933 2012.7 C
+f
+S
+n
+1909.7 2008.1 m
+1908.9 2009.2 1910.1 2009.9 1910.4 2011 C
+1911.1 2010.7 1908.9 2009.7 1909.7 2008.1 C
+f
+S
+n
+1930.1 2007.2 m
+1929.2 2006.6 1929.2 2008 1928.2 2007.9 C
+1929.1 2008.8 1928.4 2009.4 1929.4 2010 C
+1930.7 2010.2 1929.9 2008.1 1930.1 2007.2 C
+f
+S
+n
+1915 2004.3 m
+1914 2006.4 1915.7 2007.6 1916.9 2008.8 C
+1915.9 2007.5 1914.4 2006.3 1915 2004.3 C
+f
+S
+n
+1937.8 2006.2 m
+1936.4 2006.3 1934 2005.2 1933.5 2006.9 C
+1934.7 2006.1 1937.1 2006.6 1937.8 2006.2 C
+f
+S
+n
+1983.9 2006 m
+1983.3 2004.3 1980.2 2005.4 1981 2003.1 C
+1980.6 2002.7 1978.7 2001.5 1978.6 2002.8 C
+1978.8 2004.4 1982.1 2004.5 1981.7 2006.4 C
+1982.3 2007.2 1983.5 2007.2 1983.9 2006 C
+f
+S
+n
+1942.1 2003.1 m
+1942 2001.9 1940.6 2001.3 1939.7 2000.9 C
+1940.2 2001.9 1943 2001.8 1941.4 2003.3 C
+1940.9 2003.6 1939.7 2003.6 1939.7 2004.3 C
+1940.5 2003.9 1941.6 2003.7 1942.1 2003.1 C
+f
+S
+n
+vmrs
+1967.1 1998.5 m
+1967 1998.6 1966.8 1998.8 1966.6 1998.8 C
+1966.7 1999.8 1967.8 2000.2 1968.5 2000.7 C
+1969.4 2000.2 1968.2 1998.8 1967.1 1998.5 C
+[0 0 0 0.18] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1936.4 1997.6 m
+1936.7 1995.6 1935.8 1996.4 1936.1 1997.8 C
+1936.9 1997.9 1936.5 1999.2 1937.6 1999 C
+1937 1998.5 1936.8 1998 1936.4 1997.6 C
+f
+S
+n
+1975.5 1996.6 m
+1975.2 1996.7 1975.1 1996.5 1975 1996.4 C
+1975 1996.2 1975 1996.1 1975 1995.9 C
+1973.9 1996.5 1972 1995.5 1971.2 1996.8 C
+1971.2 1998.6 1977 1999.9 1975.5 1996.6 C
+[0 0 0 0] vc
+f
+S
+n
+1949.3 2097.4 m
+1950.3 2096.9 1951.2 2096.4 1952.2 2096 C
+1951.2 2096.4 1950.3 2096.9 1949.3 2097.4 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1960.8 2091.6 m
+1961.7 2091.2 1962.6 2090.9 1963.5 2090.4 C
+1962.6 2090.9 1961.7 2091.2 1960.8 2091.6 C
+f
+S
+n
+1964.4 2090 m
+1965.7 2089.2 1967 2088.5 1968.3 2087.8 C
+1967 2088.5 1965.7 2089.2 1964.4 2090 C
+f
+S
+n
+1976 2083.7 m
+1976.3 2082.3 1975.2 2079.1 1976.9 2079.4 C
+1978.8 2080.7 1980.3 2082.9 1982.2 2084.2 C
+1980.6 2083.1 1978.2 2080.2 1976 2078.9 C
+1975.6 2081.2 1977 2084.9 1973.8 2085.4 C
+1972.2 2086.1 1970.7 2087 1969 2087.6 C
+1971.4 2086.5 1974.1 2085.6 1976 2083.7 C
+f
+S
+n
+1983.9 2084.2 m
+1984.8 2083.7 1985.8 2083.2 1986.8 2082.8 C
+1985.8 2083.2 1984.8 2083.7 1983.9 2084.2 C
+f
+S
+n
+1995.4 2078.4 m
+1996.3 2078 1997.1 2077.7 1998 2077.2 C
+1997.1 2077.7 1996.3 2078 1995.4 2078.4 C
+f
+S
+n
+1999 2076.8 m
+2000.3 2076 2001.6 2075.3 2002.8 2074.6 C
+2001.6 2075.3 2000.3 2076 1999 2076.8 C
+f
+S
+n
+vmrs
+1929.6 2065.7 m
+1930.1 2065.6 1929.8 2068.6 1929.9 2070 C
+1929.8 2068.6 1930.1 2067 1929.6 2065.7 C
+[0.4 0.4 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1906.6 2049.4 m
+1906.6 2046.7 1906.6 2043.9 1906.6 2041.2 C
+1906.6 2043.9 1906.6 2046.7 1906.6 2049.4 C
+f
+S
+n
+2016 2047.5 m
+2014.8 2048 2013.5 2048.3 2012.4 2049.4 C
+2013.5 2048.3 2014.8 2048 2016 2047.5 C
+f
+S
+n
+2016.5 2047.2 m
+2017.3 2046.9 2018.1 2046.6 2018.9 2046.3 C
+2018.1 2046.6 2017.3 2046.9 2016.5 2047.2 C
+f
+S
+n
+1912.4 2028.5 m
+1911.8 2032.4 1912.4 2037.2 1911.9 2041.2 C
+1911.5 2037.2 1911.7 2032.9 1911.6 2028.8 C
+1911.6 2033.5 1911.6 2038.9 1911.6 2042.9 C
+1912.5 2042.2 1911.6 2043.9 1912.6 2043.6 C
+1912.9 2039.3 1913.1 2033.3 1912.4 2028.5 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1906.8 2040.8 m
+1906.8 2039 1906.8 2037.2 1906.8 2035.5 C
+1906.8 2037.2 1906.8 2039 1906.8 2040.8 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1905.9 2035.2 m
+1904.9 2036.4 1903.7 2037.2 1902.3 2037.4 C
+1903.7 2037.2 1904.9 2036.4 1905.9 2035.2 C
+f
+S
+n
+1906.1 2031.2 m
+1907 2031.1 1906.4 2028 1906.6 2030.7 C
+1905.5 2032.1 1904 2032.8 1902.5 2033.6 C
+1903.9 2033.2 1905 2032.1 1906.1 2031.2 C
+f
+S
+n
+1908.3 2018.7 m
+1905.2 2018.6 1907.1 2023.2 1906.6 2025.4 C
+1906.8 2023 1905.9 2019.5 1908.3 2018.7 C
+f
+S
+n
+1889.6 1998 m
+1889 2001.9 1889.6 2006.7 1889.1 2010.8 C
+1888.7 2006.7 1888.9 2002.4 1888.8 1998.3 C
+1888.8 2003 1888.8 2008.4 1888.8 2012.4 C
+1889.7 2011.7 1888.8 2013.4 1889.8 2013.2 C
+1890.1 2008.8 1890.3 2002.8 1889.6 1998 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+vmrs
+1999 2001.4 m
+2001 2000.3 2003 1999.2 2005 1998 C
+2003 1999.2 2001 2000.3 1999 2001.4 C
+[0.4 0.4 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1916.2 1986 m
+1915.7 1989.9 1916.3 1994.7 1915.7 1998.8 C
+1915.3 1994.7 1915.5 1990.4 1915.5 1986.3 C
+1915.5 1991 1915.5 1996.4 1915.5 2000.4 C
+1916.3 1999.7 1915.5 2001.4 1916.4 2001.2 C
+1916.7 1996.8 1917 1990.8 1916.2 1986 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1886.9 1989.6 m
+1887.8 1989.2 1888.7 1988.9 1889.6 1988.4 C
+1888.7 1988.9 1887.8 1989.2 1886.9 1989.6 C
+[0.4 0.4 0 0] vc
+f
+S
+n
+1892.4 1986.8 m
+1895.1 1985.1 1897.9 1983.6 1900.6 1982 C
+1897.9 1983.6 1895.1 1985.1 1892.4 1986.8 C
+f
+S
+n
+1907.3 1979.3 m
+1908.5 1978.9 1909.7 1978.5 1910.9 1978.1 C
+1909.7 1978.5 1908.5 1978.9 1907.3 1979.3 C
+f
+S
+n
+1938.5 1966.6 m
+1942.6 1970.1 1945.9 1976.4 1951.7 1979.1 C
+1946.2 1976.1 1943.1 1970.9 1938.5 1966.6 C
+f
+S
+n
+1955.1 1978.6 m
+1955.9 1978.2 1956.7 1977.8 1957.5 1977.4 C
+1956.7 1977.8 1955.9 1978.2 1955.1 1978.6 C
+f
+S
+n
+1913.6 1977.6 m
+1914.5 1977.2 1915.3 1976.9 1916.2 1976.4 C
+1915.3 1976.9 1914.5 1977.2 1913.6 1977.6 C
+f
+S
+n
+1919.1 1974.8 m
+1921.8 1973.1 1924.5 1971.6 1927.2 1970 C
+1924.5 1971.6 1921.8 1973.1 1919.1 1974.8 C
+f
+S
+n
+1963.5 1974.5 m
+1964.5 1974 1965.6 1973.4 1966.6 1972.8 C
+1965.6 1973.4 1964.5 1974 1963.5 1974.5 C
+f
+S
+n
+vmrs
+1967.8 1972.4 m
+1970 1971.2 1972.1 1970 1974.3 1968.8 C
+1972.1 1970 1970 1971.2 1967.8 1972.4 C
+[0.4 0.4 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1934 1967.3 m
+1935.2 1966.9 1936.4 1966.5 1937.6 1966.1 C
+1936.4 1966.5 1935.2 1966.9 1934 1967.3 C
+f
+S
+n
+1928.9 2061.2 m
+1928.9 2059.2 1928.9 2057.3 1928.9 2055.4 C
+1928.9 2057.3 1928.9 2059.2 1928.9 2061.2 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1917.2 2047 m
+1917.8 2046.5 1919.6 2046.8 1920 2047.2 C
+1920 2046.5 1920.9 2046.8 1921 2046.3 C
+1921.9 2047.3 1921.3 2044.1 1921.5 2044.1 C
+1919.7 2044.8 1915.7 2043.5 1916.2 2046 C
+1916.2 2048.3 1917 2045.9 1917.2 2047 C
+[0 0 0 0] vc
+f
+S
+n
+1922 2044.1 m
+1923.5 2043.2 1927 2045.4 1927.5 2042.9 C
+1927.1 2042.6 1927.3 2040.9 1927.2 2041.5 C
+1924.9 2042.3 1920.9 2040.6 1922 2044.1 C
+f
+S
+n
+1934.9 2043.9 m
+1935.2 2043.4 1934.4 2042.7 1934 2042.2 C
+1933.2 2041.8 1932.4 2042.8 1932.8 2043.2 C
+1932.9 2044 1934.3 2043.3 1934.9 2043.9 C
+f
+S
+n
+1906.1 2030.7 m
+1906.1 2028.8 1906.1 2027 1906.1 2025.2 C
+1906.1 2027 1906.1 2028.8 1906.1 2030.7 C
+[0.21 0.21 0 0] vc
+f
+S
+n
+1932.8 2018.7 m
+1932.8 2016.8 1932.8 2014.8 1932.8 2012.9 C
+1932.8 2014.8 1932.8 2016.8 1932.8 2018.7 C
+f
+S
+n
+1894.4 2016.5 m
+1895 2016 1896.8 2016.3 1897.2 2016.8 C
+1897.2 2016 1898.1 2016.3 1898.2 2015.8 C
+1899.1 2016.8 1898.5 2013.6 1898.7 2013.6 C
+1896.9 2014.4 1892.9 2013 1893.4 2015.6 C
+1893.4 2017.8 1894.2 2015.4 1894.4 2016.5 C
+[0 0 0 0] vc
+f
+S
+n
+1899.2 2013.6 m
+1900.7 2012.7 1904.2 2014.9 1904.7 2012.4 C
+1904.3 2012.1 1904.5 2010.5 1904.4 2011 C
+1902.1 2011.8 1898.1 2010.1 1899.2 2013.6 C
+f
+S
+n
+vmrs
+1912.1 2013.4 m
+1912.4 2012.9 1911.6 2012.3 1911.2 2011.7 C
+1910.4 2011.4 1909.6 2012.3 1910 2012.7 C
+1910.1 2013.5 1911.5 2012.9 1912.1 2013.4 C
+[0 0 0 0] vc
+f
+0.4 w
+2 J
+2 M
+S
+n
+1921 2004.5 m
+1921.6 2004 1923.4 2004.3 1923.9 2004.8 C
+1923.8 2004 1924.8 2004.3 1924.8 2003.8 C
+1925.7 2004.8 1925.1 2001.6 1925.3 2001.6 C
+1923.6 2002.4 1919.6 2001 1920 2003.6 C
+1920 2005.8 1920.8 2003.4 1921 2004.5 C
+f
+S
+n
+1925.8 2001.6 m
+1927.3 2000.7 1930.8 2002.9 1931.3 2000.4 C
+1930.9 2000.1 1931.1 1998.5 1931.1 1999 C
+1928.7 1999.8 1924.8 1998.1 1925.8 2001.6 C
+f
+S
+n
+1938.8 2001.4 m
+1939 2000.9 1938.2 2000.3 1937.8 1999.7 C
+1937.1 1999.4 1936.2 2000.3 1936.6 2000.7 C
+1936.7 2001.5 1938.1 2000.9 1938.8 2001.4 C
+f
+S
+n
+1908.6691 2008.1348 m
+1897.82 2010.0477 L
+1894.1735 1989.3671 L
+1905.0226 1987.4542 L
+1908.6691 2008.1348 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont
+1895.041763 1994.291153 m
+0 0 32 0 0 (l) ts
+}
+true
+[0 0 0 1]sts
+Q
+1979.2185 1991.7809 m
+1960.6353 1998.5452 L
+1953.4532 1978.8124 L
+1972.0363 1972.0481 L
+1979.2185 1991.7809 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [18.793335 -6.84082 6.84021 18.793335 0 0] makesetfont
+1955.163254 1983.510773 m
+0 0 32 0 0 (\256) ts
+}
+true
+[0 0 0 1]sts
+Q
+1952.1544 2066.5423 m
+1938.0739 2069.025 L
+1934.4274 2048.3444 L
+1948.5079 2045.8617 L
+1952.1544 2066.5423 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont
+1935.29567 2053.268433 m
+0 0 32 0 0 (") ts
+}
+true
+[0 0 0 1]sts
+Q
+1931.7231 2043.621 m
+1919.3084 2048.14 L
+1910.6898 2024.4607 L
+1923.1046 2019.9417 L
+1931.7231 2043.621 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [22.552002 -8.208984 8.208252 22.552002 0 0] makesetfont
+1912.741867 2030.098648 m
+0 0 32 0 0 (=) ts
+}
+true
+[0 0 0 1]sts
+Q
+1944 2024.5 m
+1944 2014 L
+0.8504 w
+0 J
+3.863693 M
+[0 0 0 1] vc
+false setoverprint
+S
+n
+1944.25 2019.1673 m
+1952.5 2015.9173 L
+S
+n
+1931.0787 2124.423 m
+1855.5505 2043.4285 L
+1871.0419 2013.0337 L
+1946.5701 2094.0282 L
+1931.0787 2124.423 L
+n
+q
+_bfh
+%%IncludeResource: font ZapfHumanist601BT-Bold
+_efh
+{
+f1 [22.155762 23.759277 -14.753906 28.947754 0 0] makesetfont
+1867.35347 2020.27063 m
+0 0 32 0 0 (Isabelle) ts
+}
+true
+[0 0 0 1]sts
+Q
+1933.5503 1996.9547 m
+1922.7012 1998.8677 L
+1919.0547 1978.1871 L
+1929.9038 1976.2741 L
+1933.5503 1996.9547 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont
+1919.922913 1983.111069 m
+0 0 32 0 0 (b) ts
+}
+true
+[0 0 0 1]sts
+Q
+2006.3221 2025.7184 m
+1993.8573 2027.9162 L
+1990.2108 2007.2356 L
+2002.6756 2005.0378 L
+2006.3221 2025.7184 L
+n
+q
+_bfh
+%%IncludeResource: font Symbol
+_efh
+{
+f0 [19.696045 -3.4729 3.4729 19.696045 0 0] makesetfont
+1991.07901 2012.159653 m
+0 0 32 0 0 (a) ts
+}
+true
+[0 0 0 1]sts
+Q
+vmrs
+2030.0624 2094.056 m
+1956.3187 2120.904 L
+1956.321 2095.3175 L
+2030.0647 2068.4695 L
+2030.0624 2094.056 L
+n
+q
+_bfh
+%%IncludeResource: font ZapfHumanist601BT-Bold
+_efh
+{
+f1 [22.898804 -8.336792 -0.002197 24.368408 0 0] makesetfont
+1956.320496 2101.409561 m
+0 0 32 0 0 ( S/H) ts
+}
+true
+[0 0 0 1]sts
+Q
+vmr
+vmr
+end
+%%Trailer
+%%DocumentNeededResources: font Symbol
+%%+ font ZapfHumanist601BT-Bold
+%%DocumentFonts: Symbol
+%%+ ZapfHumanist601BT-Bold
+%%DocumentNeededFonts: Symbol
+%%+ ZapfHumanist601BT-Bold
Binary file doc-src/gfx/isabelle_sledgehammer.pdf has changed
--- a/doc-src/manual.bib Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/manual.bib Sun May 23 10:38:11 2010 +0100
@@ -228,13 +228,19 @@
title="Introduction to Functional Programming using Haskell",
publisher=PH,year=1998}
-@inproceedings{blanchette-nipkow-2009,
- title = "Nitpick: A Counterexample Generator for Higher-Order Logic Based on a Relational Model Finder (Extended Abstract)",
+@inproceedings{blanchette-nipkow-2010,
+ title = "Nitpick: A Counterexample Generator for Higher-Order Logic Based on a Relational Model Finder",
author = "Jasmin Christian Blanchette and Tobias Nipkow",
- booktitle = "{TAP} 2009: Short Papers",
- editor = "Catherine Dubois",
- publisher = "ETH Technical Report 630",
- year = 2009}
+ crossref = {itp2010}}
+
+@inproceedings{boehme-nipkow-2010,
+ author={Sascha B\"ohme and Tobias Nipkow},
+ title={Sledgehammer: Judgement Day},
+ booktitle={Automated Reasoning: IJCAR 2010},
+ editor={J. Giesl and R. H\"ahnle},
+ publisher=Springer,
+ series=LNCS,
+ year=2010}
@Article{boyer86,
author = {Robert Boyer and Ewing Lusk and William McCune and Ross
@@ -620,6 +626,11 @@
pages = {203--220},
crossref = {tphols96}}
+@misc{metis,
+ author = "Joe Hurd",
+ title = "Metis Theorem Prover",
+ note = "\url{http://www.gilith.com/software/metis/}"}
+
%J
@article{haskell-revised-report,
@@ -1256,6 +1267,15 @@
publisher = {Addison-Wesley},
year = 1990}
+@article{riazanov-voronkov-2002,
+ author = "Alexander Riazanov and Andrei Voronkov",
+ title = "The Design and Implementation of {Vampire}",
+ journal = "Journal of AI Communications",
+ year = 2002,
+ volume = 15,
+ number ="2/3",
+ pages = "91--110"}
+
@book{Rosen-DMA,author={Kenneth H. Rosen},
title={Discrete Mathematics and Its Applications},
publisher={McGraw-Hill},year=1998}
@@ -1287,6 +1307,15 @@
number = 4
}
+@article{schulz-2002,
+ author = "Stephan Schulz",
+ title = "E---A Brainiac Theorem Prover",
+ journal = "Journal of AI Communications",
+ year = 2002,
+ volume = 15,
+ number ="2/3",
+ pages = "111--126"}
+
@misc{sledgehammer-2009,
key = "Sledgehammer",
title = "The {S}ledgehammer: Let Automatic Theorem Provers
@@ -1305,6 +1334,17 @@
year = 1972,
publisher = {Dover}}
+@inproceedings{sutcliffe-2000,
+ author = "Geoff Sutcliffe",
+ title = "System Description: {SystemOnTPTP}",
+ editor = "J. G. Carbonell and J. Siekmann",
+ booktitle = {Automated Deduction --- {CADE}-17 International Conference},
+ series = "Lecture Notes in Artificial Intelligence",
+ volume = {1831},
+ pages = "406--410",
+ year = 2000,
+ publisher = Springer}
+
@InCollection{szasz93,
author = {Nora Szasz},
title = {A Machine Checked Proof that {Ackermann's} Function is not
@@ -1411,6 +1451,11 @@
note = {\url{http://x-symbol.sourceforge.net}}
}
+@misc{weidenbach-et-al-2009,
+ author = "Christoph Weidenbach and Dilyana Dimova and Arnaud Fietzke and Rohit Kumar and Martin Suda and Patrick Wischnewski",
+ title = "{SPASS} Version 3.5",
+ note = {\url{http://www.spass-prover.org/publications/spass.pdf}}}
+
@manual{isabelle-sys,
author = {Markus Wenzel and Stefan Berghofer},
title = {The {Isabelle} System Manual},
@@ -1750,6 +1795,14 @@
% editor =
% volume = 4732,
+@Proceedings{itp2010,
+ title = {Interactive Theorem Proving: {ITP}-10},
+ booktitle = {Interactive Theorem Proving: {ITP}-10},
+ editor = "Matt Kaufmann and Lawrence Paulson",
+ publisher = Springer,
+ series = LNCS,
+ year = 2010}
+
@unpublished{classes_modules,
title = {{ML} Modules and {Haskell} Type Classes: A Constructive Comparison},
author = {Stefan Wehr et. al.}
--- a/doc-src/rail.ML Sun May 23 10:37:43 2010 +0100
+++ b/doc-src/rail.ML Sun May 23 10:38:11 2010 +0100
@@ -59,9 +59,9 @@
(* [(kind, (markup, check, style))*)
Symtab.make [
("syntax", ("", no_check, true)),
- ("command", ("isacommand", K (is_some o OuterKeyword.command_keyword), true)),
- ("keyword", ("isakeyword", K OuterKeyword.is_keyword, true)),
- ("element", ("isakeyword", K OuterKeyword.is_keyword, true)),
+ ("command", ("isacommand", K (is_some o Keyword.command_keyword), true)),
+ ("keyword", ("isakeyword", K Keyword.is_keyword, true)),
+ ("element", ("isakeyword", K Keyword.is_keyword, true)),
("method", ("", thy_check Method.intern Method.defined, true)),
("attribute", ("", thy_check Attrib.intern Attrib.defined, true)),
("fact", ("", no_check, true)),
@@ -473,7 +473,7 @@
|> parse
|> print;
-val _ = ThyOutput.antiquotation "rail" (Scan.lift ( OuterParse.position Args.name ))
+val _ = ThyOutput.antiquotation "rail" (Scan.lift (Parse.position Args.name))
(fn {context = ctxt,...} => fn txt => process txt ctxt);
end;
--- a/doc/Contents Sun May 23 10:37:43 2010 +0100
+++ b/doc/Contents Sun May 23 10:38:11 2010 +0100
@@ -6,7 +6,8 @@
classes Tutorial on Type Classes
functions Tutorial on Function Definitions
codegen Tutorial on Code Generation
- nitpick User's Guide to Nitpick in Isabelle/HOL
+ nitpick User's Guide to Nitpick
+ sledgehammer User's Guide to Sledgehammer
sugar LaTeX Sugar for Isabelle documents
Reference Manuals
--- a/lib/html/isabelle.css Sun May 23 10:37:43 2010 +0100
+++ b/lib/html/isabelle.css Sun May 23 10:38:11 2010 +0100
@@ -1,4 +1,4 @@
-/* css style file for Isabelle XHTML/XML output */
+/* style file for Isabelle XHTML/XML output */
body { background-color: #FFFFFF; }
--- a/lib/scripts/isabelle-platform Sun May 23 10:37:43 2010 +0100
+++ b/lib/scripts/isabelle-platform Sun May 23 10:38:11 2010 +0100
@@ -1,3 +1,4 @@
+# -*- shell-script -*- :mode=shellscript:
#
# determine general hardware and operating system type for Isabelle
#
--- a/src/FOL/ex/Iff_Oracle.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/FOL/ex/Iff_Oracle.thy Sun May 23 10:38:11 2010 +0100
@@ -52,7 +52,7 @@
subsection {* Oracle as proof method *}
method_setup iff = {*
- Scan.lift OuterParse.nat >> (fn n => fn ctxt =>
+ Scan.lift Parse.nat >> (fn n => fn ctxt =>
SIMPLE_METHOD
(HEADGOAL (Tactic.rtac (iff_oracle (ProofContext.theory_of ctxt, n)))
handle Fail _ => no_tac))
--- a/src/FOL/ex/LocaleTest.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/FOL/ex/LocaleTest.thy Sun May 23 10:38:11 2010 +0100
@@ -533,7 +533,7 @@
grefl: "gle(x, x)" gless_def: "gless(x, y) <-> gle(x, y) & x ~= y"
grefl': "gle'(x, x)" gless'_def: "gless'(x, y) <-> gle'(x, y) & x ~= y"
-text {* Setup *}
+text {* Mixin not applied to locales further up the hierachy. *}
locale mixin = reflexive
begin
@@ -548,6 +548,16 @@
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
+thm le.less_def (* mixin not applied *)
+lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y" by (rule le.less_def)
+thm le.less_thm (* mixin applied *)
+lemma "gless(x, y) <-> gle(x, y) & x ~= y" by (rule le.less_thm)
+
+lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y"
+ by (rule le.less_def)
+lemma "gless(x, y) <-> gle(x, y) & x ~= y"
+ by (rule le.less_thm)
+
text {* Mixin propagated along the locale hierarchy *}
locale mixin2 = mixin
--- a/src/FOLP/simp.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/FOLP/simp.ML Sun May 23 10:38:11 2010 +0100
@@ -431,7 +431,7 @@
are represented by rules, generalized over their parameters*)
fun add_asms(ss,thm,a,anet,ats,cs) =
let val As = strip_varify(nth_subgoal i thm, a, []);
- val thms = map (trivial o cterm_of(Thm.theory_of_thm thm)) As;
+ val thms = map (Thm.trivial o cterm_of(Thm.theory_of_thm thm)) As;
val new_rws = maps mk_rew_rules thms;
val rwrls = map mk_trans (maps mk_rew_rules thms);
val anet' = fold_rev lhs_insert_thm rwrls anet;
--- a/src/HOL/Big_Operators.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Big_Operators.thy Sun May 23 10:38:11 2010 +0100
@@ -673,7 +673,7 @@
proof induct
case empty thus ?case by simp
next
- case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg)
+ case (insert x A) thus ?case by auto
qed
next
case False thus ?thesis by (simp add: setsum_def)
--- a/src/HOL/Boogie/Tools/boogie_commands.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Sun May 23 10:38:11 2010 +0100
@@ -263,31 +263,31 @@
fun scan_arg f = Args.parens f
fun scan_opt n = Scan.optional (scan_arg (Args.$$$ n >> K true)) false
-val vc_offsets = Scan.optional (Args.bracks (OuterParse.list1
- (OuterParse.string --| Args.colon -- OuterParse.nat))) []
+val vc_offsets = Scan.optional (Args.bracks (Parse.list1
+ (Parse.string --| Args.colon -- Parse.nat))) []
val _ =
- OuterSyntax.command "boogie_open"
+ Outer_Syntax.command "boogie_open"
"Open a new Boogie environment and load a Boogie-generated .b2i file."
- OuterKeyword.thy_decl
- (scan_opt "quiet" -- OuterParse.name -- vc_offsets >>
+ Keyword.thy_decl
+ (scan_opt "quiet" -- Parse.name -- vc_offsets >>
(Toplevel.theory o boogie_open))
-val vc_name = OuterParse.name
+val vc_name = Parse.name
-val vc_label = OuterParse.name
+val vc_label = Parse.name
val vc_labels = Scan.repeat1 vc_label
val vc_paths =
- OuterParse.nat -- (Args.$$$ "-" |-- OuterParse.nat) >> (op upto) ||
- OuterParse.nat >> single
+ Parse.nat -- (Args.$$$ "-" |-- Parse.nat) >> (op upto) ||
+ Parse.nat >> single
val vc_opts =
scan_arg
(scan_val "assertion" vc_label >> VC_Single ||
scan_val "examine" vc_labels >> VC_Examine ||
- scan_val "take" ((OuterParse.list vc_paths >> flat) -- Scan.option (
+ scan_val "take" ((Parse.list vc_paths >> flat) -- Scan.option (
scan_val "without" vc_labels >> pair false ||
scan_val "and_also" vc_labels >> pair true) >> VC_Take) ||
scan_val "only" vc_labels >> VC_Only ||
@@ -295,9 +295,9 @@
Scan.succeed VC_Complete
val _ =
- OuterSyntax.command "boogie_vc"
+ Outer_Syntax.command "boogie_vc"
"Enter into proof mode for a specific verification condition."
- OuterKeyword.thy_goal
+ Keyword.thy_goal
(vc_name -- vc_opts >> (fn args =>
(Toplevel.print o Toplevel.theory_to_proof (boogie_vc args))))
@@ -305,7 +305,7 @@
val quick_timeout = 5
val default_timeout = 20
-fun timeout name = Scan.optional (scan_val name OuterParse.nat)
+fun timeout name = Scan.optional (scan_val name Parse.nat)
val status_test =
scan_arg (
@@ -328,18 +328,18 @@
f (Toplevel.theory_of state))
val _ =
- OuterSyntax.improper_command "boogie_status"
+ Outer_Syntax.improper_command "boogie_status"
"Show the name and state of all loaded verification conditions."
- OuterKeyword.diag
+ Keyword.diag
(status_test >> status_cmd ||
status_vc >> status_cmd ||
Scan.succeed (status_cmd boogie_status))
val _ =
- OuterSyntax.command "boogie_end"
+ Outer_Syntax.command "boogie_end"
"Close the current Boogie environment."
- OuterKeyword.thy_decl
+ Keyword.thy_decl
(Scan.succeed (Toplevel.theory boogie_end))
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML Sun May 23 10:38:11 2010 +0100
@@ -25,8 +25,8 @@
val label_eqs = [assert_at_def, block_at_def]
fun unfold_labels_tac ctxt =
- let val unfold = More_Conv.rewrs_conv label_eqs
- in CONVERSION (More_Conv.top_sweep_conv (K unfold) ctxt) end
+ let val unfold = Conv.rewrs_conv label_eqs
+ in CONVERSION (Conv.top_sweep_conv (K unfold) ctxt) end
fun boogie_tac ctxt rules =
unfold_labels_tac ctxt
--- a/src/HOL/Decision_Procs/Approximation.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy Sun May 23 10:38:11 2010 +0100
@@ -3209,12 +3209,96 @@
interpret_floatarith_divide interpret_floatarith_diff interpret_floatarith_tan interpret_floatarith_powr interpret_floatarith_log
interpret_floatarith_sin
-code_reflect Float_Arith
- datatypes float = Float
- and floatarith = Add | Minus | Mult | Inverse | Cos | Arctan
- | Abs | Max | Min | Pi | Sqrt | Exp | Ln | Power | Var | Num
- and form = Bound | Assign | Less | LessEqual | AtLeastAtMost
- functions approx_form approx_tse_form approx' approx_form_eval
+oracle approximation_oracle = {* fn (thy, t) =>
+let
+
+ fun bad t = error ("Bad term: " ^ Syntax.string_of_term_global thy t);
+
+ fun term_of_bool True = @{term True}
+ | term_of_bool False = @{term False};
+
+ fun term_of_float (@{code Float} (k, l)) =
+ @{term Float} $ HOLogic.mk_number @{typ int} k $ HOLogic.mk_number @{typ int} l;
+
+ fun term_of_float_float_option NONE = @{term "None :: (float \<times> float) option"}
+ | term_of_float_float_option (SOME ff) = @{term "Some :: float \<times> float \<Rightarrow> _"}
+ $ HOLogic.mk_prod (pairself term_of_float ff);
+
+ val term_of_float_float_option_list =
+ HOLogic.mk_list @{typ "(float \<times> float) option"} o map term_of_float_float_option;
+
+ fun nat_of_term t = HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t);
+
+ fun float_of_term (@{term Float} $ k $ l) =
+ @{code Float} (snd (HOLogic.dest_number k), snd (HOLogic.dest_number l))
+ | float_of_term t = bad t;
+
+ fun floatarith_of_term (@{term Add} $ a $ b) = @{code Add} (floatarith_of_term a, floatarith_of_term b)
+ | floatarith_of_term (@{term Minus} $ a) = @{code Minus} (floatarith_of_term a)
+ | floatarith_of_term (@{term Mult} $ a $ b) = @{code Mult} (floatarith_of_term a, floatarith_of_term b)
+ | floatarith_of_term (@{term Inverse} $ a) = @{code Inverse} (floatarith_of_term a)
+ | floatarith_of_term (@{term Cos} $ a) = @{code Cos} (floatarith_of_term a)
+ | floatarith_of_term (@{term Arctan} $ a) = @{code Arctan} (floatarith_of_term a)
+ | floatarith_of_term (@{term Abs} $ a) = @{code Abs} (floatarith_of_term a)
+ | floatarith_of_term (@{term Max} $ a $ b) = @{code Max} (floatarith_of_term a, floatarith_of_term b)
+ | floatarith_of_term (@{term Min} $ a $ b) = @{code Min} (floatarith_of_term a, floatarith_of_term b)
+ | floatarith_of_term @{term Pi} = @{code Pi}
+ | floatarith_of_term (@{term Sqrt} $ a) = @{code Sqrt} (floatarith_of_term a)
+ | floatarith_of_term (@{term Exp} $ a) = @{code Exp} (floatarith_of_term a)
+ | floatarith_of_term (@{term Ln} $ a) = @{code Ln} (floatarith_of_term a)
+ | floatarith_of_term (@{term Power} $ a $ n) =
+ @{code Power} (floatarith_of_term a, nat_of_term n)
+ | floatarith_of_term (@{term Var} $ n) = @{code Var} (nat_of_term n)
+ | floatarith_of_term (@{term Num} $ m) = @{code Num} (float_of_term m)
+ | floatarith_of_term t = bad t;
+
+ fun form_of_term (@{term Bound} $ a $ b $ c $ p) = @{code Bound}
+ (floatarith_of_term a, floatarith_of_term b, floatarith_of_term c, form_of_term p)
+ | form_of_term (@{term Assign} $ a $ b $ p) = @{code Assign}
+ (floatarith_of_term a, floatarith_of_term b, form_of_term p)
+ | form_of_term (@{term Less} $ a $ b) = @{code Less}
+ (floatarith_of_term a, floatarith_of_term b)
+ | form_of_term (@{term LessEqual} $ a $ b) = @{code LessEqual}
+ (floatarith_of_term a, floatarith_of_term b)
+ | form_of_term (@{term AtLeastAtMost} $ a $ b $ c) = @{code AtLeastAtMost}
+ (floatarith_of_term a, floatarith_of_term b, floatarith_of_term c)
+ | form_of_term t = bad t;
+
+ fun float_float_option_of_term @{term "None :: (float \<times> float) option"} = NONE
+ | float_float_option_of_term (@{term "Some :: float \<times> float \<Rightarrow> _"} $ ff) =
+ SOME (pairself float_of_term (HOLogic.dest_prod ff))
+ | float_float_option_of_term (@{term approx'} $ n $ a $ ffs) = @{code approx'}
+ (nat_of_term n) (floatarith_of_term a) (float_float_option_list_of_term ffs)
+ | float_float_option_of_term t = bad t
+ and float_float_option_list_of_term
+ (@{term "replicate :: _ \<Rightarrow> (float \<times> float) option \<Rightarrow> _"} $ n $ @{term "None :: (float \<times> float) option"}) =
+ @{code replicate} (nat_of_term n) NONE
+ | float_float_option_list_of_term (@{term approx_form_eval} $ n $ p $ ffs) =
+ @{code approx_form_eval} (nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs)
+ | float_float_option_list_of_term t = map float_float_option_of_term
+ (HOLogic.dest_list t);
+
+ val nat_list_of_term = map nat_of_term o HOLogic.dest_list ;
+
+ fun bool_of_term (@{term approx_form} $ n $ p $ ffs $ ms) = @{code approx_form}
+ (nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs) (nat_list_of_term ms)
+ | bool_of_term (@{term approx_tse_form} $ m $ n $ q $ p) =
+ @{code approx_tse_form} (nat_of_term m) (nat_of_term n) (nat_of_term q) (form_of_term p)
+ | bool_of_term t = bad t;
+
+ fun eval t = case fastype_of t
+ of @{typ bool} =>
+ (term_of_bool o bool_of_term) t
+ | @{typ "(float \<times> float) option"} =>
+ (term_of_float_float_option o float_float_option_of_term) t
+ | @{typ "(float \<times> float) option list"} =>
+ (term_of_float_float_option_list o float_float_option_list_of_term) t
+ | _ => bad t;
+
+ val normalize = eval o Envir.beta_norm o Pattern.eta_long [];
+
+in Thm.cterm_of thy (Logic.mk_equals (t, normalize t)) end
+*}
ML {*
fun reorder_bounds_tac prems i =
@@ -3246,9 +3330,17 @@
fold prepend_prem order all_tac
end
+ fun approximation_conv ctxt ct =
+ approximation_oracle (ProofContext.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt));
+
+ fun approximate ctxt t =
+ approximation_oracle (ProofContext.theory_of ctxt, t)
+ |> Thm.prop_of |> Logic.dest_equals |> snd;
+
(* Should be in HOL.thy ? *)
- fun gen_eval_tac conv ctxt = CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)
- THEN' rtac TrueI
+ fun gen_eval_tac conv ctxt = CONVERSION
+ (Object_Logic.judgment_conv (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
+ THEN' rtac TrueI
val form_equations = PureThy.get_thms @{theory} "interpret_form_equations";
@@ -3310,13 +3402,13 @@
by auto
method_setup approximation = {*
- Scan.lift (OuterParse.nat)
+ Scan.lift Parse.nat
--
Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon)
- |-- OuterParse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift OuterParse.nat)) []
+ |-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) []
--
Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon)
- |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift OuterParse.nat))
+ |-- (free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat))
>>
(fn ((prec, splitting), taylor) => fn ctxt =>
SIMPLE_METHOD' (fn i =>
@@ -3327,7 +3419,7 @@
THEN DETERM (TRY (filter_prems_tac (K false) i))
THEN DETERM (Reflection.genreify_tac ctxt form_equations NONE i)
THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
- THEN gen_eval_tac eval_oracle ctxt i))
+ THEN gen_eval_tac (approximation_conv ctxt) ctxt i))
*} "real number approximation"
ML {*
@@ -3435,7 +3527,7 @@
|> (fn (data, xs) =>
mk_approx_form_eval prec data (HOLogic.mk_list @{typ "(float * float) option"}
(map (fn _ => @{term "None :: (float * float) option"}) (HOLogic.dest_list xs)))
- |> Codegen.eval_term @{theory}
+ |> approximate ctxt
|> HOLogic.dest_list
|> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term)
|> map (fn (elem, s) => @{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ elem $ mk_result prec (dest_ivl s))
@@ -3448,7 +3540,7 @@
|> HOLogic.dest_eq |> snd
|> dest_interpret |> fst
|> mk_approx' prec
- |> Codegen.eval_term @{theory}
+ |> approximate ctxt
|> dest_ivl
|> mk_result prec
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Sun May 23 10:38:11 2010 +0100
@@ -681,7 +681,7 @@
else ("Nox",[])
| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
-fun xnormalize_conv ctxt [] ct = reflexive ct
+fun xnormalize_conv ctxt [] ct = Thm.reflexive ct
| xnormalize_conv ctxt (vs as (x::_)) ct =
case term_of ct of
Const(@{const_name Orderings.less},_)$_$Const(@{const_name Groups.zero},_) =>
@@ -696,8 +696,8 @@
(Thm.capply @{cterm "Trueprop"}
(if neg then Thm.capply (Thm.capply clt c) cz
else Thm.capply (Thm.capply clt cz) c))
- val cth = equal_elim (symmetric cthp) TrueI
- val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
(if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
@@ -719,12 +719,12 @@
(Thm.capply @{cterm "Trueprop"}
(if neg then Thm.capply (Thm.capply clt c) cz
else Thm.capply (Thm.capply clt cz) c))
- val cth = equal_elim (symmetric cthp) TrueI
- val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
(if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
val rth = th
in rth end
- | _ => reflexive ct)
+ | _ => Thm.reflexive ct)
| Const(@{const_name Orderings.less_eq},_)$_$Const(@{const_name Groups.zero},_) =>
@@ -740,8 +740,8 @@
(Thm.capply @{cterm "Trueprop"}
(if neg then Thm.capply (Thm.capply clt c) cz
else Thm.capply (Thm.capply clt cz) c))
- val cth = equal_elim (symmetric cthp) TrueI
- val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val th = Thm.implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
(if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
@@ -764,12 +764,12 @@
(Thm.capply @{cterm "Trueprop"}
(if neg then Thm.capply (Thm.capply clt c) cz
else Thm.capply (Thm.capply clt cz) c))
- val cth = equal_elim (symmetric cthp) TrueI
- val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
(if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
val rth = th
in rth end
- | _ => reflexive ct)
+ | _ => Thm.reflexive ct)
| Const("op =",_)$_$Const(@{const_name Groups.zero},_) =>
(case whatis x (Thm.dest_arg1 ct) of
@@ -782,8 +782,8 @@
val cthp = Simplifier.rewrite (simpset_of ctxt)
(Thm.capply @{cterm "Trueprop"}
(Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
- val cth = equal_elim (symmetric cthp) TrueI
- val th = implies_elim
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val th = Thm.implies_elim
(instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
@@ -804,11 +804,11 @@
val cthp = Simplifier.rewrite (simpset_of ctxt)
(Thm.capply @{cterm "Trueprop"}
(Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
- val cth = equal_elim (symmetric cthp) TrueI
- val rth = implies_elim
+ val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
+ val rth = Thm.implies_elim
(instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
in rth end
- | _ => reflexive ct);
+ | _ => Thm.reflexive ct);
local
val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
@@ -823,7 +823,7 @@
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
(Semiring_Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
- val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+ val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
in rth end
| Const(@{const_name Orderings.less_eq},_)$a$b =>
let val (ca,cb) = Thm.dest_binop ct
@@ -832,7 +832,7 @@
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
(Semiring_Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
- val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+ val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
in rth end
| Const("op =",_)$a$b =>
@@ -842,10 +842,10 @@
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
(Semiring_Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
- val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+ val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
in rth end
| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
-| _ => reflexive ct
+| _ => Thm.reflexive ct
end;
fun classfield_whatis phi =
--- a/src/HOL/Decision_Procs/cooper_tac.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Decision_Procs/cooper_tac.ML Sun May 23 10:38:11 2010 +0100
@@ -106,7 +106,7 @@
[simp_tac mod_div_simpset 1, simp_tac simpset0 1,
TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1),
TRY (simp_tac simpset3 1), TRY (simp_tac cooper_ss 1)]
- (trivial ct))
+ (Thm.trivial ct))
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
(* The result of the quantifier elimination *)
val (th, tac) = case (prop_of pre_thm) of
--- a/src/HOL/Decision_Procs/ferrack_tac.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML Sun May 23 10:38:11 2010 +0100
@@ -87,7 +87,7 @@
val pre_thm = Seq.hd (EVERY
[simp_tac simpset0 1,
TRY (simp_tac (Simplifier.context ctxt ferrack_ss) 1)]
- (trivial ct))
+ (Thm.trivial ct))
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
(* The result of the quantifier elimination *)
val (th, tac) = case prop_of pre_thm of
--- a/src/HOL/Decision_Procs/mir_tac.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Sun May 23 10:38:11 2010 +0100
@@ -128,7 +128,7 @@
val pre_thm = Seq.hd (EVERY
[simp_tac mod_div_simpset 1, simp_tac simpset0 1,
TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)]
- (trivial ct))
+ (Thm.trivial ct))
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)
(* The result of the quantifier elimination *)
val (th, tac) = case (prop_of pre_thm) of
--- a/src/HOL/Groups.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Groups.thy Sun May 23 10:38:11 2010 +0100
@@ -605,7 +605,7 @@
then show ?thesis by simp
qed
-lemma add_nonneg_nonneg:
+lemma add_nonneg_nonneg [simp]:
assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"
proof -
have "0 + 0 \<le> a + b"
--- a/src/HOL/HOL.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/HOL.thy Sun May 23 10:38:11 2010 +0100
@@ -29,7 +29,6 @@
"~~/src/Tools/induct.ML"
("~~/src/Tools/induct_tacs.ML")
("Tools/recfun_codegen.ML")
- "~~/src/Tools/more_conv.ML"
begin
setup {* Intuitionistic.method_setup @{binding iprover} *}
@@ -1283,12 +1282,12 @@
else let
val abs_g'= Abs (n,xT,g');
val g'x = abs_g'$x;
- val g_g'x = symmetric (beta_conversion false (cterm_of thy g'x));
+ val g_g'x = Thm.symmetric (Thm.beta_conversion false (cterm_of thy g'x));
val rl = cterm_instantiate
[(f_Let_folded, cterm_of thy f), (x_Let_folded, cx),
(g_Let_folded, cterm_of thy abs_g')]
@{thm Let_folded};
- in SOME (rl OF [transitive fx_g g_g'x])
+ in SOME (rl OF [Thm.transitive fx_g g_g'x])
end)
end
| _ => NONE)
--- a/src/HOL/IMP/Natural.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/IMP/Natural.thy Sun May 23 10:38:11 2010 +0100
@@ -25,6 +25,16 @@
notation (xsymbols)
update ("_/[_ \<mapsto> /_]" [900,0,0] 900)
+text {* Disable conflicting syntax from HOL Map theory. *}
+
+no_syntax
+ "_maplet" :: "['a, 'a] => maplet" ("_ /|->/ _")
+ "_maplets" :: "['a, 'a] => maplet" ("_ /[|->]/ _")
+ "" :: "maplet => maplets" ("_")
+ "_Maplets" :: "[maplet, maplets] => maplets" ("_,/ _")
+ "_MapUpd" :: "['a ~=> 'b, maplets] => 'a ~=> 'b" ("_/'(_')" [900,0]900)
+ "_Map" :: "maplets => 'a ~=> 'b" ("(1[_])")
+
text {*
The big-step execution relation @{text evalc} is defined inductively:
*}
@@ -111,7 +121,7 @@
in the same @{text s'}}. Formally:
*}
definition
- equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" ("_ \<sim> _") where
+ equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" ("_ \<sim> _" [56, 56] 55) where
"c \<sim> c' = (\<forall>s s'. \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s')"
text {*
--- a/src/HOL/Import/import.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Import/import.ML Sun May 23 10:38:11 2010 +0100
@@ -40,7 +40,7 @@
val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
val thm = snd (ProofKernel.to_isa_thm hol4thm)
val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
- val thm = equal_elim rew thm
+ val thm = Thm.equal_elim rew thm
val prew = ProofKernel.rewrite_hol4_term prem thy
val prem' = #2 (Logic.dest_equals (prop_of prew))
val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm)
@@ -53,7 +53,7 @@
val _ = if prem' aconv (prop_of thm)
then message "import: Terms match up"
else message "import: Terms DO NOT match up"
- val thm' = equal_elim (symmetric prew) thm
+ val thm' = Thm.equal_elim (Thm.symmetric prew) thm
val res = Thm.bicompose true (false,thm',0) 1 th
in
res
--- a/src/HOL/Import/import_syntax.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Import/import_syntax.ML Sun May 23 10:38:11 2010 +0100
@@ -4,26 +4,23 @@
signature HOL4_IMPORT_SYNTAX =
sig
- type token = OuterLex.token
- type command = token list -> (theory -> theory) * token list
-
- val import_segment: token list -> (theory -> theory) * token list
- val import_theory : token list -> (theory -> theory) * token list
- val end_import : token list -> (theory -> theory) * token list
-
- val setup_theory : token list -> (theory -> theory) * token list
- val end_setup : token list -> (theory -> theory) * token list
+ val import_segment: (theory -> theory) parser
+ val import_theory : (theory -> theory) parser
+ val end_import : (theory -> theory) parser
+
+ val setup_theory : (theory -> theory) parser
+ val end_setup : (theory -> theory) parser
+
+ val thm_maps : (theory -> theory) parser
+ val ignore_thms : (theory -> theory) parser
+ val type_maps : (theory -> theory) parser
+ val def_maps : (theory -> theory) parser
+ val const_maps : (theory -> theory) parser
+ val const_moves : (theory -> theory) parser
+ val const_renames : (theory -> theory) parser
- val thm_maps : token list -> (theory -> theory) * token list
- val ignore_thms : token list -> (theory -> theory) * token list
- val type_maps : token list -> (theory -> theory) * token list
- val def_maps : token list -> (theory -> theory) * token list
- val const_maps : token list -> (theory -> theory) * token list
- val const_moves : token list -> (theory -> theory) * token list
- val const_renames : token list -> (theory -> theory) * token list
-
- val skip_import_dir : token list -> (theory -> theory) * token list
- val skip_import : token list -> (theory -> theory) * token list
+ val skip_import_dir : (theory -> theory) parser
+ val skip_import : (theory -> theory) parser
val setup : unit -> unit
end
@@ -31,28 +28,23 @@
structure HOL4ImportSyntax :> HOL4_IMPORT_SYNTAX =
struct
-type token = OuterLex.token
-type command = token list -> (theory -> theory) * token list
-
-local structure P = OuterParse and K = OuterKeyword in
-
(* Parsers *)
-val import_segment = P.name >> set_import_segment
+val import_segment = Parse.name >> set_import_segment
-val import_theory = P.name >> (fn thyname =>
+val import_theory = Parse.name >> (fn thyname =>
fn thy =>
thy |> set_generating_thy thyname
|> Sign.add_path thyname
|> add_dump (";setup_theory " ^ thyname))
fun do_skip_import_dir s = (ImportRecorder.set_skip_import_dir (SOME s); I)
-val skip_import_dir : command = P.string >> do_skip_import_dir
+val skip_import_dir = Parse.string >> do_skip_import_dir
val lower = String.map Char.toLower
fun do_skip_import s = (ImportRecorder.set_skip_import (case lower s of "on" => true | "off" => false | _ => Scan.fail ()); I)
-val skip_import : command = P.short_ident >> do_skip_import
+val skip_import = Parse.short_ident >> do_skip_import
fun end_import toks =
Scan.succeed
@@ -84,7 +76,7 @@
|> add_dump ";end_setup"
end) toks
-val ignore_thms = Scan.repeat1 P.name
+val ignore_thms = Scan.repeat1 Parse.name
>> (fn ignored =>
fn thy =>
let
@@ -93,7 +85,7 @@
Library.foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored)
end)
-val thm_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+val thm_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
>> (fn thmmaps =>
fn thy =>
let
@@ -102,7 +94,7 @@
Library.foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps)
end)
-val type_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+val type_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
>> (fn typmaps =>
fn thy =>
let
@@ -111,7 +103,7 @@
Library.foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps)
end)
-val def_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+val def_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
>> (fn defmaps =>
fn thy =>
let
@@ -120,7 +112,7 @@
Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
end)
-val const_renames = Scan.repeat1 (P.name --| P.$$$ ">" -- P.name)
+val const_renames = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.name)
>> (fn renames =>
fn thy =>
let
@@ -129,7 +121,7 @@
Library.foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames)
end)
-val const_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+val const_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
>> (fn constmaps =>
fn thy =>
let
@@ -139,7 +131,7 @@
| (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
end)
-val const_moves = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+val const_moves = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
>> (fn constmaps =>
fn thy =>
let
@@ -160,18 +152,18 @@
(Scan.make_lexicon (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
Scan.empty_lexicon)
val get_lexes = fn () => !lexes
- val token_source = OuterLex.source {do_recover = NONE} get_lexes Position.start symb_source
- val token_list = filter_out (OuterLex.is_kind OuterLex.Space) (Source.exhaust token_source)
- val import_segmentP = OuterParse.$$$ "import_segment" |-- import_segment
- val type_mapsP = OuterParse.$$$ "type_maps" |-- type_maps
- val def_mapsP = OuterParse.$$$ "def_maps" |-- def_maps
- val const_mapsP = OuterParse.$$$ "const_maps" |-- const_maps
- val const_movesP = OuterParse.$$$ "const_moves" |-- const_moves
- val const_renamesP = OuterParse.$$$ "const_renames" |-- const_renames
- val ignore_thmsP = OuterParse.$$$ "ignore_thms" |-- ignore_thms
- val thm_mapsP = OuterParse.$$$ "thm_maps" |-- thm_maps
+ val token_source = Token.source {do_recover = NONE} get_lexes Position.start symb_source
+ val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
+ val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
+ val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
+ val def_mapsP = Parse.$$$ "def_maps" |-- def_maps
+ val const_mapsP = Parse.$$$ "const_maps" |-- const_maps
+ val const_movesP = Parse.$$$ "const_moves" |-- const_moves
+ val const_renamesP = Parse.$$$ "const_renames" |-- const_renames
+ val ignore_thmsP = Parse.$$$ "ignore_thms" |-- ignore_thms
+ val thm_mapsP = Parse.$$$ "thm_maps" |-- thm_maps
val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
- val importP = OuterParse.$$$ "import" |-- Scan.repeat parser --| OuterParse.$$$ "end"
+ val importP = Parse.$$$ "import" |-- Scan.repeat parser --| Parse.$$$ "end"
fun apply [] thy = thy
| apply (f::fs) thy = apply fs (f thy)
in
@@ -194,7 +186,7 @@
NONE => error "Import data already cleared - forgotten a setup_theory?"
| SOME _ => ImportData.put NONE thy
-val setup_theory = P.name
+val setup_theory = Parse.name
>>
(fn thyname =>
fn thy =>
@@ -218,64 +210,62 @@
|> Sign.parent_path
end) toks
-val set_dump = P.string -- P.string >> setup_dump
+val set_dump = Parse.string -- Parse.string >> setup_dump
fun fl_dump toks = Scan.succeed flush_dump toks
-val append_dump = (P.verbatim || P.string) >> add_dump
+val append_dump = (Parse.verbatim || Parse.string) >> add_dump
fun setup () =
- (OuterKeyword.keyword ">";
- OuterSyntax.command "import_segment"
+ (Keyword.keyword ">";
+ Outer_Syntax.command "import_segment"
"Set import segment name"
- K.thy_decl (import_segment >> Toplevel.theory);
- OuterSyntax.command "import_theory"
+ Keyword.thy_decl (import_segment >> Toplevel.theory);
+ Outer_Syntax.command "import_theory"
"Set default HOL4 theory name"
- K.thy_decl (import_theory >> Toplevel.theory);
- OuterSyntax.command "end_import"
+ Keyword.thy_decl (import_theory >> Toplevel.theory);
+ Outer_Syntax.command "end_import"
"End HOL4 import"
- K.thy_decl (end_import >> Toplevel.theory);
- OuterSyntax.command "skip_import_dir"
+ Keyword.thy_decl (end_import >> Toplevel.theory);
+ Outer_Syntax.command "skip_import_dir"
"Sets caching directory for skipping importing"
- K.thy_decl (skip_import_dir >> Toplevel.theory);
- OuterSyntax.command "skip_import"
+ Keyword.thy_decl (skip_import_dir >> Toplevel.theory);
+ Outer_Syntax.command "skip_import"
"Switches skipping importing on or off"
- K.thy_decl (skip_import >> Toplevel.theory);
- OuterSyntax.command "setup_theory"
+ Keyword.thy_decl (skip_import >> Toplevel.theory);
+ Outer_Syntax.command "setup_theory"
"Setup HOL4 theory replaying"
- K.thy_decl (setup_theory >> Toplevel.theory);
- OuterSyntax.command "end_setup"
+ Keyword.thy_decl (setup_theory >> Toplevel.theory);
+ Outer_Syntax.command "end_setup"
"End HOL4 setup"
- K.thy_decl (end_setup >> Toplevel.theory);
- OuterSyntax.command "setup_dump"
+ Keyword.thy_decl (end_setup >> Toplevel.theory);
+ Outer_Syntax.command "setup_dump"
"Setup the dump file name"
- K.thy_decl (set_dump >> Toplevel.theory);
- OuterSyntax.command "append_dump"
+ Keyword.thy_decl (set_dump >> Toplevel.theory);
+ Outer_Syntax.command "append_dump"
"Add text to dump file"
- K.thy_decl (append_dump >> Toplevel.theory);
- OuterSyntax.command "flush_dump"
+ Keyword.thy_decl (append_dump >> Toplevel.theory);
+ Outer_Syntax.command "flush_dump"
"Write the dump to file"
- K.thy_decl (fl_dump >> Toplevel.theory);
- OuterSyntax.command "ignore_thms"
+ Keyword.thy_decl (fl_dump >> Toplevel.theory);
+ Outer_Syntax.command "ignore_thms"
"Theorems to ignore in next HOL4 theory import"
- K.thy_decl (ignore_thms >> Toplevel.theory);
- OuterSyntax.command "type_maps"
+ Keyword.thy_decl (ignore_thms >> Toplevel.theory);
+ Outer_Syntax.command "type_maps"
"Map HOL4 type names to existing Isabelle/HOL type names"
- K.thy_decl (type_maps >> Toplevel.theory);
- OuterSyntax.command "def_maps"
+ Keyword.thy_decl (type_maps >> Toplevel.theory);
+ Outer_Syntax.command "def_maps"
"Map HOL4 constant names to their primitive definitions"
- K.thy_decl (def_maps >> Toplevel.theory);
- OuterSyntax.command "thm_maps"
+ Keyword.thy_decl (def_maps >> Toplevel.theory);
+ Outer_Syntax.command "thm_maps"
"Map HOL4 theorem names to existing Isabelle/HOL theorem names"
- K.thy_decl (thm_maps >> Toplevel.theory);
- OuterSyntax.command "const_renames"
+ Keyword.thy_decl (thm_maps >> Toplevel.theory);
+ Outer_Syntax.command "const_renames"
"Rename HOL4 const names"
- K.thy_decl (const_renames >> Toplevel.theory);
- OuterSyntax.command "const_moves"
+ Keyword.thy_decl (const_renames >> Toplevel.theory);
+ Outer_Syntax.command "const_moves"
"Rename HOL4 const names to other HOL4 constants"
- K.thy_decl (const_moves >> Toplevel.theory);
- OuterSyntax.command "const_maps"
+ Keyword.thy_decl (const_moves >> Toplevel.theory);
+ Outer_Syntax.command "const_maps"
"Map HOL4 const names to existing Isabelle/HOL const names"
- K.thy_decl (const_maps >> Toplevel.theory));
+ Keyword.thy_decl (const_maps >> Toplevel.theory));
end
-
-end
--- a/src/HOL/Import/proof_kernel.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Import/proof_kernel.ML Sun May 23 10:38:11 2010 +0100
@@ -189,7 +189,7 @@
else Delimfix (Syntax.escape c)
fun quotename c =
- if Syntax.is_identifier c andalso not (OuterKeyword.is_keyword c) then c else quote c
+ if Syntax.is_identifier c andalso not (Keyword.is_keyword c) then c else quote c
fun simple_smart_string_of_cterm ct =
let
@@ -1013,12 +1013,12 @@
local
val th = thm "not_def"
val thy = theory_of_thm th
- val pp = reflexive (cterm_of thy (Const("Trueprop",boolT-->propT)))
+ val pp = Thm.reflexive (cterm_of thy (Const("Trueprop",boolT-->propT)))
in
-val not_elim_thm = combination pp th
+val not_elim_thm = Thm.combination pp th
end
-val not_intro_thm = symmetric not_elim_thm
+val not_intro_thm = Thm.symmetric not_elim_thm
val abs_thm = thm "ext"
val trans_thm = thm "trans"
val symmetry_thm = thm "sym"
@@ -1039,7 +1039,7 @@
end
fun implies_elim_all th =
- Library.foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th)
+ Library.foldl (fn (th,p) => Thm.implies_elim th (Thm.assume p)) (th,cprems_of th)
fun norm_hyps th =
th |> beta_eta_thm
@@ -1054,7 +1054,7 @@
val clc = Thm.cterm_of sg lc
val cvty = ctyp_of_term cv
val th1 = implies_elim_all th
- val th2 = beta_eta_thm (forall_intr cv th1)
+ val th2 = beta_eta_thm (Thm.forall_intr cv th1)
val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm))
val c = prop_of th3
val vname = fst(dest_Free v)
@@ -1072,7 +1072,7 @@
fun rearrange sg tm th =
let
val tm' = Envir.beta_eta_contract tm
- fun find [] n = Thm.permute_prems 0 1 (implies_intr (Thm.cterm_of sg tm) th)
+ fun find [] n = Thm.permute_prems 0 1 (Thm.implies_intr (Thm.cterm_of sg tm) th)
| find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
then Thm.permute_prems n 1 th
else find ps (n+1)
@@ -1258,7 +1258,7 @@
fun get_isabelle_thm thyname thmname hol4conc thy =
let
val (info,hol4conc') = disamb_term hol4conc
- val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
+ val i2h_conc = Thm.symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
val isaconc =
case concl_of i2h_conc of
Const("==",_) $ lhs $ _ => lhs
@@ -1268,7 +1268,7 @@
message "Modified conclusion:";
if_debug prin isaconc)
- fun mk_res th = HOLThm(rens_of info,equal_elim i2h_conc th)
+ fun mk_res th = HOLThm (rens_of info, Thm.equal_elim i2h_conc th)
in
case get_hol4_mapping thyname thmname thy of
SOME (SOME thmname) =>
@@ -1320,7 +1320,7 @@
fun warn () =
let
val (info,hol4conc') = disamb_term hol4conc
- val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
+ val i2h_conc = Thm.symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
in
case concl_of i2h_conc of
Const("==",_) $ lhs $ _ =>
@@ -1369,7 +1369,7 @@
let
val (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
val rew = rewrite_hol4_term (concl_of th) thy
- val th = equal_elim rew th
+ val th = Thm.equal_elim rew th
val thy' = add_hol4_pending thyname thmname args thy
val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
val th = disambiguate_frees th
@@ -1683,7 +1683,7 @@
val (info',v') = disamb_term_from info v
fun strip 0 _ th = th
| strip n (p::ps) th =
- strip (n-1) ps (implies_elim th (assume p))
+ strip (n-1) ps (Thm.implies_elim th (Thm.assume p))
| strip _ _ _ = raise ERR "CHOOSE" "strip error"
val cv = cterm_of thy v'
val th2 = norm_hyps th2
@@ -1697,7 +1697,7 @@
val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
- val th23 = beta_eta_thm (forall_intr cv th22)
+ val th23 = beta_eta_thm (Thm.forall_intr cv th22)
val th11 = implies_elim_all (beta_eta_thm th1)
val th' = th23 COMP (th11 COMP choose_thm')
val th = implies_intr_hyps th'
@@ -1796,7 +1796,7 @@
| _ => raise ERR "mk_ABS" "Bad conclusion"
val (fd,fr) = dom_rng (type_of f)
val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
- val th2 = forall_intr cv th1
+ val th2 = Thm.forall_intr cv th1
val th3 = th2 COMP abs_thm'
val res = implies_intr_hyps th3
in
@@ -1867,7 +1867,7 @@
_ $ (Const("op -->",_) $ a $ Const("False",_)) => a
| _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
val ca = cterm_of thy a
- val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
+ val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
val res = HOLThm(rens,implies_intr_hyps th2)
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1884,7 +1884,7 @@
_ $ (Const("Not",_) $ a) => a
| _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
val ca = cterm_of thy a
- val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
+ val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
val res = HOLThm(rens,implies_intr_hyps th2)
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1902,9 +1902,9 @@
val prems = prems_of th
val th1 = beta_eta_thm th
val th2 = implies_elim_all th1
- val th3 = implies_intr (cterm_of thy (HOLogic.mk_Trueprop tm')) th2
+ val th3 = Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop tm')) th2
val th4 = th3 COMP disch_thm
- val res = HOLThm(rens_of info',implies_intr_hyps th4)
+ val res = HOLThm (rens_of info', implies_intr_hyps th4)
val _ = message "RESULT:"
val _ = if_debug pth res
in
@@ -2032,7 +2032,7 @@
val res' = Thm.unvarify_global res
val hth = HOLThm(rens,res')
val rew = rewrite_hol4_term (concl_of res') thy'
- val th = equal_elim rew res'
+ val th = Thm.equal_elim rew res'
fun handle_const (name,thy) =
let
val defname = Thm.def_name name
@@ -2112,7 +2112,7 @@
val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
val rew = rewrite_hol4_term (concl_of td_th) thy4
- val th = equal_elim rew (Thm.transfer thy4 td_th)
+ val th = Thm.equal_elim rew (Thm.transfer thy4 td_th)
val c = case HOLogic.dest_Trueprop (prop_of th) of
Const("Ex",exT) $ P =>
let
--- a/src/HOL/Import/shuffler.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Import/shuffler.ML Sun May 23 10:38:11 2010 +0100
@@ -95,12 +95,12 @@
val cQ = cert Q
val cPQ = cert PQ
val cPPQ = cert PPQ
- val th1 = assume cPQ |> implies_intr_list [cPQ,cP]
- val th3 = assume cP
- val th4 = implies_elim_list (assume cPPQ) [th3,th3]
+ val th1 = Thm.assume cPQ |> implies_intr_list [cPQ,cP]
+ val th3 = Thm.assume cP
+ val th4 = implies_elim_list (Thm.assume cPPQ) [th3,th3]
|> implies_intr_list [cPPQ,cP]
in
- equal_intr th4 th1 |> Drule.export_without_context
+ Thm.equal_intr th4 th1 |> Drule.export_without_context
end
val imp_comm =
@@ -115,12 +115,12 @@
val cQ = cert Q
val cPQR = cert PQR
val cQPR = cert QPR
- val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ]
+ val th1 = implies_elim_list (Thm.assume cPQR) [Thm.assume cP,Thm.assume cQ]
|> implies_intr_list [cPQR,cQ,cP]
- val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP]
+ val th2 = implies_elim_list (Thm.assume cQPR) [Thm.assume cQ,Thm.assume cP]
|> implies_intr_list [cQPR,cP,cQ]
in
- equal_intr th1 th2 |> Drule.export_without_context
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
end
val def_norm =
@@ -134,20 +134,20 @@
val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0)))
val cPQ = cert (Logic.mk_equals(P,Q))
val cv = cert v
- val rew = assume cvPQ
- |> forall_elim cv
- |> abstract_rule "v" cv
+ val rew = Thm.assume cvPQ
+ |> Thm.forall_elim cv
+ |> Thm.abstract_rule "v" cv
val (lhs,rhs) = Logic.dest_equals(concl_of rew)
- val th1 = transitive (transitive
- (eta_conversion (cert lhs) |> symmetric)
+ val th1 = Thm.transitive (Thm.transitive
+ (Thm.eta_conversion (cert lhs) |> Thm.symmetric)
rew)
- (eta_conversion (cert rhs))
- |> implies_intr cvPQ
- val th2 = combination (assume cPQ) (reflexive cv)
- |> forall_intr cv
- |> implies_intr cPQ
+ (Thm.eta_conversion (cert rhs))
+ |> Thm.implies_intr cvPQ
+ val th2 = Thm.combination (Thm.assume cPQ) (Thm.reflexive cv)
+ |> Thm.forall_intr cv
+ |> Thm.implies_intr cPQ
in
- equal_intr th1 th2 |> Drule.export_without_context
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
end
val all_comm =
@@ -164,16 +164,16 @@
val cr = cert rhs
val cx = cert x
val cy = cert y
- val th1 = assume cr
+ val th1 = Thm.assume cr
|> forall_elim_list [cy,cx]
|> forall_intr_list [cx,cy]
- |> implies_intr cr
- val th2 = assume cl
+ |> Thm.implies_intr cr
+ val th2 = Thm.assume cl
|> forall_elim_list [cx,cy]
|> forall_intr_list [cy,cx]
- |> implies_intr cl
+ |> Thm.implies_intr cl
in
- equal_intr th1 th2 |> Drule.export_without_context
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
end
val equiv_comm =
@@ -184,10 +184,10 @@
val u = Free("u",T)
val ctu = cert (Logic.mk_equals(t,u))
val cut = cert (Logic.mk_equals(u,t))
- val th1 = assume ctu |> symmetric |> implies_intr ctu
- val th2 = assume cut |> symmetric |> implies_intr cut
+ val th1 = Thm.assume ctu |> Thm.symmetric |> Thm.implies_intr ctu
+ val th2 = Thm.assume cut |> Thm.symmetric |> Thm.implies_intr cut
in
- equal_intr th1 th2 |> Drule.export_without_context
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
end
(* This simplification procedure rewrites !!x y. P x y
@@ -217,7 +217,7 @@
val lhs = list_all ([xv,yv],t)
val rhs = list_all ([yv,xv],swap_bound 0 t)
val rew = Logic.mk_equals (lhs,rhs)
- val init = trivial (cterm_of thy rew)
+ val init = Thm.trivial (cterm_of thy rew)
in
(all_comm RS init handle e => (message "rew_th"; OldGoals.print_exn e))
end
@@ -232,10 +232,10 @@
then
let
val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
- val t_th = reflexive (cterm_of thy t)
- val newt_th = reflexive (cterm_of thy newt)
+ val t_th = Thm.reflexive (cterm_of thy t)
+ val newt_th = Thm.reflexive (cterm_of thy newt)
in
- SOME (transitive t_th newt_th)
+ SOME (Thm.transitive t_th newt_th)
end
else NONE
| _ => error "norm_term (quant_rewrite) internal error"
@@ -294,7 +294,7 @@
fun eta_contract thy assumes origt =
let
val (typet,Tinst) = freeze_thaw_term origt
- val (init,thaw) = Drule.legacy_freeze_thaw (reflexive (cterm_of thy typet))
+ val (init,thaw) = Drule.legacy_freeze_thaw (Thm.reflexive (cterm_of thy typet))
val final = inst_tfrees thy Tinst o thaw
val t = #1 (Logic.dest_equals (prop_of init))
val _ =
@@ -322,18 +322,18 @@
val v = Free (Name.variant (Term.add_free_names t []) "v", xT)
val cv = cert v
val ct = cert t
- val th = (assume ct)
- |> forall_elim cv
- |> abstract_rule x cv
- val ext_th = eta_conversion (cert (Abs(x,xT,P)))
- val th' = transitive (symmetric ext_th) th
+ val th = (Thm.assume ct)
+ |> Thm.forall_elim cv
+ |> Thm.abstract_rule x cv
+ val ext_th = Thm.eta_conversion (cert (Abs(x,xT,P)))
+ val th' = Thm.transitive (Thm.symmetric ext_th) th
val cu = cert (prop_of th')
- val uth = combination (assume cu) (reflexive cv)
- val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v)))
- |> transitive uth
- |> forall_intr cv
- |> implies_intr cu
- val rew_th = equal_intr (th' |> implies_intr ct) uth'
+ val uth = Thm.combination (Thm.assume cu) (Thm.reflexive cv)
+ val uth' = (Thm.beta_conversion false (cert (Abs(x,xT,Q) $ v)))
+ |> Thm.transitive uth
+ |> Thm.forall_intr cv
+ |> Thm.implies_intr cu
+ val rew_th = Thm.equal_intr (th' |> Thm.implies_intr ct) uth'
val res = final rew_th
val lhs = (#1 (Logic.dest_equals (prop_of res)))
in
@@ -345,7 +345,7 @@
end
fun beta_fun thy assume t =
- SOME (beta_conversion true (cterm_of thy t))
+ SOME (Thm.beta_conversion true (cterm_of thy t))
val meta_sym_rew = thm "refl"
@@ -357,7 +357,7 @@
fun eta_expand thy assumes origt =
let
val (typet,Tinst) = freeze_thaw_term origt
- val (init,thaw) = Drule.legacy_freeze_thaw (reflexive (cterm_of thy typet))
+ val (init,thaw) = Drule.legacy_freeze_thaw (Thm.reflexive (cterm_of thy typet))
val final = inst_tfrees thy Tinst o thaw
val t = #1 (Logic.dest_equals (prop_of init))
val _ =
@@ -387,20 +387,20 @@
val v = Free(vname,aT)
val cv = cert v
val ct = cert t
- val th1 = (combination (assume ct) (reflexive cv))
- |> forall_intr cv
- |> implies_intr ct
+ val th1 = (Thm.combination (Thm.assume ct) (Thm.reflexive cv))
+ |> Thm.forall_intr cv
+ |> Thm.implies_intr ct
val concl = cert (concl_of th1)
- val th2 = (assume concl)
- |> forall_elim cv
- |> abstract_rule vname cv
+ val th2 = (Thm.assume concl)
+ |> Thm.forall_elim cv
+ |> Thm.abstract_rule vname cv
val (lhs,rhs) = Logic.dest_equals (prop_of th2)
- val elhs = eta_conversion (cert lhs)
- val erhs = eta_conversion (cert rhs)
- val th2' = transitive
- (transitive (symmetric elhs) th2)
+ val elhs = Thm.eta_conversion (cert lhs)
+ val erhs = Thm.eta_conversion (cert rhs)
+ val th2' = Thm.transitive
+ (Thm.transitive (Thm.symmetric elhs) th2)
erhs
- val res = equal_intr th1 (th2' |> implies_intr concl)
+ val res = Thm.equal_intr th1 (th2' |> Thm.implies_intr concl)
val res' = final res
in
SOME res'
@@ -475,7 +475,7 @@
val (t',_) = F (t,0)
val ct = cterm_of thy t
val ct' = cterm_of thy t'
- val res = transitive (reflexive ct) (reflexive ct')
+ val res = Thm.transitive (Thm.reflexive ct) (Thm.reflexive ct')
val _ = message ("disamb_term: " ^ (string_of_thm res))
in
res
@@ -496,12 +496,12 @@
let
val rhs = Thm.rhs_of th
in
- transitive th (f rhs)
+ Thm.transitive th (f rhs)
end
val th =
t |> disamb_bound thy
|> chain (Simplifier.full_rewrite ss)
- |> chain eta_conversion
+ |> chain Thm.eta_conversion
|> Thm.strip_shyps
val _ = message ("norm_term: " ^ (string_of_thm th))
in
@@ -529,7 +529,7 @@
let
val c = prop_of th
in
- equal_elim (norm_term thy c) th
+ Thm.equal_elim (norm_term thy c) th
end
(* make_equal thy t u tries to construct the theorem t == u under the
@@ -540,7 +540,7 @@
let
val t_is_t' = norm_term thy t
val u_is_u' = norm_term thy u
- val th = transitive t_is_t' (symmetric u_is_u')
+ val th = Thm.transitive t_is_t' (Thm.symmetric u_is_u')
val _ = message ("make_equal: SOME " ^ (string_of_thm th))
in
SOME th
@@ -593,7 +593,7 @@
| process ((name,th)::thms) =
let
val norm_th = Thm.varifyT_global (norm_thm thy (close_thm (Thm.transfer thy th)))
- val triv_th = trivial rhs
+ val triv_th = Thm.trivial rhs
val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
val mod_th = case Seq.pull (Thm.bicompose false (*true*) (false,norm_th,0) 1 triv_th) of
SOME(th,_) => SOME th
@@ -602,7 +602,7 @@
case mod_th of
SOME mod_th =>
let
- val closed_th = equal_elim (symmetric rew_th) mod_th
+ val closed_th = Thm.equal_elim (Thm.symmetric rew_th) mod_th
in
message ("Shuffler.set_prop succeeded by " ^ name);
SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)
--- a/src/HOL/IsaMakefile Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/IsaMakefile Sun May 23 10:38:11 2010 +0100
@@ -60,7 +60,7 @@
HOL-Proofs-Extraction \
HOL-Proofs-Lambda \
HOL-SET_Protocol \
- HOL-SMT_Examples \
+ HOL-Word-SMT_Examples \
HOL-Statespace \
HOL-Subst \
TLA-Buffer \
@@ -128,7 +128,6 @@
$(SRC)/Tools/induct.ML \
$(SRC)/Tools/induct_tacs.ML \
$(SRC)/Tools/intuitionistic.ML \
- $(SRC)/Tools/more_conv.ML \
$(SRC)/Tools/nbe.ML \
$(SRC)/Tools/project_rule.ML \
$(SRC)/Tools/quickcheck.ML \
@@ -154,7 +153,6 @@
Inductive.thy \
Lattices.thy \
Nat.thy \
- Nitpick.thy \
Option.thy \
Orderings.thy \
Plain.thy \
@@ -162,7 +160,6 @@
Predicate.thy \
Product_Type.thy \
Record.thy \
- Refute.thy \
Relation.thy \
Rings.thy \
SAT.thy \
@@ -197,22 +194,6 @@
Tools/Function/size.ML \
Tools/Function/sum_tree.ML \
Tools/Function/termination.ML \
- Tools/Nitpick/kodkod.ML \
- Tools/Nitpick/kodkod_sat.ML \
- Tools/Nitpick/minipick.ML \
- Tools/Nitpick/nitpick.ML \
- Tools/Nitpick/nitpick_hol.ML \
- Tools/Nitpick/nitpick_isar.ML \
- Tools/Nitpick/nitpick_kodkod.ML \
- Tools/Nitpick/nitpick_model.ML \
- Tools/Nitpick/nitpick_mono.ML \
- Tools/Nitpick/nitpick_nut.ML \
- Tools/Nitpick/nitpick_peephole.ML \
- Tools/Nitpick/nitpick_preproc.ML \
- Tools/Nitpick/nitpick_rep.ML \
- Tools/Nitpick/nitpick_scope.ML \
- Tools/Nitpick/nitpick_tests.ML \
- Tools/Nitpick/nitpick_util.ML \
Tools/inductive_codegen.ML \
Tools/inductive.ML \
Tools/inductive_realizer.ML \
@@ -262,6 +243,7 @@
Map.thy \
Nat_Numeral.thy \
Nat_Transfer.thy \
+ Nitpick.thy \
Numeral_Simprocs.thy \
Presburger.thy \
Predicate_Compile.thy \
@@ -270,6 +252,7 @@
Random.thy \
Random_Sequence.thy \
Recdef.thy \
+ Refute.thy \
Semiring_Normalization.thy \
SetInterval.thy \
Sledgehammer.thy \
@@ -291,6 +274,22 @@
Tools/list_code.ML \
Tools/meson.ML \
Tools/nat_numeral_simprocs.ML \
+ Tools/Nitpick/kodkod.ML \
+ Tools/Nitpick/kodkod_sat.ML \
+ Tools/Nitpick/minipick.ML \
+ Tools/Nitpick/nitpick.ML \
+ Tools/Nitpick/nitpick_hol.ML \
+ Tools/Nitpick/nitpick_isar.ML \
+ Tools/Nitpick/nitpick_kodkod.ML \
+ Tools/Nitpick/nitpick_model.ML \
+ Tools/Nitpick/nitpick_mono.ML \
+ Tools/Nitpick/nitpick_nut.ML \
+ Tools/Nitpick/nitpick_peephole.ML \
+ Tools/Nitpick/nitpick_preproc.ML \
+ Tools/Nitpick/nitpick_rep.ML \
+ Tools/Nitpick/nitpick_scope.ML \
+ Tools/Nitpick/nitpick_tests.ML \
+ Tools/Nitpick/nitpick_util.ML \
Tools/numeral.ML \
Tools/numeral_simprocs.ML \
Tools/numeral_syntax.ML \
@@ -402,16 +401,16 @@
Library/Efficient_Nat.thy Library/Sum_Of_Squares.thy \
Library/Dlist.thy Library/Sum_Of_Squares/sos_wrapper.ML \
Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy \
- Library/Glbs.thy Library/normarith.ML Library/Executable_Set.thy \
+ Library/Glbs.thy Library/Executable_Set.thy \
Library/Infinite_Set.thy Library/FuncSet.thy \
Library/Permutations.thy Library/Bit.thy Library/FrechetDeriv.thy \
Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy \
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/Multiset.thy Library/Permutation.thy \
+ Library/Library.thy Library/List_Prefix.thy Library/More_Set.thy \
+ Library/More_List.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/Nat_Infinity.thy Library/README.html Library/State_Monad.thy \
Library/Continuity.thy Library/Order_Relation.thy \
Library/Nested_Environment.thy Library/Ramsey.thy Library/Zorn.thy \
Library/Library/ROOT.ML Library/Library/document/root.tex \
@@ -434,7 +433,7 @@
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 \
- Library/OptionalSugar.thy Library/Convex.thy \
+ Library/OptionalSugar.thy Library/Convex.thy \
Library/Predicate_Compile_Quickcheck.thy Library/SML_Quickcheck.thy
@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
@@ -587,17 +586,18 @@
HOL-Old_Number_Theory: HOL $(LOG)/HOL-Old_Number_Theory.gz
-$(LOG)/HOL-Old_Number_Theory.gz: $(OUT)/HOL Library/Permutation.thy \
- Old_Number_Theory/Primes.thy Old_Number_Theory/Fib.thy \
- Old_Number_Theory/Factorization.thy Old_Number_Theory/BijectionRel.thy \
- Old_Number_Theory/Chinese.thy Old_Number_Theory/EulerFermat.thy \
- Old_Number_Theory/IntFact.thy Old_Number_Theory/IntPrimes.thy \
- Old_Number_Theory/WilsonBij.thy Old_Number_Theory/WilsonRuss.thy \
- Old_Number_Theory/Finite2.thy Old_Number_Theory/Int2.thy \
- Old_Number_Theory/EvenOdd.thy Old_Number_Theory/Residues.thy \
- Old_Number_Theory/Euler.thy Old_Number_Theory/Gauss.thy \
+$(LOG)/HOL-Old_Number_Theory.gz: $(OUT)/HOL Library/Permutation.thy \
+ Old_Number_Theory/Primes.thy Old_Number_Theory/Fib.thy \
+ Old_Number_Theory/Factorization.thy Old_Number_Theory/BijectionRel.thy\
+ Old_Number_Theory/Chinese.thy Old_Number_Theory/EulerFermat.thy \
+ Old_Number_Theory/IntFact.thy Old_Number_Theory/IntPrimes.thy \
+ Old_Number_Theory/WilsonBij.thy Old_Number_Theory/WilsonRuss.thy \
+ Old_Number_Theory/Finite2.thy Old_Number_Theory/Int2.thy \
+ Old_Number_Theory/EvenOdd.thy Old_Number_Theory/Residues.thy \
+ Old_Number_Theory/Euler.thy Old_Number_Theory/Gauss.thy \
Old_Number_Theory/Quadratic_Reciprocity.thy Library/Infinite_Set.thy \
- Old_Number_Theory/Legacy_GCD.thy Old_Number_Theory/Pocklington.thy Old_Number_Theory/ROOT.ML
+ Old_Number_Theory/Legacy_GCD.thy Old_Number_Theory/Pocklington.thy \
+ Old_Number_Theory/ROOT.ML
@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Old_Number_Theory
@@ -712,9 +712,9 @@
HOL-Auth: HOL $(LOG)/HOL-Auth.gz
$(LOG)/HOL-Auth.gz: $(OUT)/HOL \
- Auth/Auth_Shared.thy Auth/Auth_Public.thy Auth/All_Symmetric.thy \
- Auth/Guard/Auth_Guard_Shared.thy \
- Auth/Guard/Auth_Guard_Public.thy \
+ Auth/Auth_Shared.thy Auth/Auth_Public.thy Auth/All_Symmetric.thy \
+ Auth/Guard/Auth_Guard_Shared.thy \
+ Auth/Guard/Auth_Guard_Public.thy \
Auth/Smartcard/Auth_Smartcard.thy Auth/All_Symmetric.thy \
Auth/CertifiedEmail.thy Auth/Event.thy Auth/Message.thy \
Auth/NS_Public.thy Auth/NS_Public_Bad.thy Auth/NS_Shared.thy \
@@ -1088,27 +1088,27 @@
HOL-Multivariate_Analysis: HOL $(OUT)/HOL-Multivariate_Analysis
-$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL \
- Multivariate_Analysis/ROOT.ML \
- Multivariate_Analysis/document/root.tex \
- Multivariate_Analysis/Brouwer_Fixpoint.thy \
- Multivariate_Analysis/Convex_Euclidean_Space.thy \
- Multivariate_Analysis/Derivative.thy \
- Multivariate_Analysis/Determinants.thy \
- Multivariate_Analysis/Euclidean_Space.thy \
- Multivariate_Analysis/Fashoda.thy \
- Multivariate_Analysis/Finite_Cartesian_Product.thy \
- Multivariate_Analysis/Integration.thy \
- Multivariate_Analysis/Integration.certs \
- Multivariate_Analysis/L2_Norm.thy \
- Multivariate_Analysis/Multivariate_Analysis.thy \
- Multivariate_Analysis/Operator_Norm.thy \
- Multivariate_Analysis/Path_Connected.thy \
- Multivariate_Analysis/Real_Integration.thy \
- Multivariate_Analysis/Topology_Euclidean_Space.thy \
- Multivariate_Analysis/Vec1.thy Library/Glbs.thy \
- Library/Inner_Product.thy Library/Numeral_Type.thy \
- Library/Convex.thy Library/FrechetDeriv.thy \
+$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL \
+ Multivariate_Analysis/Brouwer_Fixpoint.thy \
+ Multivariate_Analysis/Convex_Euclidean_Space.thy \
+ Multivariate_Analysis/Derivative.thy \
+ Multivariate_Analysis/Determinants.thy \
+ Multivariate_Analysis/Euclidean_Space.thy \
+ Multivariate_Analysis/Fashoda.thy \
+ Multivariate_Analysis/Finite_Cartesian_Product.thy \
+ Multivariate_Analysis/Integration.certs \
+ Multivariate_Analysis/Integration.thy \
+ Multivariate_Analysis/L2_Norm.thy \
+ Multivariate_Analysis/Multivariate_Analysis.thy \
+ Multivariate_Analysis/Operator_Norm.thy \
+ Multivariate_Analysis/Path_Connected.thy \
+ Multivariate_Analysis/ROOT.ML \
+ Multivariate_Analysis/Real_Integration.thy \
+ Multivariate_Analysis/Topology_Euclidean_Space.thy \
+ Multivariate_Analysis/document/root.tex \
+ Multivariate_Analysis/normarith.ML Multivariate_Analysis/Vec1.thy \
+ Library/Glbs.thy Library/Inner_Product.thy Library/Numeral_Type.thy \
+ Library/Convex.thy Library/FrechetDeriv.thy \
Library/Product_Vector.thy Library/Product_plus.thy
@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
@@ -1254,11 +1254,11 @@
@$(ISABELLE_TOOL) usedir $(OUT)/HOL Mirabelle
-## HOL-SMT_Examples
+## HOL-Word-SMT_Examples
-HOL-SMT_Examples: HOL-Word $(LOG)/HOL-SMT_Examples.gz
+HOL-Word-SMT_Examples: HOL-Word $(LOG)/HOL-Word-SMT_Examples.gz
-$(LOG)/HOL-SMT_Examples.gz: $(OUT)/HOL-Word SMT_Examples/ROOT.ML \
+$(LOG)/HOL-Word-SMT_Examples.gz: $(OUT)/HOL-Word SMT_Examples/ROOT.ML \
SMT_Examples/SMT_Examples.thy SMT_Examples/SMT_Examples.certs \
SMT_Examples/SMT_Word_Examples.thy SMT_Examples/SMT_Tests.thy \
SMT_Examples/SMT_Word_Examples.certs SMT_Examples/SMT_Tests.certs
@@ -1346,15 +1346,15 @@
$(LOG)/HOL-Probability.gz $(LOG)/HOL-Prolog.gz \
$(LOG)/HOL-Proofs.gz $(LOG)/HOL-Proofs-Extraction.gz \
$(LOG)/HOL-Proofs-Lambda.gz $(LOG)/HOL-SET_Protocol.gz \
- $(LOG)/HOL-SMT_Examples.gz $(LOG)/HOL-Statespace.gz \
- $(LOG)/HOL-Subst.gz $(LOG)/HOL-UNITY.gz \
- $(LOG)/HOL-Unix.gz $(LOG)/HOL-Word-Examples.gz \
- $(LOG)/HOL-Word.gz $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz \
- $(LOG)/HOL.gz $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz \
- $(LOG)/TLA-Inc.gz $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz \
- $(OUT)/HOL $(OUT)/HOL-Algebra $(OUT)/HOL-Base \
- $(OUT)/HOL-Boogie $(OUT)/HOL-Main \
- $(OUT)/HOL-Multivariate_Analysis $(OUT)/HOL-NSA \
- $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \
+ $(LOG)/HOL-Word-SMT_Examples.gz \
+ $(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz \
+ $(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz \
+ $(LOG)/HOL-Word-Examples.gz $(LOG)/HOL-Word.gz \
+ $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz $(LOG)/HOL.gz \
+ $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz $(LOG)/TLA-Inc.gz \
+ $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz $(OUT)/HOL \
+ $(OUT)/HOL-Algebra $(OUT)/HOL-Base $(OUT)/HOL-Boogie \
+ $(OUT)/HOL-Main $(OUT)/HOL-Multivariate_Analysis \
+ $(OUT)/HOL-NSA $(OUT)/HOL-Nominal $(OUT)/HOL-Plain \
$(OUT)/HOL-Probability $(OUT)/HOL-Proofs \
$(OUT)/HOL-Word $(OUT)/HOL4 $(OUT)/TLA
--- a/src/HOL/Library/AssocList.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/AssocList.thy Sun May 23 10:38:11 2010 +0100
@@ -668,6 +668,10 @@
"Mapping.lookup (Mapping xs) = map_of xs"
by (simp add: Mapping_def)
+lemma keys_Mapping [simp, code]:
+ "Mapping.keys (Mapping xs) = set (map fst xs)"
+ by (simp add: keys_def dom_map_of_conv_image_fst)
+
lemma empty_Mapping [code]:
"Mapping.empty = Mapping []"
by (rule mapping_eqI) simp
@@ -684,13 +688,9 @@
"Mapping.delete k (Mapping xs) = Mapping (delete k xs)"
by (rule mapping_eqI) (simp add: delete_conv')
-lemma keys_Mapping [code]:
- "Mapping.keys (Mapping xs) = set (map fst xs)"
- by (simp add: keys_def dom_map_of_conv_image_fst)
-
lemma ordered_keys_Mapping [code]:
"Mapping.ordered_keys (Mapping xs) = sort (remdups (map fst xs))"
- by (simp only: ordered_keys_def keys_Mapping sorted_list_of_set_sort_remdups)
+ by (simp only: ordered_keys_def keys_Mapping sorted_list_of_set_sort_remdups) simp
lemma size_Mapping [code]:
"Mapping.size (Mapping xs) = length (remdups (map fst xs))"
@@ -704,4 +704,7 @@
"Mapping.bulkload vs = Mapping (map (\<lambda>n. (n, vs ! n)) [0..<length vs])"
by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
+lemma [code, code del]:
+ "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+
end
--- a/src/HOL/Library/Dlist.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Dlist.thy Sun May 23 10:38:11 2010 +0100
@@ -3,42 +3,9 @@
header {* Lists with elements distinct as canonical example for datatype invariants *}
theory Dlist
-imports Main Fset
+imports Main More_List Fset
begin
-section {* Prelude *}
-
-text {* Without canonical argument order, higher-order things tend to get confusing quite fast: *}
-
-setup {* Sign.map_naming (Name_Space.add_path "List") *}
-
-primrec member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
- "member [] y \<longleftrightarrow> False"
- | "member (x#xs) y \<longleftrightarrow> x = y \<or> member xs y"
-
-lemma member_set:
- "member = set"
-proof (rule ext)+
- fix xs :: "'a list" and x :: 'a
- have "member xs x \<longleftrightarrow> x \<in> set xs" by (induct xs) auto
- then show "member xs x = set xs x" by (simp add: mem_def)
-qed
-
-lemma not_set_compl:
- "Not \<circ> set xs = - set xs"
- by (simp add: fun_Compl_def bool_Compl_def comp_def expand_fun_eq)
-
-primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
- "fold f [] s = s"
- | "fold f (x#xs) s = fold f xs (f x s)"
-
-lemma foldl_fold:
- "foldl f s xs = List.fold (\<lambda>x s. f s x) xs s"
- by (induct xs arbitrary: s) simp_all
-
-setup {* Sign.map_naming Name_Space.parent_path *}
-
-
section {* The type of distinct lists *}
typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
@@ -101,7 +68,10 @@
"length dxs = List.length (list_of_dlist dxs)"
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
- "fold f dxs = List.fold f (list_of_dlist dxs)"
+ "fold f dxs = More_List.fold f (list_of_dlist dxs)"
+
+definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
+ "foldr f dxs = List.foldr f (list_of_dlist dxs)"
section {* Executable version obeying invariant *}
@@ -127,6 +97,16 @@
by (simp add: filter_def)
+text {* Explicit executable conversion *}
+
+definition dlist_of_list [simp]:
+ "dlist_of_list = Dlist"
+
+lemma [code abstract]:
+ "list_of_dlist (dlist_of_list xs) = remdups xs"
+ by simp
+
+
section {* Implementation of sets by distinct lists -- canonical! *}
definition Set :: "'a dlist \<Rightarrow> 'a fset" where
@@ -143,14 +123,13 @@
declare UNIV_Set [code del]
declare insert_Set [code del]
declare remove_Set [code del]
+declare compl_Set [code del]
+declare compl_Coset [code del]
declare map_Set [code del]
declare filter_Set [code del]
declare forall_Set [code del]
declare exists_Set [code del]
declare card_Set [code del]
-declare subfset_eq_forall [code del]
-declare subfset_subfset_eq [code del]
-declare eq_fset_subfset_eq [code del]
declare inter_project [code del]
declare subtract_remove [code del]
declare union_insert [code del]
@@ -173,6 +152,14 @@
"Fset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
by (simp add: Coset_def member_set not_set_compl)
+lemma Set_dlist_of_list [code]:
+ "Fset.Set xs = Set (dlist_of_list xs)"
+ by simp
+
+lemma Coset_dlist_of_list [code]:
+ "Fset.Coset xs = Coset (dlist_of_list xs)"
+ by simp
+
lemma is_empty_Set [code]:
"Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"
by (simp add: null_def null_empty member_set)
@@ -200,6 +187,11 @@
"Fset.member (Coset dxs) = Not \<circ> member dxs"
by (simp_all add: member_def)
+lemma compl_code [code]:
+ "- Set dxs = Coset dxs"
+ "- Coset dxs = Set dxs"
+ by (simp_all add: not_set_compl member_set)
+
lemma map_code [code]:
"Fset.map f (Set dxs) = Set (map f dxs)"
by (simp add: member_set)
@@ -220,38 +212,34 @@
"Fset.card (Set dxs) = length dxs"
by (simp add: length_def member_set distinct_card)
-lemma foldl_list_of_dlist:
- "foldl f s (list_of_dlist dxs) = fold (\<lambda>x s. f s x) dxs s"
- by (simp add: foldl_fold fold_def)
-
lemma inter_code [code]:
"inf A (Set xs) = Set (filter (Fset.member A) xs)"
- "inf A (Coset xs) = fold Fset.remove xs A"
- by (simp_all only: Set_def Coset_def foldl_list_of_dlist inter_project list_of_dlist_filter)
+ "inf A (Coset xs) = foldr Fset.remove xs A"
+ by (simp_all only: Set_def Coset_def foldr_def inter_project list_of_dlist_filter)
lemma subtract_code [code]:
- "A - Set xs = fold Fset.remove xs A"
+ "A - Set xs = foldr Fset.remove xs A"
"A - Coset xs = Set (filter (Fset.member A) xs)"
- by (simp_all only: Set_def Coset_def foldl_list_of_dlist subtract_remove list_of_dlist_filter)
+ by (simp_all only: Set_def Coset_def foldr_def subtract_remove list_of_dlist_filter)
lemma union_code [code]:
- "sup (Set xs) A = fold Fset.insert xs A"
+ "sup (Set xs) A = foldr Fset.insert xs A"
"sup (Coset xs) A = Coset (filter (Not \<circ> Fset.member A) xs)"
- by (simp_all only: Set_def Coset_def foldl_list_of_dlist union_insert list_of_dlist_filter)
+ by (simp_all only: Set_def Coset_def foldr_def union_insert list_of_dlist_filter)
context complete_lattice
begin
lemma Infimum_code [code]:
- "Infimum (Set As) = fold inf As top"
- by (simp only: Set_def Infimum_inf foldl_list_of_dlist inf.commute)
+ "Infimum (Set As) = foldr inf As top"
+ by (simp only: Set_def Infimum_inf foldr_def inf.commute)
lemma Supremum_code [code]:
- "Supremum (Set As) = fold sup As bot"
- by (simp only: Set_def Supremum_sup foldl_list_of_dlist sup.commute)
+ "Supremum (Set As) = foldr sup As bot"
+ by (simp only: Set_def Supremum_sup foldr_def sup.commute)
end
-hide_const (open) member fold empty insert remove map filter null member length fold
+hide_const (open) member fold foldr empty insert remove map filter null member length fold
end
--- a/src/HOL/Library/Efficient_Nat.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy Sun May 23 10:38:11 2010 +0100
@@ -252,7 +252,7 @@
*}
code_include Haskell "Nat" {*
-newtype Nat = Nat Integer deriving (Show, Eq);
+newtype Nat = Nat Integer deriving (Eq, Show, Read);
instance Num Nat where {
fromInteger k = Nat (if k >= 0 then k else 0);
--- a/src/HOL/Library/Executable_Set.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Executable_Set.thy Sun May 23 10:38:11 2010 +0100
@@ -6,7 +6,7 @@
header {* A crude implementation of finite sets by lists -- avoid using this at any cost! *}
theory Executable_Set
-imports List_Set
+imports More_Set
begin
declare mem_def [code del]
@@ -50,8 +50,8 @@
by simp
lemma [code]:
- "x \<in> Set xs \<longleftrightarrow> member x xs"
- "x \<in> Coset xs \<longleftrightarrow> \<not> member x xs"
+ "x \<in> Set xs \<longleftrightarrow> member xs x"
+ "x \<in> Coset xs \<longleftrightarrow> \<not> member xs x"
by (simp_all add: mem_iff)
definition is_empty :: "'a set \<Rightarrow> bool" where
@@ -76,9 +76,9 @@
Code.add_signature_cmd ("is_empty", "'a set \<Rightarrow> bool")
#> Code.add_signature_cmd ("empty", "'a set")
#> Code.add_signature_cmd ("insert", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
- #> Code.add_signature_cmd ("List_Set.remove", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
+ #> Code.add_signature_cmd ("More_Set.remove", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
#> Code.add_signature_cmd ("image", "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set")
- #> Code.add_signature_cmd ("List_Set.project", "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set")
+ #> Code.add_signature_cmd ("More_Set.project", "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set")
#> Code.add_signature_cmd ("Ball", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
#> Code.add_signature_cmd ("Bex", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
#> Code.add_signature_cmd ("card", "'a set \<Rightarrow> nat")
@@ -232,36 +232,36 @@
lemma inter_project [code]:
"inter A (Set xs) = Set (List.filter (\<lambda>x. x \<in> A) xs)"
- "inter A (Coset xs) = foldl (\<lambda>A x. remove x A) A xs"
- by (simp add: inter project_def, simp add: Diff_eq [symmetric] minus_set)
+ "inter A (Coset xs) = foldr remove xs A"
+ by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
lemma subtract_remove [code]:
- "subtract (Set xs) A = foldl (\<lambda>A x. remove x A) A xs"
+ "subtract (Set xs) A = foldr remove xs A"
"subtract (Coset xs) A = Set (List.filter (\<lambda>x. x \<in> A) xs)"
- by (auto simp add: minus_set)
+ by (auto simp add: minus_set_foldr)
lemma union_insert [code]:
- "union (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
+ "union (Set xs) A = foldr insert xs A"
"union (Coset xs) A = Coset (List.filter (\<lambda>x. x \<notin> A) xs)"
- by (auto simp add: union_set)
+ by (auto simp add: union_set_foldr)
lemma Inf_inf [code]:
- "Inf (Set xs) = foldl inf (top :: 'a::complete_lattice) xs"
+ "Inf (Set xs) = foldr inf xs (top :: 'a::complete_lattice)"
"Inf (Coset []) = (bot :: 'a::complete_lattice)"
- by (simp_all add: Inf_UNIV Inf_set_fold)
+ by (simp_all add: Inf_UNIV Inf_set_foldr)
lemma Sup_sup [code]:
- "Sup (Set xs) = foldl sup (bot :: 'a::complete_lattice) xs"
+ "Sup (Set xs) = foldr sup xs (bot :: 'a::complete_lattice)"
"Sup (Coset []) = (top :: 'a::complete_lattice)"
- by (simp_all add: Sup_UNIV Sup_set_fold)
+ by (simp_all add: Sup_UNIV Sup_set_foldr)
lemma Inter_inter [code]:
- "Inter (Set xs) = foldl inter (Coset []) xs"
+ "Inter (Set xs) = foldr inter xs (Coset [])"
"Inter (Coset []) = empty"
unfolding Inter_def Inf_inf by simp_all
lemma Union_union [code]:
- "Union (Set xs) = foldl union empty xs"
+ "Union (Set xs) = foldr union xs empty"
"Union (Coset []) = Coset []"
unfolding Union_def Sup_sup by simp_all
--- a/src/HOL/Library/Fset.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Fset.thy Sun May 23 10:38:11 2010 +0100
@@ -4,7 +4,7 @@
header {* Executable finite sets *}
theory Fset
-imports List_Set
+imports More_Set More_List
begin
declare mem_def [simp]
@@ -41,9 +41,9 @@
code_datatype Set Coset
lemma member_code [code]:
- "member (Set xs) y \<longleftrightarrow> List.member y xs"
- "member (Coset xs) y \<longleftrightarrow> \<not> List.member y xs"
- by (simp_all add: mem_iff fun_Compl_def bool_Compl_def)
+ "member (Set xs) = List.member xs"
+ "member (Coset xs) = Not \<circ> List.member xs"
+ by (simp_all add: expand_fun_eq mem_iff fun_Compl_def bool_Compl_def)
lemma member_image_UNIV [simp]:
"member ` UNIV = UNIV"
@@ -105,10 +105,11 @@
end
+
subsection {* Basic operations *}
definition is_empty :: "'a fset \<Rightarrow> bool" where
- [simp]: "is_empty A \<longleftrightarrow> List_Set.is_empty (member A)"
+ [simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (member A)"
lemma is_empty_Set [code]:
"is_empty (Set xs) \<longleftrightarrow> null xs"
@@ -128,16 +129,16 @@
lemma insert_Set [code]:
"insert x (Set xs) = Set (List.insert x xs)"
"insert x (Coset xs) = Coset (removeAll x xs)"
- by (simp_all add: Set_def Coset_def set_insert)
+ by (simp_all add: Set_def Coset_def)
definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
- [simp]: "remove x A = Fset (List_Set.remove x (member A))"
+ [simp]: "remove x A = Fset (More_Set.remove x (member A))"
lemma remove_Set [code]:
"remove x (Set xs) = Set (removeAll x xs)"
"remove x (Coset xs) = Coset (List.insert x xs)"
by (simp_all add: Set_def Coset_def remove_set_compl)
- (simp add: List_Set.remove_def)
+ (simp add: More_Set.remove_def)
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
[simp]: "map f A = Fset (image f (member A))"
@@ -147,7 +148,7 @@
by (simp add: Set_def)
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
- [simp]: "filter P A = Fset (List_Set.project P (member A))"
+ [simp]: "filter P A = Fset (More_Set.project P (member A))"
lemma filter_Set [code]:
"filter P (Set xs) = Set (List.filter P xs)"
@@ -175,9 +176,17 @@
proof -
have "Finite_Set.card (set (remdups xs)) = length (remdups xs)"
by (rule distinct_card) simp
- then show ?thesis by (simp add: Set_def card_def)
+ then show ?thesis by (simp add: Set_def)
qed
+lemma compl_Set [simp, code]:
+ "- Set xs = Coset xs"
+ by (simp add: Set_def Coset_def)
+
+lemma compl_Coset [simp, code]:
+ "- Coset xs = Set xs"
+ by (simp add: Set_def Coset_def)
+
subsection {* Derived operations *}
@@ -198,39 +207,49 @@
lemma inter_project [code]:
"inf A (Set xs) = Set (List.filter (member A) xs)"
- "inf A (Coset xs) = foldl (\<lambda>A x. remove x A) A xs"
+ "inf A (Coset xs) = foldr remove xs A"
proof -
show "inf A (Set xs) = Set (List.filter (member A) xs)"
by (simp add: inter project_def Set_def)
- have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
- member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
- by (rule foldl_apply) (simp add: expand_fun_eq)
- then show "inf A (Coset xs) = foldl (\<lambda>A x. remove x A) A xs"
- by (simp add: Diff_eq [symmetric] minus_set)
+ have *: "\<And>x::'a. remove = (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member)"
+ by (simp add: expand_fun_eq)
+ have "member \<circ> fold (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member) xs =
+ fold More_Set.remove xs \<circ> member"
+ by (rule fold_apply) (simp add: expand_fun_eq)
+ then have "fold More_Set.remove xs (member A) =
+ member (fold (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member) xs A)"
+ by (simp add: expand_fun_eq)
+ then have "inf A (Coset xs) = fold remove xs A"
+ by (simp add: Diff_eq [symmetric] minus_set *)
+ moreover have "\<And>x y :: 'a. Fset.remove y \<circ> Fset.remove x = Fset.remove x \<circ> Fset.remove y"
+ by (auto simp add: More_Set.remove_def * intro: ext)
+ ultimately show "inf A (Coset xs) = foldr remove xs A"
+ by (simp add: foldr_fold)
qed
lemma subtract_remove [code]:
- "A - Set xs = foldl (\<lambda>A x. remove x A) A xs"
+ "A - Set xs = foldr remove xs A"
"A - Coset xs = Set (List.filter (member A) xs)"
-proof -
- have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
- member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
- by (rule foldl_apply) (simp add: expand_fun_eq)
- then show "A - Set xs = foldl (\<lambda>A x. remove x A) A xs"
- by (simp add: minus_set)
- show "A - Coset xs = Set (List.filter (member A) xs)"
- by (auto simp add: Coset_def Set_def)
-qed
+ by (simp_all only: diff_eq compl_Set compl_Coset inter_project)
lemma union_insert [code]:
- "sup (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
+ "sup (Set xs) A = foldr insert xs A"
"sup (Coset xs) A = Coset (List.filter (Not \<circ> member A) xs)"
proof -
- have "foldl (\<lambda>A x. Set.insert x A) (member A) xs =
- member (foldl (\<lambda>A x. Fset (Set.insert x (member A))) A xs)"
- by (rule foldl_apply) (simp add: expand_fun_eq)
- then show "sup (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
- by (simp add: union_set)
+ have *: "\<And>x::'a. insert = (\<lambda>x. Fset \<circ> Set.insert x \<circ> member)"
+ by (simp add: expand_fun_eq)
+ have "member \<circ> fold (\<lambda>x. Fset \<circ> Set.insert x \<circ> member) xs =
+ fold Set.insert xs \<circ> member"
+ by (rule fold_apply) (simp add: expand_fun_eq)
+ then have "fold Set.insert xs (member A) =
+ member (fold (\<lambda>x. Fset \<circ> Set.insert x \<circ> member) xs A)"
+ by (simp add: expand_fun_eq)
+ then have "sup (Set xs) A = fold insert xs A"
+ by (simp add: union_set *)
+ moreover have "\<And>x y :: 'a. Fset.insert y \<circ> Fset.insert x = Fset.insert x \<circ> Fset.insert y"
+ by (auto simp add: * intro: ext)
+ ultimately show "sup (Set xs) A = foldr insert xs A"
+ by (simp add: foldr_fold)
show "sup (Coset xs) A = Coset (List.filter (Not \<circ> member A) xs)"
by (auto simp add: Coset_def)
qed
@@ -242,17 +261,17 @@
[simp]: "Infimum A = Inf (member A)"
lemma Infimum_inf [code]:
- "Infimum (Set As) = foldl inf top As"
+ "Infimum (Set As) = foldr inf As top"
"Infimum (Coset []) = bot"
- by (simp_all add: Inf_set_fold Inf_UNIV)
+ by (simp_all add: Inf_set_foldr Inf_UNIV)
definition Supremum :: "'a fset \<Rightarrow> 'a" where
[simp]: "Supremum A = Sup (member A)"
lemma Supremum_sup [code]:
- "Supremum (Set As) = foldl sup bot As"
+ "Supremum (Set As) = foldr sup As bot"
"Supremum (Coset []) = top"
- by (simp_all add: Sup_set_fold Sup_UNIV)
+ by (simp_all add: Sup_set_foldr Sup_UNIV)
end
@@ -277,17 +296,17 @@
lemma is_empty_simp [simp]:
"is_empty A \<longleftrightarrow> member A = {}"
- by (simp add: List_Set.is_empty_def)
+ by (simp add: More_Set.is_empty_def)
declare is_empty_def [simp del]
lemma remove_simp [simp]:
"remove x A = Fset (member A - {x})"
- by (simp add: List_Set.remove_def)
+ by (simp add: More_Set.remove_def)
declare remove_def [simp del]
lemma filter_simp [simp]:
"filter P A = Fset {x \<in> member A. P x}"
- by (simp add: List_Set.project_def)
+ by (simp add: More_Set.project_def)
declare filter_def [simp del]
declare mem_def [simp del]
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy Sun May 23 10:38:11 2010 +0100
@@ -214,23 +214,8 @@
hence "\<exists>z. ?P z n" ..}
moreover
{assume o: "odd n"
- from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
- have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
- Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
- ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
- also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2"
- apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
- by (simp add: power2_eq_square)
- finally
- have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
- Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
- 1"
- apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
- using right_inverse[OF b']
- by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] algebra_simps)
have th0: "cmod (complex_of_real (cmod b) / b) = 1"
- apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse algebra_simps )
- by (simp add: real_sqrt_mult[symmetric] th0)
+ using b by (simp add: norm_divide)
from o have "\<exists>m. n = Suc (2*m)" by presburger+
then obtain m where m: "n = Suc (2*m)" by blast
from unimodular_reduce_norm[OF th0] o
--- a/src/HOL/Library/Lattice_Algebras.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Lattice_Algebras.thy Sun May 23 10:38:11 2010 +0100
@@ -405,21 +405,15 @@
done
}
note b = this[OF refl[of a] refl[of b]]
- note addm = add_mono[of "0::'a" _ "0::'a", simplified]
- note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified]
have xy: "- ?x <= ?y"
apply (simp)
- apply (rule_tac y="0::'a" in order_trans)
- apply (rule addm2)
- apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
- apply (rule addm)
+ apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
done
have yx: "?y <= ?x"
apply (simp add:diff_def)
- apply (rule_tac y=0 in order_trans)
- apply (rule addm2, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+)
- apply (rule addm, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+)
+ apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
+ apply (simp_all add: mult_nonneg_nonpos mult_nonpos_nonneg)
done
have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
--- a/src/HOL/Library/Library.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Library.thy Sun May 23 10:38:11 2010 +0100
@@ -34,6 +34,7 @@
ListVector
Kleene_Algebra
Mapping
+ More_List
Multiset
Nat_Infinity
Nested_Environment
@@ -61,7 +62,6 @@
Transitive_Closure_Table
Univ_Poly
While_Combinator
- Word
Zorn
begin
end
--- a/src/HOL/Library/List_Set.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,114 +0,0 @@
-
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* Relating (finite) sets and lists *}
-
-theory List_Set
-imports Main
-begin
-
-subsection {* Various additional set functions *}
-
-definition is_empty :: "'a set \<Rightarrow> bool" where
- "is_empty A \<longleftrightarrow> A = {}"
-
-definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
- "remove x A = A - {x}"
-
-lemma fun_left_comm_idem_remove:
- "fun_left_comm_idem remove"
-proof -
- have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
- show ?thesis by (simp only: fun_left_comm_idem_remove rem)
-qed
-
-lemma minus_fold_remove:
- assumes "finite A"
- shows "B - A = fold remove B A"
-proof -
- have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
- show ?thesis by (simp only: rem assms minus_fold_remove)
-qed
-
-definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
- "project P A = {a\<in>A. P a}"
-
-
-subsection {* Basic set operations *}
-
-lemma is_empty_set:
- "is_empty (set xs) \<longleftrightarrow> null xs"
- by (simp add: is_empty_def null_empty)
-
-lemma ball_set:
- "(\<forall>x\<in>set xs. P x) \<longleftrightarrow> list_all P xs"
- by (rule list_ball_code)
-
-lemma bex_set:
- "(\<exists>x\<in>set xs. P x) \<longleftrightarrow> list_ex P xs"
- by (rule list_bex_code)
-
-lemma empty_set:
- "{} = set []"
- by simp
-
-lemma insert_set_compl:
- "insert x (- set xs) = - set (removeAll x xs)"
- by auto
-
-lemma remove_set_compl:
- "remove x (- set xs) = - set (List.insert x xs)"
- by (auto simp del: mem_def simp add: remove_def List.insert_def)
-
-lemma image_set:
- "image f (set xs) = set (map f xs)"
- by simp
-
-lemma project_set:
- "project P (set xs) = set (filter P xs)"
- by (auto simp add: project_def)
-
-
-subsection {* Functorial set operations *}
-
-lemma union_set:
- "set xs \<union> A = foldl (\<lambda>A x. Set.insert x A) A xs"
-proof -
- interpret fun_left_comm_idem Set.insert
- by (fact fun_left_comm_idem_insert)
- show ?thesis by (simp add: union_fold_insert fold_set)
-qed
-
-lemma minus_set:
- "A - set xs = foldl (\<lambda>A x. remove x A) A xs"
-proof -
- interpret fun_left_comm_idem remove
- by (fact fun_left_comm_idem_remove)
- show ?thesis
- by (simp add: minus_fold_remove [of _ A] fold_set)
-qed
-
-
-subsection {* Derived set operations *}
-
-lemma member:
- "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
- by simp
-
-lemma subset_eq:
- "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
- by (fact subset_eq)
-
-lemma subset:
- "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
- by (fact less_le_not_le)
-
-lemma set_eq:
- "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
- by (fact eq_iff)
-
-lemma inter:
- "A \<inter> B = project (\<lambda>x. x \<in> A) B"
- by (auto simp add: project_def)
-
-end
\ No newline at end of file
--- a/src/HOL/Library/Mapping.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Mapping.thy Sun May 23 10:38:11 2010 +0100
@@ -6,6 +6,30 @@
imports Main
begin
+lemma remove1_idem: (*FIXME move to List.thy*)
+ assumes "x \<notin> set xs"
+ shows "remove1 x xs = xs"
+ using assms by (induct xs) simp_all
+
+lemma remove1_insort [simp]:
+ "remove1 x (insort x xs) = xs"
+ by (induct xs) simp_all
+
+lemma sorted_list_of_set_remove:
+ assumes "finite A"
+ shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)"
+proof (cases "x \<in> A")
+ case False with assms have "x \<notin> set (sorted_list_of_set A)" by simp
+ with False show ?thesis by (simp add: remove1_idem)
+next
+ case True then obtain B where A: "A = insert x B" by (rule Set.set_insert)
+ with assms show ?thesis by simp
+qed
+
+lemma sorted_list_of_set_range [simp]:
+ "sorted_list_of_set {m..<n} = [m..<n]"
+ by (rule sorted_distinct_set_unique) simp_all
+
subsection {* Type definition and primitive operations *}
datatype ('a, 'b) mapping = Mapping "'a \<rightharpoonup> 'b"
@@ -29,16 +53,26 @@
"keys m = dom (lookup m)"
definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list" where
- "ordered_keys m = sorted_list_of_set (keys m)"
+ "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool" where
- "is_empty m \<longleftrightarrow> dom (lookup m) = {}"
+ "is_empty m \<longleftrightarrow> keys m = {}"
definition size :: "('a, 'b) mapping \<Rightarrow> nat" where
- "size m = (if finite (dom (lookup m)) then card (dom (lookup m)) else 0)"
+ "size m = (if finite (keys m) then card (keys m) else 0)"
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
- "replace k v m = (if lookup m k = None then m else update k v m)"
+ "replace k v m = (if k \<in> keys m then update k v m else m)"
+
+definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
+ "default k v m = (if k \<in> keys m then m else update k v m)"
+
+definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
+ "map_entry k f m = (case lookup m k of None \<Rightarrow> m
+ | Some v \<Rightarrow> update k (f v) m)"
+
+definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
+ "map_default k v f m = map_entry k f (default k v m)"
definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping" where
"tabulate ks f = Mapping (map_of (map (\<lambda>k. (k, f k)) ks))"
@@ -58,6 +92,10 @@
shows "m = n"
using assms by simp
+lemma keys_is_none_lookup [code_inline]:
+ "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"
+ by (auto simp add: keys_def is_none_def)
+
lemma lookup_empty [simp]:
"lookup empty = Map.empty"
by (simp add: empty_def)
@@ -70,6 +108,10 @@
"lookup (delete k m) = (lookup m) (k := None)"
by (cases m) simp
+lemma lookup_map_entry [simp]:
+ "lookup (map_entry k f m) = (lookup m) (k := Option.map f (lookup m k))"
+ by (cases "lookup m k") (simp_all add: map_entry_def expand_fun_eq)
+
lemma lookup_tabulate [simp]:
"lookup (tabulate ks f) = (Some o f) |` set ks"
by (induct ks) (auto simp add: tabulate_def restrict_map_def expand_fun_eq)
@@ -97,34 +139,157 @@
by (rule mapping_eqI) simp
lemma replace_update:
- "k \<notin> dom (lookup m) \<Longrightarrow> replace k v m = m"
- "k \<in> dom (lookup m) \<Longrightarrow> replace k v m = update k v m"
- by (rule mapping_eqI, auto simp add: replace_def fun_upd_twist)+
+ "k \<notin> keys m \<Longrightarrow> replace k v m = m"
+ "k \<in> keys m \<Longrightarrow> replace k v m = update k v m"
+ by (rule mapping_eqI) (auto simp add: replace_def fun_upd_twist)+
lemma size_empty [simp]:
"size empty = 0"
- by (simp add: size_def)
+ by (simp add: size_def keys_def)
lemma size_update:
- "finite (dom (lookup m)) \<Longrightarrow> size (update k v m) =
- (if k \<in> dom (lookup m) then size m else Suc (size m))"
- by (auto simp add: size_def insert_dom)
+ "finite (keys m) \<Longrightarrow> size (update k v m) =
+ (if k \<in> keys m then size m else Suc (size m))"
+ by (auto simp add: size_def insert_dom keys_def)
lemma size_delete:
- "size (delete k m) = (if k \<in> dom (lookup m) then size m - 1 else size m)"
- by (simp add: size_def)
+ "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"
+ by (simp add: size_def keys_def)
-lemma size_tabulate:
+lemma size_tabulate [simp]:
"size (tabulate ks f) = length (remdups ks)"
- by (simp add: size_def distinct_card [of "remdups ks", symmetric] comp_def)
+ by (simp add: size_def distinct_card [of "remdups ks", symmetric] comp_def keys_def)
lemma bulkload_tabulate:
"bulkload xs = tabulate [0..<length xs] (nth xs)"
by (rule mapping_eqI) (simp add: expand_fun_eq)
-lemma is_empty_empty:
+lemma is_empty_empty: (*FIXME*)
"is_empty m \<longleftrightarrow> m = Mapping Map.empty"
- by (cases m) (simp add: is_empty_def)
+ by (cases m) (simp add: is_empty_def keys_def)
+
+lemma is_empty_empty' [simp]:
+ "is_empty empty"
+ by (simp add: is_empty_empty empty_def)
+
+lemma is_empty_update [simp]:
+ "\<not> is_empty (update k v m)"
+ by (cases m) (simp add: is_empty_empty)
+
+lemma is_empty_delete:
+ "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
+ by (cases m) (auto simp add: is_empty_empty keys_def dom_eq_empty_conv [symmetric] simp del: dom_eq_empty_conv)
+
+lemma is_empty_replace [simp]:
+ "is_empty (replace k v m) \<longleftrightarrow> is_empty m"
+ by (auto simp add: replace_def) (simp add: is_empty_def)
+
+lemma is_empty_default [simp]:
+ "\<not> is_empty (default k v m)"
+ by (auto simp add: default_def) (simp add: is_empty_def)
+
+lemma is_empty_map_entry [simp]:
+ "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"
+ by (cases "lookup m k")
+ (auto simp add: map_entry_def, simp add: is_empty_empty)
+
+lemma is_empty_map_default [simp]:
+ "\<not> is_empty (map_default k v f m)"
+ by (simp add: map_default_def)
+
+lemma keys_empty [simp]:
+ "keys empty = {}"
+ by (simp add: keys_def)
+
+lemma keys_update [simp]:
+ "keys (update k v m) = insert k (keys m)"
+ by (simp add: keys_def)
+
+lemma keys_delete [simp]:
+ "keys (delete k m) = keys m - {k}"
+ by (simp add: keys_def)
+
+lemma keys_replace [simp]:
+ "keys (replace k v m) = keys m"
+ by (auto simp add: keys_def replace_def)
+
+lemma keys_default [simp]:
+ "keys (default k v m) = insert k (keys m)"
+ by (auto simp add: keys_def default_def)
+
+lemma keys_map_entry [simp]:
+ "keys (map_entry k f m) = keys m"
+ by (auto simp add: keys_def)
+
+lemma keys_map_default [simp]:
+ "keys (map_default k v f m) = insert k (keys m)"
+ by (simp add: map_default_def)
+
+lemma keys_tabulate [simp]:
+ "keys (tabulate ks f) = set ks"
+ by (simp add: tabulate_def keys_def map_of_map_restrict o_def)
+
+lemma keys_bulkload [simp]:
+ "keys (bulkload xs) = {0..<length xs}"
+ by (simp add: keys_tabulate bulkload_tabulate)
+
+lemma distinct_ordered_keys [simp]:
+ "distinct (ordered_keys m)"
+ by (simp add: ordered_keys_def)
+
+lemma ordered_keys_infinite [simp]:
+ "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"
+ by (simp add: ordered_keys_def)
+
+lemma ordered_keys_empty [simp]:
+ "ordered_keys empty = []"
+ by (simp add: ordered_keys_def)
+
+lemma ordered_keys_update [simp]:
+ "k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m"
+ "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (update k v m) = insort k (ordered_keys m)"
+ by (simp_all add: ordered_keys_def) (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)
+
+lemma ordered_keys_delete [simp]:
+ "ordered_keys (delete k m) = remove1 k (ordered_keys m)"
+proof (cases "finite (keys m)")
+ case False then show ?thesis by simp
+next
+ case True note fin = True
+ show ?thesis
+ proof (cases "k \<in> keys m")
+ case False with fin have "k \<notin> set (sorted_list_of_set (keys m))" by simp
+ with False show ?thesis by (simp add: ordered_keys_def remove1_idem)
+ next
+ case True with fin show ?thesis by (simp add: ordered_keys_def sorted_list_of_set_remove)
+ qed
+qed
+
+lemma ordered_keys_replace [simp]:
+ "ordered_keys (replace k v m) = ordered_keys m"
+ by (simp add: replace_def)
+
+lemma ordered_keys_default [simp]:
+ "k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m"
+ "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)"
+ by (simp_all add: default_def)
+
+lemma ordered_keys_map_entry [simp]:
+ "ordered_keys (map_entry k f m) = ordered_keys m"
+ by (simp add: ordered_keys_def)
+
+lemma ordered_keys_map_default [simp]:
+ "k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m"
+ "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)"
+ by (simp_all add: map_default_def)
+
+lemma ordered_keys_tabulate [simp]:
+ "ordered_keys (tabulate ks f) = sort (remdups ks)"
+ by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups)
+
+lemma ordered_keys_bulkload [simp]:
+ "ordered_keys (bulkload ks) = [0..<length ks]"
+ by (simp add: ordered_keys_def)
subsection {* Some technical code lemmas *}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/More_List.thy Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,267 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Operations on lists beyond the standard List theory *}
+
+theory More_List
+imports Main
+begin
+
+hide_const (open) Finite_Set.fold
+
+text {* Repairing code generator setup *}
+
+declare (in lattice) Inf_fin_set_fold [code_unfold del]
+declare (in lattice) Sup_fin_set_fold [code_unfold del]
+declare (in linorder) Min_fin_set_fold [code_unfold del]
+declare (in linorder) Max_fin_set_fold [code_unfold del]
+declare (in complete_lattice) Inf_set_fold [code_unfold del]
+declare (in complete_lattice) Sup_set_fold [code_unfold del]
+declare rev_foldl_cons [code del]
+
+text {* Fold combinator with canonical argument order *}
+
+primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
+ "fold f [] = id"
+ | "fold f (x # xs) = fold f xs \<circ> f x"
+
+lemma foldl_fold:
+ "foldl f s xs = fold (\<lambda>x s. f s x) xs s"
+ by (induct xs arbitrary: s) simp_all
+
+lemma foldr_fold_rev:
+ "foldr f xs = fold f (rev xs)"
+ by (simp add: foldr_foldl foldl_fold expand_fun_eq)
+
+lemma fold_rev_conv [code_unfold]:
+ "fold f (rev xs) = foldr f xs"
+ by (simp add: foldr_fold_rev)
+
+lemma fold_cong [fundef_cong, recdef_cong]:
+ "a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x)
+ \<Longrightarrow> fold f xs a = fold g ys b"
+ by (induct ys arbitrary: a b xs) simp_all
+
+lemma fold_id:
+ assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x = id"
+ shows "fold f xs = id"
+ using assms by (induct xs) simp_all
+
+lemma fold_apply:
+ assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
+ shows "h \<circ> fold g xs = fold f xs \<circ> h"
+ using assms by (induct xs) (simp_all add: expand_fun_eq)
+
+lemma fold_invariant:
+ assumes "\<And>x. x \<in> set xs \<Longrightarrow> Q x" and "P s"
+ and "\<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s)"
+ shows "P (fold f xs s)"
+ using assms by (induct xs arbitrary: s) simp_all
+
+lemma fold_weak_invariant:
+ assumes "P s"
+ and "\<And>s x. x \<in> set xs \<Longrightarrow> P s \<Longrightarrow> P (f x s)"
+ shows "P (fold f xs s)"
+ using assms by (induct xs arbitrary: s) simp_all
+
+lemma fold_append [simp]:
+ "fold f (xs @ ys) = fold f ys \<circ> fold f xs"
+ by (induct xs) simp_all
+
+lemma fold_map [code_unfold]:
+ "fold g (map f xs) = fold (g o f) xs"
+ by (induct xs) simp_all
+
+lemma fold_rev:
+ assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
+ shows "fold f (rev xs) = fold f xs"
+ using assms by (induct xs) (simp_all del: o_apply add: fold_apply)
+
+lemma foldr_fold:
+ assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
+ shows "foldr f xs = fold f xs"
+ using assms unfolding foldr_fold_rev by (rule fold_rev)
+
+lemma fold_Cons_rev:
+ "fold Cons xs = append (rev xs)"
+ by (induct xs) simp_all
+
+lemma rev_conv_fold [code]:
+ "rev xs = fold Cons xs []"
+ by (simp add: fold_Cons_rev)
+
+lemma fold_append_concat_rev:
+ "fold append xss = append (concat (rev xss))"
+ by (induct xss) simp_all
+
+lemma concat_conv_foldr [code]:
+ "concat xss = foldr append xss []"
+ by (simp add: fold_append_concat_rev foldr_fold_rev)
+
+lemma fold_plus_listsum_rev:
+ "fold plus xs = plus (listsum (rev xs))"
+ by (induct xs) (simp_all add: add.assoc)
+
+lemma listsum_conv_foldr [code]:
+ "listsum xs = foldr plus xs 0"
+ by (fact listsum_foldr)
+
+lemma sort_key_conv_fold:
+ assumes "inj_on f (set xs)"
+ shows "sort_key f xs = fold (insort_key f) xs []"
+proof -
+ have "fold (insort_key f) (rev xs) = fold (insort_key f) xs"
+ proof (rule fold_rev, rule ext)
+ fix zs
+ fix x y
+ assume "x \<in> set xs" "y \<in> set xs"
+ with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD)
+ show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs"
+ by (induct zs) (auto dest: *)
+ qed
+ then show ?thesis by (simp add: sort_key_def foldr_fold_rev)
+qed
+
+lemma sort_conv_fold:
+ "sort xs = fold insort xs []"
+ by (rule sort_key_conv_fold) simp
+
+text {* @{const Finite_Set.fold} and @{const fold} *}
+
+lemma (in fun_left_comm) fold_set_remdups:
+ "Finite_Set.fold f y (set xs) = fold f (remdups xs) y"
+ by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb)
+
+lemma (in fun_left_comm_idem) fold_set:
+ "Finite_Set.fold f y (set xs) = fold f xs y"
+ by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm)
+
+lemma (in ab_semigroup_idem_mult) fold1_set:
+ assumes "xs \<noteq> []"
+ shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)"
+proof -
+ interpret fun_left_comm_idem times by (fact fun_left_comm_idem)
+ from assms obtain y ys where xs: "xs = y # ys"
+ by (cases xs) auto
+ show ?thesis
+ proof (cases "set ys = {}")
+ case True with xs show ?thesis by simp
+ next
+ case False
+ then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)"
+ by (simp only: finite_set fold1_eq_fold_idem)
+ with xs show ?thesis by (simp add: fold_set mult_commute)
+ qed
+qed
+
+lemma (in lattice) Inf_fin_set_fold:
+ "Inf_fin (set (x # xs)) = fold inf xs x"
+proof -
+ interpret ab_semigroup_idem_mult "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact ab_semigroup_idem_mult_inf)
+ show ?thesis
+ by (simp add: Inf_fin_def fold1_set del: set.simps)
+qed
+
+lemma (in lattice) Inf_fin_set_foldr [code_unfold]:
+ "Inf_fin (set (x # xs)) = foldr inf xs x"
+ by (simp add: Inf_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)
+
+lemma (in lattice) Sup_fin_set_fold:
+ "Sup_fin (set (x # xs)) = fold sup xs x"
+proof -
+ interpret ab_semigroup_idem_mult "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact ab_semigroup_idem_mult_sup)
+ show ?thesis
+ by (simp add: Sup_fin_def fold1_set del: set.simps)
+qed
+
+lemma (in lattice) Sup_fin_set_foldr [code_unfold]:
+ "Sup_fin (set (x # xs)) = foldr sup xs x"
+ by (simp add: Sup_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)
+
+lemma (in linorder) Min_fin_set_fold:
+ "Min (set (x # xs)) = fold min xs x"
+proof -
+ interpret ab_semigroup_idem_mult "min :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact ab_semigroup_idem_mult_min)
+ show ?thesis
+ by (simp add: Min_def fold1_set del: set.simps)
+qed
+
+lemma (in linorder) Min_fin_set_foldr [code_unfold]:
+ "Min (set (x # xs)) = foldr min xs x"
+ by (simp add: Min_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)
+
+lemma (in linorder) Max_fin_set_fold:
+ "Max (set (x # xs)) = fold max xs x"
+proof -
+ interpret ab_semigroup_idem_mult "max :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact ab_semigroup_idem_mult_max)
+ show ?thesis
+ by (simp add: Max_def fold1_set del: set.simps)
+qed
+
+lemma (in linorder) Max_fin_set_foldr [code_unfold]:
+ "Max (set (x # xs)) = foldr max xs x"
+ by (simp add: Max_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)
+
+lemma (in complete_lattice) Inf_set_fold:
+ "Inf (set xs) = fold inf xs top"
+proof -
+ interpret fun_left_comm_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact fun_left_comm_idem_inf)
+ show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute)
+qed
+
+lemma (in complete_lattice) Inf_set_foldr [code_unfold]:
+ "Inf (set xs) = foldr inf xs top"
+ by (simp add: Inf_set_fold ac_simps foldr_fold expand_fun_eq)
+
+lemma (in complete_lattice) Sup_set_fold:
+ "Sup (set xs) = fold sup xs bot"
+proof -
+ interpret fun_left_comm_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ by (fact fun_left_comm_idem_sup)
+ show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute)
+qed
+
+lemma (in complete_lattice) Sup_set_foldr [code_unfold]:
+ "Sup (set xs) = foldr sup xs bot"
+ by (simp add: Sup_set_fold ac_simps foldr_fold expand_fun_eq)
+
+lemma (in complete_lattice) INFI_set_fold:
+ "INFI (set xs) f = fold (inf \<circ> f) xs top"
+ unfolding INFI_def set_map [symmetric] Inf_set_fold fold_map ..
+
+lemma (in complete_lattice) SUPR_set_fold:
+ "SUPR (set xs) f = fold (sup \<circ> f) xs bot"
+ unfolding SUPR_def set_map [symmetric] Sup_set_fold fold_map ..
+
+text {* @{text nth_map} *}
+
+definition nth_map :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "nth_map n f xs = (if n < length xs then
+ take n xs @ [f (xs ! n)] @ drop (Suc n) xs
+ else xs)"
+
+lemma nth_map_id:
+ "n \<ge> length xs \<Longrightarrow> nth_map n f xs = xs"
+ by (simp add: nth_map_def)
+
+lemma nth_map_unfold:
+ "n < length xs \<Longrightarrow> nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs"
+ by (simp add: nth_map_def)
+
+lemma nth_map_Nil [simp]:
+ "nth_map n f [] = []"
+ by (simp add: nth_map_def)
+
+lemma nth_map_zero [simp]:
+ "nth_map 0 f (x # xs) = f x # xs"
+ by (simp add: nth_map_def)
+
+lemma nth_map_Suc [simp]:
+ "nth_map (Suc n) f (x # xs) = x # nth_map n f xs"
+ by (simp add: nth_map_def)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/More_Set.thy Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,137 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Relating (finite) sets and lists *}
+
+theory More_Set
+imports Main More_List
+begin
+
+subsection {* Various additional set functions *}
+
+definition is_empty :: "'a set \<Rightarrow> bool" where
+ "is_empty A \<longleftrightarrow> A = {}"
+
+definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "remove x A = A - {x}"
+
+lemma fun_left_comm_idem_remove:
+ "fun_left_comm_idem remove"
+proof -
+ have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+ show ?thesis by (simp only: fun_left_comm_idem_remove rem)
+qed
+
+lemma minus_fold_remove:
+ assumes "finite A"
+ shows "B - A = Finite_Set.fold remove B A"
+proof -
+ have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+ show ?thesis by (simp only: rem assms minus_fold_remove)
+qed
+
+definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "project P A = {a\<in>A. P a}"
+
+
+subsection {* Basic set operations *}
+
+lemma is_empty_set:
+ "is_empty (set xs) \<longleftrightarrow> null xs"
+ by (simp add: is_empty_def null_empty)
+
+lemma ball_set:
+ "(\<forall>x\<in>set xs. P x) \<longleftrightarrow> list_all P xs"
+ by (rule list_ball_code)
+
+lemma bex_set:
+ "(\<exists>x\<in>set xs. P x) \<longleftrightarrow> list_ex P xs"
+ by (rule list_bex_code)
+
+lemma empty_set:
+ "{} = set []"
+ by simp
+
+lemma insert_set_compl:
+ "insert x (- set xs) = - set (removeAll x xs)"
+ by auto
+
+lemma remove_set_compl:
+ "remove x (- set xs) = - set (List.insert x xs)"
+ by (auto simp del: mem_def simp add: remove_def List.insert_def)
+
+lemma image_set:
+ "image f (set xs) = set (map f xs)"
+ by simp
+
+lemma project_set:
+ "project P (set xs) = set (filter P xs)"
+ by (auto simp add: project_def)
+
+
+subsection {* Functorial set operations *}
+
+lemma union_set:
+ "set xs \<union> A = fold Set.insert xs A"
+proof -
+ interpret fun_left_comm_idem Set.insert
+ by (fact fun_left_comm_idem_insert)
+ show ?thesis by (simp add: union_fold_insert fold_set)
+qed
+
+lemma union_set_foldr:
+ "set xs \<union> A = foldr Set.insert xs A"
+proof -
+ have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y"
+ by (auto intro: ext)
+ then show ?thesis by (simp add: union_set foldr_fold)
+qed
+
+lemma minus_set:
+ "A - set xs = fold remove xs A"
+proof -
+ interpret fun_left_comm_idem remove
+ by (fact fun_left_comm_idem_remove)
+ show ?thesis
+ by (simp add: minus_fold_remove [of _ A] fold_set)
+qed
+
+lemma minus_set_foldr:
+ "A - set xs = foldr remove xs A"
+proof -
+ have "\<And>x y :: 'a. remove y \<circ> remove x = remove x \<circ> remove y"
+ by (auto simp add: remove_def intro: ext)
+ then show ?thesis by (simp add: minus_set foldr_fold)
+qed
+
+
+subsection {* Derived set operations *}
+
+lemma member:
+ "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
+ by simp
+
+lemma subset_eq:
+ "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
+ by (fact subset_eq)
+
+lemma subset:
+ "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
+ by (fact less_le_not_le)
+
+lemma set_eq:
+ "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
+ by (fact eq_iff)
+
+lemma inter:
+ "A \<inter> B = project (\<lambda>x. x \<in> A) B"
+ by (auto simp add: project_def)
+
+
+subsection {* Various lemmas *}
+
+lemma not_set_compl:
+ "Not \<circ> set xs = - set xs"
+ by (simp add: fun_Compl_def bool_Compl_def comp_def expand_fun_eq)
+
+end
--- a/src/HOL/Library/Multiset.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Multiset.thy Sun May 23 10:38:11 2010 +0100
@@ -826,7 +826,8 @@
This lemma shows which properties suffice to show that a function
@{text "f"} with @{text "f xs = ys"} behaves like sort.
*}
-lemma properties_for_sort:
+
+lemma (in linorder) properties_for_sort:
"multiset_of ys = multiset_of xs \<Longrightarrow> sorted ys \<Longrightarrow> sort xs = ys"
proof (induct xs arbitrary: ys)
case Nil then show ?case by simp
--- a/src/HOL/Library/Quicksort.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Quicksort.thy Sun May 23 10:38:11 2010 +0100
@@ -2,7 +2,7 @@
Copyright 1994 TU Muenchen
*)
-header{*Quicksort*}
+header {* Quicksort *}
theory Quicksort
imports Main Multiset
@@ -12,22 +12,27 @@
begin
fun quicksort :: "'a list \<Rightarrow> 'a list" where
-"quicksort [] = []" |
-"quicksort (x#xs) = quicksort([y\<leftarrow>xs. ~ x\<le>y]) @ [x] @ quicksort([y\<leftarrow>xs. x\<le>y])"
+ "quicksort [] = []"
+| "quicksort (x#xs) = quicksort [y\<leftarrow>xs. \<not> x\<le>y] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]"
+
+lemma [code]:
+ "quicksort [] = []"
+ "quicksort (x#xs) = quicksort [y\<leftarrow>xs. y<x] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]"
+ by (simp_all add: not_le)
lemma quicksort_permutes [simp]:
"multiset_of (quicksort xs) = multiset_of xs"
-by (induct xs rule: quicksort.induct) (auto simp: union_ac)
+ by (induct xs rule: quicksort.induct) (simp_all add: ac_simps)
lemma set_quicksort [simp]: "set (quicksort xs) = set xs"
-by(simp add: set_count_greater_0)
+ by (simp add: set_count_greater_0)
-lemma sorted_quicksort: "sorted(quicksort xs)"
-apply (induct xs rule: quicksort.induct)
- apply simp
-apply (simp add:sorted_Cons sorted_append not_le less_imp_le)
-apply (metis leD le_cases le_less_trans)
-done
+lemma sorted_quicksort: "sorted (quicksort xs)"
+ by (induct xs rule: quicksort.induct) (auto simp add: sorted_Cons sorted_append not_le less_imp_le)
+
+theorem quicksort_sort [code_unfold]:
+ "sort = quicksort"
+ by (rule ext, rule properties_for_sort) (fact quicksort_permutes sorted_quicksort)+
end
--- a/src/HOL/Library/RBT.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/RBT.thy Sun May 23 10:38:11 2010 +0100
@@ -136,7 +136,7 @@
lemma lookup_map_entry [simp]:
"lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"
- by (simp add: map_entry_def lookup_RBT lookup_map_entry lookup_impl_of)
+ by (simp add: map_entry_def lookup_RBT RBT_Impl.lookup_map_entry lookup_impl_of)
lemma lookup_map [simp]:
"lookup (map f t) k = Option.map (f k) (lookup t k)"
@@ -191,7 +191,11 @@
by (rule mapping_eqI) simp
lemma delete_Mapping [code]:
- "Mapping.delete k (Mapping xs) = Mapping (delete k xs)"
+ "Mapping.delete k (Mapping t) = Mapping (delete k t)"
+ by (rule mapping_eqI) simp
+
+lemma map_entry_Mapping [code]:
+ "Mapping.map_entry k f (Mapping t) = Mapping (map_entry k f t)"
by (rule mapping_eqI) simp
lemma keys_Mapping [code]:
@@ -218,7 +222,8 @@
"Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"
by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
-lemma [code, code del]: "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+lemma [code, code del]:
+ "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
lemma eq_Mapping [code]:
"HOL.eq (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
--- a/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML Sun May 23 10:38:11 2010 +0100
@@ -144,11 +144,11 @@
val setup =
Method.setup @{binding sos}
- (Scan.lift (Scan.option OuterParse.xname)
+ (Scan.lift (Scan.option Parse.xname)
>> (sos_solver print_cert o Sum_Of_Squares.Prover o call_solver))
"prove universal problems over the reals using sums of squares" #>
Method.setup @{binding sos_cert}
- (Scan.lift OuterParse.string
+ (Scan.lift Parse.string
>> (fn cert => fn ctxt =>
sos_solver ignore
(Sum_Of_Squares.Certificate (PositivstellensatzTools.cert_to_pss_tree ctxt cert)) ctxt))
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Sun May 23 10:38:11 2010 +0100
@@ -1272,7 +1272,7 @@
fun subst_conv eqs t =
let
val t' = fold (Thm.cabs o Thm.lhs_of) eqs t
- in Conv.fconv_rule (Thm.beta_conversion true) (fold (C combination) eqs (reflexive t'))
+ in Conv.fconv_rule (Thm.beta_conversion true) (fold (C Thm.combination) eqs (Thm.reflexive t'))
end
(* A wrapper that tries to substitute away variables first. *)
--- a/src/HOL/Library/Word.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2314 +0,0 @@
-(* Title: HOL/Library/Word.thy
- Author: Sebastian Skalberg, TU Muenchen
-*)
-
-header {* Binary Words *}
-
-theory Word
-imports Main
-begin
-
-subsection {* Auxilary Lemmas *}
-
-lemma max_le [intro!]: "[| x \<le> z; y \<le> z |] ==> max x y \<le> z"
- by (simp add: max_def)
-
-lemma max_mono:
- fixes x :: "'a::linorder"
- assumes mf: "mono f"
- shows "max (f x) (f y) \<le> f (max x y)"
-proof -
- from mf and le_maxI1 [of x y]
- have fx: "f x \<le> f (max x y)" by (rule monoD)
- from mf and le_maxI2 [of y x]
- have fy: "f y \<le> f (max x y)" by (rule monoD)
- from fx and fy
- show "max (f x) (f y) \<le> f (max x y)" by auto
-qed
-
-declare zero_le_power [intro]
- and zero_less_power [intro]
-
-lemma int_nat_two_exp: "2 ^ k = int (2 ^ k)"
- by (simp add: zpower_int [symmetric])
-
-
-subsection {* Bits *}
-
-datatype bit =
- Zero ("\<zero>")
- | One ("\<one>")
-
-primrec bitval :: "bit => nat" where
- "bitval \<zero> = 0"
- | "bitval \<one> = 1"
-
-primrec bitnot :: "bit => bit" where
- bitnot_zero: "(bitnot \<zero>) = \<one>"
- | bitnot_one : "(bitnot \<one>) = \<zero>"
-
-primrec bitand :: "bit => bit => bit" (infixr "bitand" 35) where
- bitand_zero: "(\<zero> bitand y) = \<zero>"
- | bitand_one: "(\<one> bitand y) = y"
-
-primrec bitor :: "bit => bit => bit" (infixr "bitor" 30) where
- bitor_zero: "(\<zero> bitor y) = y"
- | bitor_one: "(\<one> bitor y) = \<one>"
-
-primrec bitxor :: "bit => bit => bit" (infixr "bitxor" 30) where
- bitxor_zero: "(\<zero> bitxor y) = y"
- | bitxor_one: "(\<one> bitxor y) = (bitnot y)"
-
-notation (xsymbols)
- bitnot ("\<not>\<^sub>b _" [40] 40) and
- bitand (infixr "\<and>\<^sub>b" 35) and
- bitor (infixr "\<or>\<^sub>b" 30) and
- bitxor (infixr "\<oplus>\<^sub>b" 30)
-
-notation (HTML output)
- bitnot ("\<not>\<^sub>b _" [40] 40) and
- bitand (infixr "\<and>\<^sub>b" 35) and
- bitor (infixr "\<or>\<^sub>b" 30) and
- bitxor (infixr "\<oplus>\<^sub>b" 30)
-
-lemma bitnot_bitnot [simp]: "(bitnot (bitnot b)) = b"
- by (cases b) simp_all
-
-lemma bitand_cancel [simp]: "(b bitand b) = b"
- by (cases b) simp_all
-
-lemma bitor_cancel [simp]: "(b bitor b) = b"
- by (cases b) simp_all
-
-lemma bitxor_cancel [simp]: "(b bitxor b) = \<zero>"
- by (cases b) simp_all
-
-
-subsection {* Bit Vectors *}
-
-text {* First, a couple of theorems expressing case analysis and
-induction principles for bit vectors. *}
-
-lemma bit_list_cases:
- assumes empty: "w = [] ==> P w"
- and zero: "!!bs. w = \<zero> # bs ==> P w"
- and one: "!!bs. w = \<one> # bs ==> P w"
- shows "P w"
-proof (cases w)
- assume "w = []"
- thus ?thesis by (rule empty)
-next
- fix b bs
- assume [simp]: "w = b # bs"
- show "P w"
- proof (cases b)
- assume "b = \<zero>"
- hence "w = \<zero> # bs" by simp
- thus ?thesis by (rule zero)
- next
- assume "b = \<one>"
- hence "w = \<one> # bs" by simp
- thus ?thesis by (rule one)
- qed
-qed
-
-lemma bit_list_induct:
- assumes empty: "P []"
- and zero: "!!bs. P bs ==> P (\<zero>#bs)"
- and one: "!!bs. P bs ==> P (\<one>#bs)"
- shows "P w"
-proof (induct w, simp_all add: empty)
- fix b bs
- assume "P bs"
- then show "P (b#bs)"
- by (cases b) (auto intro!: zero one)
-qed
-
-definition
- bv_msb :: "bit list => bit" where
- "bv_msb w = (if w = [] then \<zero> else hd w)"
-
-definition
- bv_extend :: "[nat,bit,bit list]=>bit list" where
- "bv_extend i b w = (replicate (i - length w) b) @ w"
-
-definition
- bv_not :: "bit list => bit list" where
- "bv_not w = map bitnot w"
-
-lemma bv_length_extend [simp]: "length w \<le> i ==> length (bv_extend i b w) = i"
- by (simp add: bv_extend_def)
-
-lemma bv_not_Nil [simp]: "bv_not [] = []"
- by (simp add: bv_not_def)
-
-lemma bv_not_Cons [simp]: "bv_not (b#bs) = (bitnot b) # bv_not bs"
- by (simp add: bv_not_def)
-
-lemma bv_not_bv_not [simp]: "bv_not (bv_not w) = w"
- by (rule bit_list_induct [of _ w]) simp_all
-
-lemma bv_msb_Nil [simp]: "bv_msb [] = \<zero>"
- by (simp add: bv_msb_def)
-
-lemma bv_msb_Cons [simp]: "bv_msb (b#bs) = b"
- by (simp add: bv_msb_def)
-
-lemma bv_msb_bv_not [simp]: "0 < length w ==> bv_msb (bv_not w) = (bitnot (bv_msb w))"
- by (cases w) simp_all
-
-lemma bv_msb_one_length [simp,intro]: "bv_msb w = \<one> ==> 0 < length w"
- by (cases w) simp_all
-
-lemma length_bv_not [simp]: "length (bv_not w) = length w"
- by (induct w) simp_all
-
-definition
- bv_to_nat :: "bit list => nat" where
- "bv_to_nat = foldl (%bn b. 2 * bn + bitval b) 0"
-
-lemma bv_to_nat_Nil [simp]: "bv_to_nat [] = 0"
- by (simp add: bv_to_nat_def)
-
-lemma bv_to_nat_helper [simp]: "bv_to_nat (b # bs) = bitval b * 2 ^ length bs + bv_to_nat bs"
-proof -
- let ?bv_to_nat' = "foldl (\<lambda>bn b. 2 * bn + bitval b)"
- have helper: "\<And>base. ?bv_to_nat' base bs = base * 2 ^ length bs + ?bv_to_nat' 0 bs"
- proof (induct bs)
- case Nil
- show ?case by simp
- next
- case (Cons x xs base)
- show ?case
- apply (simp only: foldl.simps)
- apply (subst Cons [of "2 * base + bitval x"])
- apply simp
- apply (subst Cons [of "bitval x"])
- apply (simp add: add_mult_distrib)
- done
- qed
- show ?thesis by (simp add: bv_to_nat_def) (rule helper)
-qed
-
-lemma bv_to_nat0 [simp]: "bv_to_nat (\<zero>#bs) = bv_to_nat bs"
- by simp
-
-lemma bv_to_nat1 [simp]: "bv_to_nat (\<one>#bs) = 2 ^ length bs + bv_to_nat bs"
- by simp
-
-lemma bv_to_nat_upper_range: "bv_to_nat w < 2 ^ length w"
-proof (induct w, simp_all)
- fix b bs
- assume "bv_to_nat bs < 2 ^ length bs"
- show "bitval b * 2 ^ length bs + bv_to_nat bs < 2 * 2 ^ length bs"
- proof (cases b, simp_all)
- have "bv_to_nat bs < 2 ^ length bs" by fact
- also have "... < 2 * 2 ^ length bs" by auto
- finally show "bv_to_nat bs < 2 * 2 ^ length bs" by simp
- next
- have "bv_to_nat bs < 2 ^ length bs" by fact
- hence "2 ^ length bs + bv_to_nat bs < 2 ^ length bs + 2 ^ length bs" by arith
- also have "... = 2 * (2 ^ length bs)" by simp
- finally show "bv_to_nat bs < 2 ^ length bs" by simp
- qed
-qed
-
-lemma bv_extend_longer [simp]:
- assumes wn: "n \<le> length w"
- shows "bv_extend n b w = w"
- by (simp add: bv_extend_def wn)
-
-lemma bv_extend_shorter [simp]:
- assumes wn: "length w < n"
- shows "bv_extend n b w = bv_extend n b (b#w)"
-proof -
- from wn
- have s: "n - Suc (length w) + 1 = n - length w"
- by arith
- have "bv_extend n b w = replicate (n - length w) b @ w"
- by (simp add: bv_extend_def)
- also have "... = replicate (n - Suc (length w) + 1) b @ w"
- by (subst s) rule
- also have "... = (replicate (n - Suc (length w)) b @ replicate 1 b) @ w"
- by (subst replicate_add) rule
- also have "... = replicate (n - Suc (length w)) b @ b # w"
- by simp
- also have "... = bv_extend n b (b#w)"
- by (simp add: bv_extend_def)
- finally show "bv_extend n b w = bv_extend n b (b#w)" .
-qed
-
-primrec rem_initial :: "bit => bit list => bit list" where
- "rem_initial b [] = []"
- | "rem_initial b (x#xs) = (if b = x then rem_initial b xs else x#xs)"
-
-lemma rem_initial_length: "length (rem_initial b w) \<le> length w"
- by (rule bit_list_induct [of _ w],simp_all (no_asm),safe,simp_all)
-
-lemma rem_initial_equal:
- assumes p: "length (rem_initial b w) = length w"
- shows "rem_initial b w = w"
-proof -
- have "length (rem_initial b w) = length w --> rem_initial b w = w"
- proof (induct w, simp_all, clarify)
- fix xs
- assume "length (rem_initial b xs) = length xs --> rem_initial b xs = xs"
- assume f: "length (rem_initial b xs) = Suc (length xs)"
- with rem_initial_length [of b xs]
- show "rem_initial b xs = b#xs"
- by auto
- qed
- from this and p show ?thesis ..
-qed
-
-lemma bv_extend_rem_initial: "bv_extend (length w) b (rem_initial b w) = w"
-proof (induct w, simp_all, safe)
- fix xs
- assume ind: "bv_extend (length xs) b (rem_initial b xs) = xs"
- from rem_initial_length [of b xs]
- have [simp]: "Suc (length xs) - length (rem_initial b xs) =
- 1 + (length xs - length (rem_initial b xs))"
- by arith
- have "bv_extend (Suc (length xs)) b (rem_initial b xs) =
- replicate (Suc (length xs) - length (rem_initial b xs)) b @ (rem_initial b xs)"
- by (simp add: bv_extend_def)
- also have "... =
- replicate (1 + (length xs - length (rem_initial b xs))) b @ rem_initial b xs"
- by simp
- also have "... =
- (replicate 1 b @ replicate (length xs - length (rem_initial b xs)) b) @ rem_initial b xs"
- by (subst replicate_add) (rule refl)
- also have "... = b # bv_extend (length xs) b (rem_initial b xs)"
- by (auto simp add: bv_extend_def [symmetric])
- also have "... = b # xs"
- by (simp add: ind)
- finally show "bv_extend (Suc (length xs)) b (rem_initial b xs) = b # xs" .
-qed
-
-lemma rem_initial_append1:
- assumes "rem_initial b xs ~= []"
- shows "rem_initial b (xs @ ys) = rem_initial b xs @ ys"
- using assms by (induct xs) auto
-
-lemma rem_initial_append2:
- assumes "rem_initial b xs = []"
- shows "rem_initial b (xs @ ys) = rem_initial b ys"
- using assms by (induct xs) auto
-
-definition
- norm_unsigned :: "bit list => bit list" where
- "norm_unsigned = rem_initial \<zero>"
-
-lemma norm_unsigned_Nil [simp]: "norm_unsigned [] = []"
- by (simp add: norm_unsigned_def)
-
-lemma norm_unsigned_Cons0 [simp]: "norm_unsigned (\<zero>#bs) = norm_unsigned bs"
- by (simp add: norm_unsigned_def)
-
-lemma norm_unsigned_Cons1 [simp]: "norm_unsigned (\<one>#bs) = \<one>#bs"
- by (simp add: norm_unsigned_def)
-
-lemma norm_unsigned_idem [simp]: "norm_unsigned (norm_unsigned w) = norm_unsigned w"
- by (rule bit_list_induct [of _ w],simp_all)
-
-fun
- nat_to_bv_helper :: "nat => bit list => bit list"
-where
- "nat_to_bv_helper n bs = (if n = 0 then bs
- else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \<zero> else \<one>)#bs))"
-
-definition
- nat_to_bv :: "nat => bit list" where
- "nat_to_bv n = nat_to_bv_helper n []"
-
-lemma nat_to_bv0 [simp]: "nat_to_bv 0 = []"
- by (simp add: nat_to_bv_def)
-
-lemmas [simp del] = nat_to_bv_helper.simps
-
-lemma n_div_2_cases:
- assumes zero: "(n::nat) = 0 ==> R"
- and div : "[| n div 2 < n ; 0 < n |] ==> R"
- shows "R"
-proof (cases "n = 0")
- assume "n = 0"
- thus R by (rule zero)
-next
- assume "n ~= 0"
- hence "0 < n" by simp
- hence "n div 2 < n" by arith
- from this and `0 < n` show R by (rule div)
-qed
-
-lemma int_wf_ge_induct:
- assumes ind : "!!i::int. (!!j. [| k \<le> j ; j < i |] ==> P j) ==> P i"
- shows "P i"
-proof (rule wf_induct_rule [OF wf_int_ge_less_than])
- fix x
- assume ih: "(\<And>y\<Colon>int. (y, x) \<in> int_ge_less_than k \<Longrightarrow> P y)"
- thus "P x"
- by (rule ind) (simp add: int_ge_less_than_def)
-qed
-
-lemma unfold_nat_to_bv_helper:
- "nat_to_bv_helper b l = nat_to_bv_helper b [] @ l"
-proof -
- have "\<forall>l. nat_to_bv_helper b l = nat_to_bv_helper b [] @ l"
- proof (induct b rule: less_induct)
- fix n
- assume ind: "!!j. j < n \<Longrightarrow> \<forall> l. nat_to_bv_helper j l = nat_to_bv_helper j [] @ l"
- show "\<forall>l. nat_to_bv_helper n l = nat_to_bv_helper n [] @ l"
- proof
- fix l
- show "nat_to_bv_helper n l = nat_to_bv_helper n [] @ l"
- proof (cases "n < 0")
- assume "n < 0"
- thus ?thesis
- by (simp add: nat_to_bv_helper.simps)
- next
- assume "~n < 0"
- show ?thesis
- proof (rule n_div_2_cases [of n])
- assume [simp]: "n = 0"
- show ?thesis
- apply (simp only: nat_to_bv_helper.simps [of n])
- apply simp
- done
- next
- assume n2n: "n div 2 < n"
- assume [simp]: "0 < n"
- hence n20: "0 \<le> n div 2"
- by arith
- from ind [of "n div 2"] and n2n n20
- have ind': "\<forall>l. nat_to_bv_helper (n div 2) l = nat_to_bv_helper (n div 2) [] @ l"
- by blast
- show ?thesis
- apply (simp only: nat_to_bv_helper.simps [of n])
- apply (cases "n=0")
- apply simp
- apply (simp only: if_False)
- apply simp
- apply (subst spec [OF ind',of "\<zero>#l"])
- apply (subst spec [OF ind',of "\<one>#l"])
- apply (subst spec [OF ind',of "[\<one>]"])
- apply (subst spec [OF ind',of "[\<zero>]"])
- apply simp
- done
- qed
- qed
- qed
- qed
- thus ?thesis ..
-qed
-
-lemma nat_to_bv_non0 [simp]: "n\<noteq>0 ==> nat_to_bv n = nat_to_bv (n div 2) @ [if n mod 2 = 0 then \<zero> else \<one>]"
-proof -
- assume [simp]: "n\<noteq>0"
- show ?thesis
- apply (subst nat_to_bv_def [of n])
- apply (simp only: nat_to_bv_helper.simps [of n])
- apply (subst unfold_nat_to_bv_helper)
- using prems
- apply (simp)
- apply (subst nat_to_bv_def [of "n div 2"])
- apply auto
- done
-qed
-
-lemma bv_to_nat_dist_append:
- "bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2"
-proof -
- have "\<forall>l2. bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2"
- proof (induct l1, simp_all)
- fix x xs
- assume ind: "\<forall>l2. bv_to_nat (xs @ l2) = bv_to_nat xs * 2 ^ length l2 + bv_to_nat l2"
- show "\<forall>l2::bit list. bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
- proof
- fix l2
- show "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
- proof -
- have "(2::nat) ^ (length xs + length l2) = 2 ^ length xs * 2 ^ length l2"
- by (induct ("length xs")) simp_all
- hence "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 =
- bitval x * 2 ^ length xs * 2 ^ length l2 + bv_to_nat xs * 2 ^ length l2"
- by simp
- also have "... = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
- by (simp add: ring_distribs)
- finally show ?thesis by simp
- qed
- qed
- qed
- thus ?thesis ..
-qed
-
-lemma bv_nat_bv [simp]: "bv_to_nat (nat_to_bv n) = n"
-proof (induct n rule: less_induct)
- fix n
- assume ind: "!!j. j < n \<Longrightarrow> bv_to_nat (nat_to_bv j) = j"
- show "bv_to_nat (nat_to_bv n) = n"
- proof (rule n_div_2_cases [of n])
- assume "n = 0" then show ?thesis by simp
- next
- assume nn: "n div 2 < n"
- assume n0: "0 < n"
- from ind and nn
- have ind': "bv_to_nat (nat_to_bv (n div 2)) = n div 2" by blast
- from n0 have n0': "n \<noteq> 0" by simp
- show ?thesis
- apply (subst nat_to_bv_def)
- apply (simp only: nat_to_bv_helper.simps [of n])
- apply (simp only: n0' if_False)
- apply (subst unfold_nat_to_bv_helper)
- apply (subst bv_to_nat_dist_append)
- apply (fold nat_to_bv_def)
- apply (simp add: ind' split del: split_if)
- apply (cases "n mod 2 = 0")
- proof (simp_all)
- assume "n mod 2 = 0"
- with mod_div_equality [of n 2]
- show "n div 2 * 2 = n" by simp
- next
- assume "n mod 2 = Suc 0"
- with mod_div_equality [of n 2]
- show "Suc (n div 2 * 2) = n" by arith
- qed
- qed
-qed
-
-lemma bv_to_nat_type [simp]: "bv_to_nat (norm_unsigned w) = bv_to_nat w"
- by (rule bit_list_induct) simp_all
-
-lemma length_norm_unsigned_le [simp]: "length (norm_unsigned w) \<le> length w"
- by (rule bit_list_induct) simp_all
-
-lemma bv_to_nat_rew_msb: "bv_msb w = \<one> ==> bv_to_nat w = 2 ^ (length w - 1) + bv_to_nat (tl w)"
- by (rule bit_list_cases [of w]) simp_all
-
-lemma norm_unsigned_result: "norm_unsigned xs = [] \<or> bv_msb (norm_unsigned xs) = \<one>"
-proof (rule length_induct [of _ xs])
- fix xs :: "bit list"
- assume ind: "\<forall>ys. length ys < length xs --> norm_unsigned ys = [] \<or> bv_msb (norm_unsigned ys) = \<one>"
- show "norm_unsigned xs = [] \<or> bv_msb (norm_unsigned xs) = \<one>"
- proof (rule bit_list_cases [of xs],simp_all)
- fix bs
- assume [simp]: "xs = \<zero>#bs"
- from ind
- have "length bs < length xs --> norm_unsigned bs = [] \<or> bv_msb (norm_unsigned bs) = \<one>" ..
- thus "norm_unsigned bs = [] \<or> bv_msb (norm_unsigned bs) = \<one>" by simp
- qed
-qed
-
-lemma norm_empty_bv_to_nat_zero:
- assumes nw: "norm_unsigned w = []"
- shows "bv_to_nat w = 0"
-proof -
- have "bv_to_nat w = bv_to_nat (norm_unsigned w)" by simp
- also have "... = bv_to_nat []" by (subst nw) (rule refl)
- also have "... = 0" by simp
- finally show ?thesis .
-qed
-
-lemma bv_to_nat_lower_limit:
- assumes w0: "0 < bv_to_nat w"
- shows "2 ^ (length (norm_unsigned w) - 1) \<le> bv_to_nat w"
-proof -
- from w0 and norm_unsigned_result [of w]
- have msbw: "bv_msb (norm_unsigned w) = \<one>"
- by (auto simp add: norm_empty_bv_to_nat_zero)
- have "2 ^ (length (norm_unsigned w) - 1) \<le> bv_to_nat (norm_unsigned w)"
- by (subst bv_to_nat_rew_msb [OF msbw],simp)
- thus ?thesis by simp
-qed
-
-lemmas [simp del] = nat_to_bv_non0
-
-lemma norm_unsigned_length [intro!]: "length (norm_unsigned w) \<le> length w"
-by (subst norm_unsigned_def,rule rem_initial_length)
-
-lemma norm_unsigned_equal:
- "length (norm_unsigned w) = length w ==> norm_unsigned w = w"
-by (simp add: norm_unsigned_def,rule rem_initial_equal)
-
-lemma bv_extend_norm_unsigned: "bv_extend (length w) \<zero> (norm_unsigned w) = w"
-by (simp add: norm_unsigned_def,rule bv_extend_rem_initial)
-
-lemma norm_unsigned_append1 [simp]:
- "norm_unsigned xs \<noteq> [] ==> norm_unsigned (xs @ ys) = norm_unsigned xs @ ys"
-by (simp add: norm_unsigned_def,rule rem_initial_append1)
-
-lemma norm_unsigned_append2 [simp]:
- "norm_unsigned xs = [] ==> norm_unsigned (xs @ ys) = norm_unsigned ys"
-by (simp add: norm_unsigned_def,rule rem_initial_append2)
-
-lemma bv_to_nat_zero_imp_empty:
- "bv_to_nat w = 0 \<Longrightarrow> norm_unsigned w = []"
-by (atomize (full), induct w rule: bit_list_induct) simp_all
-
-lemma bv_to_nat_nzero_imp_nempty:
- "bv_to_nat w \<noteq> 0 \<Longrightarrow> norm_unsigned w \<noteq> []"
-by (induct w rule: bit_list_induct) simp_all
-
-lemma nat_helper1:
- assumes ass: "nat_to_bv (bv_to_nat w) = norm_unsigned w"
- shows "nat_to_bv (2 * bv_to_nat w + bitval x) = norm_unsigned (w @ [x])"
-proof (cases x)
- assume [simp]: "x = \<one>"
- have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\<one>] =
- nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\<one>]"
- by (simp add: add_commute)
- also have "... = nat_to_bv (bv_to_nat w) @ [\<one>]"
- by (subst div_add1_eq) simp
- also have "... = norm_unsigned w @ [\<one>]"
- by (subst ass) (rule refl)
- also have "... = norm_unsigned (w @ [\<one>])"
- by (cases "norm_unsigned w") simp_all
- finally have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\<one>] = norm_unsigned (w @ [\<one>])" .
- then show ?thesis by (simp add: nat_to_bv_non0)
-next
- assume [simp]: "x = \<zero>"
- show ?thesis
- proof (cases "bv_to_nat w = 0")
- assume "bv_to_nat w = 0"
- thus ?thesis
- by (simp add: bv_to_nat_zero_imp_empty)
- next
- assume "bv_to_nat w \<noteq> 0"
- thus ?thesis
- apply simp
- apply (subst nat_to_bv_non0)
- apply simp
- apply auto
- apply (subst ass)
- apply (cases "norm_unsigned w")
- apply (simp_all add: norm_empty_bv_to_nat_zero)
- done
- qed
-qed
-
-lemma nat_helper2: "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \<one> # xs"
-proof -
- have "\<forall>xs. nat_to_bv (2 ^ length (rev xs) + bv_to_nat (rev xs)) = \<one> # (rev xs)" (is "\<forall>xs. ?P xs")
- proof
- fix xs
- show "?P xs"
- proof (rule length_induct [of _ xs])
- fix xs :: "bit list"
- assume ind: "\<forall>ys. length ys < length xs --> ?P ys"
- show "?P xs"
- proof (cases xs)
- assume "xs = []"
- then show ?thesis by (simp add: nat_to_bv_non0)
- next
- fix y ys
- assume [simp]: "xs = y # ys"
- show ?thesis
- apply simp
- apply (subst bv_to_nat_dist_append)
- apply simp
- proof -
- have "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) =
- nat_to_bv (2 * (2 ^ length ys + bv_to_nat (rev ys)) + bitval y)"
- by (simp add: add_ac mult_ac)
- also have "... = nat_to_bv (2 * (bv_to_nat (\<one>#rev ys)) + bitval y)"
- by simp
- also have "... = norm_unsigned (\<one>#rev ys) @ [y]"
- proof -
- from ind
- have "nat_to_bv (2 ^ length (rev ys) + bv_to_nat (rev ys)) = \<one> # rev ys"
- by auto
- hence [simp]: "nat_to_bv (2 ^ length ys + bv_to_nat (rev ys)) = \<one> # rev ys"
- by simp
- show ?thesis
- apply (subst nat_helper1)
- apply simp_all
- done
- qed
- also have "... = (\<one>#rev ys) @ [y]"
- by simp
- also have "... = \<one> # rev ys @ [y]"
- by simp
- finally show "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) =
- \<one> # rev ys @ [y]" .
- qed
- qed
- qed
- qed
- hence "nat_to_bv (2 ^ length (rev (rev xs)) + bv_to_nat (rev (rev xs))) =
- \<one> # rev (rev xs)" ..
- thus ?thesis by simp
-qed
-
-lemma nat_bv_nat [simp]: "nat_to_bv (bv_to_nat w) = norm_unsigned w"
-proof (rule bit_list_induct [of _ w],simp_all)
- fix xs
- assume "nat_to_bv (bv_to_nat xs) = norm_unsigned xs"
- have "bv_to_nat xs = bv_to_nat (norm_unsigned xs)" by simp
- have "bv_to_nat xs < 2 ^ length xs"
- by (rule bv_to_nat_upper_range)
- show "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \<one> # xs"
- by (rule nat_helper2)
-qed
-
-lemma bv_to_nat_qinj:
- assumes one: "bv_to_nat xs = bv_to_nat ys"
- and len: "length xs = length ys"
- shows "xs = ys"
-proof -
- from one
- have "nat_to_bv (bv_to_nat xs) = nat_to_bv (bv_to_nat ys)"
- by simp
- hence xsys: "norm_unsigned xs = norm_unsigned ys"
- by simp
- have "xs = bv_extend (length xs) \<zero> (norm_unsigned xs)"
- by (simp add: bv_extend_norm_unsigned)
- also have "... = bv_extend (length ys) \<zero> (norm_unsigned ys)"
- by (simp add: xsys len)
- also have "... = ys"
- by (simp add: bv_extend_norm_unsigned)
- finally show ?thesis .
-qed
-
-lemma norm_unsigned_nat_to_bv [simp]:
- "norm_unsigned (nat_to_bv n) = nat_to_bv n"
-proof -
- have "norm_unsigned (nat_to_bv n) = nat_to_bv (bv_to_nat (norm_unsigned (nat_to_bv n)))"
- by (subst nat_bv_nat) simp
- also have "... = nat_to_bv n" by simp
- finally show ?thesis .
-qed
-
-lemma length_nat_to_bv_upper_limit:
- assumes nk: "n \<le> 2 ^ k - 1"
- shows "length (nat_to_bv n) \<le> k"
-proof (cases "n = 0")
- case True
- thus ?thesis
- by (simp add: nat_to_bv_def nat_to_bv_helper.simps)
-next
- case False
- hence n0: "0 < n" by simp
- show ?thesis
- proof (rule ccontr)
- assume "~ length (nat_to_bv n) \<le> k"
- hence "k < length (nat_to_bv n)" by simp
- hence "k \<le> length (nat_to_bv n) - 1" by arith
- hence "(2::nat) ^ k \<le> 2 ^ (length (nat_to_bv n) - 1)" by simp
- also have "... = 2 ^ (length (norm_unsigned (nat_to_bv n)) - 1)" by simp
- also have "... \<le> bv_to_nat (nat_to_bv n)"
- by (rule bv_to_nat_lower_limit) (simp add: n0)
- also have "... = n" by simp
- finally have "2 ^ k \<le> n" .
- with n0 have "2 ^ k - 1 < n" by arith
- with nk show False by simp
- qed
-qed
-
-lemma length_nat_to_bv_lower_limit:
- assumes nk: "2 ^ k \<le> n"
- shows "k < length (nat_to_bv n)"
-proof (rule ccontr)
- assume "~ k < length (nat_to_bv n)"
- hence lnk: "length (nat_to_bv n) \<le> k" by simp
- have "n = bv_to_nat (nat_to_bv n)" by simp
- also have "... < 2 ^ length (nat_to_bv n)"
- by (rule bv_to_nat_upper_range)
- also from lnk have "... \<le> 2 ^ k" by simp
- finally have "n < 2 ^ k" .
- with nk show False by simp
-qed
-
-
-subsection {* Unsigned Arithmetic Operations *}
-
-definition
- bv_add :: "[bit list, bit list ] => bit list" where
- "bv_add w1 w2 = nat_to_bv (bv_to_nat w1 + bv_to_nat w2)"
-
-lemma bv_add_type1 [simp]: "bv_add (norm_unsigned w1) w2 = bv_add w1 w2"
- by (simp add: bv_add_def)
-
-lemma bv_add_type2 [simp]: "bv_add w1 (norm_unsigned w2) = bv_add w1 w2"
- by (simp add: bv_add_def)
-
-lemma bv_add_returntype [simp]: "norm_unsigned (bv_add w1 w2) = bv_add w1 w2"
- by (simp add: bv_add_def)
-
-lemma bv_add_length: "length (bv_add w1 w2) \<le> Suc (max (length w1) (length w2))"
-proof (unfold bv_add_def,rule length_nat_to_bv_upper_limit)
- from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2]
- have "bv_to_nat w1 + bv_to_nat w2 \<le> (2 ^ length w1 - 1) + (2 ^ length w2 - 1)"
- by arith
- also have "... \<le>
- max (2 ^ length w1 - 1) (2 ^ length w2 - 1) + max (2 ^ length w1 - 1) (2 ^ length w2 - 1)"
- by (rule add_mono,safe intro!: le_maxI1 le_maxI2)
- also have "... = 2 * max (2 ^ length w1 - 1) (2 ^ length w2 - 1)" by simp
- also have "... \<le> 2 ^ Suc (max (length w1) (length w2)) - 2"
- proof (cases "length w1 \<le> length w2")
- assume w1w2: "length w1 \<le> length w2"
- hence "(2::nat) ^ length w1 \<le> 2 ^ length w2" by simp
- hence "(2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1" by arith
- with w1w2 show ?thesis
- by (simp add: diff_mult_distrib2 split: split_max)
- next
- assume [simp]: "~ (length w1 \<le> length w2)"
- have "~ ((2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1)"
- proof
- assume "(2::nat) ^ length w1 - 1 \<le> 2 ^ length w2 - 1"
- hence "((2::nat) ^ length w1 - 1) + 1 \<le> (2 ^ length w2 - 1) + 1"
- by (rule add_right_mono)
- hence "(2::nat) ^ length w1 \<le> 2 ^ length w2" by simp
- hence "length w1 \<le> length w2" by simp
- thus False by simp
- qed
- thus ?thesis
- by (simp add: diff_mult_distrib2 split: split_max)
- qed
- finally show "bv_to_nat w1 + bv_to_nat w2 \<le> 2 ^ Suc (max (length w1) (length w2)) - 1"
- by arith
-qed
-
-definition
- bv_mult :: "[bit list, bit list ] => bit list" where
- "bv_mult w1 w2 = nat_to_bv (bv_to_nat w1 * bv_to_nat w2)"
-
-lemma bv_mult_type1 [simp]: "bv_mult (norm_unsigned w1) w2 = bv_mult w1 w2"
- by (simp add: bv_mult_def)
-
-lemma bv_mult_type2 [simp]: "bv_mult w1 (norm_unsigned w2) = bv_mult w1 w2"
- by (simp add: bv_mult_def)
-
-lemma bv_mult_returntype [simp]: "norm_unsigned (bv_mult w1 w2) = bv_mult w1 w2"
- by (simp add: bv_mult_def)
-
-lemma bv_mult_length: "length (bv_mult w1 w2) \<le> length w1 + length w2"
-proof (unfold bv_mult_def,rule length_nat_to_bv_upper_limit)
- from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2]
- have h: "bv_to_nat w1 \<le> 2 ^ length w1 - 1 \<and> bv_to_nat w2 \<le> 2 ^ length w2 - 1"
- by arith
- have "bv_to_nat w1 * bv_to_nat w2 \<le> (2 ^ length w1 - 1) * (2 ^ length w2 - 1)"
- apply (cut_tac h)
- apply (rule mult_mono)
- apply auto
- done
- also have "... < 2 ^ length w1 * 2 ^ length w2"
- by (rule mult_strict_mono,auto)
- also have "... = 2 ^ (length w1 + length w2)"
- by (simp add: power_add)
- finally show "bv_to_nat w1 * bv_to_nat w2 \<le> 2 ^ (length w1 + length w2) - 1"
- by arith
-qed
-
-subsection {* Signed Vectors *}
-
-primrec norm_signed :: "bit list => bit list" where
- norm_signed_Nil: "norm_signed [] = []"
- | norm_signed_Cons: "norm_signed (b#bs) =
- (case b of
- \<zero> => if norm_unsigned bs = [] then [] else b#norm_unsigned bs
- | \<one> => b#rem_initial b bs)"
-
-lemma norm_signed0 [simp]: "norm_signed [\<zero>] = []"
- by simp
-
-lemma norm_signed1 [simp]: "norm_signed [\<one>] = [\<one>]"
- by simp
-
-lemma norm_signed01 [simp]: "norm_signed (\<zero>#\<one>#xs) = \<zero>#\<one>#xs"
- by simp
-
-lemma norm_signed00 [simp]: "norm_signed (\<zero>#\<zero>#xs) = norm_signed (\<zero>#xs)"
- by simp
-
-lemma norm_signed10 [simp]: "norm_signed (\<one>#\<zero>#xs) = \<one>#\<zero>#xs"
- by simp
-
-lemma norm_signed11 [simp]: "norm_signed (\<one>#\<one>#xs) = norm_signed (\<one>#xs)"
- by simp
-
-lemmas [simp del] = norm_signed_Cons
-
-definition
- int_to_bv :: "int => bit list" where
- "int_to_bv n = (if 0 \<le> n
- then norm_signed (\<zero>#nat_to_bv (nat n))
- else norm_signed (bv_not (\<zero>#nat_to_bv (nat (-n- 1)))))"
-
-lemma int_to_bv_ge0 [simp]: "0 \<le> n ==> int_to_bv n = norm_signed (\<zero> # nat_to_bv (nat n))"
- by (simp add: int_to_bv_def)
-
-lemma int_to_bv_lt0 [simp]:
- "n < 0 ==> int_to_bv n = norm_signed (bv_not (\<zero>#nat_to_bv (nat (-n- 1))))"
- by (simp add: int_to_bv_def)
-
-lemma norm_signed_idem [simp]: "norm_signed (norm_signed w) = norm_signed w"
-proof (rule bit_list_induct [of _ w], simp_all)
- fix xs
- assume eq: "norm_signed (norm_signed xs) = norm_signed xs"
- show "norm_signed (norm_signed (\<zero>#xs)) = norm_signed (\<zero>#xs)"
- proof (rule bit_list_cases [of xs],simp_all)
- fix ys
- assume "xs = \<zero>#ys"
- from this [symmetric] and eq
- show "norm_signed (norm_signed (\<zero>#ys)) = norm_signed (\<zero>#ys)"
- by simp
- qed
-next
- fix xs
- assume eq: "norm_signed (norm_signed xs) = norm_signed xs"
- show "norm_signed (norm_signed (\<one>#xs)) = norm_signed (\<one>#xs)"
- proof (rule bit_list_cases [of xs],simp_all)
- fix ys
- assume "xs = \<one>#ys"
- from this [symmetric] and eq
- show "norm_signed (norm_signed (\<one>#ys)) = norm_signed (\<one>#ys)"
- by simp
- qed
-qed
-
-definition
- bv_to_int :: "bit list => int" where
- "bv_to_int w =
- (case bv_msb w of \<zero> => int (bv_to_nat w)
- | \<one> => - int (bv_to_nat (bv_not w) + 1))"
-
-lemma bv_to_int_Nil [simp]: "bv_to_int [] = 0"
- by (simp add: bv_to_int_def)
-
-lemma bv_to_int_Cons0 [simp]: "bv_to_int (\<zero>#bs) = int (bv_to_nat bs)"
- by (simp add: bv_to_int_def)
-
-lemma bv_to_int_Cons1 [simp]: "bv_to_int (\<one>#bs) = - int (bv_to_nat (bv_not bs) + 1)"
- by (simp add: bv_to_int_def)
-
-lemma bv_to_int_type [simp]: "bv_to_int (norm_signed w) = bv_to_int w"
-proof (rule bit_list_induct [of _ w], simp_all)
- fix xs
- assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs"
- show "bv_to_int (norm_signed (\<zero>#xs)) = int (bv_to_nat xs)"
- proof (rule bit_list_cases [of xs], simp_all)
- fix ys
- assume [simp]: "xs = \<zero>#ys"
- from ind
- show "bv_to_int (norm_signed (\<zero>#ys)) = int (bv_to_nat ys)"
- by simp
- qed
-next
- fix xs
- assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs"
- show "bv_to_int (norm_signed (\<one>#xs)) = -1 - int (bv_to_nat (bv_not xs))"
- proof (rule bit_list_cases [of xs], simp_all)
- fix ys
- assume [simp]: "xs = \<one>#ys"
- from ind
- show "bv_to_int (norm_signed (\<one>#ys)) = -1 - int (bv_to_nat (bv_not ys))"
- by simp
- qed
-qed
-
-lemma bv_to_int_upper_range: "bv_to_int w < 2 ^ (length w - 1)"
-proof (rule bit_list_cases [of w],simp_all)
- fix bs
- from bv_to_nat_upper_range
- show "int (bv_to_nat bs) < 2 ^ length bs"
- by (simp add: int_nat_two_exp)
-next
- fix bs
- have "-1 - int (bv_to_nat (bv_not bs)) \<le> 0" by simp
- also have "... < 2 ^ length bs" by (induct bs) simp_all
- finally show "-1 - int (bv_to_nat (bv_not bs)) < 2 ^ length bs" .
-qed
-
-lemma bv_to_int_lower_range: "- (2 ^ (length w - 1)) \<le> bv_to_int w"
-proof (rule bit_list_cases [of w],simp_all)
- fix bs :: "bit list"
- have "- (2 ^ length bs) \<le> (0::int)" by (induct bs) simp_all
- also have "... \<le> int (bv_to_nat bs)" by simp
- finally show "- (2 ^ length bs) \<le> int (bv_to_nat bs)" .
-next
- fix bs
- from bv_to_nat_upper_range [of "bv_not bs"]
- show "- (2 ^ length bs) \<le> -1 - int (bv_to_nat (bv_not bs))"
- by (simp add: int_nat_two_exp)
-qed
-
-lemma int_bv_int [simp]: "int_to_bv (bv_to_int w) = norm_signed w"
-proof (rule bit_list_cases [of w],simp)
- fix xs
- assume [simp]: "w = \<zero>#xs"
- show ?thesis
- apply simp
- apply (subst norm_signed_Cons [of "\<zero>" "xs"])
- apply simp
- using norm_unsigned_result [of xs]
- apply safe
- apply (rule bit_list_cases [of "norm_unsigned xs"])
- apply simp_all
- done
-next
- fix xs
- assume [simp]: "w = \<one>#xs"
- show ?thesis
- apply (simp del: int_to_bv_lt0)
- apply (rule bit_list_induct [of _ xs])
- apply simp
- apply (subst int_to_bv_lt0)
- apply (subgoal_tac "- int (bv_to_nat (bv_not (\<zero> # bs))) + -1 < 0 + 0")
- apply simp
- apply (rule add_le_less_mono)
- apply simp
- apply simp
- apply (simp del: bv_to_nat1 bv_to_nat_helper)
- apply simp
- done
-qed
-
-lemma bv_int_bv [simp]: "bv_to_int (int_to_bv i) = i"
- by (cases "0 \<le> i") simp_all
-
-lemma bv_msb_norm [simp]: "bv_msb (norm_signed w) = bv_msb w"
- by (rule bit_list_cases [of w]) (simp_all add: norm_signed_Cons)
-
-lemma norm_signed_length: "length (norm_signed w) \<le> length w"
- apply (cases w, simp_all)
- apply (subst norm_signed_Cons)
- apply (case_tac a, simp_all)
- apply (rule rem_initial_length)
- done
-
-lemma norm_signed_equal: "length (norm_signed w) = length w ==> norm_signed w = w"
-proof (rule bit_list_cases [of w], simp_all)
- fix xs
- assume "length (norm_signed (\<zero>#xs)) = Suc (length xs)"
- thus "norm_signed (\<zero>#xs) = \<zero>#xs"
- by (simp add: norm_signed_Cons norm_unsigned_equal [THEN eqTrueI]
- split: split_if_asm)
-next
- fix xs
- assume "length (norm_signed (\<one>#xs)) = Suc (length xs)"
- thus "norm_signed (\<one>#xs) = \<one>#xs"
- apply (simp add: norm_signed_Cons)
- apply (rule rem_initial_equal)
- apply assumption
- done
-qed
-
-lemma bv_extend_norm_signed: "bv_msb w = b ==> bv_extend (length w) b (norm_signed w) = w"
-proof (rule bit_list_cases [of w],simp_all)
- fix xs
- show "bv_extend (Suc (length xs)) \<zero> (norm_signed (\<zero>#xs)) = \<zero>#xs"
- proof (simp add: norm_signed_def,auto)
- assume "norm_unsigned xs = []"
- hence xx: "rem_initial \<zero> xs = []"
- by (simp add: norm_unsigned_def)
- have "bv_extend (Suc (length xs)) \<zero> (\<zero>#rem_initial \<zero> xs) = \<zero>#xs"
- apply (simp add: bv_extend_def replicate_app_Cons_same)
- apply (fold bv_extend_def)
- apply (rule bv_extend_rem_initial)
- done
- thus "bv_extend (Suc (length xs)) \<zero> [\<zero>] = \<zero>#xs"
- by (simp add: xx)
- next
- show "bv_extend (Suc (length xs)) \<zero> (\<zero>#norm_unsigned xs) = \<zero>#xs"
- apply (simp add: norm_unsigned_def)
- apply (simp add: bv_extend_def replicate_app_Cons_same)
- apply (fold bv_extend_def)
- apply (rule bv_extend_rem_initial)
- done
- qed
-next
- fix xs
- show "bv_extend (Suc (length xs)) \<one> (norm_signed (\<one>#xs)) = \<one>#xs"
- apply (simp add: norm_signed_Cons)
- apply (simp add: bv_extend_def replicate_app_Cons_same)
- apply (fold bv_extend_def)
- apply (rule bv_extend_rem_initial)
- done
-qed
-
-lemma bv_to_int_qinj:
- assumes one: "bv_to_int xs = bv_to_int ys"
- and len: "length xs = length ys"
- shows "xs = ys"
-proof -
- from one
- have "int_to_bv (bv_to_int xs) = int_to_bv (bv_to_int ys)" by simp
- hence xsys: "norm_signed xs = norm_signed ys" by simp
- hence xsys': "bv_msb xs = bv_msb ys"
- proof -
- have "bv_msb xs = bv_msb (norm_signed xs)" by simp
- also have "... = bv_msb (norm_signed ys)" by (simp add: xsys)
- also have "... = bv_msb ys" by simp
- finally show ?thesis .
- qed
- have "xs = bv_extend (length xs) (bv_msb xs) (norm_signed xs)"
- by (simp add: bv_extend_norm_signed)
- also have "... = bv_extend (length ys) (bv_msb ys) (norm_signed ys)"
- by (simp add: xsys xsys' len)
- also have "... = ys"
- by (simp add: bv_extend_norm_signed)
- finally show ?thesis .
-qed
-
-lemma int_to_bv_returntype [simp]: "norm_signed (int_to_bv w) = int_to_bv w"
- by (simp add: int_to_bv_def)
-
-lemma bv_to_int_msb0: "0 \<le> bv_to_int w1 ==> bv_msb w1 = \<zero>"
- by (rule bit_list_cases,simp_all)
-
-lemma bv_to_int_msb1: "bv_to_int w1 < 0 ==> bv_msb w1 = \<one>"
- by (rule bit_list_cases,simp_all)
-
-lemma bv_to_int_lower_limit_gt0:
- assumes w0: "0 < bv_to_int w"
- shows "2 ^ (length (norm_signed w) - 2) \<le> bv_to_int w"
-proof -
- from w0
- have "0 \<le> bv_to_int w" by simp
- hence [simp]: "bv_msb w = \<zero>" by (rule bv_to_int_msb0)
- have "2 ^ (length (norm_signed w) - 2) \<le> bv_to_int (norm_signed w)"
- proof (rule bit_list_cases [of w])
- assume "w = []"
- with w0 show ?thesis by simp
- next
- fix w'
- assume weq: "w = \<zero> # w'"
- thus ?thesis
- proof (simp add: norm_signed_Cons,safe)
- assume "norm_unsigned w' = []"
- with weq and w0 show False
- by (simp add: norm_empty_bv_to_nat_zero)
- next
- assume w'0: "norm_unsigned w' \<noteq> []"
- have "0 < bv_to_nat w'"
- proof (rule ccontr)
- assume "~ (0 < bv_to_nat w')"
- hence "bv_to_nat w' = 0"
- by arith
- hence "norm_unsigned w' = []"
- by (simp add: bv_to_nat_zero_imp_empty)
- with w'0
- show False by simp
- qed
- with bv_to_nat_lower_limit [of w']
- show "2 ^ (length (norm_unsigned w') - Suc 0) \<le> int (bv_to_nat w')"
- by (simp add: int_nat_two_exp)
- qed
- next
- fix w'
- assume "w = \<one> # w'"
- from w0 have "bv_msb w = \<zero>" by simp
- with prems show ?thesis by simp
- qed
- also have "... = bv_to_int w" by simp
- finally show ?thesis .
-qed
-
-lemma norm_signed_result: "norm_signed w = [] \<or> norm_signed w = [\<one>] \<or> bv_msb (norm_signed w) \<noteq> bv_msb (tl (norm_signed w))"
- apply (rule bit_list_cases [of w],simp_all)
- apply (case_tac "bs",simp_all)
- apply (case_tac "a",simp_all)
- apply (simp add: norm_signed_Cons)
- apply safe
- apply simp
-proof -
- fix l
- assume msb: "\<zero> = bv_msb (norm_unsigned l)"
- assume "norm_unsigned l \<noteq> []"
- with norm_unsigned_result [of l]
- have "bv_msb (norm_unsigned l) = \<one>" by simp
- with msb show False by simp
-next
- fix xs
- assume p: "\<one> = bv_msb (tl (norm_signed (\<one> # xs)))"
- have "\<one> \<noteq> bv_msb (tl (norm_signed (\<one> # xs)))"
- by (rule bit_list_induct [of _ xs],simp_all)
- with p show False by simp
-qed
-
-lemma bv_to_int_upper_limit_lem1:
- assumes w0: "bv_to_int w < -1"
- shows "bv_to_int w < - (2 ^ (length (norm_signed w) - 2))"
-proof -
- from w0
- have "bv_to_int w < 0" by simp
- hence msbw [simp]: "bv_msb w = \<one>"
- by (rule bv_to_int_msb1)
- have "bv_to_int w = bv_to_int (norm_signed w)" by simp
- also from norm_signed_result [of w]
- have "... < - (2 ^ (length (norm_signed w) - 2))"
- proof safe
- assume "norm_signed w = []"
- hence "bv_to_int (norm_signed w) = 0" by simp
- with w0 show ?thesis by simp
- next
- assume "norm_signed w = [\<one>]"
- hence "bv_to_int (norm_signed w) = -1" by simp
- with w0 show ?thesis by simp
- next
- assume "bv_msb (norm_signed w) \<noteq> bv_msb (tl (norm_signed w))"
- hence msb_tl: "\<one> \<noteq> bv_msb (tl (norm_signed w))" by simp
- show "bv_to_int (norm_signed w) < - (2 ^ (length (norm_signed w) - 2))"
- proof (rule bit_list_cases [of "norm_signed w"])
- assume "norm_signed w = []"
- hence "bv_to_int (norm_signed w) = 0" by simp
- with w0 show ?thesis by simp
- next
- fix w'
- assume nw: "norm_signed w = \<zero> # w'"
- from msbw have "bv_msb (norm_signed w) = \<one>" by simp
- with nw show ?thesis by simp
- next
- fix w'
- assume weq: "norm_signed w = \<one> # w'"
- show ?thesis
- proof (rule bit_list_cases [of w'])
- assume w'eq: "w' = []"
- from w0 have "bv_to_int (norm_signed w) < -1" by simp
- with w'eq and weq show ?thesis by simp
- next
- fix w''
- assume w'eq: "w' = \<zero> # w''"
- show ?thesis
- apply (simp add: weq w'eq)
- apply (subgoal_tac "- int (bv_to_nat (bv_not w'')) + -1 < 0 + 0")
- apply (simp add: int_nat_two_exp)
- apply (rule add_le_less_mono)
- apply simp_all
- done
- next
- fix w''
- assume w'eq: "w' = \<one> # w''"
- with weq and msb_tl show ?thesis by simp
- qed
- qed
- qed
- finally show ?thesis .
-qed
-
-lemma length_int_to_bv_upper_limit_gt0:
- assumes w0: "0 < i"
- and wk: "i \<le> 2 ^ (k - 1) - 1"
- shows "length (int_to_bv i) \<le> k"
-proof (rule ccontr)
- from w0 wk
- have k1: "1 < k"
- by (cases "k - 1",simp_all)
- assume "~ length (int_to_bv i) \<le> k"
- hence "k < length (int_to_bv i)" by simp
- hence "k \<le> length (int_to_bv i) - 1" by arith
- hence a: "k - 1 \<le> length (int_to_bv i) - 2" by arith
- hence "(2::int) ^ (k - 1) \<le> 2 ^ (length (int_to_bv i) - 2)" by simp
- also have "... \<le> i"
- proof -
- have "2 ^ (length (norm_signed (int_to_bv i)) - 2) \<le> bv_to_int (int_to_bv i)"
- proof (rule bv_to_int_lower_limit_gt0)
- from w0 show "0 < bv_to_int (int_to_bv i)" by simp
- qed
- thus ?thesis by simp
- qed
- finally have "2 ^ (k - 1) \<le> i" .
- with wk show False by simp
-qed
-
-lemma pos_length_pos:
- assumes i0: "0 < bv_to_int w"
- shows "0 < length w"
-proof -
- from norm_signed_result [of w]
- have "0 < length (norm_signed w)"
- proof (auto)
- assume ii: "norm_signed w = []"
- have "bv_to_int (norm_signed w) = 0" by (subst ii) simp
- hence "bv_to_int w = 0" by simp
- with i0 show False by simp
- next
- assume ii: "norm_signed w = []"
- assume jj: "bv_msb w \<noteq> \<zero>"
- have "\<zero> = bv_msb (norm_signed w)"
- by (subst ii) simp
- also have "... \<noteq> \<zero>"
- by (simp add: jj)
- finally show False by simp
- qed
- also have "... \<le> length w"
- by (rule norm_signed_length)
- finally show ?thesis .
-qed
-
-lemma neg_length_pos:
- assumes i0: "bv_to_int w < -1"
- shows "0 < length w"
-proof -
- from norm_signed_result [of w]
- have "0 < length (norm_signed w)"
- proof (auto)
- assume ii: "norm_signed w = []"
- have "bv_to_int (norm_signed w) = 0"
- by (subst ii) simp
- hence "bv_to_int w = 0" by simp
- with i0 show False by simp
- next
- assume ii: "norm_signed w = []"
- assume jj: "bv_msb w \<noteq> \<zero>"
- have "\<zero> = bv_msb (norm_signed w)" by (subst ii) simp
- also have "... \<noteq> \<zero>" by (simp add: jj)
- finally show False by simp
- qed
- also have "... \<le> length w"
- by (rule norm_signed_length)
- finally show ?thesis .
-qed
-
-lemma length_int_to_bv_lower_limit_gt0:
- assumes wk: "2 ^ (k - 1) \<le> i"
- shows "k < length (int_to_bv i)"
-proof (rule ccontr)
- have "0 < (2::int) ^ (k - 1)"
- by (rule zero_less_power) simp
- also have "... \<le> i" by (rule wk)
- finally have i0: "0 < i" .
- have lii0: "0 < length (int_to_bv i)"
- apply (rule pos_length_pos)
- apply (simp,rule i0)
- done
- assume "~ k < length (int_to_bv i)"
- hence "length (int_to_bv i) \<le> k" by simp
- with lii0
- have a: "length (int_to_bv i) - 1 \<le> k - 1"
- by arith
- have "i < 2 ^ (length (int_to_bv i) - 1)"
- proof -
- have "i = bv_to_int (int_to_bv i)"
- by simp
- also have "... < 2 ^ (length (int_to_bv i) - 1)"
- by (rule bv_to_int_upper_range)
- finally show ?thesis .
- qed
- also have "(2::int) ^ (length (int_to_bv i) - 1) \<le> 2 ^ (k - 1)" using a
- by simp
- finally have "i < 2 ^ (k - 1)" .
- with wk show False by simp
-qed
-
-lemma length_int_to_bv_upper_limit_lem1:
- assumes w1: "i < -1"
- and wk: "- (2 ^ (k - 1)) \<le> i"
- shows "length (int_to_bv i) \<le> k"
-proof (rule ccontr)
- from w1 wk
- have k1: "1 < k" by (cases "k - 1") simp_all
- assume "~ length (int_to_bv i) \<le> k"
- hence "k < length (int_to_bv i)" by simp
- hence "k \<le> length (int_to_bv i) - 1" by arith
- hence a: "k - 1 \<le> length (int_to_bv i) - 2" by arith
- have "i < - (2 ^ (length (int_to_bv i) - 2))"
- proof -
- have "i = bv_to_int (int_to_bv i)"
- by simp
- also have "... < - (2 ^ (length (norm_signed (int_to_bv i)) - 2))"
- by (rule bv_to_int_upper_limit_lem1,simp,rule w1)
- finally show ?thesis by simp
- qed
- also have "... \<le> -(2 ^ (k - 1))"
- proof -
- have "(2::int) ^ (k - 1) \<le> 2 ^ (length (int_to_bv i) - 2)" using a by simp
- thus ?thesis by simp
- qed
- finally have "i < -(2 ^ (k - 1))" .
- with wk show False by simp
-qed
-
-lemma length_int_to_bv_lower_limit_lem1:
- assumes wk: "i < -(2 ^ (k - 1))"
- shows "k < length (int_to_bv i)"
-proof (rule ccontr)
- from wk have "i \<le> -(2 ^ (k - 1)) - 1" by simp
- also have "... < -1"
- proof -
- have "0 < (2::int) ^ (k - 1)"
- by (rule zero_less_power) simp
- hence "-((2::int) ^ (k - 1)) < 0" by simp
- thus ?thesis by simp
- qed
- finally have i1: "i < -1" .
- have lii0: "0 < length (int_to_bv i)"
- apply (rule neg_length_pos)
- apply (simp, rule i1)
- done
- assume "~ k < length (int_to_bv i)"
- hence "length (int_to_bv i) \<le> k"
- by simp
- with lii0 have a: "length (int_to_bv i) - 1 \<le> k - 1" by arith
- hence "(2::int) ^ (length (int_to_bv i) - 1) \<le> 2 ^ (k - 1)" by simp
- hence "-((2::int) ^ (k - 1)) \<le> - (2 ^ (length (int_to_bv i) - 1))" by simp
- also have "... \<le> i"
- proof -
- have "- (2 ^ (length (int_to_bv i) - 1)) \<le> bv_to_int (int_to_bv i)"
- by (rule bv_to_int_lower_range)
- also have "... = i"
- by simp
- finally show ?thesis .
- qed
- finally have "-(2 ^ (k - 1)) \<le> i" .
- with wk show False by simp
-qed
-
-
-subsection {* Signed Arithmetic Operations *}
-
-subsubsection {* Conversion from unsigned to signed *}
-
-definition
- utos :: "bit list => bit list" where
- "utos w = norm_signed (\<zero> # w)"
-
-lemma utos_type [simp]: "utos (norm_unsigned w) = utos w"
- by (simp add: utos_def norm_signed_Cons)
-
-lemma utos_returntype [simp]: "norm_signed (utos w) = utos w"
- by (simp add: utos_def)
-
-lemma utos_length: "length (utos w) \<le> Suc (length w)"
- by (simp add: utos_def norm_signed_Cons)
-
-lemma bv_to_int_utos: "bv_to_int (utos w) = int (bv_to_nat w)"
-proof (simp add: utos_def norm_signed_Cons, safe)
- assume "norm_unsigned w = []"
- hence "bv_to_nat (norm_unsigned w) = 0" by simp
- thus "bv_to_nat w = 0" by simp
-qed
-
-
-subsubsection {* Unary minus *}
-
-definition
- bv_uminus :: "bit list => bit list" where
- "bv_uminus w = int_to_bv (- bv_to_int w)"
-
-lemma bv_uminus_type [simp]: "bv_uminus (norm_signed w) = bv_uminus w"
- by (simp add: bv_uminus_def)
-
-lemma bv_uminus_returntype [simp]: "norm_signed (bv_uminus w) = bv_uminus w"
- by (simp add: bv_uminus_def)
-
-lemma bv_uminus_length: "length (bv_uminus w) \<le> Suc (length w)"
-proof -
- have "1 < -bv_to_int w \<or> -bv_to_int w = 1 \<or> -bv_to_int w = 0 \<or> -bv_to_int w = -1 \<or> -bv_to_int w < -1"
- by arith
- thus ?thesis
- proof safe
- assume p: "1 < - bv_to_int w"
- have lw: "0 < length w"
- apply (rule neg_length_pos)
- using p
- apply simp
- done
- show ?thesis
- proof (simp add: bv_uminus_def,rule length_int_to_bv_upper_limit_gt0,simp_all)
- from prems show "bv_to_int w < 0" by simp
- next
- have "-(2^(length w - 1)) \<le> bv_to_int w"
- by (rule bv_to_int_lower_range)
- hence "- bv_to_int w \<le> 2^(length w - 1)" by simp
- also from lw have "... < 2 ^ length w" by simp
- finally show "- bv_to_int w < 2 ^ length w" by simp
- qed
- next
- assume p: "- bv_to_int w = 1"
- hence lw: "0 < length w" by (cases w) simp_all
- from p
- show ?thesis
- apply (simp add: bv_uminus_def)
- using lw
- apply (simp (no_asm) add: nat_to_bv_non0)
- done
- next
- assume "- bv_to_int w = 0"
- thus ?thesis by (simp add: bv_uminus_def)
- next
- assume p: "- bv_to_int w = -1"
- thus ?thesis by (simp add: bv_uminus_def)
- next
- assume p: "- bv_to_int w < -1"
- show ?thesis
- apply (simp add: bv_uminus_def)
- apply (rule length_int_to_bv_upper_limit_lem1)
- apply (rule p)
- apply simp
- proof -
- have "bv_to_int w < 2 ^ (length w - 1)"
- by (rule bv_to_int_upper_range)
- also have "... \<le> 2 ^ length w" by simp
- finally show "bv_to_int w \<le> 2 ^ length w" by simp
- qed
- qed
-qed
-
-lemma bv_uminus_length_utos: "length (bv_uminus (utos w)) \<le> Suc (length w)"
-proof -
- have "-bv_to_int (utos w) = 0 \<or> -bv_to_int (utos w) = -1 \<or> -bv_to_int (utos w) < -1"
- by (simp add: bv_to_int_utos, arith)
- thus ?thesis
- proof safe
- assume "-bv_to_int (utos w) = 0"
- thus ?thesis by (simp add: bv_uminus_def)
- next
- assume "-bv_to_int (utos w) = -1"
- thus ?thesis by (simp add: bv_uminus_def)
- next
- assume p: "-bv_to_int (utos w) < -1"
- show ?thesis
- apply (simp add: bv_uminus_def)
- apply (rule length_int_to_bv_upper_limit_lem1)
- apply (rule p)
- apply (simp add: bv_to_int_utos)
- using bv_to_nat_upper_range [of w]
- apply (simp add: int_nat_two_exp)
- done
- qed
-qed
-
-definition
- bv_sadd :: "[bit list, bit list ] => bit list" where
- "bv_sadd w1 w2 = int_to_bv (bv_to_int w1 + bv_to_int w2)"
-
-lemma bv_sadd_type1 [simp]: "bv_sadd (norm_signed w1) w2 = bv_sadd w1 w2"
- by (simp add: bv_sadd_def)
-
-lemma bv_sadd_type2 [simp]: "bv_sadd w1 (norm_signed w2) = bv_sadd w1 w2"
- by (simp add: bv_sadd_def)
-
-lemma bv_sadd_returntype [simp]: "norm_signed (bv_sadd w1 w2) = bv_sadd w1 w2"
- by (simp add: bv_sadd_def)
-
-lemma adder_helper:
- assumes lw: "0 < max (length w1) (length w2)"
- shows "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \<le> 2 ^ max (length w1) (length w2)"
-proof -
- have "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \<le>
- 2 ^ (max (length w1) (length w2) - 1) + 2 ^ (max (length w1) (length w2) - 1)"
- by (auto simp:max_def)
- also have "... = 2 ^ max (length w1) (length w2)"
- proof -
- from lw
- show ?thesis
- apply simp
- apply (subst power_Suc [symmetric])
- apply simp
- done
- qed
- finally show ?thesis .
-qed
-
-lemma bv_sadd_length: "length (bv_sadd w1 w2) \<le> Suc (max (length w1) (length w2))"
-proof -
- let ?Q = "bv_to_int w1 + bv_to_int w2"
-
- have helper: "?Q \<noteq> 0 ==> 0 < max (length w1) (length w2)"
- proof -
- assume p: "?Q \<noteq> 0"
- show "0 < max (length w1) (length w2)"
- proof (simp add: less_max_iff_disj,rule)
- assume [simp]: "w1 = []"
- show "w2 \<noteq> []"
- proof (rule ccontr,simp)
- assume [simp]: "w2 = []"
- from p show False by simp
- qed
- qed
- qed
-
- have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith
- thus ?thesis
- proof safe
- assume "?Q = 0"
- thus ?thesis
- by (simp add: bv_sadd_def)
- next
- assume "?Q = -1"
- thus ?thesis
- by (simp add: bv_sadd_def)
- next
- assume p: "0 < ?Q"
- show ?thesis
- apply (simp add: bv_sadd_def)
- apply (rule length_int_to_bv_upper_limit_gt0)
- apply (rule p)
- proof simp
- from bv_to_int_upper_range [of w2]
- have "bv_to_int w2 \<le> 2 ^ (length w2 - 1)"
- by simp
- with bv_to_int_upper_range [of w1]
- have "bv_to_int w1 + bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))"
- by (rule zadd_zless_mono)
- also have "... \<le> 2 ^ max (length w1) (length w2)"
- apply (rule adder_helper)
- apply (rule helper)
- using p
- apply simp
- done
- finally show "?Q < 2 ^ max (length w1) (length w2)" .
- qed
- next
- assume p: "?Q < -1"
- show ?thesis
- apply (simp add: bv_sadd_def)
- apply (rule length_int_to_bv_upper_limit_lem1,simp_all)
- apply (rule p)
- proof -
- have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \<le> (2::int) ^ max (length w1) (length w2)"
- apply (rule adder_helper)
- apply (rule helper)
- using p
- apply simp
- done
- hence "-((2::int) ^ max (length w1) (length w2)) \<le> - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))"
- by simp
- also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \<le> ?Q"
- apply (rule add_mono)
- apply (rule bv_to_int_lower_range [of w1])
- apply (rule bv_to_int_lower_range [of w2])
- done
- finally show "- (2^max (length w1) (length w2)) \<le> ?Q" .
- qed
- qed
-qed
-
-definition
- bv_sub :: "[bit list, bit list] => bit list" where
- "bv_sub w1 w2 = bv_sadd w1 (bv_uminus w2)"
-
-lemma bv_sub_type1 [simp]: "bv_sub (norm_signed w1) w2 = bv_sub w1 w2"
- by (simp add: bv_sub_def)
-
-lemma bv_sub_type2 [simp]: "bv_sub w1 (norm_signed w2) = bv_sub w1 w2"
- by (simp add: bv_sub_def)
-
-lemma bv_sub_returntype [simp]: "norm_signed (bv_sub w1 w2) = bv_sub w1 w2"
- by (simp add: bv_sub_def)
-
-lemma bv_sub_length: "length (bv_sub w1 w2) \<le> Suc (max (length w1) (length w2))"
-proof (cases "bv_to_int w2 = 0")
- assume p: "bv_to_int w2 = 0"
- show ?thesis
- proof (simp add: bv_sub_def bv_sadd_def bv_uminus_def p)
- have "length (norm_signed w1) \<le> length w1"
- by (rule norm_signed_length)
- also have "... \<le> max (length w1) (length w2)"
- by (rule le_maxI1)
- also have "... \<le> Suc (max (length w1) (length w2))"
- by arith
- finally show "length (norm_signed w1) \<le> Suc (max (length w1) (length w2))" .
- qed
-next
- assume "bv_to_int w2 \<noteq> 0"
- hence "0 < length w2" by (cases w2,simp_all)
- hence lmw: "0 < max (length w1) (length w2)" by arith
-
- let ?Q = "bv_to_int w1 - bv_to_int w2"
-
- have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith
- thus ?thesis
- proof safe
- assume "?Q = 0"
- thus ?thesis
- by (simp add: bv_sub_def bv_sadd_def bv_uminus_def)
- next
- assume "?Q = -1"
- thus ?thesis
- by (simp add: bv_sub_def bv_sadd_def bv_uminus_def)
- next
- assume p: "0 < ?Q"
- show ?thesis
- apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def)
- apply (rule length_int_to_bv_upper_limit_gt0)
- apply (rule p)
- proof simp
- from bv_to_int_lower_range [of w2]
- have v2: "- bv_to_int w2 \<le> 2 ^ (length w2 - 1)" by simp
- have "bv_to_int w1 + - bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))"
- apply (rule zadd_zless_mono)
- apply (rule bv_to_int_upper_range [of w1])
- apply (rule v2)
- done
- also have "... \<le> 2 ^ max (length w1) (length w2)"
- apply (rule adder_helper)
- apply (rule lmw)
- done
- finally show "?Q < 2 ^ max (length w1) (length w2)" by simp
- qed
- next
- assume p: "?Q < -1"
- show ?thesis
- apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def)
- apply (rule length_int_to_bv_upper_limit_lem1)
- apply (rule p)
- proof simp
- have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \<le> (2::int) ^ max (length w1) (length w2)"
- apply (rule adder_helper)
- apply (rule lmw)
- done
- hence "-((2::int) ^ max (length w1) (length w2)) \<le> - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))"
- by simp
- also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \<le> bv_to_int w1 + -bv_to_int w2"
- apply (rule add_mono)
- apply (rule bv_to_int_lower_range [of w1])
- using bv_to_int_upper_range [of w2]
- apply simp
- done
- finally show "- (2^max (length w1) (length w2)) \<le> ?Q" by simp
- qed
- qed
-qed
-
-definition
- bv_smult :: "[bit list, bit list] => bit list" where
- "bv_smult w1 w2 = int_to_bv (bv_to_int w1 * bv_to_int w2)"
-
-lemma bv_smult_type1 [simp]: "bv_smult (norm_signed w1) w2 = bv_smult w1 w2"
- by (simp add: bv_smult_def)
-
-lemma bv_smult_type2 [simp]: "bv_smult w1 (norm_signed w2) = bv_smult w1 w2"
- by (simp add: bv_smult_def)
-
-lemma bv_smult_returntype [simp]: "norm_signed (bv_smult w1 w2) = bv_smult w1 w2"
- by (simp add: bv_smult_def)
-
-lemma bv_smult_length: "length (bv_smult w1 w2) \<le> length w1 + length w2"
-proof -
- let ?Q = "bv_to_int w1 * bv_to_int w2"
-
- have lmw: "?Q \<noteq> 0 ==> 0 < length w1 \<and> 0 < length w2" by auto
-
- have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith
- thus ?thesis
- proof (safe dest!: iffD1 [OF mult_eq_0_iff])
- assume "bv_to_int w1 = 0"
- thus ?thesis by (simp add: bv_smult_def)
- next
- assume "bv_to_int w2 = 0"
- thus ?thesis by (simp add: bv_smult_def)
- next
- assume p: "?Q = -1"
- show ?thesis
- apply (simp add: bv_smult_def p)
- apply (cut_tac lmw)
- apply arith
- using p
- apply simp
- done
- next
- assume p: "0 < ?Q"
- thus ?thesis
- proof (simp add: zero_less_mult_iff,safe)
- assume bi1: "0 < bv_to_int w1"
- assume bi2: "0 < bv_to_int w2"
- show ?thesis
- apply (simp add: bv_smult_def)
- apply (rule length_int_to_bv_upper_limit_gt0)
- apply (rule p)
- proof simp
- have "?Q < 2 ^ (length w1 - 1) * 2 ^ (length w2 - 1)"
- apply (rule mult_strict_mono)
- apply (rule bv_to_int_upper_range)
- apply (rule bv_to_int_upper_range)
- apply (rule zero_less_power)
- apply simp
- using bi2
- apply simp
- done
- also have "... \<le> 2 ^ (length w1 + length w2 - Suc 0)"
- apply simp
- apply (subst zpower_zadd_distrib [symmetric])
- apply simp
- done
- finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" .
- qed
- next
- assume bi1: "bv_to_int w1 < 0"
- assume bi2: "bv_to_int w2 < 0"
- show ?thesis
- apply (simp add: bv_smult_def)
- apply (rule length_int_to_bv_upper_limit_gt0)
- apply (rule p)
- proof simp
- have "-bv_to_int w1 * -bv_to_int w2 \<le> 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)"
- apply (rule mult_mono)
- using bv_to_int_lower_range [of w1]
- apply simp
- using bv_to_int_lower_range [of w2]
- apply simp
- apply (rule zero_le_power,simp)
- using bi2
- apply simp
- done
- hence "?Q \<le> 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)"
- by simp
- also have "... < 2 ^ (length w1 + length w2 - Suc 0)"
- apply simp
- apply (subst zpower_zadd_distrib [symmetric])
- apply simp
- apply (cut_tac lmw)
- apply arith
- apply (cut_tac p)
- apply arith
- done
- finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" .
- qed
- qed
- next
- assume p: "?Q < -1"
- show ?thesis
- apply (subst bv_smult_def)
- apply (rule length_int_to_bv_upper_limit_lem1)
- apply (rule p)
- proof simp
- have "(2::int) ^ (length w1 - 1) * 2 ^(length w2 - 1) \<le> 2 ^ (length w1 + length w2 - Suc 0)"
- apply simp
- apply (subst zpower_zadd_distrib [symmetric])
- apply simp
- done
- hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \<le> -(2^(length w1 - 1) * 2 ^ (length w2 - 1))"
- by simp
- also have "... \<le> ?Q"
- proof -
- from p
- have q: "bv_to_int w1 * bv_to_int w2 < 0"
- by simp
- thus ?thesis
- proof (simp add: mult_less_0_iff,safe)
- assume bi1: "0 < bv_to_int w1"
- assume bi2: "bv_to_int w2 < 0"
- have "-bv_to_int w2 * bv_to_int w1 \<le> ((2::int)^(length w2 - 1)) * (2 ^ (length w1 - 1))"
- apply (rule mult_mono)
- using bv_to_int_lower_range [of w2]
- apply simp
- using bv_to_int_upper_range [of w1]
- apply simp
- apply (rule zero_le_power,simp)
- using bi1
- apply simp
- done
- hence "-?Q \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))"
- by (simp add: zmult_ac)
- thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \<le> ?Q"
- by simp
- next
- assume bi1: "bv_to_int w1 < 0"
- assume bi2: "0 < bv_to_int w2"
- have "-bv_to_int w1 * bv_to_int w2 \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))"
- apply (rule mult_mono)
- using bv_to_int_lower_range [of w1]
- apply simp
- using bv_to_int_upper_range [of w2]
- apply simp
- apply (rule zero_le_power,simp)
- using bi2
- apply simp
- done
- hence "-?Q \<le> ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))"
- by (simp add: zmult_ac)
- thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \<le> ?Q"
- by simp
- qed
- qed
- finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \<le> ?Q" .
- qed
- qed
-qed
-
-lemma bv_msb_one: "bv_msb w = \<one> ==> bv_to_nat w \<noteq> 0"
-by (cases w) simp_all
-
-lemma bv_smult_length_utos: "length (bv_smult (utos w1) w2) \<le> length w1 + length w2"
-proof -
- let ?Q = "bv_to_int (utos w1) * bv_to_int w2"
-
- have lmw: "?Q \<noteq> 0 ==> 0 < length (utos w1) \<and> 0 < length w2" by auto
-
- have "0 < ?Q \<or> ?Q = 0 \<or> ?Q = -1 \<or> ?Q < -1" by arith
- thus ?thesis
- proof (safe dest!: iffD1 [OF mult_eq_0_iff])
- assume "bv_to_int (utos w1) = 0"
- thus ?thesis by (simp add: bv_smult_def)
- next
- assume "bv_to_int w2 = 0"
- thus ?thesis by (simp add: bv_smult_def)
- next
- assume p: "0 < ?Q"
- thus ?thesis
- proof (simp add: zero_less_mult_iff,safe)
- assume biw2: "0 < bv_to_int w2"
- show ?thesis
- apply (simp add: bv_smult_def)
- apply (rule length_int_to_bv_upper_limit_gt0)
- apply (rule p)
- proof simp
- have "?Q < 2 ^ length w1 * 2 ^ (length w2 - 1)"
- apply (rule mult_strict_mono)
- apply (simp add: bv_to_int_utos int_nat_two_exp)
- apply (rule bv_to_nat_upper_range)
- apply (rule bv_to_int_upper_range)
- apply (rule zero_less_power,simp)
- using biw2
- apply simp
- done
- also have "... \<le> 2 ^ (length w1 + length w2 - Suc 0)"
- apply simp
- apply (subst zpower_zadd_distrib [symmetric])
- apply simp
- apply (cut_tac lmw)
- apply arith
- using p
- apply auto
- done
- finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" .
- qed
- next
- assume "bv_to_int (utos w1) < 0"
- thus ?thesis by (simp add: bv_to_int_utos)
- qed
- next
- assume p: "?Q = -1"
- thus ?thesis
- apply (simp add: bv_smult_def)
- apply (cut_tac lmw)
- apply arith
- apply simp
- done
- next
- assume p: "?Q < -1"
- show ?thesis
- apply (subst bv_smult_def)
- apply (rule length_int_to_bv_upper_limit_lem1)
- apply (rule p)
- proof simp
- have "(2::int) ^ length w1 * 2 ^(length w2 - 1) \<le> 2 ^ (length w1 + length w2 - Suc 0)"
- apply simp
- apply (subst zpower_zadd_distrib [symmetric])
- apply simp
- apply (cut_tac lmw)
- apply arith
- apply (cut_tac p)
- apply arith
- done
- hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \<le> -(2^ length w1 * 2 ^ (length w2 - 1))"
- by simp
- also have "... \<le> ?Q"
- proof -
- from p
- have q: "bv_to_int (utos w1) * bv_to_int w2 < 0"
- by simp
- thus ?thesis
- proof (simp add: mult_less_0_iff,safe)
- assume bi1: "0 < bv_to_int (utos w1)"
- assume bi2: "bv_to_int w2 < 0"
- have "-bv_to_int w2 * bv_to_int (utos w1) \<le> ((2::int)^(length w2 - 1)) * (2 ^ length w1)"
- apply (rule mult_mono)
- using bv_to_int_lower_range [of w2]
- apply simp
- apply (simp add: bv_to_int_utos)
- using bv_to_nat_upper_range [of w1]
- apply (simp add: int_nat_two_exp)
- apply (rule zero_le_power,simp)
- using bi1
- apply simp
- done
- hence "-?Q \<le> ((2::int)^length w1) * (2 ^ (length w2 - 1))"
- by (simp add: zmult_ac)
- thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \<le> ?Q"
- by simp
- next
- assume bi1: "bv_to_int (utos w1) < 0"
- thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \<le> ?Q"
- by (simp add: bv_to_int_utos)
- qed
- qed
- finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \<le> ?Q" .
- qed
- qed
-qed
-
-lemma bv_smult_sym: "bv_smult w1 w2 = bv_smult w2 w1"
- by (simp add: bv_smult_def zmult_ac)
-
-subsection {* Structural operations *}
-
-definition
- bv_select :: "[bit list,nat] => bit" where
- "bv_select w i = w ! (length w - 1 - i)"
-
-definition
- bv_chop :: "[bit list,nat] => bit list * bit list" where
- "bv_chop w i = (let len = length w in (take (len - i) w,drop (len - i) w))"
-
-definition
- bv_slice :: "[bit list,nat*nat] => bit list" where
- "bv_slice w = (\<lambda>(b,e). fst (bv_chop (snd (bv_chop w (b+1))) e))"
-
-lemma bv_select_rev:
- assumes notnull: "n < length w"
- shows "bv_select w n = rev w ! n"
-proof -
- have "\<forall>n. n < length w --> bv_select w n = rev w ! n"
- proof (rule length_induct [of _ w],auto simp add: bv_select_def)
- fix xs :: "bit list"
- fix n
- assume ind: "\<forall>ys::bit list. length ys < length xs --> (\<forall>n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n)"
- assume notx: "n < length xs"
- show "xs ! (length xs - Suc n) = rev xs ! n"
- proof (cases xs)
- assume "xs = []"
- with notx show ?thesis by simp
- next
- fix y ys
- assume [simp]: "xs = y # ys"
- show ?thesis
- proof (auto simp add: nth_append)
- assume noty: "n < length ys"
- from spec [OF ind,of ys]
- have "\<forall>n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n"
- by simp
- hence "n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" ..
- from this and noty
- have "ys ! (length ys - Suc n) = rev ys ! n" ..
- thus "(y # ys) ! (length ys - n) = rev ys ! n"
- by (simp add: nth_Cons' noty linorder_not_less [symmetric])
- next
- assume "~ n < length ys"
- hence x: "length ys \<le> n" by simp
- from notx have "n < Suc (length ys)" by simp
- hence "n \<le> length ys" by simp
- with x have "length ys = n" by simp
- thus "y = [y] ! (n - length ys)" by simp
- qed
- qed
- qed
- then have "n < length w --> bv_select w n = rev w ! n" ..
- from this and notnull show ?thesis ..
-qed
-
-lemma bv_chop_append: "bv_chop (w1 @ w2) (length w2) = (w1,w2)"
- by (simp add: bv_chop_def Let_def)
-
-lemma append_bv_chop_id: "fst (bv_chop w l) @ snd (bv_chop w l) = w"
- by (simp add: bv_chop_def Let_def)
-
-lemma bv_chop_length_fst [simp]: "length (fst (bv_chop w i)) = length w - i"
- by (simp add: bv_chop_def Let_def)
-
-lemma bv_chop_length_snd [simp]: "length (snd (bv_chop w i)) = min i (length w)"
- by (simp add: bv_chop_def Let_def)
-
-lemma bv_slice_length [simp]: "[| j \<le> i; i < length w |] ==> length (bv_slice w (i,j)) = i - j + 1"
- by (auto simp add: bv_slice_def)
-
-definition
- length_nat :: "nat => nat" where
- [code del]: "length_nat x = (LEAST n. x < 2 ^ n)"
-
-lemma length_nat: "length (nat_to_bv n) = length_nat n"
- apply (simp add: length_nat_def)
- apply (rule Least_equality [symmetric])
- prefer 2
- apply (rule length_nat_to_bv_upper_limit)
- apply arith
- apply (rule ccontr)
-proof -
- assume "~ n < 2 ^ length (nat_to_bv n)"
- hence "2 ^ length (nat_to_bv n) \<le> n" by simp
- hence "length (nat_to_bv n) < length (nat_to_bv n)"
- by (rule length_nat_to_bv_lower_limit)
- thus False by simp
-qed
-
-lemma length_nat_0 [simp]: "length_nat 0 = 0"
- by (simp add: length_nat_def Least_equality)
-
-lemma length_nat_non0:
- assumes n0: "n \<noteq> 0"
- shows "length_nat n = Suc (length_nat (n div 2))"
- apply (simp add: length_nat [symmetric])
- apply (subst nat_to_bv_non0 [of n])
- apply (simp_all add: n0)
- done
-
-definition
- length_int :: "int => nat" where
- "length_int x =
- (if 0 < x then Suc (length_nat (nat x))
- else if x = 0 then 0
- else Suc (length_nat (nat (-x - 1))))"
-
-lemma length_int: "length (int_to_bv i) = length_int i"
-proof (cases "0 < i")
- assume i0: "0 < i"
- hence "length (int_to_bv i) =
- length (norm_signed (\<zero> # norm_unsigned (nat_to_bv (nat i))))" by simp
- also from norm_unsigned_result [of "nat_to_bv (nat i)"]
- have "... = Suc (length_nat (nat i))"
- apply safe
- apply (simp del: norm_unsigned_nat_to_bv)
- apply (drule norm_empty_bv_to_nat_zero)
- using prems
- apply simp
- apply (cases "norm_unsigned (nat_to_bv (nat i))")
- apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat i)"])
- using prems
- apply simp
- apply simp
- using prems
- apply (simp add: length_nat [symmetric])
- done
- finally show ?thesis
- using i0
- by (simp add: length_int_def)
-next
- assume "~ 0 < i"
- hence i0: "i \<le> 0" by simp
- show ?thesis
- proof (cases "i = 0")
- assume "i = 0"
- thus ?thesis by (simp add: length_int_def)
- next
- assume "i \<noteq> 0"
- with i0 have i0: "i < 0" by simp
- hence "length (int_to_bv i) =
- length (norm_signed (\<one> # bv_not (norm_unsigned (nat_to_bv (nat (- i) - 1)))))"
- by (simp add: int_to_bv_def nat_diff_distrib)
- also from norm_unsigned_result [of "nat_to_bv (nat (- i) - 1)"]
- have "... = Suc (length_nat (nat (- i) - 1))"
- apply safe
- apply (simp del: norm_unsigned_nat_to_bv)
- apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat (-i) - Suc 0)"])
- using prems
- apply simp
- apply (cases "- i - 1 = 0")
- apply simp
- apply (simp add: length_nat [symmetric])
- apply (cases "norm_unsigned (nat_to_bv (nat (- i) - 1))")
- apply simp
- apply simp
- done
- finally
- show ?thesis
- using i0 by (simp add: length_int_def nat_diff_distrib del: int_to_bv_lt0)
- qed
-qed
-
-lemma length_int_0 [simp]: "length_int 0 = 0"
- by (simp add: length_int_def)
-
-lemma length_int_gt0: "0 < i ==> length_int i = Suc (length_nat (nat i))"
- by (simp add: length_int_def)
-
-lemma length_int_lt0: "i < 0 ==> length_int i = Suc (length_nat (nat (- i) - 1))"
- by (simp add: length_int_def nat_diff_distrib)
-
-lemma bv_chopI: "[| w = w1 @ w2 ; i = length w2 |] ==> bv_chop w i = (w1,w2)"
- by (simp add: bv_chop_def Let_def)
-
-lemma bv_sliceI: "[| j \<le> i ; i < length w ; w = w1 @ w2 @ w3 ; Suc i = length w2 + j ; j = length w3 |] ==> bv_slice w (i,j) = w2"
- apply (simp add: bv_slice_def)
- apply (subst bv_chopI [of "w1 @ w2 @ w3" w1 "w2 @ w3"])
- apply simp
- apply simp
- apply simp
- apply (subst bv_chopI [of "w2 @ w3" w2 w3],simp_all)
- done
-
-lemma bv_slice_bv_slice:
- assumes ki: "k \<le> i"
- and ij: "i \<le> j"
- and jl: "j \<le> l"
- and lw: "l < length w"
- shows "bv_slice w (j,i) = bv_slice (bv_slice w (l,k)) (j-k,i-k)"
-proof -
- def w1 == "fst (bv_chop w (Suc l))"
- and w2 == "fst (bv_chop (snd (bv_chop w (Suc l))) (Suc j))"
- and w3 == "fst (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)"
- and w4 == "fst (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)"
- and w5 == "snd (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)"
- note w_defs = this
-
- have w_def: "w = w1 @ w2 @ w3 @ w4 @ w5"
- by (simp add: w_defs append_bv_chop_id)
-
- from ki ij jl lw
- show ?thesis
- apply (subst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"])
- apply simp_all
- apply (rule w_def)
- apply (simp add: w_defs)
- apply (simp add: w_defs)
- apply (subst bv_sliceI [where ?j = k and ?i = l and ?w = w and ?w1.0 = w1 and ?w2.0 = "w2 @ w3 @ w4" and ?w3.0 = w5])
- apply simp_all
- apply (rule w_def)
- apply (simp add: w_defs)
- apply (simp add: w_defs)
- apply (subst bv_sliceI [where ?j = "i-k" and ?i = "j-k" and ?w = "w2 @ w3 @ w4" and ?w1.0 = w2 and ?w2.0 = w3 and ?w3.0 = w4])
- apply simp_all
- apply (simp_all add: w_defs)
- done
-qed
-
-lemma bv_to_nat_extend [simp]: "bv_to_nat (bv_extend n \<zero> w) = bv_to_nat w"
- apply (simp add: bv_extend_def)
- apply (subst bv_to_nat_dist_append)
- apply simp
- apply (induct ("n - length w"))
- apply simp_all
- done
-
-lemma bv_msb_extend_same [simp]: "bv_msb w = b ==> bv_msb (bv_extend n b w) = b"
- apply (simp add: bv_extend_def)
- apply (cases "n - length w")
- apply simp_all
- done
-
-lemma bv_to_int_extend [simp]:
- assumes a: "bv_msb w = b"
- shows "bv_to_int (bv_extend n b w) = bv_to_int w"
-proof (cases "bv_msb w")
- assume [simp]: "bv_msb w = \<zero>"
- with a have [simp]: "b = \<zero>" by simp
- show ?thesis by (simp add: bv_to_int_def)
-next
- assume [simp]: "bv_msb w = \<one>"
- with a have [simp]: "b = \<one>" by simp
- show ?thesis
- apply (simp add: bv_to_int_def)
- apply (simp add: bv_extend_def)
- apply (induct ("n - length w"), simp_all)
- done
-qed
-
-lemma length_nat_mono [simp]: "x \<le> y ==> length_nat x \<le> length_nat y"
-proof (rule ccontr)
- assume xy: "x \<le> y"
- assume "~ length_nat x \<le> length_nat y"
- hence lxly: "length_nat y < length_nat x"
- by simp
- hence "length_nat y < (LEAST n. x < 2 ^ n)"
- by (simp add: length_nat_def)
- hence "~ x < 2 ^ length_nat y"
- by (rule not_less_Least)
- hence xx: "2 ^ length_nat y \<le> x"
- by simp
- have yy: "y < 2 ^ length_nat y"
- apply (simp add: length_nat_def)
- apply (rule LeastI)
- apply (subgoal_tac "y < 2 ^ y",assumption)
- apply (cases "0 \<le> y")
- apply (induct y,simp_all)
- done
- with xx have "y < x" by simp
- with xy show False by simp
-qed
-
-lemma length_nat_mono_int [simp]: "x \<le> y ==> length_nat x \<le> length_nat y"
- by (rule length_nat_mono) arith
-
-lemma length_nat_pos [simp,intro!]: "0 < x ==> 0 < length_nat x"
- by (simp add: length_nat_non0)
-
-lemma length_int_mono_gt0: "[| 0 \<le> x ; x \<le> y |] ==> length_int x \<le> length_int y"
- by (cases "x = 0") (simp_all add: length_int_gt0 nat_le_eq_zle)
-
-lemma length_int_mono_lt0: "[| x \<le> y ; y \<le> 0 |] ==> length_int y \<le> length_int x"
- by (cases "y = 0") (simp_all add: length_int_lt0)
-
-lemmas [simp] = length_nat_non0
-
-lemma "nat_to_bv (number_of Int.Pls) = []"
- by simp
-
-primrec fast_bv_to_nat_helper :: "[bit list, int] => int" where
- fast_bv_to_nat_Nil: "fast_bv_to_nat_helper [] k = k"
- | fast_bv_to_nat_Cons: "fast_bv_to_nat_helper (b#bs) k =
- fast_bv_to_nat_helper bs ((bit_case Int.Bit0 Int.Bit1 b) k)"
-
-declare fast_bv_to_nat_helper.simps [code del]
-
-lemma fast_bv_to_nat_Cons0: "fast_bv_to_nat_helper (\<zero>#bs) bin =
- fast_bv_to_nat_helper bs (Int.Bit0 bin)"
- by simp
-
-lemma fast_bv_to_nat_Cons1: "fast_bv_to_nat_helper (\<one>#bs) bin =
- fast_bv_to_nat_helper bs (Int.Bit1 bin)"
- by simp
-
-lemma fast_bv_to_nat_def:
- "bv_to_nat bs == number_of (fast_bv_to_nat_helper bs Int.Pls)"
-proof (simp add: bv_to_nat_def)
- have "\<forall> bin. \<not> (neg (number_of bin :: int)) \<longrightarrow> (foldl (%bn b. 2 * bn + bitval b) (number_of bin) bs) = number_of (fast_bv_to_nat_helper bs bin)"
- apply (induct bs,simp)
- apply (case_tac a,simp_all)
- done
- thus "foldl (\<lambda>bn b. 2 * bn + bitval b) 0 bs \<equiv> number_of (fast_bv_to_nat_helper bs Int.Pls)"
- by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric])
-qed
-
-declare fast_bv_to_nat_Cons [simp del]
-declare fast_bv_to_nat_Cons0 [simp]
-declare fast_bv_to_nat_Cons1 [simp]
-
-setup {*
-(*comments containing lcp are the removal of fast_bv_of_nat*)
-let
- fun is_const_bool (Const("True",_)) = true
- | is_const_bool (Const("False",_)) = true
- | is_const_bool _ = false
- fun is_const_bit (Const("Word.bit.Zero",_)) = true
- | is_const_bit (Const("Word.bit.One",_)) = true
- | is_const_bit _ = false
- fun num_is_usable (Const(@{const_name Int.Pls},_)) = true
- | num_is_usable (Const(@{const_name Int.Min},_)) = false
- | num_is_usable (Const(@{const_name Int.Bit0},_) $ w) =
- num_is_usable w
- | num_is_usable (Const(@{const_name Int.Bit1},_) $ w) =
- num_is_usable w
- | num_is_usable _ = false
- fun vec_is_usable (Const("List.list.Nil",_)) = true
- | vec_is_usable (Const("List.list.Cons",_) $ b $ bs) =
- vec_is_usable bs andalso is_const_bit b
- | vec_is_usable _ = false
- (*lcp** val fast1_th = PureThy.get_thm thy "Word.fast_nat_to_bv_def"*)
- val fast2_th = @{thm "Word.fast_bv_to_nat_def"};
- (*lcp** fun f sg thms (Const("Word.nat_to_bv",_) $ (Const(@{const_name Int.number_of},_) $ t)) =
- if num_is_usable t
- then SOME (Drule.cterm_instantiate [(cterm_of sg (Var (("w", 0), @{typ int})), cterm_of sg t)] fast1_th)
- else NONE
- | f _ _ _ = NONE *)
- fun g sg thms (Const("Word.bv_to_nat",_) $ (t as (Const("List.list.Cons",_) $ _ $ _))) =
- if vec_is_usable t then
- SOME (Drule.cterm_instantiate [(cterm_of sg (Var(("bs",0),Type("List.list",[Type("Word.bit",[])]))),cterm_of sg t)] fast2_th)
- else NONE
- | g _ _ _ = NONE
- (*lcp** val simproc1 = Simplifier.simproc thy "nat_to_bv" ["Word.nat_to_bv (number_of w)"] f *)
- val simproc2 = Simplifier.simproc @{theory} "bv_to_nat" ["Word.bv_to_nat (x # xs)"] g
-in
- Simplifier.map_simpset (fn ss => ss addsimprocs [(*lcp*simproc1,*)simproc2])
-end*}
-
-declare bv_to_nat1 [simp del]
-declare bv_to_nat_helper [simp del]
-
-definition
- bv_mapzip :: "[bit => bit => bit,bit list, bit list] => bit list" where
- "bv_mapzip f w1 w2 =
- (let g = bv_extend (max (length w1) (length w2)) \<zero>
- in map (split f) (zip (g w1) (g w2)))"
-
-lemma bv_length_bv_mapzip [simp]:
- "length (bv_mapzip f w1 w2) = max (length w1) (length w2)"
- by (simp add: bv_mapzip_def Let_def split: split_max)
-
-lemma bv_mapzip_Nil [simp]: "bv_mapzip f [] [] = []"
- by (simp add: bv_mapzip_def Let_def)
-
-lemma bv_mapzip_Cons [simp]: "length w1 = length w2 ==>
- bv_mapzip f (x#w1) (y#w2) = f x y # bv_mapzip f w1 w2"
- by (simp add: bv_mapzip_def Let_def)
-
-end
--- a/src/HOL/Library/normarith.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,416 +0,0 @@
-(* Title: Library/normarith.ML
- Author: Amine Chaieb, University of Cambridge
- Description: A simple decision procedure for linear problems in euclidean space
-*)
-
- (* Now the norm procedure for euclidean spaces *)
-
-
-signature NORM_ARITH =
-sig
- val norm_arith : Proof.context -> conv
- val norm_arith_tac : Proof.context -> int -> tactic
-end
-
-structure NormArith : NORM_ARITH =
-struct
-
- open Conv;
- val bool_eq = op = : bool *bool -> bool
- fun dest_ratconst t = case term_of t of
- Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
- | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
- | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
- fun is_ratconst t = can dest_ratconst t
- fun augment_norm b t acc = case term_of t of
- Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
- | _ => acc
- fun find_normedterms t acc = case term_of t of
- @{term "op + :: real => _"}$_$_ =>
- find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
- | @{term "op * :: real => _"}$_$n =>
- if not (is_ratconst (Thm.dest_arg1 t)) then acc else
- augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero)
- (Thm.dest_arg t) acc
- | _ => augment_norm true t acc
-
- val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
- fun cterm_lincomb_cmul c t =
- if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
- fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
- fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
- fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
-
- val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
- fun int_lincomb_cmul c t =
- if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
- fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
- fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
- fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
-
-fun vector_lincomb t = case term_of t of
- Const(@{const_name plus}, _) $ _ $ _ =>
- cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name minus}, _) $ _ $ _ =>
- cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name scaleR}, _)$_$_ =>
- cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
- | Const(@{const_name uminus}, _)$_ =>
- cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
-(* FIXME: how should we handle numerals?
- | Const(@ {const_name vec},_)$_ =>
- let
- val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0
- handle TERM _=> false)
- in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one)
- else FuncUtil.Ctermfunc.empty
- end
-*)
- | _ => FuncUtil.Ctermfunc.onefunc (t,Rat.one)
-
- fun vector_lincombs ts =
- fold_rev
- (fn t => fn fns => case AList.lookup (op aconvc) fns t of
- NONE =>
- let val f = vector_lincomb t
- in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of
- SOME (_,f') => (t,f') :: fns
- | NONE => (t,f) :: fns
- end
- | SOME _ => fns) ts []
-
-fun replacenegnorms cv t = case term_of t of
- @{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
-| @{term "op * :: real => _"}$_$_ =>
- if dest_ratconst (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
-| _ => reflexive t
-fun flip v eq =
- if FuncUtil.Ctermfunc.defined eq v
- then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.Ctermfunc.apply eq v)) eq else eq
-fun allsubsets s = case s of
- [] => [[]]
-|(a::t) => let val res = allsubsets t in
- map (cons a) res @ res end
-fun evaluate env lin =
- FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x))
- lin Rat.zero
-
-fun solve (vs,eqs) = case (vs,eqs) of
- ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one))
- |(_,eq::oeqs) =>
- (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
- [] => NONE
- | v::_ =>
- if FuncUtil.Intfunc.defined eq v
- then
- let
- val c = FuncUtil.Intfunc.apply eq v
- val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
- fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn
- else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
- in (case solve (remove (op =) v vs, map eliminate oeqs) of
- NONE => NONE
- | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln))
- end
- else NONE)
-
-fun combinations k l = if k = 0 then [[]] else
- case l of
- [] => []
-| h::t => map (cons h) (combinations (k - 1) t) @ combinations k t
-
-
-fun forall2 p l1 l2 = case (l1,l2) of
- ([],[]) => true
- | (h1::t1,h2::t2) => p h1 h2 andalso forall2 p t1 t2
- | _ => false;
-
-
-fun vertices vs eqs =
- let
- fun vertex cmb = case solve(vs,cmb) of
- NONE => NONE
- | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v Rat.zero) vs)
- val rawvs = map_filter vertex (combinations (length vs) eqs)
- val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs
- in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset []
- end
-
-fun subsumes l m = forall2 (fn x => fn y => Rat.abs x <=/ Rat.abs y) l m
-
-fun subsume todo dun = case todo of
- [] => dun
-|v::ovs =>
- let val dun' = if exists (fn w => subsumes w v) dun then dun
- else v::(filter (fn w => not(subsumes v w)) dun)
- in subsume ovs dun'
- end;
-
-fun match_mp PQ P = P RS PQ;
-
-fun cterm_of_rat x =
-let val (a, b) = Rat.quotient_of_rat x
-in
- if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
- else Thm.capply (Thm.capply @{cterm "op / :: real => _"}
- (Numeral.mk_cnumber @{ctyp "real"} a))
- (Numeral.mk_cnumber @{ctyp "real"} b)
-end;
-
-fun norm_cmul_rule c th = instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
-
-fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
-
- (* I think here the static context should be sufficient!! *)
-fun inequality_canon_rule ctxt =
- let
- (* FIXME : Should be computed statically!! *)
- val real_poly_conv =
- Semiring_Normalizer.semiring_normalize_wrapper ctxt
- (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
- in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv then_conv real_poly_conv)))
-end;
-
- fun absc cv ct = case term_of ct of
- Abs (v,_, _) =>
- let val (x,t) = Thm.dest_abs (SOME v) ct
- in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t)
- end
- | _ => all_conv ct;
-
-fun sub_conv cv ct = (comb_conv cv else_conv absc cv) ct;
-fun botc1 conv ct =
- ((sub_conv (botc1 conv)) then_conv (conv else_conv all_conv)) ct;
-
- fun rewrs_conv eqs ct = first_conv (map rewr_conv eqs) ct;
- val apply_pth1 = rewr_conv @{thm pth_1};
- val apply_pth2 = rewr_conv @{thm pth_2};
- val apply_pth3 = rewr_conv @{thm pth_3};
- val apply_pth4 = rewrs_conv @{thms pth_4};
- val apply_pth5 = rewr_conv @{thm pth_5};
- val apply_pth6 = rewr_conv @{thm pth_6};
- val apply_pth7 = rewrs_conv @{thms pth_7};
- val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
- val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.field_comp_conv);
- val apply_ptha = rewr_conv @{thm pth_a};
- val apply_pthb = rewrs_conv @{thms pth_b};
- val apply_pthc = rewrs_conv @{thms pth_c};
- val apply_pthd = try_conv (rewr_conv @{thm pth_d});
-
-fun headvector t = case t of
- Const(@{const_name plus}, _)$
- (Const(@{const_name scaleR}, _)$l$v)$r => v
- | Const(@{const_name scaleR}, _)$l$v => v
- | _ => error "headvector: non-canonical term"
-
-fun vector_cmul_conv ct =
- ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.field_comp_conv) else_conv
- (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct
-
-fun vector_add_conv ct = apply_pth7 ct
- handle CTERM _ =>
- (apply_pth8 ct
- handle CTERM _ =>
- (case term_of ct of
- Const(@{const_name plus},_)$lt$rt =>
- let
- val l = headvector lt
- val r = headvector rt
- in (case Term_Ord.fast_term_ord (l,r) of
- LESS => (apply_pthb then_conv arg_conv vector_add_conv
- then_conv apply_pthd) ct
- | GREATER => (apply_pthc then_conv arg_conv vector_add_conv
- then_conv apply_pthd) ct
- | EQUAL => (apply_pth9 then_conv
- ((apply_ptha then_conv vector_add_conv) else_conv
- arg_conv vector_add_conv then_conv apply_pthd)) ct)
- end
- | _ => reflexive ct))
-
-fun vector_canon_conv ct = case term_of ct of
- Const(@{const_name plus},_)$_$_ =>
- let
- val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
- val lth = vector_canon_conv l
- val rth = vector_canon_conv r
- val th = Drule.binop_cong_rule p lth rth
- in fconv_rule (arg_conv vector_add_conv) th end
-
-| Const(@{const_name scaleR}, _)$_$_ =>
- let
- val (p,r) = Thm.dest_comb ct
- val rth = Drule.arg_cong_rule p (vector_canon_conv r)
- in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth
- end
-
-| Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct
-
-| Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct
-
-(* FIXME
-| Const(@{const_name vec},_)$n =>
- let val n = Thm.dest_arg ct
- in if is_ratconst n andalso not (dest_ratconst n =/ Rat.zero)
- then reflexive ct else apply_pth1 ct
- end
-*)
-| _ => apply_pth1 ct
-
-fun norm_canon_conv ct = case term_of ct of
- Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct
- | _ => raise CTERM ("norm_canon_conv", [ct])
-
-fun fold_rev2 f [] [] z = z
- | fold_rev2 f (x::xs) (y::ys) z = f x y (fold_rev2 f xs ys z)
- | fold_rev2 f _ _ _ = raise UnequalLengths;
-
-fun int_flip v eq =
- if FuncUtil.Intfunc.defined eq v
- then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq;
-
-local
- val pth_zero = @{thm norm_zero}
- val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
- pth_zero
- val concl = Thm.dest_arg o cprop_of
- fun real_vector_combo_prover ctxt translator (nubs,ges,gts) =
- let
- (* FIXME: Should be computed statically!!*)
- val real_poly_conv =
- Semiring_Normalizer.semiring_normalize_wrapper ctxt
- (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
- val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
- val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) []
- val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check"
- else ()
- val dests = distinct (op aconvc) (map snd rawdests)
- val srcfuns = map vector_lincomb sources
- val destfuns = map vector_lincomb dests
- val vvs = fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) []
- val n = length srcfuns
- val nvs = 1 upto n
- val srccombs = srcfuns ~~ nvs
- fun consider d =
- let
- fun coefficients x =
- let
- val inp = if FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x))
- else FuncUtil.Intfunc.empty
- in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp
- end
- val equations = map coefficients vvs
- val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,Rat.one)) nvs
- fun plausiblevertices f =
- let
- val flippedequations = map (fold_rev int_flip f) equations
- val constraints = flippedequations @ inequalities
- val rawverts = vertices nvs constraints
- fun check_solution v =
- let
- val f = fold_rev2 (curry FuncUtil.Intfunc.update) nvs v (FuncUtil.Intfunc.onefunc (0, Rat.one))
- in forall (fn e => evaluate f e =/ Rat.zero) flippedequations
- end
- val goodverts = filter check_solution rawverts
- val signfixups = map (fn n => if member (op =) f n then ~1 else 1) nvs
- in map (map2 (fn s => fn c => Rat.rat_of_int s */ c) signfixups) goodverts
- end
- val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) []
- in subsume allverts []
- end
- fun compute_ineq v =
- let
- val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE
- else SOME(norm_cmul_rule v t))
- (v ~~ nubs)
- fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
- in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
- end
- val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
- map (inequality_canon_rule ctxt) nubs @ ges
- val zerodests = filter
- (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
-
- in fst (RealArith.real_linear_prover translator
- (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
- zerodests,
- map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
- arg_conv (arg_conv real_poly_conv))) ges',
- map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
- arg_conv (arg_conv real_poly_conv))) gts))
- end
-in val real_vector_combo_prover = real_vector_combo_prover
-end;
-
-local
- val pth = @{thm norm_imp_pos_and_ge}
- val norm_mp = match_mp pth
- val concl = Thm.dest_arg o cprop_of
- fun conjunct1 th = th RS @{thm conjunct1}
- fun conjunct2 th = th RS @{thm conjunct2}
-fun real_vector_ineq_prover ctxt translator (ges,gts) =
- let
-(* val _ = error "real_vector_ineq_prover: pause" *)
- val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) []
- val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
- val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
- fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
- fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
- fun mk_equals l r = Thm.capply (Thm.capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
- val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
- val replace_conv = try_conv (rewrs_conv asl)
- val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
- val ges' =
- fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths)
- asl (map replace_rule ges)
- val gts' = map replace_rule gts
- val nubs = map (conjunct2 o norm_mp) asl
- val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
- val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1))
- val th11 = hd (Variable.export ctxt' ctxt [fold implies_intr shs th1])
- val cps = map (swap o Thm.dest_equals) (cprems_of th11)
- val th12 = instantiate ([], cps) th11
- val th13 = fold Thm.elim_implies (map (reflexive o snd) cps) th12;
- in hd (Variable.export ctxt' ctxt [th13])
- end
-in val real_vector_ineq_prover = real_vector_ineq_prover
-end;
-
-local
- val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
- fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
- fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
- (* FIXME: Lookup in the context every time!!! Fix this !!!*)
- fun splitequation ctxt th acc =
- let
- val real_poly_neg_conv = #neg
- (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
- (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
- val (th1,th2) = conj_pair(rawrule th)
- in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
- end
-in fun real_vector_prover ctxt _ translator (eqs,ges,gts) =
- (real_vector_ineq_prover ctxt translator
- (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial)
-end;
-
- fun init_conv ctxt =
- Simplifier.rewrite (Simplifier.context ctxt
- (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
- then_conv Numeral_Simprocs.field_comp_conv
- then_conv nnf_conv
-
- fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
- fun norm_arith ctxt ct =
- let
- val ctxt' = Variable.declare_term (term_of ct) ctxt
- val th = init_conv ctxt' ct
- in equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (symmetric th))
- (pure ctxt' (Thm.rhs_of th))
- end
-
- fun norm_arith_tac ctxt =
- clarify_tac HOL_cs THEN'
- Object_Logic.full_atomize_tac THEN'
- CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
-
-end;
--- a/src/HOL/Library/positivstellensatz.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/positivstellensatz.ML Sun May 23 10:38:11 2010 +0100
@@ -182,12 +182,12 @@
(* Some useful derived rules *)
fun deduct_antisym_rule tha thb =
- equal_intr (implies_intr (cprop_of thb) tha)
- (implies_intr (cprop_of tha) thb);
+ Thm.equal_intr (Thm.implies_intr (cprop_of thb) tha)
+ (Thm.implies_intr (cprop_of tha) thb);
fun prove_hyp tha thb =
if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb))
- then equal_elim (symmetric (deduct_antisym_rule tha thb)) tha else thb;
+ then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;
val pth = @{lemma "(((x::real) < y) == (y - x > 0))" and "((x <= y) == (y - x >= 0))" and
"((x = y) == (x - y = 0))" and "((~(x < y)) == (x - y >= 0))" and
@@ -375,7 +375,7 @@
val skolemize_conv = Simplifier.rewrite (Simplifier.context ctxt skolemize_ss)
val weak_dnf_ss = HOL_basic_ss addsimps weak_dnf_simps
val weak_dnf_conv = Simplifier.rewrite (Simplifier.context ctxt weak_dnf_ss)
- fun eqT_elim th = equal_elim (symmetric th) @{thm TrueI}
+ fun eqT_elim th = Thm.equal_elim (Thm.symmetric th) @{thm TrueI}
fun oprconv cv ct =
let val g = Thm.dest_fun2 ct
in if g aconvc @{cterm "op <= :: real => _"}
@@ -387,7 +387,7 @@
let
val th' = (Thm.instantiate (Thm.match (Thm.lhs_of th, ct)) th
handle MATCH => raise CTERM ("real_ineq_conv", [ct]))
- in transitive th' (oprconv poly_conv (Thm.rhs_of th'))
+ in Thm.transitive th' (oprconv poly_conv (Thm.rhs_of th'))
end
val [real_lt_conv, real_le_conv, real_eq_conv,
real_not_lt_conv, real_not_le_conv, _] =
@@ -446,10 +446,10 @@
let val (p,q) = Thm.dest_binop (concl th)
val c = concl th1
val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"
- in implies_elim (implies_elim
- (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th)
- (implies_intr (Thm.capply @{cterm Trueprop} p) th1))
- (implies_intr (Thm.capply @{cterm Trueprop} q) th2)
+ in Thm.implies_elim (Thm.implies_elim
+ (Thm.implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th)
+ (Thm.implies_intr (Thm.capply @{cterm Trueprop} p) th1))
+ (Thm.implies_intr (Thm.capply @{cterm Trueprop} q) th2)
end
fun overall cert_choice dun ths = case ths of
[] =>
@@ -468,8 +468,8 @@
overall cert_choice dun (th1::th2::oths) end
else if is_disj ct then
let
- val (th1, cert1) = overall (Left::cert_choice) dun (assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg1 ct))::oths)
- val (th2, cert2) = overall (Right::cert_choice) dun (assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg ct))::oths)
+ val (th1, cert1) = overall (Left::cert_choice) dun (Thm.assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg1 ct))::oths)
+ val (th2, cert2) = overall (Right::cert_choice) dun (Thm.assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg ct))::oths)
in (disj_cases th th1 th2, Branch (cert1, cert2)) end
else overall cert_choice (th::dun) oths
end
@@ -487,12 +487,12 @@
val th' = Drule.binop_cong_rule @{cterm "op |"}
(Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
(Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
- in transitive th th'
+ in Thm.transitive th th'
end
fun equal_implies_1_rule PQ =
let
val P = Thm.lhs_of PQ
- in implies_intr P (equal_elim PQ (assume P))
+ in Thm.implies_intr P (Thm.equal_elim PQ (Thm.assume P))
end
(* FIXME!!! Copied from groebner.ml *)
val strip_exists =
@@ -507,7 +507,7 @@
| Var ((s,_),_) => s
| _ => "x"
- fun mk_forall x th = Drule.arg_cong_rule (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" }) (abstract_rule (name_of x) x th)
+ fun mk_forall x th = Drule.arg_cong_rule (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" }) (Thm.abstract_rule (name_of x) x th)
val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
@@ -523,12 +523,12 @@
(instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o cprop_of) th'] exE)
val pv = (Thm.rhs_of o Thm.beta_conversion true)
(Thm.capply @{cterm Trueprop} (Thm.capply p v))
- val th1 = forall_intr v (implies_intr pv th')
- in implies_elim (implies_elim th0 th) th1 end
+ val th1 = Thm.forall_intr v (Thm.implies_intr pv th')
+ in Thm.implies_elim (Thm.implies_elim th0 th) th1 end
| _ => raise THM ("choose",0,[th, th'])
fun simple_choose v th =
- choose v (assume ((Thm.capply @{cterm Trueprop} o mk_ex v) ((Thm.dest_arg o hd o #hyps o Thm.crep_thm) th))) th
+ choose v (Thm.assume ((Thm.capply @{cterm Trueprop} o mk_ex v) ((Thm.dest_arg o hd o #hyps o Thm.crep_thm) th))) th
val strip_forall =
let fun h (acc, t) =
@@ -558,11 +558,11 @@
val (evs,bod) = strip_exists tm0
val (avs,ibod) = strip_forall bod
val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
- val (th2, certs) = overall [] [] [specl avs (assume (Thm.rhs_of th1))]
- val th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (Thm.capply @{cterm Trueprop} bod))) th2)
- in (Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (assume nct)) th3), certs)
+ val (th2, certs) = overall [] [] [specl avs (Thm.assume (Thm.rhs_of th1))]
+ val th3 = fold simple_choose evs (prove_hyp (Thm.equal_elim th1 (Thm.assume (Thm.capply @{cterm Trueprop} bod))) th2)
+ in (Drule.implies_intr_hyps (prove_hyp (Thm.equal_elim th0 (Thm.assume nct)) th3), certs)
end
- in (implies_elim (instantiate' [] [SOME ct] pth_final) th, certificates)
+ in (Thm.implies_elim (instantiate' [] [SOME ct] pth_final) th, certificates)
end
in f
end;
@@ -715,10 +715,10 @@
fun eliminate_construct p c tm =
let
val t = find_cterm p tm
- val th0 = (symmetric o beta_conversion false) (Thm.capply (Thm.cabs t tm) t)
+ val th0 = (Thm.symmetric o Thm.beta_conversion false) (Thm.capply (Thm.cabs t tm) t)
val (p,ax) = (Thm.dest_comb o Thm.rhs_of) th0
- in fconv_rule(arg_conv(binop_conv (arg_conv (beta_conversion false))))
- (transitive th0 (c p ax))
+ in fconv_rule(arg_conv(binop_conv (arg_conv (Thm.beta_conversion false))))
+ (Thm.transitive th0 (c p ax))
end
val elim_abs = eliminate_construct is_abs
--- a/src/HOL/Library/reflection.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Library/reflection.ML Sun May 23 10:38:11 2010 +0100
@@ -133,7 +133,7 @@
(AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) bds))
in (([(ta, ctxt')],
fn ([th], bds) =>
- (hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI]),
+ (hd (Variable.export ctxt' ctxt [(Thm.forall_intr (cert x) th) COMP allI]),
let val (bsT,asT) = the(AList.lookup Type.could_unify bds (HOLogic.listT xT))
in AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) bds
end)),
--- a/src/HOL/List.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/List.thy Sun May 23 10:38:11 2010 +0100
@@ -2532,6 +2532,23 @@
"distinct xs \<Longrightarrow> listsum xs = Setsum(set xs)"
by (induct xs) simp_all
+lemma listsum_eq_0_nat_iff_nat[simp]:
+ "listsum ns = (0::nat) \<longleftrightarrow> (\<forall> n \<in> set ns. n = 0)"
+by(simp add: listsum)
+
+lemma elem_le_listsum_nat: "k<size ns \<Longrightarrow> ns!k <= listsum(ns::nat list)"
+apply(induct ns arbitrary: k)
+ apply simp
+apply(fastsimp simp add:nth_Cons split: nat.split)
+done
+
+lemma listsum_update_nat: "k<size ns \<Longrightarrow>
+ listsum (ns[k := (n::nat)]) = listsum ns + n - ns!k"
+apply(induct ns arbitrary:k)
+ apply (auto split:nat.split)
+apply(drule elem_le_listsum_nat)
+apply arith
+done
text{* Some syntactic sugar for summing a function over a list: *}
@@ -4474,11 +4491,8 @@
subsubsection {* Generation of efficient code *}
-primrec
- member :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infixl "mem" 55)
-where
- "x mem [] \<longleftrightarrow> False"
- | "x mem (y#ys) \<longleftrightarrow> x = y \<or> x mem ys"
+definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
+ mem_iff [code_post]: "member xs x \<longleftrightarrow> x \<in> set xs"
primrec
null:: "'a list \<Rightarrow> bool"
@@ -4534,7 +4548,7 @@
| "concat_map f (x#xs) = f x @ concat_map f xs"
text {*
- Only use @{text mem} for generating executable code. Otherwise use
+ Only use @{text member} for generating executable code. Otherwise use
@{prop "x : set xs"} instead --- it is much easier to reason about.
The same is true for @{const list_all} and @{const list_ex}: write
@{text "\<forall>x\<in>set xs"} and @{text "\<exists>x\<in>set xs"} instead because the HOL
@@ -4566,12 +4580,20 @@
show ?case by (induct xs) (auto simp add: Cons aux)
qed
-lemma mem_iff [code_post]:
- "x mem xs \<longleftrightarrow> x \<in> set xs"
-by (induct xs) auto
-
lemmas in_set_code [code_unfold] = mem_iff [symmetric]
+lemma member_simps [simp, code]:
+ "member [] y \<longleftrightarrow> False"
+ "member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
+ by (auto simp add: mem_iff)
+
+lemma member_set:
+ "member = set"
+ by (simp add: expand_fun_eq mem_iff mem_def)
+
+abbreviation (input) mem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infixl "mem" 55) where
+ "x mem xs \<equiv> member xs x" -- "for backward compatibility"
+
lemma empty_null:
"xs = [] \<longleftrightarrow> null xs"
by (cases xs) simp_all
--- a/src/HOL/MacLaurin.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/MacLaurin.thy Sun May 23 10:38:11 2010 +0100
@@ -383,6 +383,11 @@
text{*It is unclear why so many variant results are needed.*}
+lemma sin_expansion_lemma:
+ "sin (x + real (Suc m) * pi / 2) =
+ cos (x + real (m) * pi / 2)"
+by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto)
+
lemma Maclaurin_sin_expansion2:
"\<exists>t. abs t \<le> abs x &
sin x =
@@ -394,7 +399,7 @@
and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl)
apply safe
apply (simp (no_asm))
-apply (simp (no_asm))
+apply (simp (no_asm) add: sin_expansion_lemma)
apply (case_tac "n", clarify, simp, simp add: lemma_STAR_sin)
apply (rule ccontr, simp)
apply (drule_tac x = x in spec, simp)
@@ -414,7 +419,6 @@
apply (blast intro: elim:);
done
-
lemma Maclaurin_sin_expansion3:
"[| n > 0; 0 < x |] ==>
\<exists>t. 0 < t & t < x &
@@ -426,7 +430,7 @@
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl)
apply safe
apply simp
-apply (simp (no_asm))
+apply (simp (no_asm) add: sin_expansion_lemma)
apply (erule ssubst)
apply (rule_tac x = t in exI, simp)
apply (rule setsum_cong[OF refl])
@@ -444,7 +448,7 @@
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl)
apply safe
apply simp
-apply (simp (no_asm))
+apply (simp (no_asm) add: sin_expansion_lemma)
apply (erule ssubst)
apply (rule_tac x = t in exI, simp)
apply (rule setsum_cong[OF refl])
@@ -459,6 +463,10 @@
(if even m then -1 ^ (m div 2)/(real (fact m)) else 0) * 0 ^ m) = 1"
by (induct "n", auto)
+lemma cos_expansion_lemma:
+ "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
+by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto)
+
lemma Maclaurin_cos_expansion:
"\<exists>t. abs t \<le> abs x &
cos x =
@@ -470,7 +478,7 @@
apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
apply safe
apply (simp (no_asm))
-apply (simp (no_asm))
+apply (simp (no_asm) add: cos_expansion_lemma)
apply (case_tac "n", simp)
apply (simp del: setsum_op_ivl_Suc)
apply (rule ccontr, simp)
@@ -493,7 +501,7 @@
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl)
apply safe
apply simp
-apply (simp (no_asm))
+apply (simp (no_asm) add: cos_expansion_lemma)
apply (erule ssubst)
apply (rule_tac x = t in exI, simp)
apply (rule setsum_cong[OF refl])
@@ -512,7 +520,7 @@
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl)
apply safe
apply simp
-apply (simp (no_asm))
+apply (simp (no_asm) add: cos_expansion_lemma)
apply (erule ssubst)
apply (rule_tac x = t in exI, simp)
apply (rule setsum_cong[OF refl])
--- a/src/HOL/Matrix/cplex/matrixlp.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Matrix/cplex/matrixlp.ML Sun May 23 10:38:11 2010 +0100
@@ -81,8 +81,8 @@
fun matrix_simplify th =
let
val simp_th = matrix_compute (cprop_of th)
- val th = Thm.strip_shyps (equal_elim simp_th th)
- fun removeTrue th = removeTrue (implies_elim th TrueI) handle _ => th (* FIXME avoid handle _ *)
+ val th = Thm.strip_shyps (Thm.equal_elim simp_th th)
+ fun removeTrue th = removeTrue (Thm.implies_elim th TrueI) handle _ => th (* FIXME avoid handle _ *)
in
removeTrue th
end
--- a/src/HOL/Metis_Examples/BigO.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Metis_Examples/BigO.thy Sun May 23 10:38:11 2010 +0100
@@ -28,7 +28,7 @@
(*** Now various verions with an increasing shrink factor ***)
-sledgehammer_params [shrink_factor = 1]
+sledgehammer_params [isar_proof, isar_shrink_factor = 1]
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom).
ALL x. (abs (h x)) <= (c * (abs (f x))))
@@ -58,7 +58,7 @@
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F4)
qed
-sledgehammer_params [shrink_factor = 2]
+sledgehammer_params [isar_proof, isar_shrink_factor = 2]
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom).
ALL x. (abs (h x)) <= (c * (abs (f x))))
@@ -80,7 +80,7 @@
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F2)
qed
-sledgehammer_params [shrink_factor = 3]
+sledgehammer_params [isar_proof, isar_shrink_factor = 3]
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom).
ALL x. (abs (h x)) <= (c * (abs (f x))))
@@ -99,7 +99,7 @@
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis A1 abs_mult abs_ge_zero)
qed
-sledgehammer_params [shrink_factor = 4]
+sledgehammer_params [isar_proof, isar_shrink_factor = 4]
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom).
ALL x. (abs (h x)) <= (c * (abs (f x))))
@@ -118,7 +118,7 @@
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis abs_mult)
qed
-sledgehammer_params [shrink_factor = 1]
+sledgehammer_params [isar_proof, isar_shrink_factor = 1]
lemma bigo_alt_def: "O(f) =
{h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
--- a/src/HOL/Metis_Examples/TransClosure.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Metis_Examples/TransClosure.thy Sun May 23 10:38:11 2010 +0100
@@ -39,7 +39,7 @@
lemma "\<lbrakk>f c = Intg x; \<forall>y. f b = Intg y \<longrightarrow> y \<noteq> x; (a, b) \<in> R\<^sup>*; (b,c) \<in> R\<^sup>*\<rbrakk>
\<Longrightarrow> \<exists>c. (b, c) \<in> R \<and> (a, c) \<in> R\<^sup>*"
-(* sledgehammer [isar_proof, shrink_factor = 2] *)
+(* sledgehammer [isar_proof, isar_shrink_factor = 2] *)
proof -
assume A1: "f c = Intg x"
assume A2: "\<forall>y. f b = Intg y \<longrightarrow> y \<noteq> x"
--- a/src/HOL/Metis_Examples/set.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Metis_Examples/set.thy Sun May 23 10:38:11 2010 +0100
@@ -16,7 +16,7 @@
lemma "P(n::nat) ==> ~P(0) ==> n ~= 0"
by metis
-sledgehammer_params [shrink_factor = 1]
+sledgehammer_params [isar_proof, isar_shrink_factor = 1]
(*multiple versions of this example*)
lemma (*equal_union: *)
@@ -85,7 +85,7 @@
by (metis 31 29)
qed
-sledgehammer_params [shrink_factor = 2]
+sledgehammer_params [isar_proof, isar_shrink_factor = 2]
lemma (*equal_union: *)
"(X = Y \<union> Z) =
@@ -128,7 +128,7 @@
by (metis 18 17)
qed
-sledgehammer_params [shrink_factor = 3]
+sledgehammer_params [isar_proof, isar_shrink_factor = 3]
lemma (*equal_union: *)
"(X = Y \<union> Z) =
@@ -163,7 +163,7 @@
(*Example included in TPHOLs paper*)
-sledgehammer_params [shrink_factor = 4]
+sledgehammer_params [isar_proof, isar_shrink_factor = 4]
lemma (*equal_union: *)
"(X = Y \<union> Z) =
--- a/src/HOL/MicroJava/BV/BVExample.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy Sun May 23 10:38:11 2010 +0100
@@ -449,7 +449,7 @@
(\<lambda>(ss, w).
let p = some_elem w in propa f (step p (ss ! p)) ss (w - {p}))
(ss, w)"
- unfolding iter_def List_Set.is_empty_def some_elem_def ..
+ unfolding iter_def More_Set.is_empty_def some_elem_def ..
lemma JVM_sup_unfold [code]:
"JVMType.sup S m n = lift2 (Opt.sup
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Sun May 23 10:38:11 2010 +0100
@@ -330,14 +330,14 @@
fun thms_of_name ctxt name =
let
- val lex = OuterKeyword.get_lexicons
+ val lex = Keyword.get_lexicons
val get = maps (ProofContext.get_fact ctxt o fst)
in
Source.of_string name
|> Symbol.source {do_recover=false}
- |> OuterLex.source {do_recover=SOME false} lex Position.start
- |> OuterLex.source_proper
- |> Source.source OuterLex.stopper (SpecParse.xthms1 >> get) NONE
+ |> Token.source {do_recover=SOME false} lex Position.start
+ |> Token.source_proper
+ |> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE
|> Source.exhaust
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/normarith.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,402 @@
+(* Title: Library/normarith.ML
+ Author: Amine Chaieb, University of Cambridge
+
+Simple decision procedure for linear problems in Euclidean space.
+*)
+
+signature NORM_ARITH =
+sig
+ val norm_arith : Proof.context -> conv
+ val norm_arith_tac : Proof.context -> int -> tactic
+end
+
+structure NormArith : NORM_ARITH =
+struct
+
+ open Conv;
+ val bool_eq = op = : bool *bool -> bool
+ fun dest_ratconst t = case term_of t of
+ Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
+ | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
+ | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
+ fun is_ratconst t = can dest_ratconst t
+ fun augment_norm b t acc = case term_of t of
+ Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
+ | _ => acc
+ fun find_normedterms t acc = case term_of t of
+ @{term "op + :: real => _"}$_$_ =>
+ find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
+ | @{term "op * :: real => _"}$_$n =>
+ if not (is_ratconst (Thm.dest_arg1 t)) then acc else
+ augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero)
+ (Thm.dest_arg t) acc
+ | _ => augment_norm true t acc
+
+ val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
+ fun cterm_lincomb_cmul c t =
+ if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
+ fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
+ fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
+ fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
+
+ val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
+ fun int_lincomb_cmul c t =
+ if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
+ fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
+ fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
+ fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
+
+fun vector_lincomb t = case term_of t of
+ Const(@{const_name plus}, _) $ _ $ _ =>
+ cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name minus}, _) $ _ $ _ =>
+ cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name scaleR}, _)$_$_ =>
+ cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
+ | Const(@{const_name uminus}, _)$_ =>
+ cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
+(* FIXME: how should we handle numerals?
+ | Const(@ {const_name vec},_)$_ =>
+ let
+ val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0
+ handle TERM _=> false)
+ in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one)
+ else FuncUtil.Ctermfunc.empty
+ end
+*)
+ | _ => FuncUtil.Ctermfunc.onefunc (t,Rat.one)
+
+ fun vector_lincombs ts =
+ fold_rev
+ (fn t => fn fns => case AList.lookup (op aconvc) fns t of
+ NONE =>
+ let val f = vector_lincomb t
+ in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of
+ SOME (_,f') => (t,f') :: fns
+ | NONE => (t,f) :: fns
+ end
+ | SOME _ => fns) ts []
+
+fun replacenegnorms cv t = case term_of t of
+ @{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
+| @{term "op * :: real => _"}$_$_ =>
+ if dest_ratconst (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else Thm.reflexive t
+| _ => Thm.reflexive t
+fun flip v eq =
+ if FuncUtil.Ctermfunc.defined eq v
+ then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.Ctermfunc.apply eq v)) eq else eq
+fun allsubsets s = case s of
+ [] => [[]]
+|(a::t) => let val res = allsubsets t in
+ map (cons a) res @ res end
+fun evaluate env lin =
+ FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x))
+ lin Rat.zero
+
+fun solve (vs,eqs) = case (vs,eqs) of
+ ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one))
+ |(_,eq::oeqs) =>
+ (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
+ [] => NONE
+ | v::_ =>
+ if FuncUtil.Intfunc.defined eq v
+ then
+ let
+ val c = FuncUtil.Intfunc.apply eq v
+ val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
+ fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn
+ else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
+ in (case solve (remove (op =) v vs, map eliminate oeqs) of
+ NONE => NONE
+ | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln))
+ end
+ else NONE)
+
+fun combinations k l = if k = 0 then [[]] else
+ case l of
+ [] => []
+| h::t => map (cons h) (combinations (k - 1) t) @ combinations k t
+
+
+fun forall2 p l1 l2 = case (l1,l2) of
+ ([],[]) => true
+ | (h1::t1,h2::t2) => p h1 h2 andalso forall2 p t1 t2
+ | _ => false;
+
+
+fun vertices vs eqs =
+ let
+ fun vertex cmb = case solve(vs,cmb) of
+ NONE => NONE
+ | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v Rat.zero) vs)
+ val rawvs = map_filter vertex (combinations (length vs) eqs)
+ val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs
+ in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset []
+ end
+
+fun subsumes l m = forall2 (fn x => fn y => Rat.abs x <=/ Rat.abs y) l m
+
+fun subsume todo dun = case todo of
+ [] => dun
+|v::ovs =>
+ let val dun' = if exists (fn w => subsumes w v) dun then dun
+ else v::(filter (fn w => not(subsumes v w)) dun)
+ in subsume ovs dun'
+ end;
+
+fun match_mp PQ P = P RS PQ;
+
+fun cterm_of_rat x =
+let val (a, b) = Rat.quotient_of_rat x
+in
+ if b = 1 then Numeral.mk_cnumber @{ctyp "real"} a
+ else Thm.capply (Thm.capply @{cterm "op / :: real => _"}
+ (Numeral.mk_cnumber @{ctyp "real"} a))
+ (Numeral.mk_cnumber @{ctyp "real"} b)
+end;
+
+fun norm_cmul_rule c th = instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm});
+
+fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm};
+
+ (* I think here the static context should be sufficient!! *)
+fun inequality_canon_rule ctxt =
+ let
+ (* FIXME : Should be computed statically!! *)
+ val real_poly_conv =
+ Semiring_Normalizer.semiring_normalize_wrapper ctxt
+ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
+ in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv then_conv real_poly_conv)))
+end;
+
+ val apply_pth1 = rewr_conv @{thm pth_1};
+ val apply_pth2 = rewr_conv @{thm pth_2};
+ val apply_pth3 = rewr_conv @{thm pth_3};
+ val apply_pth4 = rewrs_conv @{thms pth_4};
+ val apply_pth5 = rewr_conv @{thm pth_5};
+ val apply_pth6 = rewr_conv @{thm pth_6};
+ val apply_pth7 = rewrs_conv @{thms pth_7};
+ val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left})));
+ val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.field_comp_conv);
+ val apply_ptha = rewr_conv @{thm pth_a};
+ val apply_pthb = rewrs_conv @{thms pth_b};
+ val apply_pthc = rewrs_conv @{thms pth_c};
+ val apply_pthd = try_conv (rewr_conv @{thm pth_d});
+
+fun headvector t = case t of
+ Const(@{const_name plus}, _)$
+ (Const(@{const_name scaleR}, _)$l$v)$r => v
+ | Const(@{const_name scaleR}, _)$l$v => v
+ | _ => error "headvector: non-canonical term"
+
+fun vector_cmul_conv ct =
+ ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.field_comp_conv) else_conv
+ (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct
+
+fun vector_add_conv ct = apply_pth7 ct
+ handle CTERM _ =>
+ (apply_pth8 ct
+ handle CTERM _ =>
+ (case term_of ct of
+ Const(@{const_name plus},_)$lt$rt =>
+ let
+ val l = headvector lt
+ val r = headvector rt
+ in (case Term_Ord.fast_term_ord (l,r) of
+ LESS => (apply_pthb then_conv arg_conv vector_add_conv
+ then_conv apply_pthd) ct
+ | GREATER => (apply_pthc then_conv arg_conv vector_add_conv
+ then_conv apply_pthd) ct
+ | EQUAL => (apply_pth9 then_conv
+ ((apply_ptha then_conv vector_add_conv) else_conv
+ arg_conv vector_add_conv then_conv apply_pthd)) ct)
+ end
+ | _ => Thm.reflexive ct))
+
+fun vector_canon_conv ct = case term_of ct of
+ Const(@{const_name plus},_)$_$_ =>
+ let
+ val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
+ val lth = vector_canon_conv l
+ val rth = vector_canon_conv r
+ val th = Drule.binop_cong_rule p lth rth
+ in fconv_rule (arg_conv vector_add_conv) th end
+
+| Const(@{const_name scaleR}, _)$_$_ =>
+ let
+ val (p,r) = Thm.dest_comb ct
+ val rth = Drule.arg_cong_rule p (vector_canon_conv r)
+ in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth
+ end
+
+| Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct
+
+| Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct
+
+(* FIXME
+| Const(@{const_name vec},_)$n =>
+ let val n = Thm.dest_arg ct
+ in if is_ratconst n andalso not (dest_ratconst n =/ Rat.zero)
+ then Thm.reflexive ct else apply_pth1 ct
+ end
+*)
+| _ => apply_pth1 ct
+
+fun norm_canon_conv ct = case term_of ct of
+ Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct
+ | _ => raise CTERM ("norm_canon_conv", [ct])
+
+fun fold_rev2 f [] [] z = z
+ | fold_rev2 f (x::xs) (y::ys) z = f x y (fold_rev2 f xs ys z)
+ | fold_rev2 f _ _ _ = raise UnequalLengths;
+
+fun int_flip v eq =
+ if FuncUtil.Intfunc.defined eq v
+ then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq;
+
+local
+ val pth_zero = @{thm norm_zero}
+ val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
+ pth_zero
+ val concl = Thm.dest_arg o cprop_of
+ fun real_vector_combo_prover ctxt translator (nubs,ges,gts) =
+ let
+ (* FIXME: Should be computed statically!!*)
+ val real_poly_conv =
+ Semiring_Normalizer.semiring_normalize_wrapper ctxt
+ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
+ val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
+ val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) []
+ val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check"
+ else ()
+ val dests = distinct (op aconvc) (map snd rawdests)
+ val srcfuns = map vector_lincomb sources
+ val destfuns = map vector_lincomb dests
+ val vvs = fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) []
+ val n = length srcfuns
+ val nvs = 1 upto n
+ val srccombs = srcfuns ~~ nvs
+ fun consider d =
+ let
+ fun coefficients x =
+ let
+ val inp = if FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x))
+ else FuncUtil.Intfunc.empty
+ in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp
+ end
+ val equations = map coefficients vvs
+ val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,Rat.one)) nvs
+ fun plausiblevertices f =
+ let
+ val flippedequations = map (fold_rev int_flip f) equations
+ val constraints = flippedequations @ inequalities
+ val rawverts = vertices nvs constraints
+ fun check_solution v =
+ let
+ val f = fold_rev2 (curry FuncUtil.Intfunc.update) nvs v (FuncUtil.Intfunc.onefunc (0, Rat.one))
+ in forall (fn e => evaluate f e =/ Rat.zero) flippedequations
+ end
+ val goodverts = filter check_solution rawverts
+ val signfixups = map (fn n => if member (op =) f n then ~1 else 1) nvs
+ in map (map2 (fn s => fn c => Rat.rat_of_int s */ c) signfixups) goodverts
+ end
+ val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) []
+ in subsume allverts []
+ end
+ fun compute_ineq v =
+ let
+ val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE
+ else SOME(norm_cmul_rule v t))
+ (v ~~ nubs)
+ fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
+ in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
+ end
+ val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
+ map (inequality_canon_rule ctxt) nubs @ ges
+ val zerodests = filter
+ (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
+
+ in fst (RealArith.real_linear_prover translator
+ (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
+ zerodests,
+ map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
+ arg_conv (arg_conv real_poly_conv))) ges',
+ map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
+ arg_conv (arg_conv real_poly_conv))) gts))
+ end
+in val real_vector_combo_prover = real_vector_combo_prover
+end;
+
+local
+ val pth = @{thm norm_imp_pos_and_ge}
+ val norm_mp = match_mp pth
+ val concl = Thm.dest_arg o cprop_of
+ fun conjunct1 th = th RS @{thm conjunct1}
+ fun conjunct2 th = th RS @{thm conjunct2}
+fun real_vector_ineq_prover ctxt translator (ges,gts) =
+ let
+(* val _ = error "real_vector_ineq_prover: pause" *)
+ val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) []
+ val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
+ val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
+ fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
+ fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
+ fun mk_equals l r = Thm.capply (Thm.capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
+ val asl = map2 (fn (t,_) => fn n => Thm.assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
+ val replace_conv = try_conv (rewrs_conv asl)
+ val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
+ val ges' =
+ fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths)
+ asl (map replace_rule ges)
+ val gts' = map replace_rule gts
+ val nubs = map (conjunct2 o norm_mp) asl
+ val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
+ val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1))
+ val th11 = hd (Variable.export ctxt' ctxt [fold Thm.implies_intr shs th1])
+ val cps = map (swap o Thm.dest_equals) (cprems_of th11)
+ val th12 = instantiate ([], cps) th11
+ val th13 = fold Thm.elim_implies (map (Thm.reflexive o snd) cps) th12;
+ in hd (Variable.export ctxt' ctxt [th13])
+ end
+in val real_vector_ineq_prover = real_vector_ineq_prover
+end;
+
+local
+ val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
+ fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
+ fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
+ (* FIXME: Lookup in the context every time!!! Fix this !!!*)
+ fun splitequation ctxt th acc =
+ let
+ val real_poly_neg_conv = #neg
+ (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
+ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
+ val (th1,th2) = conj_pair(rawrule th)
+ in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
+ end
+in fun real_vector_prover ctxt _ translator (eqs,ges,gts) =
+ (real_vector_ineq_prover ctxt translator
+ (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial)
+end;
+
+ fun init_conv ctxt =
+ Simplifier.rewrite (Simplifier.context ctxt
+ (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})))
+ then_conv Numeral_Simprocs.field_comp_conv
+ then_conv nnf_conv
+
+ fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
+ fun norm_arith ctxt ct =
+ let
+ val ctxt' = Variable.declare_term (term_of ct) ctxt
+ val th = init_conv ctxt' ct
+ in Thm.equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (Thm.symmetric th))
+ (pure ctxt' (Thm.rhs_of th))
+ end
+
+ fun norm_arith_tac ctxt =
+ clarify_tac HOL_cs THEN'
+ Object_Logic.full_atomize_tac THEN'
+ CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
+
+end;
--- a/src/HOL/NSA/Examples/NSPrimes.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/NSA/Examples/NSPrimes.thy Sun May 23 10:38:11 2010 +0100
@@ -13,9 +13,7 @@
text{*These can be used to derive an alternative proof of the infinitude of
primes by considering a property of nonstandard sets.*}
-definition
- hdvd :: "[hypnat, hypnat] => bool" (infixl "hdvd" 50) where
- [transfer_unfold]: "(M::hypnat) hdvd N = ( *p2* (op dvd)) M N"
+declare dvd_def [transfer_refold]
definition
starprime :: "hypnat set" where
@@ -49,14 +47,14 @@
(* Goldblatt: Exercise 5.11(2) - p. 57 *)
-lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::hypnat) <= M --> m hdvd N)"
+lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::hypnat) <= M --> m dvd N)"
by (transfer, rule dvd_by_all)
lemmas hdvd_by_all2 = hdvd_by_all [THEN spec, standard]
(* Goldblatt: Exercise 5.11(2) - p. 57 *)
lemma hypnat_dvd_all_hypnat_of_nat:
- "\<exists>(N::hypnat). 0 < N & (\<forall>n \<in> -{0::nat}. hypnat_of_nat(n) hdvd N)"
+ "\<exists>(N::hypnat). 0 < N & (\<forall>n \<in> -{0::nat}. hypnat_of_nat(n) dvd N)"
apply (cut_tac hdvd_by_all)
apply (drule_tac x = whn in spec, auto)
apply (rule exI, auto)
@@ -70,7 +68,7 @@
(* Goldblatt: Exercise 5.11(3a) - p 57 *)
lemma starprime:
- "starprime = {p. 1 < p & (\<forall>m. m hdvd p --> m = 1 | m = p)}"
+ "starprime = {p. 1 < p & (\<forall>m. m dvd p --> m = 1 | m = p)}"
by (transfer, auto simp add: prime_def)
lemma prime_two: "prime 2"
@@ -88,13 +86,11 @@
apply (rule_tac x = n in exI, auto)
apply (drule conjI [THEN not_prime_ex_mk], auto)
apply (drule_tac x = m in spec, auto)
-apply (rule_tac x = ka in exI)
-apply (auto intro: dvd_mult2)
done
(* Goldblatt Exercise 5.11(3b) - p 57 *)
lemma hyperprime_factor_exists [rule_format]:
- "!!n. 1 < n ==> (\<exists>k \<in> starprime. k hdvd n)"
+ "!!n. 1 < n ==> (\<exists>k \<in> starprime. k dvd n)"
by (transfer, simp add: prime_factor_exists)
(* Goldblatt Exercise 3.10(1) - p. 29 *)
@@ -257,14 +253,14 @@
subsection{*Existence of Infinitely Many Primes: a Nonstandard Proof*}
-lemma lemma_not_dvd_hypnat_one: "~ (\<forall>n \<in> - {0}. hypnat_of_nat n hdvd 1)"
+lemma lemma_not_dvd_hypnat_one: "~ (\<forall>n \<in> - {0}. hypnat_of_nat n dvd 1)"
apply auto
apply (rule_tac x = 2 in bexI)
apply (transfer, auto)
done
declare lemma_not_dvd_hypnat_one [simp]
-lemma lemma_not_dvd_hypnat_one2: "\<exists>n \<in> - {0}. ~ hypnat_of_nat n hdvd 1"
+lemma lemma_not_dvd_hypnat_one2: "\<exists>n \<in> - {0}. ~ hypnat_of_nat n dvd 1"
apply (cut_tac lemma_not_dvd_hypnat_one)
apply (auto simp del: lemma_not_dvd_hypnat_one)
done
@@ -314,13 +310,13 @@
by (cut_tac hypnat_of_nat_one_not_prime, simp)
declare hypnat_one_not_prime [simp]
-lemma hdvd_diff: "!!k m n. [| k hdvd m; k hdvd n |] ==> k hdvd (m - n)"
-by (transfer, rule dvd_diff)
+lemma hdvd_diff: "!!k m n :: hypnat. [| k dvd m; k dvd n |] ==> k dvd (m - n)"
+by (transfer, rule dvd_diff_nat)
lemma dvd_one_eq_one: "x dvd (1::nat) ==> x = 1"
by (unfold dvd_def, auto)
-lemma hdvd_one_eq_one: "!!x. x hdvd 1 ==> x = 1"
+lemma hdvd_one_eq_one: "!!x. x dvd (1::hypnat) ==> x = 1"
by (transfer, rule dvd_one_eq_one)
theorem not_finite_prime: "~ finite {p. prime p}"
--- a/src/HOL/NSA/NSCA.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/NSA/NSCA.thy Sun May 23 10:38:11 2010 +0100
@@ -279,13 +279,9 @@
"\<lbrakk>x \<in> Infinitesimal; y \<in> Infinitesimal\<rbrakk> \<Longrightarrow> HComplex x y \<in> Infinitesimal"
apply (rule Infinitesimal_hcmod_iff [THEN iffD2])
apply (simp add: hcmod_i)
-apply (rule Infinitesimal_sqrt)
apply (rule Infinitesimal_add)
apply (erule Infinitesimal_hrealpow, simp)
apply (erule Infinitesimal_hrealpow, simp)
-apply (rule add_nonneg_nonneg)
-apply (rule zero_le_power2)
-apply (rule zero_le_power2)
done
lemma hcomplex_Infinitesimal_iff:
--- a/src/HOL/Nat.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nat.thy Sun May 23 10:38:11 2010 +0100
@@ -1294,15 +1294,11 @@
begin
lemma zero_le_imp_of_nat: "0 \<le> of_nat m"
- apply (induct m, simp_all)
- apply (erule order_trans)
- apply (rule ord_le_eq_trans [OF _ add_commute])
- apply (rule less_add_one [THEN less_imp_le])
- done
+ by (induct m) simp_all
lemma less_imp_of_nat_less: "m < n \<Longrightarrow> of_nat m < of_nat n"
apply (induct m n rule: diff_induct, simp_all)
- apply (insert add_less_le_mono [OF zero_less_one zero_le_imp_of_nat], force)
+ apply (rule add_pos_nonneg [OF zero_less_one zero_le_imp_of_nat])
done
lemma of_nat_less_imp_less: "of_nat m < of_nat n \<Longrightarrow> m < n"
--- a/src/HOL/Nat_Numeral.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nat_Numeral.thy Sun May 23 10:38:11 2010 +0100
@@ -120,9 +120,9 @@
"a\<twosuperior> = 0 \<longleftrightarrow> a = 0"
unfolding power2_eq_square by simp
-lemma power2_eq_1_iff [simp]:
+lemma power2_eq_1_iff:
"a\<twosuperior> = 1 \<longleftrightarrow> a = 1 \<or> a = - 1"
- unfolding power2_eq_square by simp
+ unfolding power2_eq_square by (rule square_eq_1_iff)
end
--- a/src/HOL/Nitpick.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nitpick.thy Sun May 23 10:38:11 2010 +0100
@@ -52,11 +52,11 @@
Alternative definitions.
*}
-lemma If_def [nitpick_def]:
+lemma If_def [nitpick_def, no_atp]:
"(if P then Q else R) \<equiv> (P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> R)"
by (rule eq_reflection) (rule if_bool_eq_conj)
-lemma Ex1_def [nitpick_def]:
+lemma Ex1_def [nitpick_def, no_atp]:
"Ex1 P \<equiv> \<exists>x. P = {x}"
apply (rule eq_reflection)
apply (simp add: Ex1_def expand_set_eq)
@@ -69,14 +69,14 @@
apply (erule_tac x = y in allE)
by (auto simp: mem_def)
-lemma rtrancl_def [nitpick_def]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
+lemma rtrancl_def [nitpick_def, no_atp]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
by simp
-lemma rtranclp_def [nitpick_def]:
+lemma rtranclp_def [nitpick_def, no_atp]:
"rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
by (rule eq_reflection) (auto dest: rtranclpD)
-lemma tranclp_def [nitpick_def]:
+lemma tranclp_def [nitpick_def, no_atp]:
"tranclp r a b \<equiv> trancl (split r) (a, b)"
by (simp add: trancl_def Collect_def mem_def)
@@ -110,18 +110,18 @@
\textit{special\_level} optimization.
*}
-lemma The_psimp [nitpick_psimp]:
+lemma The_psimp [nitpick_psimp, no_atp]:
"P = {x} \<Longrightarrow> The P = x"
by (subgoal_tac "{x} = (\<lambda>y. y = x)") (auto simp: mem_def)
-lemma Eps_psimp [nitpick_psimp]:
+lemma Eps_psimp [nitpick_psimp, no_atp]:
"\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
apply (case_tac "P (Eps P)")
apply auto
apply (erule contrapos_np)
by (rule someI)
-lemma unit_case_def [nitpick_def]:
+lemma unit_case_def [nitpick_def, no_atp]:
"unit_case x u \<equiv> x"
apply (subgoal_tac "u = ()")
apply (simp only: unit.cases)
@@ -129,14 +129,14 @@
declare unit.cases [nitpick_simp del]
-lemma nat_case_def [nitpick_def]:
+lemma nat_case_def [nitpick_def, no_atp]:
"nat_case x f n \<equiv> if n = 0 then x else f (n - 1)"
apply (rule eq_reflection)
by (case_tac n) auto
declare nat.cases [nitpick_simp del]
-lemma list_size_simp [nitpick_simp]:
+lemma list_size_simp [nitpick_simp, no_atp]:
"list_size f xs = (if xs = [] then 0
else Suc (f (hd xs) + list_size f (tl xs)))"
"size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
--- a/src/HOL/Nominal/nominal_atoms.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML Sun May 23 10:38:11 2010 +0100
@@ -1004,10 +1004,10 @@
(* syntax und parsing *)
-structure P = OuterParse and K = OuterKeyword;
+structure P = Parse and K = Keyword;
val _ =
- OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl
- (Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls));
+ Outer_Syntax.command "atom_decl" "declare new kinds of atoms" Keyword.thy_decl
+ (Scan.repeat1 Parse.name >> (Toplevel.theory o create_nom_typedecls));
end;
--- a/src/HOL/Nominal/nominal_datatype.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML Sun May 23 10:38:11 2010 +0100
@@ -602,7 +602,7 @@
(fn _ => EVERY
[indtac rep_induct [] 1,
ALLGOALS (simp_tac (global_simpset_of thy4 addsimps
- (symmetric perm_fun_def :: abs_perm))),
+ (Thm.symmetric perm_fun_def :: abs_perm))),
ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
length new_type_names));
@@ -927,7 +927,7 @@
simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
constr_defs @ perm_closed_thms)) 1,
TRY (simp_tac (HOL_basic_ss addsimps
- (symmetric perm_fun_def :: abs_perm)) 1),
+ (Thm.symmetric perm_fun_def :: abs_perm)) 1),
TRY (simp_tac (HOL_basic_ss addsimps
(perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
perm_closed_thms)) 1)])
@@ -1077,7 +1077,7 @@
(indtac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
EVERY (map (fn (prem, r) => (EVERY
[REPEAT (eresolve_tac Abs_inverse_thms' 1),
- simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
+ simp_tac (HOL_basic_ss addsimps [Thm.symmetric r]) 1,
DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
(prems ~~ constr_defs))]);
@@ -1641,7 +1641,7 @@
fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
supports_fresh) 1,
simp_tac (HOL_basic_ss addsimps
- [supports_def, symmetric fresh_def, fresh_prod]) 1,
+ [supports_def, Thm.symmetric fresh_def, fresh_prod]) 1,
REPEAT_DETERM (resolve_tac [allI, impI] 1),
REPEAT_DETERM (etac conjE 1),
rtac unique 1,
@@ -1655,7 +1655,7 @@
rtac rec_prem 1,
simp_tac (HOL_ss addsimps (fs_name ::
supp_prod :: finite_Un :: finite_prems)) 1,
- simp_tac (HOL_ss addsimps (symmetric fresh_def ::
+ simp_tac (HOL_ss addsimps (Thm.symmetric fresh_def ::
fresh_prod :: fresh_prems)) 1]
end))
end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
@@ -1746,7 +1746,7 @@
asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
(finite_thss ~~ finite_ctxt_ths) @
maps (fn ((_, idxss), elim) => maps (fn idxs =>
- [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
+ [full_simp_tac (HOL_ss addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1,
REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
rtac ex1I 1,
(resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
@@ -2076,11 +2076,10 @@
(* FIXME: The following stuff should be exported by Datatype *)
-local structure P = OuterParse and K = OuterKeyword in
-
val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_mixfix --
- (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
+ Scan.option (Parse.$$$ "(" |-- Parse.name --| Parse.$$$ ")") --
+ Parse.type_args -- Parse.name -- Parse.opt_mixfix --
+ (Parse.$$$ "=" |-- Parse.enum1 "|" (Parse.name -- Scan.repeat Parse.typ -- Parse.opt_mixfix));
fun mk_datatype args =
let
@@ -2090,9 +2089,7 @@
in add_nominal_datatype Datatype.default_config names specs end;
val _ =
- OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-end;
+ Outer_Syntax.command "nominal_datatype" "define inductive datatypes" Keyword.thy_decl
+ (Parse.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
end
--- a/src/HOL/Nominal/nominal_induct.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_induct.ML Sun May 23 10:38:11 2010 +0100
@@ -138,8 +138,6 @@
local
-structure P = OuterParse;
-
val avoidingN = "avoiding";
val fixingN = "arbitrary"; (* to be consistent with induct; hopefully this changes again *)
val ruleN = "rule";
@@ -165,14 +163,14 @@
Scan.repeat (unless_more_args free)) [];
val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |--
- P.and_list' (Scan.repeat (unless_more_args free))) [];
+ Parse.and_list' (Scan.repeat (unless_more_args free))) [];
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms;
in
val nominal_induct_method =
- Args.mode Induct.no_simpN -- (P.and_list' (Scan.repeat (unless_more_args def_inst)) --
+ Args.mode Induct.no_simpN -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
avoiding -- fixing -- rule_spec) >>
(fn (no_simp, (((x, y), z), w)) => fn ctxt =>
RAW_METHOD_CASES (fn facts =>
--- a/src/HOL/Nominal/nominal_inductive.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML Sun May 23 10:38:11 2010 +0100
@@ -672,23 +672,20 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = OuterKeyword.keyword "avoids";
+val _ = Keyword.keyword "avoids";
val _ =
- OuterSyntax.local_theory_to_proof "nominal_inductive"
- "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
- (P.xname -- Scan.optional (P.$$$ "avoids" |-- P.and_list1 (P.name --
- (P.$$$ ":" |-- Scan.repeat1 P.name))) [] >> (fn (name, avoids) =>
+ Outer_Syntax.local_theory_to_proof "nominal_inductive"
+ "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes"
+ Keyword.thy_goal
+ (Parse.xname -- Scan.optional (Parse.$$$ "avoids" |-- Parse.and_list1 (Parse.name --
+ (Parse.$$$ ":" |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) =>
prove_strong_ind name avoids));
val _ =
- OuterSyntax.local_theory "equivariance"
- "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
- (P.xname -- Scan.optional (P.$$$ "[" |-- P.list1 P.name --| P.$$$ "]") [] >>
+ Outer_Syntax.local_theory "equivariance"
+ "prove equivariance for inductive predicate involving nominal datatypes" Keyword.thy_decl
+ (Parse.xname -- Scan.optional (Parse.$$$ "[" |-- Parse.list1 Parse.name --| Parse.$$$ "]") [] >>
(fn (name, atoms) => prove_eqvt name atoms));
-end;
-
end
--- a/src/HOL/Nominal/nominal_inductive2.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML Sun May 23 10:38:11 2010 +0100
@@ -485,17 +485,14 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.local_theory_to_proof "nominal_inductive2"
- "prove strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
- (P.xname --
- Scan.option (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) --
- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name --
- (P.$$$ ":" |-- P.and_list1 P.term))) []) >> (fn ((name, rule_name), avoids) =>
+ Outer_Syntax.local_theory_to_proof "nominal_inductive2"
+ "prove strong induction theorem for inductive predicate involving nominal datatypes"
+ Keyword.thy_goal
+ (Parse.xname --
+ Scan.option (Parse.$$$ "(" |-- Parse.!!! (Parse.name --| Parse.$$$ ")")) --
+ (Scan.optional (Parse.$$$ "avoids" |-- Parse.enum1 "|" (Parse.name --
+ (Parse.$$$ ":" |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) =>
prove_strong_ind name rule_name avoids));
-end;
-
end
--- a/src/HOL/Nominal/nominal_permeq.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_permeq.ML Sun May 23 10:38:11 2010 +0100
@@ -266,7 +266,7 @@
(* intros, then applies eqvt_simp_tac *)
fun supports_tac_i tactical ss i =
let
- val simps = [supports_def,symmetric fresh_def,fresh_prod]
+ val simps = [supports_def, Thm.symmetric fresh_def, fresh_prod]
in
EVERY [tactical ("unfolding of supports ", simp_tac (HOL_basic_ss addsimps simps) i),
tactical ("stripping of foralls ", REPEAT_DETERM (rtac allI i)),
@@ -329,7 +329,7 @@
val goal = List.nth(cprems_of st, i-1)
val fin_supp = dynamic_thms st ("fin_supp")
val fresh_atm = dynamic_thms st ("fresh_atm")
- val ss1 = ss addsimps [symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
+ val ss1 = ss addsimps [Thm.symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
val ss2 = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
in
case Logic.strip_assums_concl (term_of goal) of
--- a/src/HOL/Nominal/nominal_primrec.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML Sun May 23 10:38:11 2010 +0100
@@ -393,28 +393,24 @@
(* outer syntax *)
-local structure P = OuterParse in
-
-val freshness_context = P.reserved "freshness_context";
-val invariant = P.reserved "invariant";
+val freshness_context = Parse.reserved "freshness_context";
+val invariant = Parse.reserved "invariant";
-fun unless_flag scan = Scan.unless ((freshness_context || invariant) -- P.$$$ ":") scan;
+fun unless_flag scan = Scan.unless ((freshness_context || invariant) -- Parse.$$$ ":") scan;
-val parser1 = (freshness_context -- P.$$$ ":") |-- unless_flag P.term >> SOME;
-val parser2 = (invariant -- P.$$$ ":") |--
- (Scan.repeat1 (unless_flag P.term) >> SOME) -- Scan.optional parser1 NONE ||
+val parser1 = (freshness_context -- Parse.$$$ ":") |-- unless_flag Parse.term >> SOME;
+val parser2 = (invariant -- Parse.$$$ ":") |--
+ (Scan.repeat1 (unless_flag Parse.term) >> SOME) -- Scan.optional parser1 NONE ||
(parser1 >> pair NONE);
val options =
- Scan.optional (P.$$$ "(" |-- P.!!! (parser2 --| P.$$$ ")")) (NONE, NONE);
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (parser2 --| Parse.$$$ ")")) (NONE, NONE);
val _ =
- OuterSyntax.local_theory_to_proof "nominal_primrec"
- "define primitive recursive functions on nominal datatypes" OuterKeyword.thy_goal
- (options -- P.fixes -- P.for_fixes -- SpecParse.where_alt_specs
+ Outer_Syntax.local_theory_to_proof "nominal_primrec"
+ "define primitive recursive functions on nominal datatypes" Keyword.thy_goal
+ (options -- Parse.fixes -- Parse.for_fixes -- Parse_Spec.where_alt_specs
>> (fn ((((invs, fctxt), fixes), params), specs) =>
add_primrec_cmd invs fctxt fixes params specs));
end;
-end;
-
--- a/src/HOL/Orderings.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Orderings.thy Sun May 23 10:38:11 2010 +0100
@@ -422,8 +422,8 @@
(** Diagnostic command **)
val _ =
- OuterSyntax.improper_command "print_orders"
- "print order structures available to transitivity reasoner" OuterKeyword.diag
+ Outer_Syntax.improper_command "print_orders"
+ "print order structures available to transitivity reasoner" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
Toplevel.keep (print_structures o Toplevel.context_of)));
--- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Sun May 23 10:38:11 2010 +0100
@@ -13,7 +13,7 @@
"greater_than_index xs = (\<forall>i x. nth_el' xs i = Some x --> x > i)"
code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index .
-ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name specialised_nth_el'P} *}
+ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_nth_el'P} *}
thm greater_than_index.equation
@@ -42,7 +42,7 @@
thm max_of_my_SucP.equation
-ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name specialised_max_natP} *}
+ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_max_natP} *}
values "{x. max_of_my_SucP x 6}"
--- a/src/HOL/Probability/Borel.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Probability/Borel.thy Sun May 23 10:38:11 2010 +0100
@@ -397,8 +397,7 @@
{ 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 order_refl order_trans less_le
- xt1(7) zero_less_divide_1_iff) }
+ less_le_trans 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]
--- a/src/HOL/Probability/Caratheodory.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Probability/Caratheodory.thy Sun May 23 10:38:11 2010 +0100
@@ -167,9 +167,9 @@
{l. l \<in> sets M & (\<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x)}"
proof -
have [simp]: "!!l x. l \<in> sets M \<Longrightarrow> x \<in> sets M \<Longrightarrow> (space M - l) \<inter> x = x - l"
- by (metis Diff_eq Int_Diff Int_absorb1 Int_commute sets_into_space)
+ by (metis Int_Diff Int_absorb1 Int_commute sets_into_space)
show ?thesis
- by (auto simp add: lambda_system_def) (metis Diff_Compl Int_commute)+
+ by (auto simp add: lambda_system_def) (metis Int_commute)+
qed
lemma (in algebra) lambda_system_empty:
@@ -352,7 +352,7 @@
proof (auto simp add: disjointed_def)
fix n
show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> sets M" using UNION_in_sets
- by (metis A Diff UNIV_I disjointed_def image_subset_iff)
+ by (metis A Diff UNIV_I image_subset_iff)
qed
lemma sigma_algebra_disjoint_iff:
@@ -476,7 +476,7 @@
lambda_system_positive lambda_system_additive lambda_system_increasing
A' oms outer_measure_space_def disj)
have U_in: "(\<Union>i. A i) \<in> sets M"
- by (metis A countable_UN image_subset_iff lambda_system_sets)
+ by (metis A'' countable_UN)
have U_eq: "f (\<Union>i. A i) = suminf (f o A)"
proof (rule antisym)
show "f (\<Union>i. A i) \<le> suminf (f \<circ> A)"
@@ -500,7 +500,7 @@
apply (auto simp add: )
apply (subst abs_of_nonneg)
apply (metis A'' Int UNIV_I a image_subset_iff)
- apply (blast intro: increasingD [OF inc] a)
+ apply (blast intro: increasingD [OF inc])
done
show ?thesis
proof (rule antisym)
@@ -523,15 +523,15 @@
have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
by (metis A'' UNION_in_sets)
have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
- by (blast intro: increasingD [OF inc] A'' Int UNION_in_sets a)
+ by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system M f"
using algebra.UNION_in_sets [OF lambda_system_algebra [OF pos]]
by (simp add: A)
hence eq_fa: "f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i)) = f a"
- by (simp add: lambda_system_eq UNION_in Diff_Compl a)
+ by (simp add: lambda_system_eq UNION_in)
have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
- by (blast intro: increasingD [OF inc] Diff UNION_eq_Union_image
- UNION_in U_in a)
+ by (blast intro: increasingD [OF inc] UNION_eq_Union_image
+ UNION_in U_in)
thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {0..<n} \<le> f a - f (a - (\<Union>i. A i))"
using eq_fa
by (simp add: suminf_le [OF summ] lambda_system_strong_sum pos
@@ -541,9 +541,9 @@
by arith
next
have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
- by (blast intro: increasingD [OF inc] a U_in)
+ by (blast intro: increasingD [OF inc] U_in)
also have "... \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))"
- by (blast intro: subadditiveD [OF sa] Int Diff U_in)
+ by (blast intro: subadditiveD [OF sa] U_in)
finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" .
qed
qed
@@ -610,11 +610,11 @@
fix x y
assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
have "f x \<le> f x + f (y-x)" using posf
- by (simp add: positive_def) (metis Diff xy)
+ by (simp add: positive_def) (metis Diff xy(1,2))
also have "... = f (x \<union> (y-x))" using addf
- by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy)
+ by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
also have "... = f y"
- by (metis Un_Diff_cancel Un_absorb1 xy)
+ by (metis Un_Diff_cancel Un_absorb1 xy(3))
finally show "f x \<le> f y" .
qed
@@ -663,7 +663,7 @@
by (metis UN_extend_simps(4) s seq)
qed
hence "f s = suminf (\<lambda>i. f (A i \<inter> s))"
- by (metis Int_commute UN_simps(4) seq sums_iff)
+ using seq [symmetric] by (simp add: sums_iff)
also have "... \<le> suminf (f \<circ> A)"
proof (rule summable_le [OF _ _ sm])
show "\<forall>n. f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
@@ -704,7 +704,7 @@
apply (auto simp add: increasing_def)
apply (rule Inf_greatest, metis emptyE inf_measure_nonempty top posf)
apply (rule Inf_lower)
-apply (clarsimp simp add: measure_set_def, blast)
+apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
apply (blast intro: inf_measure_pos0 posf)
done
@@ -743,7 +743,7 @@
apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA)
done
show ?thesis
- by (blast intro: Inf_lower y order_trans [OF _ ley] inf_measure_pos0 posf)
+ by (blast intro: y order_trans [OF _ ley] inf_measure_pos0 posf)
qed
lemma (in algebra) inf_measure_close:
--- a/src/HOL/Probability/Lebesgue.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Probability/Lebesgue.thy Sun May 23 10:38:11 2010 +0100
@@ -76,15 +76,14 @@
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"
+ have "pos_part f x - neg_part f x = 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
+ hence "(\<lambda> x. pos_part f x - neg_part f x) = (\<lambda>x. f x)" by auto
+ hence f: "(\<lambda> x. pos_part f x - neg_part f x) = f" 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
+ show ?thesis unfolding f by fast
qed
context measure_space
--- a/src/HOL/Probability/Measure.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Probability/Measure.thy Sun May 23 10:38:11 2010 +0100
@@ -143,7 +143,7 @@
shows "A \<union> B \<in> smallest_ccdi_sets M"
proof -
have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)"
- by (simp add: range_binaryset_eq A B empty_sets smallest_ccdi_sets.Basic)
+ by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic)
have di: "disjoint_family (binaryset A B)" using disj
by (simp add: disjoint_family_on_def binaryset_def Int_commute)
from Disj [OF ra di]
@@ -277,9 +277,8 @@
apply (blast intro: smallest_ccdi_sets.Disj)
done
hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M"
- by auto
- (metis sigma_algebra.sigma_sets_subset algebra.simps(1)
- algebra.simps(2) subsetD)
+ by clarsimp
+ (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto)
also have "... \<subseteq> C"
proof
fix x
@@ -344,7 +343,7 @@
show "algebra M"
by (blast intro: sigma_algebra.axioms local.sigma_algebra_axioms)
show "positive M (measure M)"
- by (simp add: positive_def empty_measure positive)
+ by (simp add: positive_def positive)
qed
lemma (in measure_space) measure_additive:
@@ -410,7 +409,7 @@
have "(measure M \<circ> A) n =
setsum (measure M \<circ> (\<lambda>i. A (Suc i) - A i)) {0..<n}"
proof (induct n)
- case 0 thus ?case by (auto simp add: A0 empty_measure)
+ case 0 thus ?case by (auto simp add: A0)
next
case (Suc m)
have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
@@ -440,7 +439,7 @@
have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
by (metis A Diff range_subsetD)
have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
- by (blast intro: countable_UN range_subsetD [OF A])
+ by (blast intro: range_subsetD [OF A])
have "(measure M o (\<lambda>i. A (Suc i) - A i)) sums measure M (\<Union>i. A (Suc i) - A i)"
by (rule measure_countably_additive)
(auto simp add: disjoint_family_Suc ASuc A1 A2)
@@ -486,7 +485,8 @@
assume y: "y \<in> sets b"
with a b f
have "space a - f -` y = f -` (space b - y)"
- by (auto simp add: sigma_algebra_iff2) (blast intro: ba)
+ by (auto, clarsimp simp add: sigma_algebra_iff2)
+ (drule ba, blast intro: ba)
hence "space a - (f -` y) \<in> (vimage f) ` sets b"
by auto
(metis b y subsetD equalityE imageI sigma_algebra.sigma_sets_eq
@@ -788,7 +788,7 @@
show ?thesis
proof (rule measurable_sigma)
show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2
- by (force simp add: prod_sets_def sigma_algebra_iff algebra_def)
+ by (auto simp add: prod_sets_def sigma_algebra_iff dest: algebra.space_closed)
next
show "f \<in> space M \<rightarrow> space a1 \<times> space a2"
by (rule prod_final [OF fn1 fn2])
@@ -843,13 +843,13 @@
next
case (insert x t)
hence x: "x \<in> s" and t: "t \<subseteq> s"
- by (simp_all add: insert_subset)
+ by simp_all
hence ts: "t \<in> sets M" using insert
by (metis Diff_insert_absorb Diff ssets)
have "measure M (insert x t) = measure M ({x} \<union> t)"
by simp
also have "... = measure M {x} + measure M t"
- by (simp add: measure_additive insert insert_disjoint ssets x ts
+ by (simp add: measure_additive insert ssets x ts
del: Un_insert_left)
also have "... = (\<Sum>x\<in>insert x t. measure M {x})"
by (simp add: x t ts insert)
--- a/src/HOL/Probability/Sigma_Algebra.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy Sun May 23 10:38:11 2010 +0100
@@ -75,8 +75,8 @@
assumes a: "range a \<subseteq> sets M"
shows "(\<Inter>i::nat. a i) \<in> sets M"
proof -
- have "space M - (\<Union>i. space M - a i) \<in> sets M" using a
- by (blast intro: countable_UN compl_sets a)
+ from a have "\<forall>i. a i \<in> sets M" by fast
+ hence "space M - (\<Union>i. space M - a i) \<in> sets M" by blast
moreover
have "(\<Inter>i. a i) = space M - (\<Union>i. space M - a i)" using space_closed a
by blast
@@ -161,7 +161,7 @@
lemma sigma_sets_Un:
"a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
apply (simp add: Un_range_binary range_binary_eq)
-apply (metis Union COMBK_def binary_def fun_upd_apply)
+apply (rule Union, simp add: binary_def COMBK_def fun_upd_apply)
done
lemma sigma_sets_Inter:
@@ -210,7 +210,7 @@
"sigma_sets (space M) (sets M) = sets M"
proof
show "sets M \<subseteq> sigma_sets (space M) (sets M)"
- by (metis Set.subsetI sigma_sets.Basic space_closed)
+ by (metis Set.subsetI sigma_sets.Basic)
next
show "sigma_sets (space M) (sets M) \<subseteq> sets M"
by (metis sigma_sets_subset subset_refl)
@@ -221,7 +221,7 @@
apply (auto simp add: sigma_algebra_def sigma_algebra_axioms_def
algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un)
apply (blast dest: sigma_sets_into_sp)
- apply (blast intro: sigma_sets.intros)
+ apply (rule sigma_sets.Union, fast)
done
lemma sigma_algebra_sigma:
--- a/src/HOL/Prolog/prolog.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Prolog/prolog.ML Sun May 23 10:38:11 2010 +0100
@@ -82,7 +82,7 @@
val (prems, Const("Trueprop", _)$concl) = rep_goal
(#3(dest_state (st,i)));
*)
- val subgoal = #3(dest_state (st,i));
+ val subgoal = #3(Thm.dest_state (st,i));
val prems = Logic.strip_assums_hyp subgoal;
val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal);
fun drot_tac k i = DETERM (rotate_tac k i);
--- a/src/HOL/Quotient.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Quotient.thy Sun May 23 10:38:11 2010 +0100
@@ -618,15 +618,13 @@
lemma let_prs:
assumes q1: "Quotient R1 Abs1 Rep1"
and q2: "Quotient R2 Abs2 Rep2"
- shows "Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f)) = Let x f"
- using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
+ shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let"
+ using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
+ by (auto simp add: expand_fun_eq)
lemma let_rsp:
- assumes q1: "Quotient R1 Abs1 Rep1"
- and a1: "(R1 ===> R2) f g"
- and a2: "R1 x y"
- shows "R2 ((Let x f)::'c) ((Let y g)::'c)"
- using apply_rsp[OF q1 a1] a2 by auto
+ shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let"
+ by auto
locale quot_type =
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
@@ -716,8 +714,8 @@
declare [[map "fun" = (fun_map, fun_rel)]]
lemmas [quot_thm] = fun_quotient
-lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp
-lemmas [quot_preserve] = if_prs o_prs
+lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp
+lemmas [quot_preserve] = if_prs o_prs let_prs
lemmas [quot_equiv] = identity_equivp
--- a/src/HOL/RComplete.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/RComplete.thy Sun May 23 10:38:11 2010 +0100
@@ -117,10 +117,12 @@
lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)"
by (drule reals_Archimedean6) auto
+text {* TODO: delete *}
lemma reals_Archimedean_6b_int:
"0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
unfolding real_of_int_def by (rule floor_exists)
+text {* TODO: delete *}
lemma reals_Archimedean_6c_int:
"r < 0 ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
unfolding real_of_int_def by (rule floor_exists)
@@ -204,9 +206,6 @@
"(real (m::int) \<le> (number_of n)) = (m \<le> number_of n)"
by (simp add: linorder_not_less [symmetric])
-lemma floor_real_of_nat_zero: "floor (real (0::nat)) = 0"
-by auto (* delete? *)
-
lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
unfolding real_of_nat_def by simp
@@ -259,10 +258,6 @@
apply (auto intro: floor_eq3)
done
-lemma floor_number_of_eq:
- "floor(number_of n :: real) = (number_of n :: int)"
- by (rule floor_number_of) (* already declared [simp] *)
-
lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
unfolding real_of_int_def using floor_correct [of r] by simp
@@ -284,68 +279,21 @@
lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
unfolding real_of_int_def by (rule le_floor_iff)
-lemma le_floor_eq_number_of:
- "(number_of n <= floor x) = (number_of n <= x)"
- by (rule number_of_le_floor) (* already declared [simp] *)
-
-lemma le_floor_eq_zero: "(0 <= floor x) = (0 <= x)"
- by (rule zero_le_floor) (* already declared [simp] *)
-
-lemma le_floor_eq_one: "(1 <= floor x) = (1 <= x)"
- by (rule one_le_floor) (* already declared [simp] *)
-
lemma floor_less_eq: "(floor x < a) = (x < real a)"
unfolding real_of_int_def by (rule floor_less_iff)
-lemma floor_less_eq_number_of:
- "(floor x < number_of n) = (x < number_of n)"
- by (rule floor_less_number_of) (* already declared [simp] *)
-
-lemma floor_less_eq_zero: "(floor x < 0) = (x < 0)"
- by (rule floor_less_zero) (* already declared [simp] *)
-
-lemma floor_less_eq_one: "(floor x < 1) = (x < 1)"
- by (rule floor_less_one) (* already declared [simp] *)
-
lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
unfolding real_of_int_def by (rule less_floor_iff)
-lemma less_floor_eq_number_of:
- "(number_of n < floor x) = (number_of n + 1 <= x)"
- by (rule number_of_less_floor) (* already declared [simp] *)
-
-lemma less_floor_eq_zero: "(0 < floor x) = (1 <= x)"
- by (rule zero_less_floor) (* already declared [simp] *)
-
-lemma less_floor_eq_one: "(1 < floor x) = (2 <= x)"
- by (rule one_less_floor) (* already declared [simp] *)
-
lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
unfolding real_of_int_def by (rule floor_le_iff)
-lemma floor_le_eq_number_of:
- "(floor x <= number_of n) = (x < number_of n + 1)"
- by (rule floor_le_number_of) (* already declared [simp] *)
-
-lemma floor_le_eq_zero: "(floor x <= 0) = (x < 1)"
- by (rule floor_le_zero) (* already declared [simp] *)
-
-lemma floor_le_eq_one: "(floor x <= 1) = (x < 2)"
- by (rule floor_le_one) (* already declared [simp] *)
-
lemma floor_add [simp]: "floor (x + real a) = floor x + a"
unfolding real_of_int_def by (rule floor_add_of_int)
lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
unfolding real_of_int_def by (rule floor_diff_of_int)
-lemma floor_subtract_number_of: "floor (x - number_of n) =
- floor x - number_of n"
- by (rule floor_diff_number_of) (* already declared [simp] *)
-
-lemma floor_subtract_one: "floor (x - 1) = floor x - 1"
- by (rule floor_diff_one) (* already declared [simp] *)
-
lemma le_mult_floor:
assumes "0 \<le> (a :: real)" and "0 \<le> b"
shows "floor a * floor b \<le> floor (a * b)"
@@ -361,9 +309,6 @@
lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
unfolding real_of_nat_def by simp
-lemma ceiling_real_of_nat_zero: "ceiling (real (0::nat)) = 0"
-by auto (* delete? *)
-
lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r"
unfolding real_of_int_def by simp
@@ -389,10 +334,6 @@
lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
unfolding real_of_int_def using ceiling_unique [of n x] by simp
-lemma ceiling_number_of_eq:
- "ceiling (number_of n :: real) = (number_of n)"
- by (rule ceiling_number_of) (* already declared [simp] *)
-
lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
unfolding real_of_int_def using ceiling_correct [of r] by simp
@@ -408,68 +349,21 @@
lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
unfolding real_of_int_def by (rule ceiling_le_iff)
-lemma ceiling_le_eq_number_of:
- "(ceiling x <= number_of n) = (x <= number_of n)"
- by (rule ceiling_le_number_of) (* already declared [simp] *)
-
-lemma ceiling_le_zero_eq: "(ceiling x <= 0) = (x <= 0)"
- by (rule ceiling_le_zero) (* already declared [simp] *)
-
-lemma ceiling_le_eq_one: "(ceiling x <= 1) = (x <= 1)"
- by (rule ceiling_le_one) (* already declared [simp] *)
-
lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
unfolding real_of_int_def by (rule less_ceiling_iff)
-lemma less_ceiling_eq_number_of:
- "(number_of n < ceiling x) = (number_of n < x)"
- by (rule number_of_less_ceiling) (* already declared [simp] *)
-
-lemma less_ceiling_eq_zero: "(0 < ceiling x) = (0 < x)"
- by (rule zero_less_ceiling) (* already declared [simp] *)
-
-lemma less_ceiling_eq_one: "(1 < ceiling x) = (1 < x)"
- by (rule one_less_ceiling) (* already declared [simp] *)
-
lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
unfolding real_of_int_def by (rule ceiling_less_iff)
-lemma ceiling_less_eq_number_of:
- "(ceiling x < number_of n) = (x <= number_of n - 1)"
- by (rule ceiling_less_number_of) (* already declared [simp] *)
-
-lemma ceiling_less_eq_zero: "(ceiling x < 0) = (x <= -1)"
- by (rule ceiling_less_zero) (* already declared [simp] *)
-
-lemma ceiling_less_eq_one: "(ceiling x < 1) = (x <= 0)"
- by (rule ceiling_less_one) (* already declared [simp] *)
-
lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
unfolding real_of_int_def by (rule le_ceiling_iff)
-lemma le_ceiling_eq_number_of:
- "(number_of n <= ceiling x) = (number_of n - 1 < x)"
- by (rule number_of_le_ceiling) (* already declared [simp] *)
-
-lemma le_ceiling_eq_zero: "(0 <= ceiling x) = (-1 < x)"
- by (rule zero_le_ceiling) (* already declared [simp] *)
-
-lemma le_ceiling_eq_one: "(1 <= ceiling x) = (0 < x)"
- by (rule one_le_ceiling) (* already declared [simp] *)
-
lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
unfolding real_of_int_def by (rule ceiling_add_of_int)
lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
unfolding real_of_int_def by (rule ceiling_diff_of_int)
-lemma ceiling_subtract_number_of: "ceiling (x - number_of n) =
- ceiling x - number_of n"
- by (rule ceiling_diff_number_of) (* already declared [simp] *)
-
-lemma ceiling_subtract_one: "ceiling (x - 1) = ceiling x - 1"
- by (rule ceiling_diff_one) (* already declared [simp] *)
-
subsection {* Versions for the natural numbers *}
--- a/src/HOL/RealDef.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/RealDef.thy Sun May 23 10:38:11 2010 +0100
@@ -1055,6 +1055,7 @@
lemmas real_le_refl = order_refl [of "w::real", standard]
lemmas real_le_antisym = order_antisym [of "z::real" "w", standard]
lemmas real_le_trans = order_trans [of "i::real" "j" "k", standard]
+lemmas real_le_linear = linear [of "z::real" "w", standard]
lemmas real_le_eq_diff = le_iff_diff_le_0 [of "x::real" "y", standard]
lemmas real_add_left_mono = add_left_mono [of "x::real" "y" "z", standard]
lemmas real_mult_order = mult_pos_pos [of "x::real" "y", standard]
@@ -1619,14 +1620,6 @@
"(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
by (rule sum_power2_eq_zero_iff)
-text {* FIXME: declare this [simp] for all types, or not at all *}
-lemma realpow_two_le_add_order [simp]: "(0::real) \<le> u ^ 2 + v ^ 2"
-by (rule sum_power2_ge_zero)
-
-text {* FIXME: declare this [simp] for all types, or not at all *}
-lemma realpow_two_le_add_order2 [simp]: "(0::real) \<le> u ^ 2 + v ^ 2 + w ^ 2"
-by (intro add_nonneg_nonneg zero_le_power2)
-
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
by (rule_tac y = 0 in order_trans, auto)
--- a/src/HOL/Rings.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Rings.thy Sun May 23 10:38:11 2010 +0100
@@ -349,7 +349,7 @@
class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
begin
-lemma square_eq_1_iff [simp]:
+lemma square_eq_1_iff:
"x * x = 1 \<longleftrightarrow> x = 1 \<or> x = - 1"
proof -
have "(x - 1) * (x + 1) = x * x - 1"
@@ -956,7 +956,7 @@
show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
by (auto simp add: abs_if not_less)
(auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric],
- auto intro: add_nonneg_nonneg, auto intro!: less_imp_le add_neg_neg)
+ auto intro!: less_imp_le add_neg_neg)
qed (auto simp add: abs_if)
lemma zero_le_square [simp]: "0 \<le> a * a"
--- a/src/HOL/Statespace/distinct_tree_prover.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Statespace/distinct_tree_prover.ML Sun May 23 10:38:11 2010 +0100
@@ -159,8 +159,6 @@
in mtch env (t,ct) end;
-fun mp prem rule = implies_elim rule prem;
-
fun discharge prems rule =
let
val thy = theory_of_thm (hd prems);
@@ -172,7 +170,7 @@
val insts' = map (fn (idxn,ct) => (cterm_of thy (Var (idxn,typ_of (ctyp_of_term ct))),ct))
insts;
val rule' = Thm.instantiate (tyinsts',insts') rule;
- in fold mp prems rule' end;
+ in fold Thm.elim_implies prems rule' end;
local
--- a/src/HOL/Statespace/state_fun.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Statespace/state_fun.ML Sun May 23 10:38:11 2010 +0100
@@ -245,7 +245,7 @@
(fn _ => rtac meta_ext 1 THEN
simp_tac ss1 1);
val eq2 = Simplifier.asm_full_rewrite ss2 (Thm.dest_equals_rhs (cprop_of eq1));
- in SOME (transitive eq1 eq2) end
+ in SOME (Thm.transitive eq1 eq2) end
| _ => NONE)
end
| _ => NONE));
--- a/src/HOL/Statespace/state_space.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Statespace/state_space.ML Sun May 23 10:38:11 2010 +0100
@@ -32,10 +32,9 @@
(string * typ) list -> theory -> theory
val statespace_decl :
- OuterParse.token list ->
((string list * bstring) *
((string list * xstring * (bstring * bstring) list) list *
- (bstring * string) list)) * OuterParse.token list
+ (bstring * string) list)) parser
val neq_x_y : Proof.context -> term -> term -> thm option
@@ -668,37 +667,33 @@
(*** outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val type_insts =
- P.typ >> single ||
- P.$$$ "(" |-- P.!!! (P.list1 P.typ --| P.$$$ ")")
+ Parse.typ >> single ||
+ Parse.$$$ "(" |-- Parse.!!! (Parse.list1 Parse.typ --| Parse.$$$ ")")
-val comp = P.name -- (P.$$$ "::" |-- P.!!! P.typ);
+val comp = Parse.name -- (Parse.$$$ "::" |-- Parse.!!! Parse.typ);
fun plus1_unless test scan =
- scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
+ scan ::: Scan.repeat (Parse.$$$ "+" |-- Scan.unless test (Parse.!!! scan));
-val mapsto = P.$$$ "=";
-val rename = P.name -- (mapsto |-- P.name);
-val renames = Scan.optional (P.$$$ "[" |-- P.!!! (P.list1 rename --| P.$$$ "]")) [];
+val mapsto = Parse.$$$ "=";
+val rename = Parse.name -- (mapsto |-- Parse.name);
+val renames = Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.list1 rename --| Parse.$$$ "]")) [];
-val parent = ((type_insts -- P.xname) || (P.xname >> pair [])) -- renames
+val parent = ((type_insts -- Parse.xname) || (Parse.xname >> pair [])) -- renames
>> (fn ((insts,name),renames) => (insts,name,renames))
val statespace_decl =
- P.type_args -- P.name --
- (P.$$$ "=" |--
+ Parse.type_args -- Parse.name --
+ (Parse.$$$ "=" |--
((Scan.repeat1 comp >> pair []) ||
(plus1_unless comp parent --
- Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 comp)) [])))
+ Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 comp)) [])))
val statespace_command =
- OuterSyntax.command "statespace" "define state space" K.thy_decl
+ Outer_Syntax.command "statespace" "define state space" Keyword.thy_decl
(statespace_decl >> (fn ((args,name),(parents,comps)) =>
Toplevel.theory (define_statespace args name parents comps)))
-end;
-
end;
\ No newline at end of file
--- a/src/HOL/Tools/ATP_Manager/atp_manager.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML Sun May 23 10:38:11 2010 +0100
@@ -20,11 +20,11 @@
explicit_apply: bool,
respect_no_atp: bool,
relevance_threshold: real,
- convergence: real,
+ relevance_convergence: real,
theory_relevant: bool option,
- follow_defs: bool,
+ defs_relevant: bool,
isar_proof: bool,
- shrink_factor: int,
+ isar_shrink_factor: int,
timeout: Time.time,
minimize_timeout: Time.time}
type problem =
@@ -34,8 +34,8 @@
axiom_clauses: (thm * (string * int)) list option,
filtered_clauses: (thm * (string * int)) list option}
datatype failure =
- Unprovable | TimedOut | OutOfResources | OldSpass | MalformedOutput |
- UnknownError
+ Unprovable | TimedOut | OutOfResources | OldSpass | MalformedInput |
+ MalformedOutput | UnknownError
type prover_result =
{outcome: failure option,
message: string,
@@ -79,11 +79,11 @@
explicit_apply: bool,
respect_no_atp: bool,
relevance_threshold: real,
- convergence: real,
+ relevance_convergence: real,
theory_relevant: bool option,
- follow_defs: bool,
+ defs_relevant: bool,
isar_proof: bool,
- shrink_factor: int,
+ isar_shrink_factor: int,
timeout: Time.time,
minimize_timeout: Time.time}
@@ -95,8 +95,8 @@
filtered_clauses: (thm * (string * int)) list option};
datatype failure =
- Unprovable | TimedOut | OutOfResources | OldSpass | MalformedOutput |
- UnknownError
+ Unprovable | TimedOut | OutOfResources | OldSpass | MalformedInput |
+ MalformedOutput | UnknownError
type prover_result =
{outcome: failure option,
--- a/src/HOL/Tools/ATP_Manager/atp_systems.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/ATP_Manager/atp_systems.ML Sun May 23 10:38:11 2010 +0100
@@ -46,7 +46,7 @@
(* prover configuration *)
type prover_config =
- {home: string,
+ {home_var: string,
executable: string,
arguments: Time.time -> string,
proof_delims: (string * string) list,
@@ -100,6 +100,9 @@
(Path.variable "ISABELLE_HOME_USER" ::
map Path.basic ["etc", "components"]))) ^
"\" on a line of its own."
+ | string_for_failure MalformedInput =
+ "Internal Sledgehammer error: The ATP problem is malformed. Please report \
+ \this to the Isabelle developers."
| string_for_failure MalformedOutput = "Error: The ATP output is malformed."
| string_for_failure UnknownError = "Error: An unknown ATP error occurred."
@@ -110,10 +113,10 @@
(j :: hd shape) :: tl shape
end
-fun generic_prover overlord get_facts prepare write_file home executable args
- proof_delims known_failures name
- ({debug, full_types, explicit_apply, isar_proof, shrink_factor, ...}
- : params) minimize_command
+fun generic_prover overlord get_facts prepare write_file home_var executable
+ args proof_delims known_failures name
+ ({debug, full_types, explicit_apply, isar_proof, isar_shrink_factor,
+ ...} : params) minimize_command
({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses}
: problem) =
let
@@ -151,6 +154,7 @@
else error ("No such directory: " ^ the_dest_dir ^ ".")
end;
+ val home = getenv home_var
val command = Path.explode (home ^ "/" ^ executable)
(* write out problem file and call prover *)
fun command_line probfile =
@@ -186,6 +190,8 @@
if File.exists command then
write_file full_types explicit_apply probfile clauses
|> pair (apfst split_time' (bash_output (command_line probfile)))
+ else if home = "" then
+ error ("The environment variable " ^ quote home_var ^ " is not set.")
else
error ("Bad executable: " ^ Path.implode command ^ ".");
@@ -215,7 +221,8 @@
case outcome of
NONE =>
proof_text isar_proof
- (pool, debug, full_types, shrink_factor, ctxt, conjecture_shape)
+ (pool, debug, full_types, isar_shrink_factor, ctxt,
+ conjecture_shape)
(minimize_command, proof, internal_thm_names, th, subgoal)
| SOME failure => (string_for_failure failure ^ "\n", [])
in
@@ -231,17 +238,18 @@
(* generic TPTP-based provers *)
fun generic_tptp_prover
- (name, {home, executable, arguments, proof_delims, known_failures,
+ (name, {home_var, executable, arguments, proof_delims, known_failures,
max_axiom_clauses, prefers_theory_relevant})
(params as {debug, overlord, respect_no_atp, relevance_threshold,
- convergence, theory_relevant, follow_defs, isar_proof, ...})
+ relevance_convergence, theory_relevant, defs_relevant,
+ isar_proof, ...})
minimize_command timeout =
generic_prover overlord
- (get_relevant_facts respect_no_atp relevance_threshold convergence
- follow_defs max_axiom_clauses
+ (get_relevant_facts respect_no_atp relevance_threshold
+ relevance_convergence defs_relevant max_axiom_clauses
(the_default prefers_theory_relevant theory_relevant))
(prepare_clauses false)
- (write_tptp_file (debug andalso overlord)) home
+ (write_tptp_file (debug andalso overlord)) home_var
executable (arguments timeout) proof_delims known_failures name params
minimize_command
@@ -257,7 +265,7 @@
(* Vampire requires an explicit time limit. *)
val vampire_config : prover_config =
- {home = getenv "VAMPIRE_HOME",
+ {home_var = "VAMPIRE_HOME",
executable = "vampire",
arguments = fn timeout =>
"--output_syntax tptp --mode casc -t " ^
@@ -281,7 +289,7 @@
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
val e_config : prover_config =
- {home = getenv "E_HOME",
+ {home_var = "E_HOME",
executable = "eproof",
arguments = fn timeout =>
"--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^
@@ -304,16 +312,16 @@
(* SPASS *)
fun generic_dfg_prover
- (name, {home, executable, arguments, proof_delims, known_failures,
+ (name, {home_var, executable, arguments, proof_delims, known_failures,
max_axiom_clauses, prefers_theory_relevant})
- (params as {overlord, respect_no_atp, relevance_threshold, convergence,
- theory_relevant, follow_defs, ...})
+ (params as {overlord, respect_no_atp, relevance_threshold,
+ relevance_convergence, theory_relevant, defs_relevant, ...})
minimize_command timeout =
generic_prover overlord
- (get_relevant_facts respect_no_atp relevance_threshold convergence
- follow_defs max_axiom_clauses
+ (get_relevant_facts respect_no_atp relevance_threshold
+ relevance_convergence defs_relevant max_axiom_clauses
(the_default prefers_theory_relevant theory_relevant))
- (prepare_clauses true) write_dfg_file home executable
+ (prepare_clauses true) write_dfg_file home_var executable
(arguments timeout) proof_delims known_failures name params
minimize_command
@@ -322,7 +330,7 @@
(* The "-VarWeight=3" option helps the higher-order problems, probably by
counteracting the presence of "hAPP". *)
val spass_config : prover_config =
- {home = getenv "SPASS_HOME",
+ {home_var = "SPASS_HOME",
executable = "SPASS",
arguments = fn timeout =>
"-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
@@ -331,7 +339,8 @@
known_failures =
[(Unprovable, "SPASS beiseite: Completion found"),
(TimedOut, "SPASS beiseite: Ran out of time"),
- (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded")],
+ (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
+ (MalformedInput, "Undefined symbol")],
max_axiom_clauses = 40,
prefers_theory_relevant = true}
val spass = dfg_prover "spass" spass_config
@@ -342,7 +351,7 @@
Sledgehammer) and rename "spass_tptp" "spass". *)
val spass_tptp_config =
- {home = #home spass_config,
+ {home_var = #home_var spass_config,
executable = #executable spass_config,
arguments = prefix "-TPTP " o #arguments spass_config,
proof_delims = #proof_delims spass_config,
@@ -384,7 +393,7 @@
fun remote_prover_config atp_prefix args
({proof_delims, known_failures, max_axiom_clauses,
prefers_theory_relevant, ...} : prover_config) : prover_config =
- {home = getenv "ISABELLE_ATP_MANAGER",
+ {home_var = "ISABELLE_ATP_MANAGER",
executable = "SystemOnTPTP",
arguments = fn timeout =>
args ^ " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^
@@ -406,8 +415,8 @@
tptp_prover (remotify (fst spass))
(remote_prover_config "SPASS---" "-x" spass_config)
-fun maybe_remote (name, _) ({home, ...} : prover_config) =
- name |> home = "" ? remotify
+fun maybe_remote (name, _) ({home_var, ...} : prover_config) =
+ name |> getenv home_var = "" ? remotify
fun default_atps_param_value () =
space_implode " " [maybe_remote e e_config, maybe_remote spass spass_config,
--- a/src/HOL/Tools/Datatype/datatype.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML Sun May 23 10:38:11 2010 +0100
@@ -483,8 +483,8 @@
[(indtac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
REPEAT (rtac TrueI 1),
rewrite_goals_tac (mk_meta_eq @{thm choice_eq} ::
- symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
- rewrite_goals_tac (map symmetric range_eqs),
+ Thm.symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
+ rewrite_goals_tac (map Thm.symmetric range_eqs),
REPEAT (EVERY
[REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
@@ -621,7 +621,7 @@
(indtac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
EVERY (map (fn (prem, r) => (EVERY
[REPEAT (eresolve_tac Abs_inverse_thms 1),
- simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
+ simp_tac (HOL_basic_ss addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1,
DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
(prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
@@ -721,8 +721,6 @@
local
-structure P = OuterParse and K = OuterKeyword
-
fun prep_datatype_decls args =
let
val names = map
@@ -732,15 +730,16 @@
in (names, specs) end;
val parse_datatype_decl =
- (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_mixfix --
- (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix)));
+ (Scan.option (Parse.$$$ "(" |-- Parse.name --| Parse.$$$ ")") --
+ Parse.type_args -- Parse.binding -- Parse.opt_mixfix --
+ (Parse.$$$ "=" |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix)));
-val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls;
+val parse_datatype_decls = Parse.and_list1 parse_datatype_decl >> prep_datatype_decls;
in
val _ =
- OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
+ Outer_Syntax.command "datatype" "define inductive datatypes" Keyword.thy_decl
(parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs)));
end;
--- a/src/HOL/Tools/Datatype/datatype_data.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Sun May 23 10:38:11 2010 +0100
@@ -455,18 +455,11 @@
(* outer syntax *)
-local
-
-structure P = OuterParse and K = OuterKeyword
-
-in
-
val _ =
- OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
- (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
- >> (fn (alt_names, ts) => Toplevel.print
- o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+ Outer_Syntax.command "rep_datatype" "represent existing types inductively" Keyword.thy_goal
+ (Scan.option (Parse.$$$ "(" |-- Scan.repeat1 Parse.name --| Parse.$$$ ")") --
+ Scan.repeat1 Parse.term
+ >> (fn (alt_names, ts) =>
+ Toplevel.print o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
end;
-
-end;
--- a/src/HOL/Tools/Function/context_tree.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/context_tree.ML Sun May 23 10:38:11 2010 +0100
@@ -148,7 +148,7 @@
val (r, dep, branches) = find_cong_rule ctx fvar h congs_deps t
fun subtree (ctx', fixes, assumes, st) =
((fixes,
- map (assume o cterm_of (ProofContext.theory_of ctx)) assumes),
+ map (Thm.assume o cterm_of (ProofContext.theory_of ctx)) assumes),
mk_tree' ctx' st)
in
Cong (r, dep, map subtree branches)
@@ -165,7 +165,7 @@
fun inst_term t =
subst_bound(f, abstract_over (fvar, t))
- val inst_thm = forall_elim cf o forall_intr cfvar
+ val inst_thm = Thm.forall_elim cf o Thm.forall_intr cfvar
fun inst_tree_aux (Leaf t) = Leaf t
| inst_tree_aux (Cong (crule, deps, branches)) =
@@ -173,7 +173,7 @@
| inst_tree_aux (RCall (t, str)) =
RCall (inst_term t, inst_tree_aux str)
and inst_branch ((fxs, assms), str) =
- ((fxs, map (assume o cterm_of thy o inst_term o prop_of) assms),
+ ((fxs, map (Thm.assume o cterm_of thy o inst_term o prop_of) assms),
inst_tree_aux str)
in
inst_tree_aux tr
@@ -188,11 +188,11 @@
#> fold_rev (Logic.all o Free) fixes
fun export_thm thy (fixes, assumes) =
- fold_rev (implies_intr o cprop_of) assumes
- #> fold_rev (forall_intr o cterm_of thy o Free) fixes
+ fold_rev (Thm.implies_intr o cprop_of) assumes
+ #> fold_rev (Thm.forall_intr o cterm_of thy o Free) fixes
fun import_thm thy (fixes, athms) =
- fold (forall_elim o cterm_of thy o Free) fixes
+ fold (Thm.forall_elim o cterm_of thy o Free) fixes
#> fold Thm.elim_implies athms
@@ -241,7 +241,7 @@
fun rewrite_by_tree thy h ih x tr =
let
- fun rewrite_help _ _ x (Leaf t) = (reflexive (cterm_of thy t), x)
+ fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (cterm_of thy t), x)
| rewrite_help fix h_as x (RCall (_ $ arg, st)) =
let
val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *)
@@ -250,11 +250,11 @@
|> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner))))
(* (a, h a) : G *)
val inst_ih = instantiate' [] [SOME (cterm_of thy arg)] ih
- val eq = implies_elim (implies_elim inst_ih lRi) iha (* h a = f a *)
+ val eq = Thm.implies_elim (Thm.implies_elim inst_ih lRi) iha (* h a = f a *)
- val h_a'_eq_h_a = combination (reflexive (cterm_of thy h)) inner
+ val h_a'_eq_h_a = Thm.combination (Thm.reflexive (cterm_of thy h)) inner
val h_a_eq_f_a = eq RS eq_reflection
- val result = transitive h_a'_eq_h_a h_a_eq_f_a
+ val result = Thm.transitive h_a'_eq_h_a h_a_eq_f_a
in
(result, x')
end
--- a/src/HOL/Tools/Function/fun.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/fun.ML Sun May 23 10:38:11 2010 +0100
@@ -159,13 +159,10 @@
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.local_theory "fun" "define general recursive functions (short version)" K.thy_decl
+ Outer_Syntax.local_theory "fun" "define general recursive functions (short version)"
+ Keyword.thy_decl
(function_parser fun_config
- >> (fn ((config, fixes), statements) => add_fun_cmd fixes statements config));
+ >> (fn ((config, fixes), statements) => add_fun_cmd fixes statements config))
end
-
-end
--- a/src/HOL/Tools/Function/function.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/function.ML Sun May 23 10:38:11 2010 +0100
@@ -274,20 +274,19 @@
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
|> the_single |> snd
+
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
+ Outer_Syntax.local_theory_to_proof "function" "define general recursive functions"
+ Keyword.thy_goal
(function_parser default_config
>> (fn ((config, fixes), statements) => function_cmd fixes statements config))
val _ =
- OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
- (Scan.option P.term >> termination_cmd)
-
-end
+ Outer_Syntax.local_theory_to_proof "termination" "prove termination of a recursive function"
+ Keyword.thy_goal
+ (Scan.option Parse.term >> termination_cmd)
end
--- a/src/HOL/Tools/Function/function_common.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/function_common.ML Sun May 23 10:38:11 2010 +0100
@@ -341,21 +341,19 @@
local
- structure P = OuterParse and K = OuterKeyword
-
- val option_parser = P.group "option"
- ((P.reserved "sequential" >> K Sequential)
- || ((P.reserved "default" |-- P.term) >> Default)
- || (P.reserved "domintros" >> K DomIntros)
- || (P.reserved "no_partials" >> K No_Partials)
- || (P.reserved "tailrec" >> K Tailrec))
+ val option_parser = Parse.group "option"
+ ((Parse.reserved "sequential" >> K Sequential)
+ || ((Parse.reserved "default" |-- Parse.term) >> Default)
+ || (Parse.reserved "domintros" >> K DomIntros)
+ || (Parse.reserved "no_partials" >> K No_Partials)
+ || (Parse.reserved "tailrec" >> K Tailrec))
fun config_parser default =
- (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
+ (Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Parse.list1 option_parser) --| Parse.$$$ ")") [])
>> (fn opts => fold apply_opt opts default)
in
fun function_parser default_cfg =
- config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
+ config_parser default_cfg -- Parse.fixes -- Parse_Spec.where_alt_specs
end
--- a/src/HOL/Tools/Function/function_core.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/function_core.ML Sun May 23 10:38:11 2010 +0100
@@ -154,9 +154,9 @@
val rhs = inst pre_rhs
val cqs = map (cterm_of thy) qs
- val ags = map (assume o cterm_of thy) gs
+ val ags = map (Thm.assume o cterm_of thy) gs
- val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+ val case_hyp = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
in
ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
cqs = cqs, ags = ags, case_hyp = case_hyp }
@@ -188,15 +188,15 @@
(* Instantiate the GIntro thm with "f" and import into the clause context. *)
val lGI = GIntro_thm
- |> forall_elim (cert f)
- |> fold forall_elim cqs
+ |> Thm.forall_elim (cert f)
+ |> fold Thm.forall_elim cqs
|> fold Thm.elim_implies ags
fun mk_call_info (rcfix, rcassm, rcarg) RI =
let
val llRI = RI
- |> fold forall_elim cqs
- |> fold (forall_elim o cert o Free) rcfix
+ |> fold Thm.forall_elim cqs
+ |> fold (Thm.forall_elim o cert o Free) rcfix
|> fold Thm.elim_implies ags
|> fold Thm.elim_implies rcassm
@@ -242,20 +242,20 @@
val compat = lookup_compat_thm j i cts
in
compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
- |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ |> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
|> fold Thm.elim_implies agsj
|> fold Thm.elim_implies agsi
- |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
+ |> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
end
else
let
val compat = lookup_compat_thm i j cts
in
compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
- |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ |> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
|> fold Thm.elim_implies agsi
|> fold Thm.elim_implies agsj
- |> Thm.elim_implies (assume lhsi_eq_lhsj)
+ |> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
|> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
end
end
@@ -274,16 +274,16 @@
val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
val h_assums = map (fn RCInfo {h_assum, ...} =>
- assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+ Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
val (eql, _) =
Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
val replace_lemma = (eql RS meta_eq_to_obj_eq)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) h_assums
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
+ |> Thm.implies_intr (cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o cprop_of) h_assums
+ |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> fold_rev Thm.forall_intr cqs
|> Thm.close_derivation
in
replace_lemma
@@ -301,30 +301,30 @@
val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
val compat = get_compat_thm thy compat_store i j cctxi cctxj
- val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+ val Ghsj' = map (fn RCInfo {h_assum, ...} => Thm.assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
val RLj_import = RLj
- |> fold forall_elim cqsj'
+ |> fold Thm.forall_elim cqsj'
|> fold Thm.elim_implies agsj'
|> fold Thm.elim_implies Ghsj'
- val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
- val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
+ val y_eq_rhsj'h = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+ val lhsi_eq_lhsj' = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
(* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
in
(trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
- |> implies_elim RLj_import
+ |> Thm.implies_elim RLj_import
(* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
|> (fn it => trans OF [it, compat])
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
|> (fn it => trans OF [y_eq_rhsj'h, it])
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
- |> fold_rev (implies_intr o cprop_of) Ghsj'
- |> fold_rev (implies_intr o cprop_of) agsj'
+ |> fold_rev (Thm.implies_intr o cprop_of) Ghsj'
+ |> fold_rev (Thm.implies_intr o cprop_of) agsj'
(* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
- |> implies_intr (cprop_of y_eq_rhsj'h)
- |> implies_intr (cprop_of lhsi_eq_lhsj')
- |> fold_rev forall_intr (cterm_of thy h :: cqsj')
+ |> Thm.implies_intr (cprop_of y_eq_rhsj'h)
+ |> Thm.implies_intr (cprop_of lhsi_eq_lhsj')
+ |> fold_rev Thm.forall_intr (cterm_of thy h :: cqsj')
end
@@ -338,13 +338,13 @@
val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+ |> fold_rev (Thm.implies_intr o cprop_of) CCas
+ |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
val existence = fold (curry op COMP o prep_RC) RCs lGI
val P = cterm_of thy (mk_eq (y, rhsC))
- val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+ val G_lhs_y = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
val unique_clauses =
map2 (mk_uniqueness_clause thy globals compat_store clausei) clauses rep_lemmas
@@ -353,13 +353,13 @@
Thm.compose_no_flatten true (A, 0) 1 AB |> Seq.list_of |> the_single
val uniqueness = G_cases
- |> forall_elim (cterm_of thy lhs)
- |> forall_elim (cterm_of thy y)
- |> forall_elim P
+ |> Thm.forall_elim (cterm_of thy lhs)
+ |> Thm.forall_elim (cterm_of thy y)
+ |> Thm.forall_elim P
|> Thm.elim_implies G_lhs_y
|> fold elim_implies_eta unique_clauses
- |> implies_intr (cprop_of G_lhs_y)
- |> forall_intr (cterm_of thy y)
+ |> Thm.implies_intr (cprop_of G_lhs_y)
+ |> Thm.forall_intr (cterm_of thy y)
val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
@@ -368,16 +368,16 @@
|> curry (op COMP) existence
|> curry (op COMP) uniqueness
|> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
+ |> Thm.implies_intr (cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> fold_rev Thm.forall_intr cqs
val function_value =
existence
- |> implies_intr ihyp
- |> implies_intr (cprop_of case_hyp)
- |> forall_intr (cterm_of thy x)
- |> forall_elim (cterm_of thy lhs)
+ |> Thm.implies_intr ihyp
+ |> Thm.implies_intr (cprop_of case_hyp)
+ |> Thm.forall_intr (cterm_of thy x)
+ |> Thm.forall_elim (cterm_of thy lhs)
|> curry (op RS) refl
in
(exactly_one, function_value)
@@ -396,7 +396,7 @@
Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
|> cterm_of thy
- val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
+ val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
|> instantiate' [] [NONE, SOME (cterm_of thy h)]
@@ -412,17 +412,17 @@
val graph_is_function = complete
|> Thm.forall_elim_vars 0
|> fold (curry op COMP) ex1s
- |> implies_intr (ihyp)
- |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
- |> forall_intr (cterm_of thy x)
+ |> Thm.implies_intr (ihyp)
+ |> Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+ |> Thm.forall_intr (cterm_of thy x)
|> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
- |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+ |> (fn it => fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
val goalstate = Conjunction.intr graph_is_function complete
|> Thm.close_derivation
|> Goal.protect
- |> fold_rev (implies_intr o cprop_of) compat
- |> implies_intr (cprop_of complete)
+ |> fold_rev (Thm.implies_intr o cprop_of) compat
+ |> Thm.implies_intr (cprop_of complete)
in
(goalstate, values)
end
@@ -544,7 +544,7 @@
let
fun inst_term t = subst_bound(f, abstract_over (fvar, t))
in
- (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+ (rcfix, map (Thm.assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
end
@@ -562,14 +562,14 @@
val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
in
- ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
+ ((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
|> (fn it => it COMP graph_is_function)
- |> implies_intr z_smaller
- |> forall_intr (cterm_of thy z)
+ |> Thm.implies_intr z_smaller
+ |> Thm.forall_intr (cterm_of thy z)
|> (fn it => it COMP valthm)
- |> implies_intr lhs_acc
+ |> Thm.implies_intr lhs_acc
|> asm_simplify (HOL_basic_ss addsimps [f_iff])
- |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev (Thm.implies_intr o cprop_of) ags
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
end
in
@@ -588,7 +588,7 @@
val Globals {domT, x, z, a, P, D, ...} = globals
val acc_R = mk_acc domT R
- val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+ val x_D = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
val D_subset = cterm_of thy (Logic.all x
@@ -606,7 +606,7 @@
HOLogic.mk_Trueprop (P $ Bound 0)))
|> cterm_of thy
- val aihyp = assume ihyp
+ val aihyp = Thm.assume ihyp
fun prove_case clause =
let
@@ -617,15 +617,15 @@
local open Conv in
val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
val sih =
- fconv_rule (More_Conv.binder_conv
+ fconv_rule (Conv.binder_conv
(K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
end
fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
- |> forall_elim (cterm_of thy rcarg)
+ |> Thm.forall_elim (cterm_of thy rcarg)
|> Thm.elim_implies llRI
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+ |> fold_rev (Thm.implies_intr o cprop_of) CCas
+ |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
@@ -636,19 +636,19 @@
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|> cterm_of thy
- val P_lhs = assume step
- |> fold forall_elim cqs
+ val P_lhs = Thm.assume step
+ |> fold Thm.forall_elim cqs
|> Thm.elim_implies lhs_D
|> fold Thm.elim_implies ags
|> fold Thm.elim_implies P_recs
val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
|> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
- |> symmetric (* P lhs == P x *)
- |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
+ |> Thm.symmetric (* P lhs == P x *)
+ |> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
+ |> Thm.implies_intr (cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> fold_rev Thm.forall_intr cqs
in
(res, step)
end
@@ -658,33 +658,33 @@
val istep = complete_thm
|> Thm.forall_elim_vars 0
|> fold (curry op COMP) cases (* P x *)
- |> implies_intr ihyp
- |> implies_intr (cprop_of x_D)
- |> forall_intr (cterm_of thy x)
+ |> Thm.implies_intr ihyp
+ |> Thm.implies_intr (cprop_of x_D)
+ |> Thm.forall_intr (cterm_of thy x)
val subset_induct_rule =
acc_subset_induct
- |> (curry op COMP) (assume D_subset)
- |> (curry op COMP) (assume D_dcl)
- |> (curry op COMP) (assume a_D)
+ |> (curry op COMP) (Thm.assume D_subset)
+ |> (curry op COMP) (Thm.assume D_dcl)
+ |> (curry op COMP) (Thm.assume a_D)
|> (curry op COMP) istep
- |> fold_rev implies_intr steps
- |> implies_intr a_D
- |> implies_intr D_dcl
- |> implies_intr D_subset
+ |> fold_rev Thm.implies_intr steps
+ |> Thm.implies_intr a_D
+ |> Thm.implies_intr D_dcl
+ |> Thm.implies_intr D_subset
val simple_induct_rule =
subset_induct_rule
- |> forall_intr (cterm_of thy D)
- |> forall_elim (cterm_of thy acc_R)
+ |> Thm.forall_intr (cterm_of thy D)
+ |> Thm.forall_elim (cterm_of thy acc_R)
|> assume_tac 1 |> Seq.hd
|> (curry op COMP) (acc_downward
|> (instantiate' [SOME (ctyp_of thy domT)]
(map (SOME o cterm_of thy) [R, x, z]))
- |> forall_intr (cterm_of thy z)
- |> forall_intr (cterm_of thy x))
- |> forall_intr (cterm_of thy a)
- |> forall_intr (cterm_of thy P)
+ |> Thm.forall_intr (cterm_of thy z)
+ |> Thm.forall_intr (cterm_of thy x))
+ |> Thm.forall_intr (cterm_of thy a)
+ |> Thm.forall_intr (cterm_of thy P)
in
simple_induct_rule
end
@@ -736,18 +736,19 @@
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|> cterm_of thy
- val thm = assume hyp
- |> fold forall_elim cqs
+ val thm = Thm.assume hyp
+ |> fold Thm.forall_elim cqs
|> fold Thm.elim_implies ags
|> Function_Ctx_Tree.import_thm thy (fixes, assumes)
|> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
- |> cterm_of thy |> assume
+ |> cterm_of thy |> Thm.assume
val acc = thm COMP ih_case
val z_acc_local = acc
- |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
+ |> Conv.fconv_rule
+ (Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
val ethm = z_acc_local
|> Function_Ctx_Tree.export_thm thy (fixes,
@@ -785,34 +786,34 @@
HOLogic.mk_Trueprop (acc_R $ Bound 0)))
|> cterm_of thy
- val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
+ val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
val (hyps, cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([], [])
in
R_cases
- |> forall_elim (cterm_of thy z)
- |> forall_elim (cterm_of thy x)
- |> forall_elim (cterm_of thy (acc_R $ z))
- |> curry op COMP (assume R_z_x)
+ |> Thm.forall_elim (cterm_of thy z)
+ |> Thm.forall_elim (cterm_of thy x)
+ |> Thm.forall_elim (cterm_of thy (acc_R $ z))
+ |> curry op COMP (Thm.assume R_z_x)
|> fold_rev (curry op COMP) cases
- |> implies_intr R_z_x
- |> forall_intr (cterm_of thy z)
+ |> Thm.implies_intr R_z_x
+ |> Thm.forall_intr (cterm_of thy z)
|> (fn it => it COMP accI)
- |> implies_intr ihyp
- |> forall_intr (cterm_of thy x)
+ |> Thm.implies_intr ihyp
+ |> Thm.forall_intr (cterm_of thy x)
|> (fn it => Drule.compose_single(it,2,wf_induct_rule))
- |> curry op RS (assume wfR')
+ |> curry op RS (Thm.assume wfR')
|> forall_intr_vars
|> (fn it => it COMP allI)
- |> fold implies_intr hyps
- |> implies_intr wfR'
- |> forall_intr (cterm_of thy R')
- |> forall_elim (cterm_of thy (inrel_R))
+ |> fold Thm.implies_intr hyps
+ |> Thm.implies_intr wfR'
+ |> Thm.forall_intr (cterm_of thy R')
+ |> Thm.forall_elim (cterm_of thy (inrel_R))
|> curry op RS wf_in_rel
|> full_simplify (HOL_basic_ss addsimps [in_rel_def])
- |> forall_intr (cterm_of thy Rrel)
+ |> Thm.forall_intr (cterm_of thy Rrel)
end
@@ -919,11 +920,11 @@
RCss GIntro_thms) RIntro_thmss
val complete =
- mk_completeness globals clauses abstract_qglrs |> cert |> assume
+ mk_completeness globals clauses abstract_qglrs |> cert |> Thm.assume
val compat =
mk_compat_proof_obligations domT ranT fvar f abstract_qglrs
- |> map (cert #> assume)
+ |> map (cert #> Thm.assume)
val compat_store = store_compat_thms n compat
--- a/src/HOL/Tools/Function/function_lib.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/function_lib.ML Sun May 23 10:38:11 2010 +0100
@@ -5,7 +5,7 @@
Some fairly general functions that should probably go somewhere else...
*)
-structure Function_Lib =
+structure Function_Lib = (* FIXME proper signature *)
struct
fun map_option f NONE = NONE
@@ -104,7 +104,7 @@
fun forall_intr_rename (n, cv) thm =
let
- val allthm = forall_intr cv thm
+ val allthm = Thm.forall_intr cv thm
val (_ $ abs) = prop_of allthm
in
Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
--- a/src/HOL/Tools/Function/induction_schema.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/induction_schema.ML Sun May 23 10:38:11 2010 +0100
@@ -226,7 +226,7 @@
HOLogic.mk_Trueprop (P_comp $ Bound 0)))
|> cert
- val aihyp = assume ihyp
+ val aihyp = Thm.assume ihyp
(* Rule for case splitting along the sum types *)
val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
@@ -237,9 +237,9 @@
fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
let
val fxs = map Free xs
- val branch_hyp = assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat))))
+ val branch_hyp = Thm.assume (cert (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat))))
- val C_hyps = map (cert #> assume) Cs
+ val C_hyps = map (cert #> Thm.assume) Cs
val (relevant_cases, ineqss') =
(scases_idx ~~ ineqss)
@@ -248,32 +248,33 @@
fun prove_case (cidx, SchemeCase {qs, gs, lhs, rs, ...}) ineq_press =
let
- val case_hyps = map (assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs)
+ val case_hyps =
+ map (Thm.assume o cert o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs)
val cqs = map (cert o Free) qs
- val ags = map (assume o cert) gs
+ val ags = map (Thm.assume o cert) gs
val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps)
val sih = full_simplify replace_x_ss aihyp
fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) =
let
- val cGas = map (assume o cert) Gas
+ val cGas = map (Thm.assume o cert) Gas
val cGvs = map (cert o Free) Gvs
- val import = fold forall_elim (cqs @ cGvs)
+ val import = fold Thm.forall_elim (cqs @ cGvs)
#> fold Thm.elim_implies (ags @ cGas)
val ipres = pres
- |> forall_elim (cert (list_comb (P_of idx, rcargs)))
+ |> Thm.forall_elim (cert (list_comb (P_of idx, rcargs)))
|> import
in
sih
- |> forall_elim (cert (inject idx rcargs))
+ |> Thm.forall_elim (cert (inject idx rcargs))
|> Thm.elim_implies (import ineq) (* Psum rcargs *)
|> Conv.fconv_rule sum_prod_conv
|> Conv.fconv_rule ind_rulify
|> (fn th => th COMP ipres) (* P rs *)
- |> fold_rev (implies_intr o cprop_of) cGas
- |> fold_rev forall_intr cGvs
+ |> fold_rev (Thm.implies_intr o cprop_of) cGas
+ |> fold_rev Thm.forall_intr cGvs
end
val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *)
@@ -288,13 +289,13 @@
foldl1 (uncurry Conv.combination_conv)
(Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps)
- val res = assume step
- |> fold forall_elim cqs
+ val res = Thm.assume step
+ |> fold Thm.forall_elim cqs
|> fold Thm.elim_implies ags
|> fold Thm.elim_implies P_recs (* P lhs *)
|> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *)
- |> fold_rev (implies_intr o cprop_of) (ags @ case_hyps)
- |> fold_rev forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
+ |> fold_rev (Thm.implies_intr o cprop_of) (ags @ case_hyps)
+ |> fold_rev Thm.forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
in
(res, (cidx, step))
end
@@ -302,12 +303,12 @@
val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss')
val bstep = complete_thm
- |> forall_elim (cert (list_comb (P, fxs)))
- |> fold (forall_elim o cert) (fxs @ map Free ws)
+ |> Thm.forall_elim (cert (list_comb (P, fxs)))
+ |> fold (Thm.forall_elim o cert) (fxs @ map Free ws)
|> fold Thm.elim_implies C_hyps
|> fold Thm.elim_implies cases (* P xs *)
- |> fold_rev (implies_intr o cprop_of) C_hyps
- |> fold_rev (forall_intr o cert o Free) ws
+ |> fold_rev (Thm.implies_intr o cprop_of) C_hyps
+ |> fold_rev (Thm.forall_intr o cert o Free) ws
val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x))
|> Goal.init
@@ -316,8 +317,8 @@
|> Seq.hd
|> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
|> Goal.finish ctxt
- |> implies_intr (cprop_of branch_hyp)
- |> fold_rev (forall_intr o cert) fxs
+ |> Thm.implies_intr (cprop_of branch_hyp)
+ |> fold_rev (Thm.forall_intr o cert) fxs
in
(Pxs, steps)
end
@@ -328,8 +329,8 @@
val istep = sum_split_rule
|> fold (fn b => fn th => Drule.compose_single (b, 1, th)) branches
- |> implies_intr ihyp
- |> forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *)
+ |> Thm.implies_intr ihyp
+ |> Thm.forall_intr (cert x) (* "!!x. (!!y<x. P y) ==> P x" *)
val induct_rule =
@{thm "wf_induct_rule"}
@@ -353,10 +354,10 @@
val cert = cterm_of (ProofContext.theory_of ctxt)
val ineqss = mk_ineqs R scheme
- |> map (map (pairself (assume o cert)))
+ |> map (map (pairself (Thm.assume o cert)))
val complete =
- map_range (mk_completeness ctxt scheme #> cert #> assume) (length branches)
- val wf_thm = mk_wf R scheme |> cert |> assume
+ map_range (mk_completeness ctxt scheme #> cert #> Thm.assume) (length branches)
+ val wf_thm = mk_wf R scheme |> cert |> Thm.assume
val (descent, pres) = split_list (flat ineqss)
val newgoals = complete @ pres @ wf_thm :: descent
@@ -377,7 +378,7 @@
end
val res = Conjunction.intr_balanced (map_index project branches)
- |> fold_rev implies_intr (map cprop_of newgoals @ steps)
+ |> fold_rev Thm.implies_intr (map cprop_of newgoals @ steps)
|> Drule.generalize ([], [Rn])
val nbranches = length branches
--- a/src/HOL/Tools/Function/mutual.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/mutual.ML Sun May 23 10:38:11 2010 +0100
@@ -164,12 +164,12 @@
val rhs = inst pre_rhs
val cqs = map (cterm_of thy) qs
- val ags = map (assume o cterm_of thy) gs
+ val ags = map (Thm.assume o cterm_of thy) gs
- val import = fold forall_elim cqs
+ val import = fold Thm.forall_elim cqs
#> fold Thm.elim_implies ags
- val export = fold_rev (implies_intr o cprop_of) ags
+ val export = fold_rev (Thm.implies_intr o cprop_of) ags
#> fold_rev forall_intr_rename (oqnames ~~ cqs)
in
F ctxt (f, qs, gs, args, rhs) import export
@@ -184,7 +184,7 @@
val (simp, restore_cond) =
case cprems_of psimp of
[] => (psimp, I)
- | [cond] => (implies_elim psimp (assume cond), implies_intr cond)
+ | [cond] => (Thm.implies_elim psimp (Thm.assume cond), Thm.implies_intr cond)
| _ => sys_error "Too many conditions"
in
@@ -202,9 +202,9 @@
val thy = ProofContext.theory_of ctxt
val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
in
- fold (fn x => fn thm => combination thm (reflexive x)) xs thm
+ fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm
|> Conv.fconv_rule (Thm.beta_conversion true)
- |> fold_rev forall_intr xs
+ |> fold_rev Thm.forall_intr xs
|> Thm.forall_elim_vars 0
end
@@ -228,7 +228,7 @@
val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
val induct_inst =
- forall_elim (cert case_exp) induct
+ Thm.forall_elim (cert case_exp) induct
|> full_simplify SumTree.sumcase_split_ss
|> full_simplify (HOL_basic_ss addsimps all_f_defs)
@@ -238,9 +238,9 @@
val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
in
(rule
- |> forall_elim (cert inj)
+ |> Thm.forall_elim (cert inj)
|> full_simplify SumTree.sumcase_split_ss
- |> fold_rev (forall_intr o cert) (afs @ newPs),
+ |> fold_rev (Thm.forall_intr o cert) (afs @ newPs),
k + length cargTs)
end
in
@@ -255,7 +255,7 @@
val (all_f_defs, fs) =
map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
- (mk_applied_form lthy cargTs (symmetric f_def), f))
+ (mk_applied_form lthy cargTs (Thm.symmetric f_def), f))
parts
|> split_list
--- a/src/HOL/Tools/Function/pat_completeness.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Function/pat_completeness.ML Sun May 23 10:38:11 2010 +0100
@@ -24,7 +24,7 @@
fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
-fun inst_free var inst = forall_elim inst o forall_intr var
+fun inst_free var inst = Thm.forall_elim inst o Thm.forall_intr var
fun inst_case_thm thy x P thm =
let val [Pv, xv] = Term.add_vars (prop_of thm) []
@@ -69,10 +69,10 @@
let
val (_, subps) = strip_comb pat
val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
- val c_eq_pat = simplify (HOL_basic_ss addsimps (map assume eqs)) c_assum
+ val c_eq_pat = simplify (HOL_basic_ss addsimps (map Thm.assume eqs)) c_assum
in
(subps @ pats,
- fold_rev implies_intr eqs (implies_elim thm c_eq_pat))
+ fold_rev Thm.implies_intr eqs (Thm.implies_elim thm c_eq_pat))
end
@@ -82,12 +82,12 @@
let
val (avars, pvars, newidx) = invent_vars cons idx
val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
- val c_assum = assume c_hyp
+ val c_assum = Thm.assume c_hyp
val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
in
o_alg thy P newidx (avars @ vs) newpats
- |> implies_intr c_hyp
- |> fold_rev (forall_intr o cterm_of thy) avars
+ |> Thm.implies_intr c_hyp
+ |> fold_rev (Thm.forall_intr o cterm_of thy) avars
end
| constr_case _ _ _ _ _ _ = raise Match
and o_alg thy P idx [] (([], Pthm) :: _) = Pthm
@@ -119,11 +119,11 @@
|> cterm_of thy
val hyps = map2 mk_assum qss patss
- fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
+ fun inst_hyps hyp qs = fold (Thm.forall_elim o cterm_of thy) qs (Thm.assume hyp)
val assums = map2 inst_hyps hyps qss
in
o_alg thy P 2 xs (patss ~~ assums)
- |> fold_rev implies_intr hyps
+ |> fold_rev Thm.implies_intr hyps
end
fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>
@@ -147,7 +147,7 @@
val patss = map (map snd) x_pats
val complete_thm = prove_completeness thy xs thesis qss patss
- |> fold_rev (forall_intr o cterm_of thy) vs
+ |> fold_rev (Thm.forall_intr o cterm_of thy) vs
in
PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st))
end
--- a/src/HOL/Tools/Nitpick/kodkod.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Sun May 23 10:38:11 2010 +0100
@@ -1034,6 +1034,7 @@
handle IO.Io _ => "" | OS.SysErr _ => "")
|> perhaps (try (unsuffix "\r"))
|> perhaps (try (unsuffix "."))
+ |> perhaps (try (unprefix "Solve error: "))
|> perhaps (try (unprefix "Error: "))
in
if null ps then
--- a/src/HOL/Tools/Nitpick/kodkod_sat.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod_sat.ML Sun May 23 10:38:11 2010 +0100
@@ -47,9 +47,7 @@
("MiniSat_JNI", Internal (JNI, Incremental, ["MiniSat"])),
("zChaff_JNI", Internal (JNI, Batch, ["zChaff"])),
("SAT4J", Internal (Java, Incremental, ["DefaultSAT4J"])),
- ("SAT4J_Light", Internal (Java, Incremental, ["LightSAT4J"])),
- ("HaifaSat", ExternalV2 (ToStdout, "HAIFASAT_HOME", "HaifaSat", ["-p", "1"],
- "s SATISFIABLE", "v ", "s UNSATISFIABLE"))]
+ ("SAT4J_Light", Internal (Java, Incremental, ["LightSAT4J"]))]
fun dynamic_entry_for_external name dev home exec args markers =
case getenv home of
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Sun May 23 10:38:11 2010 +0100
@@ -24,8 +24,6 @@
open Nitpick_Nut
open Nitpick
-structure K = OuterKeyword and P = OuterParse
-
val auto = Unsynchronized.ref false;
val _ =
@@ -289,14 +287,14 @@
extract_params (ProofContext.init_global thy) false (default_raw_params thy)
o map (apsnd single)
-val parse_key = Scan.repeat1 P.typ_group >> space_implode " "
+val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " "
val parse_value =
- Scan.repeat1 (P.minus >> single
- || Scan.repeat1 (Scan.unless P.minus P.name)
- || P.$$$ "," |-- P.number >> prefix "," >> single) >> flat
-val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) []
+ Scan.repeat1 (Parse.minus >> single
+ || Scan.repeat1 (Scan.unless Parse.minus Parse.name)
+ || Parse.$$$ "," |-- Parse.number >> prefix "," >> single) >> flat
+val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) []
val parse_params =
- Scan.optional (P.$$$ "[" |-- P.list parse_param --| P.$$$ "]") []
+ Scan.optional (Parse.$$$ "[" |-- Parse.list parse_param --| Parse.$$$ "]") []
fun handle_exceptions ctxt f x =
f x
@@ -375,15 +373,15 @@
|> sort_strings |> cat_lines)))))
val parse_nitpick_command =
- (parse_params -- Scan.optional P.nat 1) #>> nitpick_trans
+ (parse_params -- Scan.optional Parse.nat 1) #>> nitpick_trans
val parse_nitpick_params_command = parse_params #>> nitpick_params_trans
-val _ = OuterSyntax.improper_command "nitpick"
+val _ = Outer_Syntax.improper_command "nitpick"
"try to find a counterexample for a given subgoal using Nitpick"
- K.diag parse_nitpick_command
-val _ = OuterSyntax.command "nitpick_params"
+ Keyword.diag parse_nitpick_command
+val _ = Outer_Syntax.command "nitpick_params"
"set and display the default parameters for Nitpick"
- K.thy_decl parse_nitpick_params_command
+ Keyword.thy_decl parse_nitpick_params_command
fun auto_nitpick state =
if not (!auto) then (false, state) else pick_nits [] true 1 0 state
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun May 23 10:38:11 2010 +0100
@@ -74,7 +74,8 @@
end
fun preprocess_strong_conn_constnames options gr ts thy =
- if forall (fn (Const (c, _)) => Predicate_Compile_Core.is_registered thy c) ts then
+ if forall (fn (Const (c, _)) =>
+ Predicate_Compile_Core.is_registered (ProofContext.init_global thy) c) ts then
thy
else
let
@@ -192,8 +193,6 @@
val setup = Predicate_Compile_Core.setup
-local structure P = OuterParse
-in
(* Parser for mode annotations *)
@@ -207,15 +206,15 @@
(parse_mode_tuple_expr --| Args.$$$ "=>" -- parse_mode_expr >> Fun || parse_mode_tuple_expr) xs
val mode_and_opt_proposal = parse_mode_expr --
- Scan.optional (Args.$$$ "as" |-- P.xname >> SOME) NONE
+ Scan.optional (Args.$$$ "as" |-- Parse.xname >> SOME) NONE
val opt_modes =
- Scan.optional (P.$$$ "(" |-- Args.$$$ "modes" |-- P.$$$ ":" |--
- P.enum "," mode_and_opt_proposal --| P.$$$ ")" >> SOME) NONE
+ Scan.optional (Parse.$$$ "(" |-- Args.$$$ "modes" |-- Parse.$$$ ":" |--
+ Parse.enum "," mode_and_opt_proposal --| Parse.$$$ ")" >> SOME) NONE
val opt_expected_modes =
- Scan.optional (P.$$$ "(" |-- Args.$$$ "expected_modes" |-- P.$$$ ":" |--
- P.enum "," parse_mode_expr --| P.$$$ ")" >> SOME) NONE
+ Scan.optional (Parse.$$$ "(" |-- Args.$$$ "expected_modes" |-- Parse.$$$ ":" |--
+ Parse.enum "," parse_mode_expr --| Parse.$$$ ")" >> SOME) NONE
(* Parser for options *)
@@ -224,46 +223,46 @@
val scan_bool_option = foldl1 (op ||) (map Args.$$$ bool_options)
val scan_compilation = foldl1 (op ||) (map (fn (s, c) => Args.$$$ s >> K c) compilation_names)
in
- Scan.optional (P.$$$ "[" |-- Scan.optional scan_compilation Pred
- -- P.enum "," scan_bool_option --| P.$$$ "]")
+ Scan.optional (Parse.$$$ "[" |-- Scan.optional scan_compilation Pred
+ -- Parse.enum "," scan_bool_option --| Parse.$$$ "]")
(Pred, [])
end
-val opt_print_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
+val opt_print_modes =
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
-val opt_mode = (P.$$$ "_" >> K NONE) || (parse_mode_expr >> SOME)
+val opt_mode = (Parse.$$$ "_" >> K NONE) || (parse_mode_expr >> SOME)
-val opt_param_modes = Scan.optional (P.$$$ "[" |-- Args.$$$ "mode" |-- P.$$$ ":" |--
- P.enum ", " opt_mode --| P.$$$ "]" >> SOME) NONE
+val opt_param_modes = Scan.optional (Parse.$$$ "[" |-- Args.$$$ "mode" |-- Parse.$$$ ":" |--
+ Parse.enum ", " opt_mode --| Parse.$$$ "]" >> SOME) NONE
val stats = Scan.optional (Args.$$$ "stats" >> K true) false
val value_options =
let
- val expected_values = Scan.optional (Args.$$$ "expected" |-- P.term >> SOME) NONE
+ val expected_values = Scan.optional (Args.$$$ "expected" |-- Parse.term >> SOME) NONE
val scan_compilation =
Scan.optional
(foldl1 (op ||)
- (map (fn (s, c) => Args.$$$ s -- P.enum "," P.int >> (fn (_, ps) => (c, ps)))
+ (map (fn (s, c) => Args.$$$ s -- Parse.enum "," Parse.int >> (fn (_, ps) => (c, ps)))
compilation_names))
(Pred, [])
in
- Scan.optional (P.$$$ "[" |-- (expected_values -- stats) -- scan_compilation --| P.$$$ "]")
+ Scan.optional
+ (Parse.$$$ "[" |-- (expected_values -- stats) -- scan_compilation --| Parse.$$$ "]")
((NONE, false), (Pred, []))
end
(* code_pred command and values command *)
-val _ = OuterSyntax.local_theory_to_proof "code_pred"
+val _ = Outer_Syntax.local_theory_to_proof "code_pred"
"prove equations for predicate specified by intro/elim rules"
- OuterKeyword.thy_goal
- (opt_expected_modes -- opt_modes -- scan_options -- P.term_group >> code_pred_cmd)
+ Keyword.thy_goal
+ (opt_expected_modes -- opt_modes -- scan_options -- Parse.term_group >> code_pred_cmd)
-val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag
- (opt_print_modes -- opt_param_modes -- value_options -- Scan.optional P.nat ~1 -- P.term
+val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
+ (opt_print_modes -- opt_param_modes -- value_options -- Scan.optional Parse.nat ~1 -- Parse.term
>> (fn ((((print_modes, param_modes), options), k), t) => Toplevel.keep
(Predicate_Compile_Core.values_cmd print_modes param_modes options k t)));
end
-
-end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sun May 23 10:38:11 2010 +0100
@@ -18,27 +18,29 @@
-> ((string option * bool) * (compilation * int list)) -> int -> string -> Toplevel.state -> unit
val register_predicate : (string * thm list * thm) -> theory -> theory
val register_intros : string * thm list -> theory -> theory
- val is_registered : theory -> string -> bool
- val function_name_of : compilation -> theory -> string -> mode -> string
+ val is_registered : Proof.context -> string -> bool
+ val function_name_of : compilation -> Proof.context -> string -> mode -> string
val predfun_intro_of: Proof.context -> string -> mode -> thm
val predfun_elim_of: Proof.context -> string -> mode -> thm
- val all_preds_of : theory -> string list
- val modes_of: compilation -> theory -> string -> mode list
- val all_modes_of : compilation -> theory -> (string * mode list) list
- val all_random_modes_of : theory -> (string * mode list) list
- val intros_of : theory -> string -> thm list
+ val all_preds_of : Proof.context -> string list
+ val modes_of: compilation -> Proof.context -> string -> mode list
+ val all_modes_of : compilation -> Proof.context -> (string * mode list) list
+ val all_random_modes_of : Proof.context -> (string * mode list) list
+ val intros_of : Proof.context -> string -> thm list
val add_intro : thm -> theory -> theory
val set_elim : thm -> theory -> theory
val register_alternative_function : string -> mode -> string -> theory -> theory
- val alternative_compilation_of : theory -> string -> mode ->
+ val alternative_compilation_of_global : theory -> string -> mode ->
+ (compilation_funs -> typ -> term) option
+ val alternative_compilation_of : Proof.context -> string -> mode ->
(compilation_funs -> typ -> term) option
val functional_compilation : string -> mode -> compilation_funs -> typ -> term
val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory
val force_modes_and_compilations : string ->
(mode * ((compilation_funs -> typ -> term) * bool)) list -> theory -> theory
val preprocess_intro : theory -> thm -> thm
- val print_stored_rules : theory -> unit
- val print_all_modes : compilation -> theory -> unit
+ val print_stored_rules : Proof.context -> unit
+ val print_all_modes : compilation -> Proof.context -> unit
val mk_casesrule : Proof.context -> term -> thm list -> term
val eval_ref : (unit -> term Predicate.pred) option Unsynchronized.ref
val random_eval_ref : (unit -> int * int -> term Predicate.pred * (int * int))
@@ -228,77 +230,81 @@
(* queries *)
-fun lookup_pred_data thy name =
- Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
+fun lookup_pred_data ctxt name =
+ Option.map rep_pred_data (try (Graph.get_node (PredData.get (ProofContext.theory_of ctxt))) name)
-fun the_pred_data thy name = case lookup_pred_data thy name
+fun the_pred_data ctxt name = case lookup_pred_data ctxt name
of NONE => error ("No such predicate " ^ quote name)
| SOME data => data;
val is_registered = is_some oo lookup_pred_data
-val all_preds_of = Graph.keys o PredData.get
+val all_preds_of = Graph.keys o PredData.get o ProofContext.theory_of
-fun intros_of thy = map (Thm.transfer thy) o #intros o the_pred_data thy
+val intros_of = #intros oo the_pred_data
-fun the_elim_of thy name = case #elim (the_pred_data thy name)
+fun the_elim_of ctxt name = case #elim (the_pred_data ctxt name)
of NONE => error ("No elimination rule for predicate " ^ quote name)
- | SOME thm => Thm.transfer thy thm
+ | SOME thm => thm
-val has_elim = is_some o #elim oo the_pred_data;
+val has_elim = is_some o #elim oo the_pred_data
-fun function_names_of compilation thy name =
- case AList.lookup (op =) (#function_names (the_pred_data thy name)) compilation of
+fun function_names_of compilation ctxt name =
+ case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation of
NONE => error ("No " ^ string_of_compilation compilation
^ "functions defined for predicate " ^ quote name)
| SOME fun_names => fun_names
-fun function_name_of compilation thy name mode =
+fun function_name_of compilation ctxt name mode =
case AList.lookup eq_mode
- (function_names_of compilation thy name) mode of
+ (function_names_of compilation ctxt name) mode of
NONE => error ("No " ^ string_of_compilation compilation
^ " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name)
| SOME function_name => function_name
-fun modes_of compilation thy name = map fst (function_names_of compilation thy name)
+fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name)
-fun all_modes_of compilation thy =
- map_filter (fn name => Option.map (pair name) (try (modes_of compilation thy) name))
- (all_preds_of thy)
+fun all_modes_of compilation ctxt =
+ map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
+ (all_preds_of ctxt)
val all_random_modes_of = all_modes_of Random
-fun defined_functions compilation thy name =
- AList.defined (op =) (#function_names (the_pred_data thy name)) compilation
+fun defined_functions compilation ctxt name = case lookup_pred_data ctxt name of
+ NONE => false
+ | SOME data => AList.defined (op =) (#function_names data) compilation
-fun lookup_predfun_data thy name mode =
+fun needs_random ctxt s m =
+ member (op =) (#needs_random (the_pred_data ctxt s)) m
+
+fun lookup_predfun_data ctxt name mode =
Option.map rep_predfun_data
- (AList.lookup (op =) (#predfun_data (the_pred_data thy name)) mode)
+ (AList.lookup (op =) (#predfun_data (the_pred_data ctxt name)) mode)
-fun the_predfun_data thy name mode =
- case lookup_predfun_data thy name mode of
+fun the_predfun_data ctxt name mode =
+ case lookup_predfun_data ctxt name mode of
NONE => error ("No function defined for mode " ^ string_of_mode mode ^
" of predicate " ^ name)
| SOME data => data;
-val predfun_definition_of = #definition ooo the_predfun_data o ProofContext.theory_of
+val predfun_definition_of = #definition ooo the_predfun_data
-val predfun_intro_of = #intro ooo the_predfun_data o ProofContext.theory_of
+val predfun_intro_of = #intro ooo the_predfun_data
-val predfun_elim_of = #elim ooo the_predfun_data o ProofContext.theory_of
+val predfun_elim_of = #elim ooo the_predfun_data
-val predfun_neg_intro_of = #neg_intro ooo the_predfun_data o ProofContext.theory_of
+val predfun_neg_intro_of = #neg_intro ooo the_predfun_data
(* diagnostic display functions *)
-fun print_modes options thy modes =
+fun print_modes options modes =
if show_modes options then
tracing ("Inferred modes:\n" ^
cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
(fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
else ()
-fun print_pred_mode_table string_of_entry thy pred_mode_table =
+fun print_pred_mode_table string_of_entry pred_mode_table =
let
fun print_mode pred ((pol, mode), entry) = "mode : " ^ string_of_mode mode
^ string_of_entry pred mode entry
@@ -307,35 +313,35 @@
val _ = tracing (cat_lines (map print_pred pred_mode_table))
in () end;
-fun string_of_prem thy (Prem t) =
- (Syntax.string_of_term_global thy t) ^ "(premise)"
- | string_of_prem thy (Negprem t) =
- (Syntax.string_of_term_global thy (HOLogic.mk_not t)) ^ "(negative premise)"
- | string_of_prem thy (Sidecond t) =
- (Syntax.string_of_term_global thy t) ^ "(sidecondition)"
- | string_of_prem thy _ = raise Fail "string_of_prem: unexpected input"
+fun string_of_prem ctxt (Prem t) =
+ (Syntax.string_of_term ctxt t) ^ "(premise)"
+ | string_of_prem ctxt (Negprem t) =
+ (Syntax.string_of_term ctxt (HOLogic.mk_not t)) ^ "(negative premise)"
+ | string_of_prem ctxt (Sidecond t) =
+ (Syntax.string_of_term ctxt t) ^ "(sidecondition)"
+ | string_of_prem ctxt _ = raise Fail "string_of_prem: unexpected input"
-fun string_of_clause thy pred (ts, prems) =
+fun string_of_clause ctxt pred (ts, prems) =
(space_implode " --> "
- (map (string_of_prem thy) prems)) ^ " --> " ^ pred ^ " "
- ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))
+ (map (string_of_prem ctxt) prems)) ^ " --> " ^ pred ^ " "
+ ^ (space_implode " " (map (Syntax.string_of_term ctxt) ts))
-fun print_compiled_terms options thy =
+fun print_compiled_terms options ctxt =
if show_compilation options then
- print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
+ print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term ctxt)
else K ()
-fun print_stored_rules thy =
+fun print_stored_rules ctxt =
let
- val preds = (Graph.keys o PredData.get) thy
+ val preds = Graph.keys (PredData.get (ProofContext.theory_of ctxt))
fun print pred () = let
val _ = writeln ("predicate: " ^ pred)
val _ = writeln ("introrules: ")
- val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
- (rev (intros_of thy pred)) ()
+ val _ = fold (fn thm => fn u => writeln (Display.string_of_thm ctxt thm))
+ (rev (intros_of ctxt pred)) ()
in
- if (has_elim thy pred) then
- writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
+ if (has_elim ctxt pred) then
+ writeln ("elimrule: " ^ Display.string_of_thm ctxt (the_elim_of ctxt pred))
else
writeln ("no elimrule defined")
end
@@ -343,7 +349,7 @@
fold print preds ()
end;
-fun print_all_modes compilation thy =
+fun print_all_modes compilation ctxt =
let
val _ = writeln ("Inferred modes:")
fun print (pred, modes) u =
@@ -352,7 +358,7 @@
val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
in u end
in
- fold print (all_modes_of compilation thy) ()
+ fold print (all_modes_of compilation ctxt) ()
end
(* validity checks *)
@@ -575,12 +581,12 @@
(Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
(Thm.transfer thy rule)
-fun preprocess_elim thy elimrule =
+fun preprocess_elim ctxt elimrule =
let
fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
| replace_eqs t = t
- val ctxt = ProofContext.init_global thy
+ val thy = ProofContext.theory_of ctxt
val ((_, [elimrule]), ctxt') = Variable.import false [elimrule] ctxt
val prems = Thm.prems_of elimrule
val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems))))
@@ -607,21 +613,21 @@
val no_compilation = ([], ([], []))
-fun fetch_pred_data thy name =
- case try (Inductive.the_inductive (ProofContext.init_global thy)) name of
+fun fetch_pred_data ctxt name =
+ case try (Inductive.the_inductive ctxt) name of
SOME (info as (_, result)) =>
let
fun is_intro_of intro =
let
val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
- in (fst (dest_Const const) = name) end;
+ in (fst (dest_Const const) = name) end;
+ val thy = ProofContext.theory_of ctxt
val intros =
(map (expand_tuples thy #> preprocess_intro thy) (filter is_intro_of (#intrs result)))
val index = find_index (fn s => s = name) (#names (fst info))
val pre_elim = nth (#elims result) index
val pred = nth (#preds result) index
val nparams = length (Inductive.params_of (#raw_induct result))
- val ctxt = ProofContext.init_global thy
val elim_t = mk_casesrule ctxt pred intros
val elim =
prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t
@@ -635,16 +641,16 @@
val add = (apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data))
in PredData.map (Graph.map_node name (map_pred_data add)) end
-fun is_inductive_predicate thy name =
- is_some (try (Inductive.the_inductive (ProofContext.init_global thy)) name)
+fun is_inductive_predicate ctxt name =
+ is_some (try (Inductive.the_inductive ctxt) name)
-fun depending_preds_of thy (key, value) =
+fun depending_preds_of ctxt (key, value) =
let
val intros = (#intros o rep_pred_data) value
in
fold Term.add_const_names (map Thm.prop_of intros) []
|> filter (fn c => (not (c = key)) andalso
- (is_inductive_predicate thy c orelse is_registered thy c))
+ (is_inductive_predicate ctxt c orelse is_registered ctxt c))
end;
fun add_intro thm thy =
@@ -667,7 +673,7 @@
fun register_predicate (constname, pre_intros, pre_elim) thy =
let
val intros = map (preprocess_intro thy) pre_intros
- val elim = preprocess_elim thy pre_elim
+ val elim = preprocess_elim (ProofContext.init_global thy) pre_elim
in
if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
PredData.map
@@ -725,9 +731,13 @@
(mode * (compilation_funs -> typ -> term)) list -> bool));
);
-fun alternative_compilation_of thy pred_name mode =
+fun alternative_compilation_of_global thy pred_name mode =
AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode
+fun alternative_compilation_of ctxt pred_name mode =
+ AList.lookup eq_mode
+ (Symtab.lookup_list (Alt_Compilations_Data.get (ProofContext.theory_of ctxt)) pred_name) mode
+
fun force_modes_and_compilations pred_name compilations =
let
(* thm refl is a dummy thm *)
@@ -1157,10 +1167,10 @@
in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
*)
-fun is_possible_output thy vs t =
+fun is_possible_output ctxt vs t =
forall
(fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
- (non_invertible_subterms (ProofContext.init_global thy) t)
+ (non_invertible_subterms ctxt t)
andalso
(forall (is_eqT o snd)
(inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
@@ -1176,7 +1186,7 @@
[]
end
-fun is_constructable thy vs t = forall (member (op =) vs) (term_vs t)
+fun is_constructable vs t = forall (member (op =) vs) (term_vs t)
fun missing_vars vs t = subtract (op =) vs (term_vs t)
@@ -1196,11 +1206,11 @@
SOME ms => SOME (map (fn m => (Context m , [])) ms)
| NONE => NONE)
-fun derivations_of (thy : theory) modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
+fun derivations_of (ctxt : Proof.context) modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
map_product
(fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
- (derivations_of thy modes vs t1 m1) (derivations_of thy modes vs t2 m2)
- | derivations_of thy modes vs t (m as Fun _) =
+ (derivations_of ctxt modes vs t1 m1) (derivations_of ctxt modes vs t2 m2)
+ | derivations_of ctxt modes vs t (m as Fun _) =
(*let
val (p, args) = strip_comb t
in
@@ -1216,37 +1226,37 @@
end) ms
| NONE => (if is_all_input mode then [(Context mode, [])] else []))
end*)
- (case try (all_derivations_of thy modes vs) t of
+ (case try (all_derivations_of ctxt modes vs) t of
SOME derivs =>
filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
| NONE => (if is_all_input m then [(Context m, [])] else []))
- | derivations_of thy modes vs t m =
+ | derivations_of ctxt modes vs t m =
if eq_mode (m, Input) then
[(Term Input, missing_vars vs t)]
else if eq_mode (m, Output) then
- (if is_possible_output thy vs t then [(Term Output, [])] else [])
+ (if is_possible_output ctxt vs t then [(Term Output, [])] else [])
else []
-and all_derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) =
+and all_derivations_of ctxt modes vs (Const ("Pair", _) $ t1 $ t2) =
let
- val derivs1 = all_derivations_of thy modes vs t1
- val derivs2 = all_derivations_of thy modes vs t2
+ val derivs1 = all_derivations_of ctxt modes vs t1
+ val derivs2 = all_derivations_of ctxt modes vs t2
in
map_product
(fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
derivs1 derivs2
end
- | all_derivations_of thy modes vs (t1 $ t2) =
+ | all_derivations_of ctxt modes vs (t1 $ t2) =
let
- val derivs1 = all_derivations_of thy modes vs t1
+ val derivs1 = all_derivations_of ctxt modes vs t1
in
maps (fn (d1, mvars1) =>
case mode_of d1 of
Fun (m', _) => map (fn (d2, mvars2) =>
- (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of thy modes vs t2 m')
+ (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of ctxt modes vs t2 m')
| _ => raise Fail "Something went wrong") derivs1
end
- | all_derivations_of thy modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
- | all_derivations_of thy modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
+ | all_derivations_of _ modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
+ | all_derivations_of _ modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
| all_derivations_of _ modes vs _ = raise Fail "all_derivations_of"
fun rev_option_ord ord (NONE, NONE) = EQUAL
@@ -1287,7 +1297,7 @@
EQUAL => lexl_ord ords' (x, x')
| ord => ord
-fun deriv_ord' thy pol pred modes t1 t2 ((deriv1, mvars1), (deriv2, mvars2)) =
+fun deriv_ord' ctxt pol pred modes t1 t2 ((deriv1, mvars1), (deriv2, mvars2)) =
let
(* prefer functional modes if it is a function *)
fun fun_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
@@ -1295,7 +1305,7 @@
fun is_functional t mode =
case try (fst o dest_Const o fst o strip_comb) t of
NONE => false
- | SOME c => is_some (alternative_compilation_of thy c mode)
+ | SOME c => is_some (alternative_compilation_of ctxt c mode)
in
case (is_functional t1 (head_mode_of deriv1), is_functional t2 (head_mode_of deriv2)) of
(true, true) => EQUAL
@@ -1325,7 +1335,7 @@
ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2))
end
-fun deriv_ord thy pol pred modes t = deriv_ord' thy pol pred modes t t
+fun deriv_ord ctxt pol pred modes t = deriv_ord' ctxt pol pred modes t t
fun premise_ord thy pol pred modes ((prem1, a1), (prem2, a2)) =
rev_option_ord (deriv_ord' thy pol pred modes (dest_indprem prem1) (dest_indprem prem2)) (a1, a2)
@@ -1334,25 +1344,25 @@
tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
-fun select_mode_prem (mode_analysis_options : mode_analysis_options) (thy : theory) pred
+fun select_mode_prem (mode_analysis_options : mode_analysis_options) (ctxt : Proof.context) pred
pol (modes, (pos_modes, neg_modes)) vs ps =
let
fun choose_mode_of_prem (Prem t) = partial_hd
- (sort (deriv_ord thy pol pred modes t) (all_derivations_of thy pos_modes vs t))
+ (sort (deriv_ord ctxt pol pred modes t) (all_derivations_of ctxt pos_modes vs t))
| choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
| choose_mode_of_prem (Negprem t) = partial_hd
- (sort (deriv_ord thy (not pol) pred modes t)
+ (sort (deriv_ord ctxt (not pol) pred modes t)
(filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
- (all_derivations_of thy neg_modes vs t)))
- | choose_mode_of_prem p = raise Fail ("choose_mode_of_prem: " ^ string_of_prem thy p)
+ (all_derivations_of ctxt neg_modes vs t)))
+ | choose_mode_of_prem p = raise Fail ("choose_mode_of_prem: " ^ string_of_prem ctxt p)
in
if #reorder_premises mode_analysis_options then
- partial_hd (sort (premise_ord thy pol pred modes) (ps ~~ map choose_mode_of_prem ps))
+ partial_hd (sort (premise_ord ctxt pol pred modes) (ps ~~ map choose_mode_of_prem ps))
else
SOME (hd ps, choose_mode_of_prem (hd ps))
end
-fun check_mode_clause' (mode_analysis_options : mode_analysis_options) thy pred param_vs (modes :
+fun check_mode_clause' (mode_analysis_options : mode_analysis_options) ctxt pred param_vs (modes :
(string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
let
val vTs = distinct (op =) (fold Term.add_frees (map dest_indprem ps) (fold Term.add_frees ts []))
@@ -1367,7 +1377,7 @@
val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
in (modes, modes) end
val (in_ts, out_ts) = split_mode mode ts
- val in_vs = maps (vars_of_destructable_term (ProofContext.init_global thy)) in_ts
+ val in_vs = maps (vars_of_destructable_term ctxt) in_ts
val out_vs = terms_vs out_ts
fun known_vs_after p vs = (case p of
Prem t => union (op =) vs (term_vs t)
@@ -1377,7 +1387,7 @@
fun check_mode_prems acc_ps rnd vs [] = SOME (acc_ps, vs, rnd)
| check_mode_prems acc_ps rnd vs ps =
(case
- (select_mode_prem mode_analysis_options thy pred pol (modes', (pos_modes', neg_modes')) vs ps) of
+ (select_mode_prem mode_analysis_options ctxt pred pol (modes', (pos_modes', neg_modes')) vs ps) of
SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
(known_vs_after p vs) (filter_out (equal p) ps)
| SOME (p, SOME (deriv, missing_vars)) =>
@@ -1393,7 +1403,7 @@
case check_mode_prems [] false in_vs ps of
NONE => NONE
| SOME (acc_ps, vs, rnd) =>
- if forall (is_constructable thy vs) (in_ts @ out_ts) then
+ if forall (is_constructable vs) (in_ts @ out_ts) then
SOME (ts, rev acc_ps, rnd)
else
if #use_random mode_analysis_options andalso pol then
@@ -1473,11 +1483,12 @@
fun infer_modes mode_analysis_options options compilation preds all_modes param_vs clauses thy =
let
+ val ctxt = ProofContext.init_global thy
val collect_errors = false
fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
- fun needs_random s (false, m) = ((false, m), false)
- | needs_random s (true, m) = ((true, m), member (op =) (#needs_random (the_pred_data thy s)) m)
- fun add_polarity_and_random_bit s b ms = map (fn m => needs_random s (b, m)) ms
+ fun add_needs_random s (false, m) = ((false, m), false)
+ | add_needs_random s (true, m) = ((true, m), needs_random ctxt s m)
+ fun add_polarity_and_random_bit s b ms = map (fn m => add_needs_random s (b, m)) ms
val prednames = map fst preds
(* extramodes contains all modes of all constants, should we only use the necessary ones
- what is the impact on performance? *)
@@ -1492,15 +1503,13 @@
if #infer_pos_and_neg_modes mode_analysis_options then
let
val pos_extra_modes =
- map_filter (fn name => Option.map (pair name) (try (modes_of compilation thy) name))
+ map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
relevant_prednames
- (* all_modes_of compilation thy *)
|> filter_out (fn (name, _) => member (op =) prednames name)
val neg_extra_modes =
map_filter (fn name => Option.map (pair name)
- (try (modes_of (negative_compilation_of compilation) thy) name))
+ (try (modes_of (negative_compilation_of compilation) ctxt) name))
relevant_prednames
- (*all_modes_of (negative_compilation_of compilation) thy*)
|> filter_out (fn (name, _) => member (op =) prednames name)
in
map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
@@ -1509,9 +1518,8 @@
end
else
map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
- (map_filter (fn name => Option.map (pair name) (try (modes_of compilation thy) name))
+ (map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
relevant_prednames
- (*all_modes_of compilation thy*)
|> filter_out (fn (name, _) => member (op =) prednames name))
val _ = print_extra_modes options extra_modes
val start_modes =
@@ -1521,7 +1529,7 @@
else
map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
fun iteration modes = map
- (check_modes_pred' mode_analysis_options options thy param_vs clauses
+ (check_modes_pred' mode_analysis_options options ctxt param_vs clauses
(modes @ extra_modes)) modes
val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
@@ -1532,7 +1540,7 @@
in (modes', errors @ flat new_errors) end) (start_modes, [])
else
(fixp (fn modes => map fst (iteration modes)) start_modes, []))
- val moded_clauses = map (get_modes_pred' mode_analysis_options thy param_vs clauses
+ val moded_clauses = map (get_modes_pred' mode_analysis_options ctxt param_vs clauses
(modes @ extra_modes)) modes
val thy' = fold (fn (s, ms) => if member (op =) (map fst preds) s then
set_needs_random s (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms) else I)
@@ -1671,11 +1679,11 @@
(t, Term Input) => SOME t
| (t, Term Output) => NONE
| (Const (name, T), Context mode) =>
- (case alternative_compilation_of (ProofContext.theory_of ctxt) name mode of
+ (case alternative_compilation_of ctxt name mode of
SOME alt_comp => SOME (alt_comp compfuns T)
| NONE =>
SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
- (ProofContext.theory_of ctxt) name mode,
+ ctxt name mode,
Comp_Mod.funT_of compilation_modifiers mode T)))
| (Free (s, T), Context m) =>
SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T))
@@ -1813,7 +1821,7 @@
(** switch detection analysis **)
-fun find_switch_test thy (i, is) (ts, prems) =
+fun find_switch_test ctxt (i, is) (ts, prems) =
let
val t = nth_pair is (nth ts i)
val T = fastype_of t
@@ -1821,7 +1829,7 @@
case T of
TFree _ => NONE
| Type (Tcon, _) =>
- (case Datatype_Data.get_constrs thy Tcon of
+ (case Datatype_Data.get_constrs (ProofContext.theory_of ctxt) Tcon of
NONE => NONE
| SOME cs =>
(case strip_comb t of
@@ -1830,11 +1838,11 @@
| (Const (c, T), _) => if AList.defined (op =) cs c then SOME (c, T) else NONE))
end
-fun partition_clause thy pos moded_clauses =
+fun partition_clause ctxt pos moded_clauses =
let
fun insert_list eq (key, value) = AList.map_default eq (key, []) (cons value)
fun find_switch_test' moded_clause (cases, left) =
- case find_switch_test thy pos moded_clause of
+ case find_switch_test ctxt pos moded_clause of
SOME (c, T) => (insert_list (op =) ((c, T), moded_clause) cases, left)
| NONE => (cases, moded_clause :: left)
in
@@ -1844,12 +1852,12 @@
datatype switch_tree =
Atom of moded_clause list | Node of (position * ((string * typ) * switch_tree) list) * switch_tree
-fun mk_switch_tree thy mode moded_clauses =
+fun mk_switch_tree ctxt mode moded_clauses =
let
fun select_best_switch moded_clauses input_position best_switch =
let
val ord = option_ord (rev_order o int_ord o (pairself (length o snd o snd)))
- val partition = partition_clause thy input_position moded_clauses
+ val partition = partition_clause ctxt input_position moded_clauses
val switch = if (length (fst partition) > 1) then SOME (input_position, partition) else NONE
in
case ord (switch, best_switch) of LESS => best_switch
@@ -1937,9 +1945,8 @@
(* compilation of predicates *)
-fun compile_pred options compilation_modifiers thy all_vs param_vs s T (pol, mode) moded_cls =
+fun compile_pred options compilation_modifiers ctxt all_vs param_vs s T (pol, mode) moded_cls =
let
- val ctxt = ProofContext.init_global thy
val compilation_modifiers = if pol then compilation_modifiers else
negative_comp_modifiers_of compilation_modifiers
val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
@@ -1967,7 +1974,7 @@
if detect_switches options then
the_default (mk_bot compfuns (HOLogic.mk_tupleT outTs))
(compile_switch compilation_modifiers ctxt all_vs param_vs additional_arguments
- mode in_ts' outTs (mk_switch_tree thy mode moded_cls))
+ mode in_ts' outTs (mk_switch_tree ctxt mode moded_cls))
else
let
val cl_ts =
@@ -1983,7 +1990,7 @@
end
val fun_const =
Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
- (ProofContext.theory_of ctxt) s mode, funT)
+ ctxt s mode, funT)
in
HOLogic.mk_Trueprop
(HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
@@ -2044,7 +2051,7 @@
(Free (x, T), x :: names)
end
-fun create_intro_elim_rule mode defthm mode_id funT pred thy =
+fun create_intro_elim_rule ctxt mode defthm mode_id funT pred =
let
val funtrm = Const (mode_id, funT)
val Ts = binder_types (fastype_of pred)
@@ -2072,11 +2079,11 @@
val simprules = [defthm, @{thm eval_pred},
@{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
- val introthm = Goal.prove (ProofContext.init_global thy)
+ val introthm = Goal.prove ctxt
(argnames @ hoarg_names' @ ["y"]) [] introtrm (fn _ => unfolddef_tac)
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
- val elimthm = Goal.prove (ProofContext.init_global thy)
+ val elimthm = Goal.prove ctxt
(argnames @ ["y", "P"]) [] elimtrm (fn _ => unfolddef_tac)
val opt_neg_introthm =
if is_all_input mode then
@@ -2090,7 +2097,7 @@
Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps
(@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1
THEN rtac @{thm Predicate.singleI} 1
- in SOME (Goal.prove (ProofContext.init_global thy) (argnames @ hoarg_names') []
+ in SOME (Goal.prove ctxt (argnames @ hoarg_names') []
neg_introtrm (fn _ => tac))
end
else NONE
@@ -2131,8 +2138,9 @@
val ([definition], thy') = thy |>
Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
+ val ctxt' = ProofContext.init_global thy'
val rules as ((intro, elim), _) =
- create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
+ create_intro_elim_rule ctxt' mode definition mode_cname funT (Const (name, T))
in thy'
|> set_function_name Pred name mode mode_cname
|> add_predfun_data name mode (definition, rules)
@@ -2301,10 +2309,9 @@
fun prove_sidecond ctxt t =
let
- val thy = ProofContext.theory_of ctxt
fun preds_of t nameTs = case strip_comb t of
(f as Const (name, T), args) =>
- if is_registered thy name then (name, T) :: nameTs
+ if is_registered ctxt name then (name, T) :: nameTs
else fold preds_of args nameTs
| _ => nameTs
val preds = preds_of t []
@@ -2402,10 +2409,9 @@
fun prove_one_direction options ctxt clauses preds pred mode moded_clauses =
let
- val thy = ProofContext.theory_of ctxt
val T = the (AList.lookup (op =) preds pred)
val nargs = length (binder_types T)
- val pred_case_rule = the_elim_of thy pred
+ val pred_case_rule = the_elim_of ctxt pred
in
REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
THEN print_tac options "before applying elim rule"
@@ -2495,7 +2501,7 @@
fun prove_sidecond2 options ctxt t = let
fun preds_of t nameTs = case strip_comb t of
(f as Const (name, T), args) =>
- if is_registered (ProofContext.theory_of ctxt) name then (name, T) :: nameTs
+ if is_registered ctxt name then (name, T) :: nameTs
else fold preds_of args nameTs
| _ => nameTs
val preds = preds_of t []
@@ -2512,7 +2518,7 @@
fun prove_clause2 options ctxt pred mode (ts, ps) i =
let
- val pred_intro_rule = nth (intros_of (ProofContext.theory_of ctxt) pred) (i - 1)
+ val pred_intro_rule = nth (intros_of ctxt pred) (i - 1)
val (in_ts, clause_out_ts) = split_mode mode ts;
fun prove_prems2 out_ts [] =
print_tac options "before prove_match2 - last call:"
@@ -2635,8 +2641,8 @@
map (fn (pred, modes) =>
(pred, map (fn (mode, value) => value) modes)) preds_modes_table
-fun compile_preds options comp_modifiers thy all_vs param_vs preds moded_clauses =
- map_preds_modes (fn pred => compile_pred options comp_modifiers thy all_vs param_vs pred
+fun compile_preds options comp_modifiers ctxt all_vs param_vs preds moded_clauses =
+ map_preds_modes (fn pred => compile_pred options comp_modifiers ctxt all_vs param_vs pred
(the (AList.lookup (op =) preds pred))) moded_clauses
fun prove options thy clauses preds moded_clauses compiled_terms =
@@ -2649,25 +2655,25 @@
compiled_terms
(* preparation of introduction rules into special datastructures *)
-fun dest_prem thy params t =
+fun dest_prem ctxt params t =
(case strip_comb t of
(v as Free _, ts) => if member (op =) params v then Prem t else Sidecond t
- | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem thy params t of
+ | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem ctxt params t of
Prem t => Negprem t
| Negprem _ => error ("Double negation not allowed in premise: " ^
- Syntax.string_of_term_global thy (c $ t))
+ Syntax.string_of_term ctxt (c $ t))
| Sidecond t => Sidecond (c $ t))
| (c as Const (s, _), ts) =>
- if is_registered thy s then Prem t else Sidecond t
+ if is_registered ctxt s then Prem t else Sidecond t
| _ => Sidecond t)
fun prepare_intrs options compilation thy prednames intros =
let
+ val ctxt = ProofContext.init_global thy
val intrs = map prop_of intros
val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
val (preds, intrs) = unify_consts thy preds intrs
- val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs]
- (ProofContext.init_global thy)
+ val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] ctxt
val preds = map dest_Const preds
val all_vs = terms_vs intrs
val all_modes =
@@ -2698,7 +2704,7 @@
fun add_clause intr clauses =
let
val (Const (name, T), ts) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))
- val prems = map (dest_prem thy params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr)
+ val prems = map (dest_prem ctxt params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr)
in
AList.update op = (name, these (AList.lookup op = clauses name) @
[(ts, prems)]) clauses
@@ -2755,20 +2761,19 @@
(* create code equation *)
-fun add_code_equations thy preds result_thmss =
+fun add_code_equations ctxt preds result_thmss =
let
- val ctxt = ProofContext.init_global thy
fun add_code_equation (predname, T) (pred, result_thms) =
let
val full_mode = fold_rev (curry Fun) (map (K Input) (binder_types T)) Bool
in
- if member (op =) (modes_of Pred thy predname) full_mode then
+ if member (op =) (modes_of Pred ctxt predname) full_mode then
let
val Ts = binder_types T
val arg_names = Name.variant_list []
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
val args = map2 (curry Free) arg_names Ts
- val predfun = Const (function_name_of Pred thy predname full_mode,
+ val predfun = Const (function_name_of Pred ctxt predname full_mode,
Ts ---> PredicateCompFuns.mk_predT @{typ unit})
val rhs = @{term Predicate.holds} $ (list_comb (predfun, args))
val eq_term = HOLogic.mk_Trueprop
@@ -2794,7 +2799,7 @@
define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory,
prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list
-> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
- add_code_equations : theory -> (string * typ) list
+ add_code_equations : Proof.context -> (string * typ) list
-> (string * thm list) list -> (string * thm list) list,
comp_modifiers : Comp_Mod.comp_modifiers,
use_random : bool,
@@ -2805,6 +2810,7 @@
let
fun dest_steps (Steps s) = s
val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps))
+ val ctxt = ProofContext.init_global thy
val _ = print_step options
("Starting predicate compiler (compilation: " ^ string_of_compilation compilation
^ ") for predicates " ^ commas prednames ^ "...")
@@ -2812,35 +2818,36 @@
(*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
val _ =
if show_intermediate_results options then
- tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
+ tracing (commas (map (Display.string_of_thm ctxt) (maps (intros_of ctxt) prednames)))
else ()
val (preds, all_vs, param_vs, all_modes, clauses) =
- prepare_intrs options compilation thy prednames (maps (intros_of thy) prednames)
+ prepare_intrs options compilation thy prednames (maps (intros_of ctxt) prednames)
val _ = print_step options "Infering modes..."
val ((moded_clauses, errors), thy') =
- Output.cond_timeit true "Infering modes"
+ Output.cond_timeit (!Quickcheck.timing) "Infering modes"
(fn _ => infer_modes mode_analysis_options
options compilation preds all_modes param_vs clauses thy)
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
val _ = check_expected_modes preds options modes
(*val _ = check_proposed_modes preds options modes (fst extra_modes) errors*)
- val _ = print_modes options thy' modes
+ val _ = print_modes options modes
val _ = print_step options "Defining executable functions..."
val thy'' =
Output.cond_timeit (!Quickcheck.timing) "Defining executable functions..."
(fn _ => fold (#define_functions (dest_steps steps) options preds) modes thy'
|> Theory.checkpoint)
+ val ctxt'' = ProofContext.init_global thy''
val _ = print_step options "Compiling equations..."
val compiled_terms =
Output.cond_timeit (!Quickcheck.timing) "Compiling equations...." (fn _ =>
compile_preds options
- (#comp_modifiers (dest_steps steps)) thy'' all_vs param_vs preds moded_clauses)
- val _ = print_compiled_terms options thy'' compiled_terms
+ (#comp_modifiers (dest_steps steps)) ctxt'' all_vs param_vs preds moded_clauses)
+ val _ = print_compiled_terms options ctxt'' compiled_terms
val _ = print_step options "Proving equations..."
val result_thms =
Output.cond_timeit (!Quickcheck.timing) "Proving equations...." (fn _ =>
#prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms)
- val result_thms' = #add_code_equations (dest_steps steps) thy'' preds
+ val result_thms' = #add_code_equations (dest_steps steps) ctxt'' preds
(maps_modes result_thms)
val qname = #qname (dest_steps steps)
val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
@@ -2873,8 +2880,9 @@
let
fun dest_steps (Steps s) = s
val defined = defined_functions (Comp_Mod.compilation (#comp_modifiers (dest_steps steps)))
+ val ctxt = ProofContext.init_global thy
val thy' = thy
- |> PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names)
+ |> PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names)
|> Theory.checkpoint;
fun strong_conn_of gr keys =
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
@@ -2882,7 +2890,7 @@
val thy'' = fold_rev
(fn preds => fn thy =>
- if not (forall (defined thy) preds) then
+ if not (forall (defined (ProofContext.init_global thy)) preds) then
let
val mode_analysis_options = {use_random = #use_random (dest_steps steps),
reorder_premises =
@@ -3025,15 +3033,17 @@
let
val thy = ProofContext.theory_of lthy
val const = prep_const thy raw_const
+ val ctxt = ProofContext.init_global thy
val lthy' = Local_Theory.theory (PredData.map
- (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
+ (extend (fetch_pred_data ctxt) (depending_preds_of ctxt) const)) lthy
val thy' = ProofContext.theory_of lthy'
- val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy')
+ val ctxt' = ProofContext.init_global thy'
+ val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim ctxt')
fun mk_cases const =
let
val T = Sign.the_const_type thy const
val pred = Const (const, T)
- val intros = intros_of thy' const
+ val intros = intros_of ctxt' const
in mk_casesrule lthy' pred intros end
val cases_rules = map mk_cases preds
val cases =
@@ -3083,11 +3093,11 @@
(unit -> int -> int -> int * int -> int -> (term * int) Lazy_Sequence.lazy_sequence) option)
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy compfuns param_user_modes (compilation, arguments) t_compr =
+fun analyze_compr ctxt compfuns param_user_modes (compilation, arguments) t_compr =
let
val all_modes_of = all_modes_of compilation
val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
- | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
+ | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
val (body, Ts, fp) = HOLogic.strip_psplits split;
val output_names = Name.variant_list (Term.add_free_names body [])
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
@@ -3098,9 +3108,9 @@
val (pred as Const (name, T), all_args) =
case strip_comb body of
(Const (name, T), all_args) => (Const (name, T), all_args)
- | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term_global thy head)
+ | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
in
- if defined_functions compilation thy name then
+ if defined_functions compilation ctxt name then
let
fun extract_mode (Const ("Pair", _) $ t1 $ t2) = Pair (extract_mode t1, extract_mode t2)
| extract_mode (Free (x, _)) = if member (op =) output_names x then Output else Input
@@ -3123,9 +3133,13 @@
instance_of (m1, Input) andalso instance_of (m2, Input)
| instance_of (Pair (m1, m2), Output) =
instance_of (m1, Output) andalso instance_of (m2, Output)
+ | instance_of (Input, Pair (m1, m2)) =
+ instance_of (Input, m1) andalso instance_of (Input, m2)
+ | instance_of (Output, Pair (m1, m2)) =
+ instance_of (Output, m1) andalso instance_of (Output, m2)
| instance_of _ = false
in forall instance_of (strip_fun_mode m1 ~~ strip_fun_mode m2) end
- val derivs = all_derivations_of thy (all_modes_of thy) [] body
+ val derivs = all_derivations_of ctxt (all_modes_of ctxt) [] body
|> filter (fn (d, missing_vars) =>
let
val (p_mode :: modes) = collect_context_modes d
@@ -3137,10 +3151,10 @@
|> map fst
val deriv = case derivs of
[] => error ("No mode possible for comprehension "
- ^ Syntax.string_of_term_global thy t_compr)
+ ^ Syntax.string_of_term ctxt t_compr)
| [d] => d
| d :: _ :: _ => (warning ("Multiple modes possible for comprehension "
- ^ Syntax.string_of_term_global thy t_compr); d);
+ ^ Syntax.string_of_term ctxt t_compr); d);
val (_, outargs) = split_mode (head_mode_of deriv) all_args
val additional_arguments =
case compilation of
@@ -3164,7 +3178,7 @@
| DSeq => dseq_comp_modifiers
| Pos_Random_DSeq => pos_random_dseq_comp_modifiers
| New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers
- val t_pred = compile_expr comp_modifiers (ProofContext.init_global thy)
+ val t_pred = compile_expr comp_modifiers ctxt
(body, deriv) additional_arguments;
val T_pred = dest_predT compfuns (fastype_of t_pred)
val arrange = split_lambda (HOLogic.mk_tuple outargs) output_tuple
@@ -3175,7 +3189,7 @@
error "Evaluation with values is not possible because compilation with code_pred was not invoked"
end
-fun eval thy stats param_user_modes (options as (compilation, arguments)) k t_compr =
+fun eval ctxt stats param_user_modes (options as (compilation, arguments)) k t_compr =
let
fun count xs x =
let
@@ -3190,7 +3204,7 @@
| Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
| New_Pos_Random_DSeq => New_Pos_Random_Sequence_CompFuns.compfuns
| _ => PredicateCompFuns.compfuns
- val t = analyze_compr thy compfuns param_user_modes options t_compr;
+ val t = analyze_compr ctxt compfuns param_user_modes options t_compr;
val T = dest_predT compfuns (fastype_of t);
val t' =
if stats andalso compilation = New_Pos_Random_DSeq then
@@ -3199,6 +3213,7 @@
@{term Code_Numeral.of_nat} $ (HOLogic.size_const T $ Bound 0)))) t
else
mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t
+ val thy = ProofContext.theory_of ctxt
val (ts, statistics) =
case compilation of
(* Random =>
@@ -3250,8 +3265,7 @@
fun values ctxt param_user_modes ((raw_expected, stats), comp_options) k t_compr =
let
- val thy = ProofContext.theory_of ctxt
- val ((T, ts), statistics) = eval thy stats param_user_modes comp_options k t_compr
+ val ((T, ts), statistics) = eval ctxt stats param_user_modes comp_options k t_compr
val setT = HOLogic.mk_setT T
val elems = HOLogic.mk_set T ts
val cont = Free ("...", setT)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Sun May 23 10:38:11 2010 +0100
@@ -253,8 +253,9 @@
fun obtain_specification_graph options thy t =
let
+ val ctxt = ProofContext.init_global thy
fun is_nondefining_const (c, T) = member (op =) logic_operator_names c
- fun has_code_pred_intros (c, T) = can (Predicate_Compile_Core.intros_of thy) c
+ fun has_code_pred_intros (c, T) = can (Predicate_Compile_Core.intros_of ctxt) c
fun case_consts (c, T) = is_some (Datatype.info_of_case thy c)
fun is_datatype_constructor (c, T) = is_some (Datatype.info_of_constr thy (c, T))
fun defiants_of specs =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Sun May 23 10:38:11 2010 +0100
@@ -216,12 +216,13 @@
(*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*)
val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
- val modes = Predicate_Compile_Core.modes_of compilation thy4 full_constname
+ val ctxt4 = ProofContext.init_global thy4
+ val modes = Predicate_Compile_Core.modes_of compilation ctxt4 full_constname
val output_mode = fold_rev (curry Fun) (map (K Output) (binder_types constT)) Bool
val prog =
if member eq_mode modes output_mode then
let
- val name = Predicate_Compile_Core.function_name_of compilation thy4
+ val name = Predicate_Compile_Core.function_name_of compilation ctxt4
full_constname output_mode
val T =
case compilation of
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Sun May 23 10:38:11 2010 +0100
@@ -181,7 +181,7 @@
if member (op =) ((map fst specs) @ black_list) pred_name then
thy
else
- (case try (Predicate_Compile_Core.intros_of thy) pred_name of
+ (case try (Predicate_Compile_Core.intros_of (ProofContext.init_global thy)) pred_name of
NONE => thy
| SOME [] => thy
| SOME intros =>
--- a/src/HOL/Tools/Qelim/cooper.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Sun May 23 10:38:11 2010 +0100
@@ -331,7 +331,7 @@
| t => if is_intrel t
then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop))
RS eq_reflection
- else reflexive ct;
+ else Thm.reflexive ct;
val dvdc = @{cterm "op dvd :: int => _"};
@@ -369,7 +369,7 @@
(Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"}
(Thm.capply (Thm.capply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x))
@{cterm "0::int"})))
- in equal_elim (Thm.symmetric th) TrueI end;
+ in Thm.equal_elim (Thm.symmetric th) TrueI end;
val notz =
let val tab = fold Inttab.update
(ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty
@@ -412,11 +412,11 @@
val ltx = Thm.capply (Thm.capply cmulC clt) cx
val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth)
val th' = inst' [Thm.cabs ltx (Thm.rhs_of uth), clt] unity_coeff_ex
- val thf = transitive th
- (transitive (symmetric (beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th')
+ val thf = Thm.transitive th
+ (Thm.transitive (Thm.symmetric (Thm.beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th')
val (lth,rth) = Thm.dest_comb (cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true
- ||> beta_conversion true |>> Thm.symmetric
- in transitive (transitive lth thf) rth end;
+ ||> Thm.beta_conversion true |>> Thm.symmetric
+ in Thm.transitive (Thm.transitive lth thf) rth end;
val emptyIS = @{cterm "{}::int set"};
@@ -459,7 +459,7 @@
Simplifier.rewrite lin_ss
(Thm.capply @{cterm Trueprop}
(Thm.capply (Thm.capply dvdc (Numeral.mk_cnumber @{ctyp "int"} x)) cd))
- in equal_elim (Thm.symmetric th) TrueI end;
+ in Thm.equal_elim (Thm.symmetric th) TrueI end;
val dvd =
let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in
fn ct => the (Inttab.lookup tab (term_of ct |> dest_number))
@@ -471,7 +471,7 @@
let val th = Simplifier.rewrite lin_ss
(Thm.capply @{cterm Trueprop}
(Thm.capply (Thm.capply @{cterm "op < :: int => _"} @{cterm "0::int"}) cd))
- in equal_elim (Thm.symmetric th) TrueI end;
+ in Thm.equal_elim (Thm.symmetric th) TrueI end;
(* A and B set *)
local
val insI1 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI1"}
@@ -483,7 +483,7 @@
| Const(@{const_name insert}, _) $ y $ _ =>
let val (cy,S') = Thm.dest_binop S
in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
- else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
+ else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
(provein x S')
end
end
@@ -494,14 +494,14 @@
if length (distinct (op aconv) bl) <= length (distinct (op aconv) al)
then
(bl,b0,decomp_minf,
- fn B => (map (fn th => implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp)
+ fn B => (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp)
[bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@
(map (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]))
[bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj,
bsetdisj,infDconj, infDdisj]),
cpmi)
else (al,a0,decomp_pinf,fn A =>
- (map (fn th => implies_elim (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]) th) dp)
+ (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]) th) dp)
[aseteq,asetneq,asetlt, asetle, asetgt,asetge])@
(map (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]))
[asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj,
@@ -776,7 +776,7 @@
fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
val p' = fold_rev gen ts p
- in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
+ in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end));
local
val ss1 = comp_ss
@@ -792,7 +792,7 @@
@{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}]
addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
val div_mod_ss = HOL_basic_ss addsimps @{thms simp_thms}
- @ map (symmetric o mk_meta_eq)
+ @ map (Thm.symmetric o mk_meta_eq)
[@{thm "dvd_eq_mod_eq_0"},
@{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"},
@{thm "mod_add_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}]
@@ -823,7 +823,7 @@
then oracle (ctxt, Envir.beta_norm (Pattern.eta_long [] (term_of (Thm.dest_arg p))))
else Conv.arg_conv (conv ctxt) p
val p' = Thm.rhs_of cpth
- val th = implies_intr p' (equal_elim (symmetric cpth) (assume p'))
+ val th = Thm.implies_intr p' (Thm.equal_elim (Thm.symmetric cpth) (Thm.assume p'))
in rtac th i end
handle COOPER _ => no_tac);
--- a/src/HOL/Tools/Qelim/ferrante_rackoff.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Qelim/ferrante_rackoff.ML Sun May 23 10:38:11 2010 +0100
@@ -69,7 +69,7 @@
val th0 = if mi then
instantiate ([],[(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th
else instantiate ([],[(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th
- in implies_elim (implies_elim th0 th') th'' end
+ in Thm.implies_elim (Thm.implies_elim th0 th') th'' end
in (fw true th1 th_1 th_1', fw false th2 th_2 th_2',
fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5')
end
@@ -88,7 +88,7 @@
val enth = Drule.arg_cong_rule ce nthx
val [th0,th1] = map (instantiate' [SOME cT] []) @{thms "finite.intros"}
fun ins x th =
- implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg)
+ Thm.implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg)
(Thm.cprop_of th), SOME x] th1) th
val fU = fold ins u th0
val cU = funpow 2 Thm.dest_arg (Thm.cprop_of fU)
@@ -102,7 +102,7 @@
| Const(@{const_name insert}, _) $ y $_ =>
let val (cy,S') = Thm.dest_binop S
in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
- else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
+ else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
(provein x S')
end
end
@@ -152,7 +152,7 @@
| Eq => [minf_eq, pinf_eq, nmi_eq, npi_eq, ld_eq]
| NEq => [minf_neq, pinf_neq, nmi_neq, npi_neq, ld_neq])
val tU = U t
- fun Ufw th = implies_elim th tU
+ fun Ufw th = Thm.implies_elim th tU
in [mi_th, pi_th, Ufw nmi_th, Ufw npi_th, Ufw ld_th]
end
in ([], K (mi_th, pi_th, nmi_th, npi_th, ld_th)) end)
@@ -164,7 +164,7 @@
[fU, ld_th, nmi_th, npi_th, minf_th, pinf_th]
val bex_conv =
Simplifier.rewrite (HOL_basic_ss addsimps simp_thms@(@{thms "bex_simps" (1-5)}))
- val result_th = fconv_rule (arg_conv bex_conv) (transitive enth qe_th)
+ val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th)
in result_th
end
--- a/src/HOL/Tools/Qelim/langford.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Qelim/langford.ML Sun May 23 10:38:11 2010 +0100
@@ -22,7 +22,7 @@
fun prove_finite cT u =
let val [th0,th1] = map (instantiate' [SOME cT] []) @{thms "finite.intros"}
fun ins x th =
- implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg)
+ Thm.implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg)
(Thm.cprop_of th), SOME x] th1) th
in fold ins u th0 end;
@@ -51,17 +51,17 @@
(Const (@{const_name Orderings.bot}, _),_) =>
let
val (neU,fU) = proveneF U
- in simp_rule (transitive ths (dlo_qeth_nolb OF [neU, fU])) end
+ in simp_rule (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end
| (_,Const (@{const_name Orderings.bot}, _)) =>
let
val (neL,fL) = proveneF L
- in simp_rule (transitive ths (dlo_qeth_noub OF [neL, fL])) end
+ in simp_rule (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end
| (_,_) =>
let
val (neL,fL) = proveneF L
val (neU,fU) = proveneF U
- in simp_rule (transitive ths (dlo_qeth OF [neL, neU, fL, fU]))
+ in simp_rule (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU]))
end
in qe end
| _ => error "dlo_qe : Not an existential formula";
@@ -101,8 +101,8 @@
fun conj_aci_rule eq =
let
val (l,r) = Thm.dest_equals eq
- fun tabl c = the (Termtab.lookup (mk_conj_tab (assume l)) (term_of c))
- fun tabr c = the (Termtab.lookup (mk_conj_tab (assume r)) (term_of c))
+ fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (term_of c))
+ fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (term_of c))
val ll = Thm.dest_arg l
val rr = Thm.dest_arg r
@@ -111,7 +111,7 @@
val thr = prove_conj tabr (conjuncts ll)
|> Drule.implies_intr_hyps
val eqI = instantiate' [] [SOME ll, SOME rr] @{thm iffI}
- in implies_elim (implies_elim eqI thl) thr |> mk_meta_eq end;
+ in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end;
fun contains x ct = member (op aconv) (OldTerm.term_frees (term_of ct)) (term_of x);
@@ -136,7 +136,7 @@
| _ =>
conj_aci_rule (Thm.mk_binop @{cterm "op == :: prop => _"} Pp
(Thm.capply @{cterm Trueprop} (list_conj (ndx @dx))))
- |> abstract_rule xn x |> Drule.arg_cong_rule e
+ |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e
|> Conv.fconv_rule (Conv.arg_conv
(Simplifier.rewrite
(HOL_basic_ss addsimps (simp_thms @ ex_simps))))
@@ -144,7 +144,7 @@
end
| _ => conj_aci_rule (Thm.mk_binop @{cterm "op == :: prop => _"} Pp
(Thm.capply @{cterm Trueprop} (list_conj (eqs@neqs))))
- |> abstract_rule xn x |> Drule.arg_cong_rule e
+ |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e
|> Conv.fconv_rule (Conv.arg_conv
(Simplifier.rewrite
(HOL_basic_ss addsimps (simp_thms @ ex_simps))))
@@ -167,7 +167,7 @@
@ [@{thm "all_not_ex"}, not_all, ex_disj_distrib])
in fn p =>
Qelim.gen_qelim_conv pcv pcv dnfex_conv cons
- (Thm.add_cterm_frees p []) (K reflexive) (K reflexive)
+ (Thm.add_cterm_frees p []) (K Thm.reflexive) (K Thm.reflexive)
(K (basic_dloqe () gst qe_bnds qe_nolb qe_noub gs)) p
end;
@@ -212,7 +212,7 @@
fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
val p' = fold_rev gen ts p
- in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
+ in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end));
fun cfrees ats ct =
--- a/src/HOL/Tools/Qelim/qelim.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Qelim/qelim.ML Sun May 23 10:38:11 2010 +0100
@@ -47,7 +47,7 @@
val p = Thm.dest_arg p
val ([(_,T)],[]) = Thm.match (@{cpat "All"}, Thm.dest_fun p)
val th = instantiate' [SOME T] [SOME p] all_not_ex
- in transitive th (conv env (Thm.rhs_of th))
+ in Thm.transitive th (conv env (Thm.rhs_of th))
end
| _ => atcv env p
in precv then_conv (conv env) then_conv postcv end
--- a/src/HOL/Tools/Quotient/quotient_def.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Quotient/quotient_def.ML Sun May 23 10:38:11 2010 +0100
@@ -91,20 +91,15 @@
quotient_def ((NONE, NoSyn), (Attrib.empty_binding,
(Quotient_Term.quotient_lift_const qtys (b, t) ctxt, t))) ctxt
-local
- structure P = OuterParse;
-in
-
-val quotdef_decl = (P.binding >> SOME) -- P.opt_mixfix' --| P.$$$ "where"
+val quotdef_decl = (Parse.binding >> SOME) -- Parse.opt_mixfix' --| Parse.$$$ "where"
val quotdef_parser =
Scan.optional quotdef_decl (NONE, NoSyn) --
- P.!!! (SpecParse.opt_thm_name ":" -- (P.term --| P.$$$ "is" -- P.term))
-end
+ Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| Parse.$$$ "is" -- Parse.term))
val _ =
- OuterSyntax.local_theory "quotient_definition"
+ Outer_Syntax.local_theory "quotient_definition"
"definition for constants over the quotient type"
- OuterKeyword.thy_decl (quotdef_parser >> (snd oo quotdef_cmd))
+ Keyword.thy_decl (quotdef_parser >> (snd oo quotdef_cmd))
end; (* structure *)
--- a/src/HOL/Tools/Quotient/quotient_info.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Quotient/quotient_info.ML Sun May 23 10:38:11 2010 +0100
@@ -91,9 +91,9 @@
val maps_attr_parser =
Args.context -- Scan.lift
- ((Args.name --| OuterParse.$$$ "=") --
- (OuterParse.$$$ "(" |-- Args.name --| OuterParse.$$$ "," --
- Args.name --| OuterParse.$$$ ")"))
+ ((Args.name --| Parse.$$$ "=") --
+ (Parse.$$$ "(" |-- Args.name --| Parse.$$$ "," --
+ Args.name --| Parse.$$$ ")"))
val _ = Context.>> (Context.map_theory
(Attrib.setup @{binding "map"} (maps_attr_parser >> maps_attribute)
@@ -278,8 +278,8 @@
(* setup of the printing commands *)
fun improper_command (pp_fn, cmd_name, descr_str) =
- OuterSyntax.improper_command cmd_name descr_str
- OuterKeyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
+ Outer_Syntax.improper_command cmd_name descr_str
+ Keyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
val _ = map improper_command
[(print_mapsinfo, "print_quotmaps", "prints out all map functions"),
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Sun May 23 10:38:11 2010 +0100
@@ -490,7 +490,7 @@
end
| _ => Conv.all_conv ctrm
-fun lambda_prs_conv ctxt = More_Conv.top_conv lambda_prs_simple_conv ctxt
+fun lambda_prs_conv ctxt = Conv.top_conv lambda_prs_simple_conv ctxt
fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
@@ -517,7 +517,7 @@
*)
fun clean_tac lthy =
let
- val defs = map (symmetric o #def) (qconsts_dest lthy)
+ val defs = map (Thm.symmetric o #def) (qconsts_dest lthy)
val prs = prs_rules_get lthy
val ids = id_simps_get lthy
val thms = @{thms Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
--- a/src/HOL/Tools/Quotient/quotient_typ.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML Sun May 23 10:38:11 2010 +0100
@@ -299,16 +299,16 @@
end
val quotspec_parser =
- OuterParse.and_list1
- ((OuterParse.type_args -- OuterParse.binding) --
- OuterParse.opt_mixfix -- (OuterParse.$$$ "=" |-- OuterParse.typ) --
- (OuterParse.$$$ "/" |-- OuterParse.term))
+ Parse.and_list1
+ ((Parse.type_args -- Parse.binding) --
+ Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.typ) --
+ (Parse.$$$ "/" |-- Parse.term))
-val _ = OuterKeyword.keyword "/"
+val _ = Keyword.keyword "/"
val _ =
- OuterSyntax.local_theory_to_proof "quotient_type"
+ Outer_Syntax.local_theory_to_proof "quotient_type"
"quotient type definitions (require equivalence proofs)"
- OuterKeyword.thy_goal (quotspec_parser >> quotient_type_cmd)
+ Keyword.thy_goal (quotspec_parser >> quotient_type_cmd)
end; (* structure *)
--- a/src/HOL/Tools/SMT/smt_normalize.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Sun May 23 10:38:11 2010 +0100
@@ -47,11 +47,11 @@
"distinct [x, y] == (x ~= y)"
by simp_all}
fun distinct_conv _ =
- if_true_conv is_trivial_distinct (More_Conv.rewrs_conv thms)
+ if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)
in
fun trivial_distinct ctxt =
map ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_conv distinct_conv ctxt))
+ Conv.fconv_rule (Conv.top_conv distinct_conv ctxt))
end
@@ -67,11 +67,11 @@
"(case P of True => x | False => y) == (if P then x else y)"
"(case P of False => y | True => x) == (if P then x else y)"
by (rule eq_reflection, simp)+}
- val unfold_conv = if_true_conv is_bool_case (More_Conv.rewrs_conv thms)
+ val unfold_conv = if_true_conv is_bool_case (Conv.rewrs_conv thms)
in
fun rewrite_bool_cases ctxt =
map ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_conv (K unfold_conv) ctxt))
+ Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt))
end
@@ -107,7 +107,7 @@
in
fun normalize_numerals ctxt =
map ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ??
- Conv.fconv_rule (More_Conv.top_sweep_conv pos_conv ctxt))
+ Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt))
end
@@ -193,7 +193,7 @@
local
val eta_conv = eta_expand_conv
- fun keep_conv ctxt = More_Conv.binder_conv norm_conv ctxt
+ fun keep_conv ctxt = Conv.binder_conv (norm_conv o snd) ctxt
and eta_binder_conv ctxt = Conv.arg_conv (eta_conv norm_conv ctxt)
and keep_let_conv ctxt = Conv.combination_conv
(Conv.arg_conv (norm_conv ctxt)) (Conv.abs_conv (norm_conv o snd) ctxt)
@@ -267,7 +267,7 @@
Conv.binop_conv (atomize_conv ctxt) then_conv
Conv.rewr_conv @{thm atomize_eq}
| Const (@{const_name all}, _) $ Abs _ =>
- More_Conv.binder_conv atomize_conv ctxt then_conv
+ Conv.binder_conv (atomize_conv o snd) ctxt then_conv
Conv.rewr_conv @{thm atomize_all}
| _ => Conv.all_conv) ct
--- a/src/HOL/Tools/SMT/smt_solver.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_solver.ML Sun May 23 10:38:11 2010 +0100
@@ -166,7 +166,7 @@
end
-fun trace_recon_data ctxt {typs, terms, ...} =
+fun trace_recon_data ctxt ({typs, terms, ...} : SMT_Translate.recon) =
let
fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
fun pretty_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
@@ -311,14 +311,14 @@
val setup =
Attrib.setup (Binding.name "smt_solver")
- (Scan.lift (OuterParse.$$$ "=" |-- Args.name) >>
+ (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
(Thm.declaration_attribute o K o select_solver))
"SMT solver configuration" #>
setup_timeout #>
setup_trace #>
setup_fixed_certificates #>
Attrib.setup (Binding.name "smt_certificates")
- (Scan.lift (OuterParse.$$$ "=" |-- Args.name) >>
+ (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
(Thm.declaration_attribute o K o select_certificates))
"SMT certificates" #>
Method.setup (Binding.name "smt") smt_method
@@ -353,9 +353,9 @@
Pretty.big_list "Solver-specific settings:" infos])
end
-val _ = OuterSyntax.improper_command "smt_status"
- "Show the available SMT solvers and the currently selected solver."
- OuterKeyword.diag
+val _ =
+ Outer_Syntax.improper_command "smt_status"
+ "show the available SMT solvers and the currently selected solver" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
print_setup (Context.Proof (Toplevel.context_of state)))))
--- a/src/HOL/Tools/SMT/smtlib_interface.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML Sun May 23 10:38:11 2010 +0100
@@ -81,7 +81,7 @@
fun unfold_abs_min_max_defs ctxt thm =
if exists_abs_min_max (Thm.prop_of thm)
- then Conv.fconv_rule (More_Conv.top_conv unfold_def_conv ctxt) thm
+ then Conv.fconv_rule (Conv.top_conv unfold_def_conv ctxt) thm
else thm
--- a/src/HOL/Tools/SMT/z3_proof_parser.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/z3_proof_parser.ML Sun May 23 10:38:11 2010 +0100
@@ -285,15 +285,15 @@
"4" => SOME 4 | "5" => SOME 5 | "6" => SOME 6 | "7" => SOME 7 |
"8" => SOME 8 | "9" => SOME 9 | _ => NONE)
-val digits = Scan.lift (Scan.many1 Symbol.is_ascii_digit) >> implode
-val nat_num = Scan.lift (Scan.repeat1 (Scan.some digit)) >> (fn ds =>
- fold (fn d => fn i => i * 10 + d) ds 0)
-val int_num = Scan.optional ($$ "-" >> K (fn i => ~i)) I :|--
- (fn sign => nat_num >> sign)
+fun digits st = (Scan.lift (Scan.many1 Symbol.is_ascii_digit) >> implode) st
+fun nat_num st = (Scan.lift (Scan.repeat1 (Scan.some digit)) >> (fn ds =>
+ fold (fn d => fn i => i * 10 + d) ds 0)) st
+fun int_num st = (Scan.optional ($$ "-" >> K (fn i => ~i)) I :|--
+ (fn sign => nat_num >> sign)) st
val is_char = Symbol.is_ascii_letter orf Symbol.is_ascii_digit orf
member (op =) (explode "_+*-/%~=<>$&|?!.@^#")
-val name = Scan.lift (Scan.many1 is_char) >> implode
+fun name st = (Scan.lift (Scan.many1 is_char) >> implode) st
fun sym st =
(name -- Scan.optional (bra (seps_by ($$ ":") sym)) [] >> I.Sym) st
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Sun May 23 10:38:11 2010 +0100
@@ -540,7 +540,7 @@
fun elim_unused_conv ctxt =
Conv.params_conv ~1 (K (Conv.arg_conv (Conv.arg1_conv
- (More_Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
+ (Conv.rewrs_conv [elim_all, elim_ex])))) ctxt
fun elim_unused_tac ctxt =
REPEAT_ALL_NEW (
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Sun May 23 10:38:11 2010 +0100
@@ -96,7 +96,7 @@
fun unfold_eqs _ [] = Conv.all_conv
| unfold_eqs ctxt eqs =
- More_Conv.top_sweep_conv (K (More_Conv.rewrs_conv eqs)) ctxt
+ Conv.top_sweep_conv (K (Conv.rewrs_conv eqs)) ctxt
fun match_instantiate f ct thm =
Thm.instantiate (Thm.match (f (Thm.cprop_of thm), ct)) thm
@@ -256,7 +256,7 @@
val set5 = @{lemma "x ~: set (y # ys) == x ~= y & x ~: set ys" by simp}
fun set_conv ct =
- (More_Conv.rewrs_conv [set1, set2, set3, set4] else_conv
+ (Conv.rewrs_conv [set1, set2, set3, set4] else_conv
(Conv.rewr_conv set5 then_conv Conv.arg_conv set_conv)) ct
val dist1 = @{lemma "distinct [] == ~False" by simp}
@@ -267,7 +267,7 @@
fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
in
fun unfold_distinct_conv ct =
- (More_Conv.rewrs_conv [dist1, dist2] else_conv
+ (Conv.rewrs_conv [dist1, dist2] else_conv
(Conv.rewr_conv dist3 then_conv binop_conv set_conv unfold_distinct_conv)) ct
end
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Sun May 23 10:38:11 2010 +0100
@@ -35,7 +35,7 @@
(* Useful Theorems *)
(* ------------------------------------------------------------------------- *)
val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
-val REFL_THM = incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
+val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
@@ -141,9 +141,10 @@
in
if is_conjecture then
(Metis.Thm.axiom (Metis.LiteralSet.fromList mlits),
- add_type_literals types_sorts)
+ type_literals_for_types types_sorts)
else
- let val tylits = add_type_literals (filter (not o default_sort ctxt) types_sorts)
+ let val tylits = filter_out (default_sort ctxt) types_sorts
+ |> type_literals_for_types
val mtylits = if Config.get ctxt type_lits
then map (metis_of_type_literals false) tylits else []
in
@@ -216,6 +217,8 @@
fun strip_happ args (Metis.Term.Fn(".",[t,u])) = strip_happ (u::args) t
| strip_happ args x = (x, args);
+fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
+
fun fol_type_to_isa _ (Metis.Term.Var v) =
(case strip_prefix tvar_prefix v of
SOME w => make_tvar w
@@ -364,7 +367,7 @@
in cterm_instantiate substs th end;
(* INFERENCE RULE: AXIOM *)
-fun axiom_inf thpairs th = incr_indexes 1 (lookth thpairs th);
+fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
(*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
(* INFERENCE RULE: ASSUME *)
@@ -527,7 +530,7 @@
val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*)
- val subst' = incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
+ val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
val eq_terms = map (pairself (cterm_of thy))
(ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
@@ -599,9 +602,9 @@
(*Extract TFree constraints from context to include as conjecture clauses*)
fun init_tfrees ctxt =
- let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts
- in
- add_type_literals (Vartab.fold add (#2 (Variable.constraints_of ctxt)) [])
+ let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts in
+ Vartab.fold add (#2 (Variable.constraints_of ctxt)) []
+ |> type_literals_for_types
end;
(*transform isabelle type / arity clause to metis clause *)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Sun May 23 10:38:11 2010 +0100
@@ -253,7 +253,7 @@
end
end;
-fun relevant_clauses ctxt convergence follow_defs max_new
+fun relevant_clauses ctxt relevance_convergence defs_relevant max_new
(relevance_override as {add, del, only}) thy ctab =
let
val thms_for_facts =
@@ -268,7 +268,7 @@
val (newrels, more_rejects) = take_best max_new newpairs
val new_consts = maps #2 newrels
val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
- val newp = p + (1.0-p) / convergence
+ val newp = p + (1.0 - p) / relevance_convergence
in
trace_msg (fn () => "relevant this iteration: " ^
Int.toString (length newrels));
@@ -283,7 +283,7 @@
else clause_weight ctab rel_consts consts_typs
in
if p <= weight orelse
- (follow_defs andalso defines thy (#1 clsthm) rel_consts) then
+ (defs_relevant andalso defines thy (#1 clsthm) rel_consts) then
(trace_msg (fn () => name ^ " clause " ^ Int.toString n ^
" passes: " ^ Real.toString weight);
relevant ((ax, weight) :: newrels, rejects) axs)
@@ -297,8 +297,9 @@
end
in iter end
-fun relevance_filter ctxt relevance_threshold convergence follow_defs max_new
- theory_relevant relevance_override thy axioms goals =
+fun relevance_filter ctxt relevance_threshold relevance_convergence
+ defs_relevant max_new theory_relevant relevance_override
+ thy axioms goals =
if relevance_threshold > 0.0 then
let
val const_tab = List.foldl (count_axiom_consts theory_relevant thy)
@@ -308,8 +309,9 @@
trace_msg (fn () => "Initial constants: " ^
commas (Symtab.keys goal_const_tab))
val relevant =
- relevant_clauses ctxt convergence follow_defs max_new relevance_override
- thy const_tab relevance_threshold goal_const_tab
+ relevant_clauses ctxt relevance_convergence defs_relevant max_new
+ relevance_override thy const_tab relevance_threshold
+ goal_const_tab
(map (pair_consts_typs_axiom theory_relevant thy)
axioms)
in
@@ -350,7 +352,7 @@
fun subtract_cls ax_clauses =
filter_out (Termtab.defined (mk_clause_table ax_clauses) o prop_of)
-fun all_valid_thms respect_no_atp ctxt =
+fun all_valid_thms respect_no_atp ctxt chain_ths =
let
val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
val local_facts = ProofContext.facts_of ctxt;
@@ -369,7 +371,8 @@
val name2 = Name_Space.extern full_space name;
val ths = filter_out bad_for_atp ths0;
in
- if Facts.is_concealed facts name orelse null ths orelse
+ if Facts.is_concealed facts name orelse
+ forall (member Thm.eq_thm chain_ths) ths orelse
(respect_no_atp andalso is_package_def name) then
I
else case find_first check_thms [name1, name2, name] of
@@ -394,10 +397,10 @@
(* The single-name theorems go after the multiple-name ones, so that single
names are preferred when both are available. *)
-fun name_thm_pairs respect_no_atp ctxt =
+fun name_thm_pairs respect_no_atp ctxt chain_ths =
let
val (mults, singles) =
- List.partition is_multi (all_valid_thms respect_no_atp ctxt)
+ List.partition is_multi (all_valid_thms respect_no_atp ctxt chain_ths)
val ps = [] |> fold add_single_names singles
|> fold add_multi_names mults
in ps |> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) end;
@@ -406,11 +409,11 @@
(warning ("No name for theorem " ^ Display.string_of_thm_without_context th); false)
| check_named _ = true;
-fun get_all_lemmas respect_no_atp ctxt =
+fun get_all_lemmas respect_no_atp ctxt chain_ths =
let val included_thms =
tap (fn ths => trace_msg
(fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
- (name_thm_pairs respect_no_atp ctxt)
+ (name_thm_pairs respect_no_atp ctxt chain_ths)
in
filter check_named included_thms
end;
@@ -504,24 +507,24 @@
fun is_first_order thy = forall (Meson.is_fol_term thy) o map prop_of
-fun get_relevant_facts respect_no_atp relevance_threshold convergence
- follow_defs max_new theory_relevant
+fun get_relevant_facts respect_no_atp relevance_threshold relevance_convergence
+ defs_relevant max_new theory_relevant
(relevance_override as {add, only, ...})
- (ctxt, (chain_ths, th)) goal_cls =
+ (ctxt, (chain_ths, _)) goal_cls =
if (only andalso null add) orelse relevance_threshold > 1.0 then
[]
else
let
val thy = ProofContext.theory_of ctxt
val is_FO = is_first_order thy goal_cls
- val included_cls = get_all_lemmas respect_no_atp ctxt
+ val included_cls = get_all_lemmas respect_no_atp ctxt chain_ths
|> cnf_rules_pairs thy |> make_unique
|> restrict_to_logic thy is_FO
|> remove_unwanted_clauses
in
- relevance_filter ctxt relevance_threshold convergence follow_defs max_new
- theory_relevant relevance_override thy included_cls
- (map prop_of goal_cls)
+ relevance_filter ctxt relevance_threshold relevance_convergence
+ defs_relevant max_new theory_relevant relevance_override
+ thy included_cls (map prop_of goal_cls)
end
(* prepare for passing to writer,
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML Sun May 23 10:38:11 2010 +0100
@@ -69,7 +69,7 @@
(* minimalization of thms *)
fun minimize_theorems (params as {debug, atps, full_types, minimize_timeout,
- isar_proof, shrink_factor, ...})
+ isar_proof, isar_shrink_factor, ...})
i n state name_thms_pairs =
let
val thy = Proof.theory_of state
@@ -110,7 +110,8 @@
in
(SOME min_thms,
proof_text isar_proof
- (pool, debug, full_types, shrink_factor, ctxt, conjecture_shape)
+ (pool, debug, full_types, isar_shrink_factor, ctxt,
+ conjecture_shape)
(K "", proof, internal_thm_names, goal, i) |> fst)
end
| {outcome = SOME TimedOut, ...} =>
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML Sun May 23 10:38:11 2010 +0100
@@ -209,7 +209,7 @@
Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)
end
else if loose_bvar1 (rand,0) then (*B or eta*)
- if rand = Bound 0 then eta_conversion ct
+ if rand = Bound 0 then Thm.eta_conversion ct
else (*B*)
let val crand = cterm_of thy (Abs(x,xT,rand))
val crator = cterm_of thy rator
@@ -225,7 +225,7 @@
(*Traverse a theorem, declaring abstraction function definitions. String s is the suggested
prefix for the constants.*)
fun combinators_aux ct =
- if lambda_free (term_of ct) then reflexive ct
+ if lambda_free (term_of ct) then Thm.reflexive ct
else
case term_of ct of
Abs _ =>
@@ -234,17 +234,17 @@
val u_th = combinators_aux cta
val cu = Thm.rhs_of u_th
val comb_eq = abstract (Thm.cabs cv cu)
- in transitive (abstract_rule v cv u_th) comb_eq end
+ in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end
| _ $ _ =>
let val (ct1, ct2) = Thm.dest_comb ct
- in combination (combinators_aux ct1) (combinators_aux ct2) end;
+ in Thm.combination (combinators_aux ct1) (combinators_aux ct2) end;
fun combinators th =
if lambda_free (prop_of th) then th
else
let val th = Drule.eta_contraction_rule th
val eqth = combinators_aux (cprop_of th)
- in equal_elim eqth th end
+ in Thm.equal_elim eqth th end
handle THM (msg,_,_) =>
(warning (cat_lines
["Error in the combinator translation of " ^ Display.string_of_thm_without_context th,
@@ -298,7 +298,7 @@
(*Generate Skolem functions for a theorem supplied in nnf*)
fun assume_skolem_of_def s th =
- map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);
+ map (skolem_of_def o Thm.assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);
(*** Blacklisting (FIXME: duplicated in "Sledgehammer_Fact_Filter"?) ***)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML Sun May 23 10:38:11 2010 +0100
@@ -48,7 +48,7 @@
TyLitVar of string * name |
TyLitFree of string * name
exception CLAUSE of string * term
- val add_type_literals : typ list -> type_literal list
+ val type_literals_for_types : typ list -> type_literal list
val get_tvar_strs: typ list -> string list
datatype arLit =
TConsLit of class * string * string list
@@ -331,7 +331,8 @@
| pred_of_sort (TyLitFree (s, _)) = (s, 1)
(*Given a list of sorted type variables, return a list of type literals.*)
-fun add_type_literals Ts = fold (union (op =)) (map sorts_on_typs Ts) []
+fun type_literals_for_types Ts =
+ fold (union (op =)) (map sorts_on_typs Ts) []
(*The correct treatment of TFrees like 'a in lemmas (axiom clauses) is not clear.
* Ignoring them leads to unsound proofs, since we do nothing to ensure that 'a
@@ -520,7 +521,7 @@
dfg_forall vars ("or(" ^ commas lits ^ ")") ^ ",\n" ^
string_of_clausename (cls_id,ax_name) ^ ").\n\n";
-fun string_of_arity (name, num) = "(" ^ name ^ "," ^ Int.toString num ^ ")"
+fun string_of_arity (name, arity) = "(" ^ name ^ ", " ^ Int.toString arity ^ ")"
fun string_of_preds [] = ""
| string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML Sun May 23 10:38:11 2010 +0100
@@ -52,14 +52,6 @@
open Sledgehammer_FOL_Clause
open Sledgehammer_Fact_Preprocessor
-(* Parameter "full_types" below indicates that full type information is to be
- exported.
-
- If "explicit_apply" is false, each function will be directly applied to as
- many arguments as possible, avoiding use of the "apply" operator. Use of
- "hBOOL" is also minimized.
-*)
-
fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c);
(*True if the constant ever appears outside of the top-level position in literals.
@@ -308,7 +300,7 @@
let
val (lits, pool) = pool_map (tptp_literal params) literals pool
val (tylits, pool) = pool_map (tptp_of_type_literal pos)
- (add_type_literals ctypes_sorts) pool
+ (type_literals_for_types ctypes_sorts) pool
in ((lits, tylits), pool) end
fun tptp_clause params (cls as HOLClause {axiom_name, clause_id, kind, ...})
@@ -328,7 +320,8 @@
fun dfg_type_literals params pos (HOLClause {literals, ctypes_sorts, ...}) =
pool_map (dfg_literal params) literals
- #>> rpair (map (dfg_of_type_literal pos) (add_type_literals ctypes_sorts))
+ #>> rpair (map (dfg_of_type_literal pos)
+ (type_literals_for_types ctypes_sorts))
fun get_uvars (CombConst _) vars pool = (vars, pool)
| get_uvars (CombVar (name, _)) vars pool =
@@ -539,6 +532,8 @@
(* DFG format *)
+fun dfg_tfree_predicate s = (first_field "(" s |> the |> fst, 1)
+
fun write_dfg_file full_types explicit_apply file clauses =
let
(* Some of the helper functions below are not name-pool-aware. However,
@@ -551,13 +546,16 @@
val params = (full_types, explicit_apply, cma, cnh)
val ((conjecture_clss, tfree_litss), pool) =
pool_map (dfg_clause params) conjectures pool |>> ListPair.unzip
- and problem_name = Path.implode (Path.base file)
+ val tfree_lits = union_all tfree_litss
+ val problem_name = Path.implode (Path.base file)
val (axstrs, pool) = pool_map (apfst fst oo dfg_clause params) axclauses pool
- val tfree_clss = map dfg_tfree_clause (union_all tfree_litss)
+ val tfree_clss = map dfg_tfree_clause tfree_lits
+ val tfree_preds = map dfg_tfree_predicate tfree_lits
val (helper_clauses_strs, pool) =
pool_map (apfst fst oo dfg_clause params) helper_clauses pool
val (funcs, cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
- and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
+ val ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
+ val preds = tfree_preds @ cl_preds @ ty_preds
val conjecture_offset =
length axclauses + length classrel_clauses + length arity_clauses
+ length helper_clauses
@@ -567,7 +565,7 @@
string_of_start problem_name ::
string_of_descrip problem_name ::
string_of_symbols (string_of_funcs funcs)
- (string_of_preds (cl_preds @ ty_preds)) ::
+ (string_of_preds preds) ::
"list_of_clauses(axioms, cnf).\n" ::
axstrs @
map dfg_classrel_clause classrel_clauses @
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Sun May 23 10:38:11 2010 +0100
@@ -24,8 +24,6 @@
open ATP_Systems
open Sledgehammer_Fact_Minimizer
-structure K = OuterKeyword and P = OuterParse
-
(** Proof method used in Isar proofs **)
val neg_clausify_setup =
@@ -36,7 +34,7 @@
(** Attribute for converting a theorem into clauses **)
val parse_clausify_attribute : attribute context_parser =
- Scan.lift OuterParse.nat
+ Scan.lift Parse.nat
>> (fn i => Thm.rule_attribute (fn context => fn th =>
let val thy = Context.theory_of context in
Meson.make_meta_clause (nth (cnf_axiom thy th) i)
@@ -94,11 +92,11 @@
("explicit_apply", "false"),
("respect_no_atp", "true"),
("relevance_threshold", "50"),
- ("convergence", "320"),
+ ("relevance_convergence", "320"),
("theory_relevant", "smart"),
- ("follow_defs", "false"),
+ ("defs_relevant", "false"),
("isar_proof", "false"),
- ("shrink_factor", "1"),
+ ("isar_shrink_factor", "1"),
("minimize_timeout", "5 s")]
val alias_params =
@@ -111,12 +109,12 @@
("implicit_apply", "explicit_apply"),
("ignore_no_atp", "respect_no_atp"),
("theory_irrelevant", "theory_relevant"),
- ("dont_follow_defs", "follow_defs"),
+ ("defs_irrelevant", "defs_relevant"),
("no_isar_proof", "isar_proof")]
val params_for_minimize =
["debug", "verbose", "overlord", "full_types", "explicit_apply",
- "isar_proof", "shrink_factor", "minimize_timeout"]
+ "isar_proof", "isar_shrink_factor", "minimize_timeout"]
val property_dependent_params = ["atps", "full_types", "timeout"]
@@ -194,21 +192,22 @@
val respect_no_atp = lookup_bool "respect_no_atp"
val relevance_threshold =
0.01 * Real.fromInt (lookup_int "relevance_threshold")
- val convergence = 0.01 * Real.fromInt (lookup_int "convergence")
+ val relevance_convergence =
+ 0.01 * Real.fromInt (lookup_int "relevance_convergence")
val theory_relevant = lookup_bool_option "theory_relevant"
- val follow_defs = lookup_bool "follow_defs"
+ val defs_relevant = lookup_bool "defs_relevant"
val isar_proof = lookup_bool "isar_proof"
- val shrink_factor = Int.max (1, lookup_int "shrink_factor")
+ val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
val timeout = lookup_time "timeout"
val minimize_timeout = lookup_time "minimize_timeout"
in
{debug = debug, verbose = verbose, overlord = overlord, atps = atps,
full_types = full_types, explicit_apply = explicit_apply,
respect_no_atp = respect_no_atp, relevance_threshold = relevance_threshold,
- convergence = convergence, theory_relevant = theory_relevant,
- follow_defs = follow_defs, isar_proof = isar_proof,
- shrink_factor = shrink_factor, timeout = timeout,
- minimize_timeout = minimize_timeout}
+ relevance_convergence = relevance_convergence,
+ theory_relevant = theory_relevant, defs_relevant = defs_relevant,
+ isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
+ timeout = timeout, minimize_timeout = minimize_timeout}
end
fun get_params thy = extract_params thy (default_raw_params thy)
@@ -320,13 +319,13 @@
params |> map string_for_raw_param
|> sort_strings |> cat_lines)))))
-val parse_key = Scan.repeat1 P.typ_group >> space_implode " "
-val parse_value = Scan.repeat1 P.xname
-val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) []
-val parse_params = Scan.optional (Args.bracks (P.list parse_param)) []
+val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " "
+val parse_value = Scan.repeat1 Parse.xname
+val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) []
+val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) []
val parse_fact_refs =
- Scan.repeat1 (Scan.unless (P.name -- Args.colon)
- (P.xname -- Scan.option Attrib.thm_sel)
+ Scan.repeat1 (Scan.unless (Parse.name -- Args.colon)
+ (Parse.xname -- Scan.option Attrib.thm_sel)
>> (fn (name, interval) =>
Facts.Named ((name, Position.none), interval)))
val parse_relevance_chunk =
@@ -339,18 +338,18 @@
>> merge_relevance_overrides))
(add_to_relevance_override [])
val parse_sledgehammer_command =
- (Scan.optional P.short_ident runN -- parse_params -- parse_relevance_override
- -- Scan.option P.nat) #>> sledgehammer_trans
+ (Scan.optional Parse.short_ident runN -- parse_params -- parse_relevance_override
+ -- Scan.option Parse.nat) #>> sledgehammer_trans
val parse_sledgehammer_params_command =
parse_params #>> sledgehammer_params_trans
val _ =
- OuterSyntax.improper_command sledgehammerN
- "search for first-order proof using automatic theorem provers" K.diag
+ Outer_Syntax.improper_command sledgehammerN
+ "search for first-order proof using automatic theorem provers" Keyword.diag
parse_sledgehammer_command
val _ =
- OuterSyntax.command sledgehammer_paramsN
- "set and display the default parameters for Sledgehammer" K.thy_decl
+ Outer_Syntax.command sledgehammer_paramsN
+ "set and display the default parameters for Sledgehammer" Keyword.thy_decl
parse_sledgehammer_params_command
val setup =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Sun May 23 10:38:11 2010 +0100
@@ -14,7 +14,6 @@
val invert_const: string -> string
val invert_type_const: string -> string
val num_type_args: theory -> string -> int
- val make_tvar: string -> typ
val strip_prefix: string -> string -> string option
val metis_line: int -> int -> string list -> string
val metis_proof_text:
@@ -235,26 +234,27 @@
SOME c' => c'
| NONE => c;
-fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS);
-fun make_tparam s = TypeInfer.param 0 (s, HOLogic.typeS)
-fun make_var (b,T) = Var((b,0),T);
-
(* Type variables are given the basic sort "HOL.type". Some will later be
constrained by information from type literals, or by type inference. *)
-fun type_from_node (u as IntLeaf _) = raise NODE [u]
- | type_from_node (u as StrNode (a, us)) =
- let val Ts = map type_from_node us in
+fun type_from_node _ (u as IntLeaf _) = raise NODE [u]
+ | type_from_node tfrees (u as StrNode (a, us)) =
+ let val Ts = map (type_from_node tfrees) us in
case strip_prefix tconst_prefix a of
SOME b => Type (invert_type_const b, Ts)
| NONE =>
if not (null us) then
raise NODE [u] (* only "tconst"s have type arguments *)
else case strip_prefix tfree_prefix a of
- SOME b => TFree ("'" ^ b, HOLogic.typeS)
+ SOME b =>
+ let val s = "'" ^ b in
+ TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
+ end
| NONE =>
case strip_prefix tvar_prefix a of
- SOME b => make_tvar b
- | NONE => make_tparam a (* Variable from the ATP, say "X1" *)
+ SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
+ | NONE =>
+ (* Variable from the ATP, say "X1" *)
+ TypeInfer.param 0 (a, HOLogic.typeS)
end
(*Invert the table of translations between Isabelle and ATPs*)
@@ -287,7 +287,7 @@
(* First-order translation. No types are known for variables. "HOLogic.typeT"
should allow them to be inferred.*)
-fun term_from_node thy full_types =
+fun term_from_node thy full_types tfrees =
let
fun aux opt_T args u =
case u of
@@ -298,7 +298,8 @@
| StrNode (a, us) =>
if a = type_wrapper_name then
case us of
- [term_u, typ_u] => aux (SOME (type_from_node typ_u)) args term_u
+ [term_u, typ_u] =>
+ aux (SOME (type_from_node tfrees typ_u)) args term_u
| _ => raise NODE us
else case strip_prefix const_prefix a of
SOME "equal" =>
@@ -324,7 +325,8 @@
(* Extra args from "hAPP" come after any arguments
given directly to the constant. *)
Sign.const_instance thy (c,
- map type_from_node (drop num_term_args us)))
+ map (type_from_node tfrees)
+ (drop num_term_args us)))
in list_comb (t, ts) end
| NONE => (* a free or schematic variable *)
let
@@ -335,21 +337,22 @@
SOME b => Free (b, T)
| NONE =>
case strip_prefix schematic_var_prefix a of
- SOME b => make_var (b, T)
+ SOME b => Var ((b, 0), T)
| NONE =>
(* Variable from the ATP, say "X1" *)
- make_var (fix_atp_variable_name a, T)
+ Var ((fix_atp_variable_name a, 0), T)
in list_comb (t, ts) end
in aux end
(* Type class literal applied to a type. Returns triple of polarity, class,
type. *)
-fun type_constraint_from_node pos (StrNode ("c_Not", [u])) =
- type_constraint_from_node (not pos) u
- | type_constraint_from_node pos u = case u of
+fun type_constraint_from_node pos tfrees (StrNode ("c_Not", [u])) =
+ type_constraint_from_node (not pos) tfrees u
+ | type_constraint_from_node pos tfrees u = case u of
IntLeaf _ => raise NODE [u]
| StrNode (a, us) =>
- (case (strip_prefix class_prefix a, map type_from_node us) of
+ (case (strip_prefix class_prefix a,
+ map (type_from_node tfrees) us) of
(SOME b, [T]) => (pos, b, T)
| _ => raise NODE [u])
@@ -395,24 +398,24 @@
|> clause_for_literals thy
(*Accumulate sort constraints in vt, with "real" literals in lits.*)
-fun lits_of_nodes thy full_types (vt, lits) us =
- case us of
- [] => (vt, finish_clause thy lits)
- | (u :: us) =>
- lits_of_nodes thy full_types
- (add_type_constraint (type_constraint_from_node true u) vt, lits) us
- handle NODE _ =>
- lits_of_nodes thy full_types
- (vt, term_from_node thy full_types (SOME @{typ bool})
- [] u :: lits) us
+fun lits_of_nodes thy full_types tfrees =
+ let
+ fun aux (vt, lits) [] = (vt, finish_clause thy lits)
+ | aux (vt, lits) (u :: us) =
+ aux (add_type_constraint
+ (type_constraint_from_node true tfrees u) vt, lits) us
+ handle NODE _ =>
+ aux (vt, term_from_node thy full_types tfrees (SOME @{typ bool})
+ [] u :: lits) us
+ in aux end
-(*Update TVars/TFrees with detected sort constraints.*)
-fun repair_sorts vt =
+(* Update TVars with detected sort constraints. It's not totally clear when
+ this code is necessary. *)
+fun repair_tvar_sorts vt =
let
fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
| do_type (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi))
- | do_type (TFree (x, s)) =
- TFree (x, the_default s (Vartab.lookup vt (x, ~1)))
+ | do_type (TFree z) = TFree z
fun do_term (Const (a, T)) = Const (a, do_type T)
| do_term (Free (a, T)) = Free (a, do_type T)
| do_term (Var (xi, T)) = Var (xi, do_type T)
@@ -444,45 +447,28 @@
(* Interpret a list of syntax trees as a clause, given by "real" literals and
sort constraints. "vt" holds the initial sort constraints, from the
conjecture clauses. *)
-fun clause_of_nodes ctxt full_types vt us =
+fun clause_of_nodes ctxt full_types tfrees us =
let
val thy = ProofContext.theory_of ctxt
- val (vt, t) = lits_of_nodes thy full_types (vt, []) us
- in repair_sorts vt t end
+ val (vt, t) = lits_of_nodes thy full_types tfrees (Vartab.empty, []) us
+ in repair_tvar_sorts vt t end
fun check_formula ctxt =
TypeInfer.constrain @{typ bool}
#> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
-(** Global sort constraints on TFrees (from tfree_tcs) are positive unit
- clauses. **)
-
-fun add_tfree_constraint (true, cl, TFree (a, _)) = add_var ((a, ~1), cl)
- | add_tfree_constraint _ = I
-fun tfree_constraints_of_clauses vt [] = vt
- | tfree_constraints_of_clauses vt ([lit] :: uss) =
- (tfree_constraints_of_clauses (add_tfree_constraint
- (type_constraint_from_node true lit) vt) uss
- handle NODE _ => (* Not a positive type constraint? Ignore the literal. *)
- tfree_constraints_of_clauses vt uss)
- | tfree_constraints_of_clauses vt (_ :: uss) =
- tfree_constraints_of_clauses vt uss
-
(**** Translation of TSTP files to Isar Proofs ****)
fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
| unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
-fun clauses_in_lines (Definition (_, u, us)) = u :: us
- | clauses_in_lines (Inference (_, us, _)) = us
-
-fun decode_line full_types vt (Definition (num, u, us)) ctxt =
+fun decode_line full_types tfrees (Definition (num, u, us)) ctxt =
let
- val t1 = clause_of_nodes ctxt full_types vt [u]
+ val t1 = clause_of_nodes ctxt full_types tfrees [u]
val vars = snd (strip_comb t1)
val frees = map unvarify_term vars
val unvarify_args = subst_atomic (vars ~~ frees)
- val t2 = clause_of_nodes ctxt full_types vt us
+ val t2 = clause_of_nodes ctxt full_types tfrees us
val (t1, t2) =
HOLogic.eq_const HOLogic.typeT $ t1 $ t2
|> unvarify_args |> uncombine_term |> check_formula ctxt
@@ -491,19 +477,16 @@
(Definition (num, t1, t2),
fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
end
- | decode_line full_types vt (Inference (num, us, deps)) ctxt =
+ | decode_line full_types tfrees (Inference (num, us, deps)) ctxt =
let
- val t = us |> clause_of_nodes ctxt full_types vt
+ val t = us |> clause_of_nodes ctxt full_types tfrees
|> unskolemize_term |> uncombine_term |> check_formula ctxt
in
(Inference (num, t, deps),
fold Variable.declare_term (OldTerm.term_frees t) ctxt)
end
-fun decode_lines ctxt full_types lines =
- let
- val vt = tfree_constraints_of_clauses Vartab.empty
- (map clauses_in_lines lines)
- in #1 (fold_map (decode_line full_types vt) lines ctxt) end
+fun decode_lines ctxt full_types tfrees lines =
+ fst (fold_map (decode_line full_types tfrees) lines ctxt)
fun aint_inference _ (Definition _) = true
| aint_inference t (Inference (_, t', _)) = not (t aconv t')
@@ -569,7 +552,7 @@
fun add_desired_line _ _ _ _ _ (line as Definition _) (j, lines) =
(j, line :: lines)
- | add_desired_line ctxt shrink_factor conjecture_shape thm_names frees
+ | add_desired_line ctxt isar_shrink_factor conjecture_shape thm_names frees
(Inference (num, t, deps)) (j, lines) =
(j + 1,
if is_axiom_clause_number thm_names num orelse
@@ -579,7 +562,7 @@
not (exists_subterm (is_bad_free frees) t) andalso
not (is_trivial_formula t) andalso
(null lines orelse (* last line must be kept *)
- (length deps >= 2 andalso j mod shrink_factor = 0))) then
+ (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then
Inference (num, t, deps) :: lines (* keep line *)
else
map (replace_deps_in_line (num, deps)) lines) (* drop line *)
@@ -599,8 +582,7 @@
| extract_num _ = NONE
in atp_proof |> split_lines |> map_filter (extract_num o tokens_of) end
-(* Used to label theorems chained into the Sledgehammer call (or rather
- goal?) *)
+(* Used to label theorems chained into the goal. *)
val chained_hint = "sledgehammer_chained"
fun apply_command _ 1 = "by "
@@ -674,16 +656,16 @@
forall_vars t,
ByMetis (fold (add_fact_from_dep thm_names) deps ([], [])))
-fun proof_from_atp_proof pool ctxt full_types shrink_factor atp_proof
- conjecture_shape thm_names params frees =
+fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
+ atp_proof conjecture_shape thm_names params frees =
let
val lines =
atp_proof ^ "$" (* the $ sign acts as a sentinel *)
|> parse_proof pool
- |> decode_lines ctxt full_types
+ |> decode_lines ctxt full_types tfrees
|> rpair [] |-> fold_rev (add_line conjecture_shape thm_names)
|> rpair [] |-> fold_rev add_nontrivial_line
- |> rpair (0, []) |-> fold_rev (add_desired_line ctxt shrink_factor
+ |> rpair (0, []) |-> fold_rev (add_desired_line ctxt isar_shrink_factor
conjecture_shape thm_names frees)
|> snd
in
@@ -839,7 +821,7 @@
apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
fun do_step (step as Assume (l, t)) (proof, subst, assums) =
(case AList.lookup (op aconv) assums t of
- SOME l' => (proof, (l', l) :: subst, assums)
+ SOME l' => (proof, (l, l') :: subst, assums)
| NONE => (step :: proof, subst, (t, l) :: assums))
| do_step (Have (qs, l, t, by)) (proof, subst, assums) =
(Have (qs, l, t,
@@ -981,19 +963,21 @@
do_indent 0 ^ (if n <> 1 then "next" else "qed") ^ "\n"
in do_proof end
-fun isar_proof_text (pool, debug, full_types, shrink_factor, ctxt,
+fun isar_proof_text (pool, debug, full_types, isar_shrink_factor, ctxt,
conjecture_shape)
(minimize_command, atp_proof, thm_names, goal, i) =
let
val thy = ProofContext.theory_of ctxt
val (params, hyp_ts, concl_t) = strip_subgoal goal i
val frees = fold Term.add_frees (concl_t :: hyp_ts) []
+ val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
val n = Logic.count_prems (prop_of goal)
val (one_line_proof, lemma_names) =
metis_proof_text (minimize_command, atp_proof, thm_names, goal, i)
fun isar_proof_for () =
- case proof_from_atp_proof pool ctxt full_types shrink_factor atp_proof
- conjecture_shape thm_names params frees
+ case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
+ atp_proof conjecture_shape thm_names params
+ frees
|> redirect_proof thy conjecture_shape hyp_ts concl_t
|> kill_duplicate_assumptions_in_proof
|> then_chain_proof
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Sun May 23 10:38:11 2010 +0100
@@ -102,7 +102,7 @@
let val s = unyxml y in
y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
- OuterKeyword.is_keyword s) ? quote
+ Keyword.is_keyword s) ? quote
end
fun monomorphic_term subst t =
--- a/src/HOL/Tools/TFL/casesplit.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/TFL/casesplit.ML Sun May 23 10:38:11 2010 +0100
@@ -79,8 +79,8 @@
(* beta-eta contract the theorem *)
fun beta_eta_contract thm =
let
- val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
- val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
+ val thm2 = Thm.equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
+ val thm3 = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
in thm3 end;
(* make a casethm from an induction thm *)
--- a/src/HOL/Tools/TFL/rules.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/TFL/rules.ML Sun May 23 10:38:11 2010 +0100
@@ -172,7 +172,7 @@
val [P,Q] = OldTerm.term_vars prop
val disj1 = Thm.forall_intr (Thm.cterm_of thy Q) disjI1
in
-fun DISJ1 thm tm = thm RS (forall_elim (D.drop_prop tm) disj1)
+fun DISJ1 thm tm = thm RS (Thm.forall_elim (D.drop_prop tm) disj1)
handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg;
end;
@@ -181,7 +181,7 @@
val [P,Q] = OldTerm.term_vars prop
val disj2 = Thm.forall_intr (Thm.cterm_of thy P) disjI2
in
-fun DISJ2 tm thm = thm RS (forall_elim (D.drop_prop tm) disj2)
+fun DISJ2 tm thm = thm RS (Thm.forall_elim (D.drop_prop tm) disj2)
handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg;
end;
@@ -276,11 +276,11 @@
val prop = Thm.prop_of spec
val x = hd (tl (OldTerm.term_vars prop))
val cTV = ctyp_of thy (type_of x)
- val gspec = forall_intr (cterm_of thy x) spec
+ val gspec = Thm.forall_intr (cterm_of thy x) spec
in
fun SPEC tm thm =
let val gspec' = instantiate ([(cTV, Thm.ctyp_of_term tm)], []) gspec
- in thm RS (forall_elim tm gspec') end
+ in thm RS (Thm.forall_elim tm gspec') end
end;
fun SPEC_ALL thm = fold SPEC (#1(D.strip_forall(cconcl thm))) thm;
@@ -302,7 +302,7 @@
ctm_theta s (Vartab.dest tm_theta))
in
fun GEN v th =
- let val gth = forall_intr v th
+ let val gth = Thm.forall_intr v th
val thy = Thm.theory_of_thm gth
val Const("all",_)$Abs(x,ty,rst) = Thm.prop_of gth
val P' = Abs(x,ty, HOLogic.dest_Trueprop rst) (* get rid of trueprop *)
@@ -533,9 +533,9 @@
*---------------------------------------------------------------------------*)
fun list_beta_conv tm =
- let fun rbeta th = Thm.transitive th (beta_conversion false (#2(D.dest_eq(cconcl th))))
+ let fun rbeta th = Thm.transitive th (Thm.beta_conversion false (#2(D.dest_eq(cconcl th))))
fun iter [] = Thm.reflexive tm
- | iter (v::rst) = rbeta (combination(iter rst) (Thm.reflexive v))
+ | iter (v::rst) = rbeta (Thm.combination(iter rst) (Thm.reflexive v))
in iter end;
@@ -619,7 +619,7 @@
fun VSTRUCT_ELIM tych a vstr th =
let val L = S.free_vars_lr vstr
val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr)))
- val thm1 = implies_intr bind1 (SUBS [SYM(assume bind1)] th)
+ val thm1 = Thm.implies_intr bind1 (SUBS [SYM(Thm.assume bind1)] th)
val thm2 = forall_intr_list (map tych L) thm1
val thm3 = forall_elim_list (XFILL tych a vstr) thm2
in refl RS
@@ -630,7 +630,7 @@
let val a1 = tych a
val vstr1 = tych vstr
in
- forall_intr a1
+ Thm.forall_intr a1
(if (is_Free vstr)
then cterm_instantiate [(vstr1,a1)] th
else VSTRUCT_ELIM tych a vstr th)
@@ -803,7 +803,7 @@
val names = OldTerm.add_term_names (term_of ctm, [])
val th1 = MetaSimplifier.rewrite_cterm(false,true,false)
(prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm
- val th2 = equal_elim th1 th
+ val th2 = Thm.equal_elim th1 th
in
(th2, filter_out restricted (!tc_list))
end;
--- a/src/HOL/Tools/TFL/tfl.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/TFL/tfl.ML Sun May 23 10:38:11 2010 +0100
@@ -556,8 +556,8 @@
else ()
val f_free = Free (fid, fastype_of f) (*'cos f is a Const*)
val SV' = map tych SV;
- val SVrefls = map reflexive SV'
- val def0 = (fold (fn x => fn th => R.rbeta(combination th x))
+ val SVrefls = map Thm.reflexive SV'
+ val def0 = (fold (fn x => fn th => R.rbeta(Thm.combination th x))
SVrefls def)
RS meta_eq_to_obj_eq
val def' = R.MP (R.SPEC (tych R') (R.GEN (tych R1) baz)) def0
--- a/src/HOL/Tools/choice_specification.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/choice_specification.ML Sun May 23 10:38:11 2010 +0100
@@ -183,7 +183,7 @@
fun process_single ((name,atts),rew_imp,frees) args =
let
fun undo_imps thm =
- equal_elim (symmetric rew_imp) thm
+ Thm.equal_elim (Thm.symmetric rew_imp) thm
fun add_final (arg as (thy, thm)) =
if name = ""
@@ -232,33 +232,28 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
-val opt_overloaded = P.opt_keyword "overloaded";
+val opt_name = Scan.optional (Parse.name --| Parse.$$$ ":") ""
+val opt_overloaded = Parse.opt_keyword "overloaded"
val specification_decl =
- P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop)
+ Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
+ Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop)
val _ =
- OuterSyntax.command "specification" "define constants by specification" K.thy_goal
+ Outer_Syntax.command "specification" "define constants by specification" Keyword.thy_goal
(specification_decl >> (fn (cos,alt_props) =>
Toplevel.print o (Toplevel.theory_to_proof
(process_spec NONE cos alt_props))))
val ax_specification_decl =
- P.name --
- (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop))
+ Parse.name --
+ (Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
+ Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop))
val _ =
- OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
+ Outer_Syntax.command "ax_specification" "define constants by specification" Keyword.thy_goal
(ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
Toplevel.print o (Toplevel.theory_to_proof
(process_spec (SOME axname) cos alt_props))))
end
-
-
-end
--- a/src/HOL/Tools/cnf_funcs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/cnf_funcs.ML Sun May 23 10:38:11 2010 +0100
@@ -435,9 +435,9 @@
val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_l (* ((Ex x') | y') = (Ex (x' | y')) *)
val var = new_free ()
val thm2 = make_cnfx_disj_thm (betapply (x', var)) y' (* (x' | y') = body' *)
- val thm3 = forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
- val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
- val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
+ val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
+ val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
+ val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
in
iff_trans OF [thm1, thm5] (* ((Ex x') | y') = (Ex v. body') *)
end
@@ -446,9 +446,9 @@
val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_r (* (x' | (Ex y')) = (Ex (x' | y')) *)
val var = new_free ()
val thm2 = make_cnfx_disj_thm x' (betapply (y', var)) (* (x' | y') = body' *)
- val thm3 = forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
- val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
- val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
+ val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
+ val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
+ val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
in
iff_trans OF [thm1, thm5] (* (x' | (Ex y')) = (EX v. body') *)
end
@@ -466,8 +466,8 @@
val var = new_free ()
val body = HOLogic.mk_conj (HOLogic.mk_disj (x, var), HOLogic.mk_disj (y, HOLogic.Not $ var))
val thm2 = make_cnfx_thm_from_nnf body (* (x | v) & (y | ~v) = body' *)
- val thm3 = forall_intr (cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
- val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x | v) & (y | ~v) = body' *)
+ val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
+ val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x | v) & (y | ~v) = body' *)
val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x | v) & (y | ~v)) = (EX v. body') *)
in
iff_trans OF [thm1, thm5]
--- a/src/HOL/Tools/inductive.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/inductive.ML Sun May 23 10:38:11 2010 +0100
@@ -83,8 +83,7 @@
(binding * string option * mixfix) list ->
(Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
bool -> local_theory -> inductive_result * local_theory
- val gen_ind_decl: add_ind_def -> bool ->
- OuterParse.token list -> (bool -> local_theory -> local_theory) * OuterParse.token list
+ val gen_ind_decl: add_ind_def -> bool -> (bool -> local_theory -> local_theory) parser
end;
structure Inductive: INDUCTIVE =
@@ -971,32 +970,28 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = OuterKeyword.keyword "monos";
+val _ = Keyword.keyword "monos";
fun gen_ind_decl mk_def coind =
- P.fixes -- P.for_fixes --
- Scan.optional SpecParse.where_alt_specs [] --
- Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
+ Parse.fixes -- Parse.for_fixes --
+ Scan.optional Parse_Spec.where_alt_specs [] --
+ Scan.optional (Parse.$$$ "monos" |-- Parse.!!! Parse_Spec.xthms1) []
>> (fn (((preds, params), specs), monos) =>
(snd oo gen_add_inductive mk_def true coind preds params specs monos));
val ind_decl = gen_ind_decl add_ind_def;
val _ =
- OuterSyntax.local_theory' "inductive" "define inductive predicates" K.thy_decl
+ Outer_Syntax.local_theory' "inductive" "define inductive predicates" Keyword.thy_decl
(ind_decl false);
val _ =
- OuterSyntax.local_theory' "coinductive" "define coinductive predicates" K.thy_decl
+ Outer_Syntax.local_theory' "coinductive" "define coinductive predicates" Keyword.thy_decl
(ind_decl true);
val _ =
- OuterSyntax.local_theory "inductive_cases"
- "create simplified instances of elimination rules (improper)" K.thy_script
- (P.and_list1 SpecParse.specs >> (snd oo inductive_cases));
+ Outer_Syntax.local_theory "inductive_cases"
+ "create simplified instances of elimination rules (improper)" Keyword.thy_script
+ (Parse.and_list1 Parse_Spec.specs >> (snd oo inductive_cases));
end;
-
-end;
--- a/src/HOL/Tools/inductive_codegen.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML Sun May 23 10:38:11 2010 +0100
@@ -775,7 +775,7 @@
add_codegen "inductive" inductive_codegen #>
Attrib.setup @{binding code_ind}
(Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
- Scan.option (Args.$$$ "params" |-- Args.colon |-- OuterParse.nat) >> uncurry add))
+ Scan.option (Args.$$$ "params" |-- Args.colon |-- Parse.nat) >> uncurry add))
"introduction rules for executable predicates";
(**** Quickcheck involving inductive predicates ****)
--- a/src/HOL/Tools/inductive_realizer.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML Sun May 23 10:38:11 2010 +0100
@@ -21,7 +21,7 @@
[name] => name
| _ => error ("name_of_thm: bad proof of theorem\n" ^ Display.string_of_thm_without_context thm));
-val all_simps = map (symmetric o mk_meta_eq) @{thms HOL.all_simps};
+val all_simps = map (Thm.symmetric o mk_meta_eq) @{thms HOL.all_simps};
fun prf_of thm =
let
--- a/src/HOL/Tools/inductive_set.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/inductive_set.ML Sun May 23 10:38:11 2010 +0100
@@ -114,9 +114,9 @@
let val rews = map mk_rew ts
in
if forall is_none rews then NONE
- else SOME (fold (fn th1 => fn th2 => combination th2 th1)
- (map2 (fn SOME r => K r | NONE => reflexive o cterm_of thy)
- rews ts) (reflexive (cterm_of thy h)))
+ else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1)
+ (map2 (fn SOME r => K r | NONE => Thm.reflexive o cterm_of thy)
+ rews ts) (Thm.reflexive (cterm_of thy h)))
end
| NONE => NONE)
| _ => NONE
@@ -562,18 +562,14 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def;
val _ =
- OuterSyntax.local_theory' "inductive_set" "define inductive sets" K.thy_decl
+ Outer_Syntax.local_theory' "inductive_set" "define inductive sets" Keyword.thy_decl
(ind_set_decl false);
val _ =
- OuterSyntax.local_theory' "coinductive_set" "define coinductive sets" K.thy_decl
+ Outer_Syntax.local_theory' "coinductive_set" "define coinductive sets" Keyword.thy_decl
(ind_set_decl true);
end;
-
-end;
--- a/src/HOL/Tools/int_arith.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/int_arith.ML Sun May 23 10:38:11 2010 +0100
@@ -21,7 +21,7 @@
That is, m and n consist only of 1s combined with "+", "-" and "*".
*)
-val zeroth = (symmetric o mk_meta_eq) @{thm of_int_0};
+val zeroth = (Thm.symmetric o mk_meta_eq) @{thm of_int_0};
val lhss0 = [@{cpat "0::?'a::ring"}];
@@ -35,7 +35,7 @@
make_simproc {lhss = lhss0, name = "zero_to_of_int_zero_simproc",
proc = proc0, identifier = []};
-val oneth = (symmetric o mk_meta_eq) @{thm of_int_1};
+val oneth = (Thm.symmetric o mk_meta_eq) @{thm of_int_1};
val lhss1 = [@{cpat "1::?'a::ring_1"}];
--- a/src/HOL/Tools/numeral_simprocs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML Sun May 23 10:38:11 2010 +0100
@@ -614,12 +614,12 @@
fun prove_nz ss T t =
let
- val z = instantiate_cterm ([(zT,T)],[]) zr
- val eq = instantiate_cterm ([(eqT,T)],[]) geq
+ val z = Thm.instantiate_cterm ([(zT,T)],[]) zr
+ val eq = Thm.instantiate_cterm ([(eqT,T)],[]) geq
val th = Simplifier.rewrite (ss addsimps @{thms simp_thms})
(Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"}
(Thm.capply (Thm.capply eq t) z)))
- in equal_elim (symmetric th) TrueI
+ in Thm.equal_elim (Thm.symmetric th) TrueI
end
fun proc phi ss ct =
@@ -630,7 +630,7 @@
val T = ctyp_of_term x
val [y_nz, z_nz] = map (prove_nz ss T) [y, z]
val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq
- in SOME (implies_elim (implies_elim th y_nz) z_nz)
+ in SOME (Thm.implies_elim (Thm.implies_elim th y_nz) z_nz)
end
handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE
@@ -643,13 +643,13 @@
let val (x,y) = Thm.dest_binop l val z = r
val _ = map (HOLogic.dest_number o term_of) [x,y,z]
val ynz = prove_nz ss T y
- in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz)
+ in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz)
end
| (_, Const (@{const_name Rings.divide},_)$_$_) =>
let val (x,y) = Thm.dest_binop r val z = l
val _ = map (HOLogic.dest_number o term_of) [x,y,z]
val ynz = prove_nz ss T y
- in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz)
+ in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz)
end
| _ => NONE)
end
--- a/src/HOL/Tools/primrec.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/primrec.ML Sun May 23 10:38:11 2010 +0100
@@ -307,28 +307,26 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val opt_unchecked_name =
- Scan.optional (P.$$$ "(" |-- P.!!!
- (((P.$$$ "unchecked" >> K true) -- Scan.optional P.name "" ||
- P.name >> pair false) --| P.$$$ ")")) (false, "");
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!!
+ (((Parse.$$$ "unchecked" >> K true) -- Scan.optional Parse.name "" ||
+ Parse.name >> pair false) --| Parse.$$$ ")")) (false, "");
val old_primrec_decl =
- opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop);
+ opt_unchecked_name --
+ Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop);
-val primrec_decl = P.opt_target -- P.fixes -- SpecParse.where_alt_specs;
+val primrec_decl = Parse.opt_target -- Parse.fixes -- Parse_Spec.where_alt_specs;
val _ =
- OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
+ Outer_Syntax.command "primrec" "define primitive recursive functions on datatypes"
+ Keyword.thy_decl
((primrec_decl >> (fn ((opt_target, fixes), specs) =>
Toplevel.local_theory opt_target (add_primrec_cmd fixes specs #> snd)))
|| (old_primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
Toplevel.theory (snd o
(if unchecked then OldPrimrec.add_primrec_unchecked else OldPrimrec.add_primrec)
- alt_name (map P.triple_swap eqns) o
+ alt_name (map Parse.triple_swap eqns) o
tap (fn _ => legacy_feature "Old variant of 'primrec' command -- use new syntax instead")))));
end;
-
-end;
--- a/src/HOL/Tools/recdef.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/recdef.ML Sun May 23 10:38:11 2010 +0100
@@ -289,40 +289,39 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword ["permissive", "congs", "hints"];
+val _ = List.app Keyword.keyword ["permissive", "congs", "hints"];
val hints =
- P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- Args.parse) --| P.$$$ ")") >> Args.src;
+ Parse.$$$ "(" |--
+ Parse.!!! (Parse.position (Parse.$$$ "hints" -- Args.parse) --| Parse.$$$ ")") >> Args.src;
val recdef_decl =
- Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
- P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
+ Scan.optional
+ (Parse.$$$ "(" -- Parse.!!! (Parse.$$$ "permissive" -- Parse.$$$ ")") >> K false) true --
+ Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)
-- Scan.option hints
- >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
+ >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src);
val _ =
- OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
+ Outer_Syntax.command "recdef" "define general recursive functions (TFL)" Keyword.thy_decl
(recdef_decl >> Toplevel.theory);
val defer_recdef_decl =
- P.name -- Scan.repeat1 P.prop --
- Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (SpecParse.xthms1 --| P.$$$ ")")) []
+ Parse.name -- Scan.repeat1 Parse.prop --
+ Scan.optional
+ (Parse.$$$ "(" |-- Parse.$$$ "congs" |-- Parse.!!! (Parse_Spec.xthms1 --| Parse.$$$ ")")) []
>> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
val _ =
- OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
+ Outer_Syntax.command "defer_recdef" "defer general recursive functions (TFL)" Keyword.thy_decl
(defer_recdef_decl >> Toplevel.theory);
val _ =
- OuterSyntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
- K.thy_goal
- ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.xname --
- Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
+ Outer_Syntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
+ Keyword.thy_goal
+ ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.xname --
+ Scan.option (Parse.$$$ "(" |-- Parse.nat --| Parse.$$$ ")")
>> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
end;
-
-end;
--- a/src/HOL/Tools/record.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/record.ML Sun May 23 10:38:11 2010 +0100
@@ -2171,7 +2171,7 @@
fun get_upd_acc_congs () =
let
- val symdefs = map symmetric (sel_defs @ upd_defs);
+ val symdefs = map Thm.symmetric (sel_defs @ upd_defs);
val fold_ss = HOL_basic_ss addsimps symdefs;
val ua_congs = map (Drule.export_without_context o simplify fold_ss) upd_acc_cong_assists;
in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end;
@@ -2257,17 +2257,17 @@
val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l);
val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r);
val thl =
- assume cl (*All r. P r*) (* 1 *)
- |> obj_to_meta_all (*!!r. P r*)
- |> equal_elim split_meta' (*!!n m more. P (ext n m more)*)
- |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*)
- |> implies_intr cl (* 1 ==> 2 *)
+ Thm.assume cl (*All r. P r*) (* 1 *)
+ |> obj_to_meta_all (*!!r. P r*)
+ |> Thm.equal_elim split_meta' (*!!n m more. P (ext n m more)*)
+ |> meta_to_obj_all (*All n m more. P (ext n m more)*) (* 2*)
+ |> Thm.implies_intr cl (* 1 ==> 2 *)
val thr =
- assume cr (*All n m more. P (ext n m more)*)
- |> obj_to_meta_all (*!!n m more. P (ext n m more)*)
- |> equal_elim (symmetric split_meta') (*!!r. P r*)
- |> meta_to_obj_all (*All r. P r*)
- |> implies_intr cr (* 2 ==> 1 *)
+ Thm.assume cr (*All n m more. P (ext n m more)*)
+ |> obj_to_meta_all (*!!n m more. P (ext n m more)*)
+ |> Thm.equal_elim (Thm.symmetric split_meta') (*!!r. P r*)
+ |> meta_to_obj_all (*All r. P r*)
+ |> Thm.implies_intr cr (* 2 ==> 1 *)
in thr COMP (thl COMP iffI) end;
@@ -2460,17 +2460,14 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.command "record" "define extensible record" K.thy_decl
- (P.type_args_constrained -- P.binding --
- (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const_binding)
+ Outer_Syntax.command "record" "define extensible record" Keyword.thy_decl
+ (Parse.type_args_constrained -- Parse.binding --
+ (Parse.$$$ "=" |-- Scan.option (Parse.typ --| Parse.$$$ "+") --
+ Scan.repeat1 Parse.const_binding)
>> (fn (x, (y, z)) => Toplevel.theory (add_record_cmd false x y z)));
end;
-end;
-
structure Basic_Record: BASIC_RECORD = Record;
open Basic_Record;
--- a/src/HOL/Tools/refute_isar.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/refute_isar.ML Sun May 23 10:38:11 2010 +0100
@@ -12,19 +12,16 @@
(*optional list of arguments of the form [name1=value1, name2=value2, ...]*)
-val scan_parm =
- OuterParse.name
- -- (Scan.optional (OuterParse.$$$ "=" |-- OuterParse.name) "true")
-val scan_parms = Scan.optional
- (OuterParse.$$$ "[" |-- OuterParse.list scan_parm --| OuterParse.$$$ "]") [];
+val scan_parm = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- Parse.name) "true")
+val scan_parms = Scan.optional (Parse.$$$ "[" |-- Parse.list scan_parm --| Parse.$$$ "]") [];
(* 'refute' command *)
val _ =
- OuterSyntax.improper_command "refute"
- "try to find a model that refutes a given subgoal" OuterKeyword.diag
- (scan_parms -- Scan.optional OuterParse.nat 1 >>
+ Outer_Syntax.improper_command "refute"
+ "try to find a model that refutes a given subgoal" Keyword.diag
+ (scan_parms -- Scan.optional Parse.nat 1 >>
(fn (parms, i) =>
Toplevel.keep (fn state =>
let
@@ -36,8 +33,8 @@
(* 'refute_params' command *)
val _ =
- OuterSyntax.command "refute_params"
- "show/store default parameters for the 'refute' command" OuterKeyword.thy_decl
+ Outer_Syntax.command "refute_params"
+ "show/store default parameters for the 'refute' command" Keyword.thy_decl
(scan_parms >> (fn parms =>
Toplevel.theory (fn thy =>
let
--- a/src/HOL/Tools/semiring_normalizer.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/semiring_normalizer.ML Sun May 23 10:38:11 2010 +0100
@@ -147,7 +147,7 @@
val semiring = (sr_ops, sr_rules');
val ring = (r_ops, r_rules');
val field = (f_ops, f_rules');
- val ideal' = map (symmetric o mk_meta) ideal
+ val ideal' = map (Thm.symmetric o mk_meta) ideal
in
AList.delete Thm.eq_thm key #>
cons (key, ({vars = vars, semiring = semiring,
@@ -328,16 +328,16 @@
((let val (rx,rn) = dest_pow r
val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29
val (tm1,tm2) = Thm.dest_comb(concl th1) in
- transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)
+ Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)
handle CTERM _ =>
(let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31
val (tm1,tm2) = Thm.dest_comb(concl th1) in
- transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)) end)
+ Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)) end)
handle CTERM _ =>
((let val (rx,rn) = dest_pow r
val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30
val (tm1,tm2) = Thm.dest_comb(concl th1) in
- transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)
+ Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)
handle CTERM _ => inst_thm [(cx,l)] pthm_32
))
@@ -348,7 +348,7 @@
fun monomial_deone th =
(let val (l,r) = dest_mul(concl th) in
if l aconvc one_tm
- then transitive th (inst_thm [(ca,r)] pthm_13) else th end)
+ then Thm.transitive th (inst_thm [(ca,r)] pthm_13) else th end)
handle CTERM _ => th;
(* Conversion for "(monomial)^n", where n is a numeral. *)
@@ -357,7 +357,7 @@
let
fun monomial_pow tm bod ntm =
if not(is_comb bod)
- then reflexive tm
+ then Thm.reflexive tm
else
if is_semiring_constant bod
then semiring_pow_conv tm
@@ -365,7 +365,7 @@
let
val (lopr,r) = Thm.dest_comb bod
in if not(is_comb lopr)
- then reflexive tm
+ then Thm.reflexive tm
else
let
val (opr,l) = Thm.dest_comb lopr
@@ -374,7 +374,7 @@
then
let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34
val (l,r) = Thm.dest_comb(concl th1)
- in transitive th1 (Drule.arg_cong_rule l (nat_add_conv r))
+ in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv r))
end
else
if opr aconvc mul_tm
@@ -385,9 +385,9 @@
val (x,y) = Thm.dest_comb xy
val thl = monomial_pow y l ntm
val thr = monomial_pow z r ntm
- in transitive th1 (combination (Drule.arg_cong_rule x thl) thr)
+ in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule x thl) thr)
end
- else reflexive tm
+ else Thm.reflexive tm
end
end
in fn tm =>
@@ -436,18 +436,18 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2
- val th3 = transitive th1 th2
+ val th3 = Thm.transitive th1 th2
val (tm5,tm6) = Thm.dest_comb(concl th3)
val (tm7,tm8) = Thm.dest_comb tm6
val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8
- in transitive th3 (Drule.arg_cong_rule tm5 th4)
+ in Thm.transitive th3 (Drule.arg_cong_rule tm5 th4)
end
else
let val th0 = if ord < 0 then pthm_16 else pthm_17
val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm2
- in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
+ in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
end
end)
handle CTERM _ =>
@@ -459,14 +459,14 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2
- in transitive th1 th2
+ in Thm.transitive th1 th2
end
else
if ord < 0 then
let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm2
- in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
+ in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
end
else inst_thm [(ca,l),(cb,r)] pthm_09
end)) end)
@@ -480,22 +480,22 @@
let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
- in transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2)
+ in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2)
end
else if ord > 0 then
let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm2
- in transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
+ in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4))
end
- else reflexive tm
+ else Thm.reflexive tm
end)
handle CTERM _ =>
(let val vr = powvar r
val ord = vorder vl vr
in if ord = 0 then powvar_mul_conv tm
else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09
- else reflexive tm
+ else Thm.reflexive tm
end)) end))
in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r)
end
@@ -511,8 +511,8 @@
val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
- val th2 = combination (Drule.arg_cong_rule tm3 (monomial_mul_conv tm4)) (pmm_conv tm2)
- in transitive th1 th2
+ val th2 = Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv tm4)) (pmm_conv tm2)
+ in Thm.transitive th1 th2
end)
handle CTERM _ => monomial_mul_conv tm)
end
@@ -537,11 +537,11 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4)
- val th3 = transitive th1 (Drule.fun_cong_rule th2 tm2)
+ val th3 = Thm.transitive th1 (Drule.fun_cong_rule th2 tm2)
val tm5 = concl th3
in
if (Thm.dest_arg1 tm5) aconvc zero_tm
- then transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11)
+ then Thm.transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11)
else monomial_deone th3
end
end;
@@ -603,9 +603,9 @@
val l = Thm.dest_arg lopr
in
if l aconvc zero_tm
- then transitive th (inst_thm [(ca,r)] pthm_07) else
+ then Thm.transitive th (inst_thm [(ca,r)] pthm_07) else
if r aconvc zero_tm
- then transitive th (inst_thm [(ca,l)] pthm_08) else th
+ then Thm.transitive th (inst_thm [(ca,l)] pthm_08) else th
end
end
fun padd tm =
@@ -628,14 +628,14 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4)
- in dezero_rule (transitive th1 (combination th2 (padd tm2)))
+ in dezero_rule (Thm.transitive th1 (Thm.combination th2 (padd tm2)))
end
else (* ord <> 0*)
let val th1 =
if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24
else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25
val (tm1,tm2) = Thm.dest_comb(concl th1)
- in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
+ in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
end
end
else (* not (is_add r)*)
@@ -646,13 +646,13 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2
- in dezero_rule (transitive th1 th2)
+ in dezero_rule (Thm.transitive th1 th2)
end
else (* ord <> 0*)
if ord > 0 then
let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24
val (tm1,tm2) = Thm.dest_comb(concl th1)
- in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
+ in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
end
else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27)
end
@@ -667,21 +667,21 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2
- in dezero_rule (transitive th1 th2)
+ in dezero_rule (Thm.transitive th1 th2)
end
else
- if ord > 0 then reflexive tm
+ if ord > 0 then Thm.reflexive tm
else
let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25
val (tm1,tm2) = Thm.dest_comb(concl th1)
- in dezero_rule (transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
+ in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2)))
end
end
else
let val ord = morder l r
in
if ord = 0 then monomial_add_conv tm
- else if ord > 0 then dezero_rule(reflexive tm)
+ else if ord > 0 then dezero_rule(Thm.reflexive tm)
else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27)
end
end
@@ -699,7 +699,7 @@
else
if not(is_add r) then
let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09
- in transitive th1 (polynomial_monomial_mul_conv(concl th1))
+ in Thm.transitive th1 (polynomial_monomial_mul_conv(concl th1))
end
else
let val (a,b) = dest_add l
@@ -707,8 +707,8 @@
val (tm1,tm2) = Thm.dest_comb(concl th1)
val (tm3,tm4) = Thm.dest_comb tm1
val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv tm4)
- val th3 = transitive th1 (combination th2 (pmul tm2))
- in transitive th3 (polynomial_add_conv (concl th3))
+ val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2))
+ in Thm.transitive th3 (polynomial_add_conv (concl th3))
end
end
in fn tm =>
@@ -740,9 +740,9 @@
let val th1 = num_conv n
val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38
val (tm1,tm2) = Thm.dest_comb(concl th2)
- val th3 = transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2))
- val th4 = transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3
- in transitive th4 (polynomial_mul_conv (concl th4))
+ val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2))
+ val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3
+ in Thm.transitive th4 (polynomial_mul_conv (concl th4))
end
end
in fn tm =>
@@ -755,8 +755,8 @@
let val (l,r) = Thm.dest_comb tm in
if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else
let val th1 = inst_thm [(cx',r)] neg_mul
- val th2 = transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1))
- in transitive th2 (polynomial_monomial_mul_conv (concl th2))
+ val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1))
+ in Thm.transitive th2 (polynomial_monomial_mul_conv (concl th2))
end
end;
@@ -767,51 +767,52 @@
val th1 = inst_thm [(cx',l),(cy',r)] sub_add
val (tm1,tm2) = Thm.dest_comb(concl th1)
val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv tm2)
- in transitive th1 (transitive th2 (polynomial_add_conv (concl th2)))
+ in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv (concl th2)))
end;
(* Conversion from HOL term. *)
fun polynomial_conv tm =
if is_semiring_constant tm then semiring_add_conv tm
- else if not(is_comb tm) then reflexive tm
+ else if not(is_comb tm) then Thm.reflexive tm
else
let val (lopr,r) = Thm.dest_comb tm
in if lopr aconvc neg_tm then
let val th1 = Drule.arg_cong_rule lopr (polynomial_conv r)
- in transitive th1 (polynomial_neg_conv (concl th1))
+ in Thm.transitive th1 (polynomial_neg_conv (concl th1))
end
else if lopr aconvc inverse_tm then
let val th1 = Drule.arg_cong_rule lopr (polynomial_conv r)
- in transitive th1 (semiring_mul_conv (concl th1))
+ in Thm.transitive th1 (semiring_mul_conv (concl th1))
end
else
- if not(is_comb lopr) then reflexive tm
+ if not(is_comb lopr) then Thm.reflexive tm
else
let val (opr,l) = Thm.dest_comb lopr
in if opr aconvc pow_tm andalso is_numeral r
then
let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv l)) r
- in transitive th1 (polynomial_pow_conv (concl th1))
+ in Thm.transitive th1 (polynomial_pow_conv (concl th1))
end
else if opr aconvc divide_tm
then
- let val th1 = combination (Drule.arg_cong_rule opr (polynomial_conv l))
+ let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv l))
(polynomial_conv r)
val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv)
(Thm.rhs_of th1)
- in transitive th1 th2
+ in Thm.transitive th1 th2
end
else
if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm
then
- let val th1 = combination (Drule.arg_cong_rule opr (polynomial_conv l)) (polynomial_conv r)
+ let val th1 =
+ Thm.combination (Drule.arg_cong_rule opr (polynomial_conv l)) (polynomial_conv r)
val f = if opr aconvc add_tm then polynomial_add_conv
else if opr aconvc mul_tm then polynomial_mul_conv
else polynomial_sub_conv
- in transitive th1 (f (concl th1))
+ in Thm.transitive th1 (f (concl th1))
end
- else reflexive tm
+ else Thm.reflexive tm
end
end;
in
@@ -852,7 +853,7 @@
fun semiring_normalize_ord_conv ctxt ord tm =
(case match ctxt tm of
- NONE => reflexive tm
+ NONE => Thm.reflexive tm
| SOME res => semiring_normalize_ord_wrapper ctxt res ord tm);
fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt simple_cterm_ord;
--- a/src/HOL/Tools/split_rule.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/split_rule.ML Sun May 23 10:38:11 2010 +0100
@@ -135,7 +135,7 @@
Attrib.setup @{binding split_format}
(Scan.lift
(Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute (K complete_split_rule)) ||
- OuterParse.and_list1 (Scan.repeat Args.name_source)
+ Parse.and_list1 (Scan.repeat Args.name_source)
>> (fn xss => Thm.rule_attribute (fn context =>
split_rule_goal (Context.proof_of context) xss))))
"split pair-typed subterms in premises, or function arguments" #>
--- a/src/HOL/Tools/typedef.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Tools/typedef.ML Sun May 23 10:38:11 2010 +0100
@@ -296,22 +296,19 @@
(** outer syntax **)
-local structure P = OuterParse in
-
-val _ = OuterKeyword.keyword "morphisms";
+val _ = Keyword.keyword "morphisms";
val _ =
- OuterSyntax.local_theory_to_proof "typedef" "HOL type definition (requires non-emptiness proof)"
- OuterKeyword.thy_goal
- (Scan.optional (P.$$$ "(" |--
- ((P.$$$ "open" >> K false) -- Scan.option P.binding ||
- P.binding >> (fn s => (true, SOME s))) --| P.$$$ ")") (true, NONE) --
- (P.type_args_constrained -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding))
+ Outer_Syntax.local_theory_to_proof "typedef" "HOL type definition (requires non-emptiness proof)"
+ Keyword.thy_goal
+ (Scan.optional (Parse.$$$ "(" |--
+ ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding ||
+ Parse.binding >> (fn s => (true, SOME s))) --| Parse.$$$ ")") (true, NONE) --
+ (Parse.type_args_constrained -- Parse.binding) --
+ Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) --
+ Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding))
>> (fn ((((((def, opt_name), (args, t)), mx), A), morphs)) =>
typedef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs)));
end;
-end;
-
--- a/src/HOL/Transcendental.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/Transcendental.thy Sun May 23 10:38:11 2010 +0100
@@ -1778,7 +1778,7 @@
lemma sin_pi_half [simp]: "sin(pi/2) = 1"
apply (cut_tac x = "pi/2" in sin_cos_squared_add2)
apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two])
-apply (simp add: power2_eq_square)
+apply (simp add: power2_eq_1_iff)
done
lemma cos_pi [simp]: "cos pi = -1"
@@ -2580,18 +2580,6 @@
lemma tan_60: "tan (pi / 3) = sqrt 3"
unfolding tan_def by (simp add: sin_60 cos_60)
-text{*NEEDED??*}
-lemma [simp]:
- "sin (x + 1 / 2 * real (Suc m) * pi) =
- cos (x + 1 / 2 * real (m) * pi)"
-by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto)
-
-text{*NEEDED??*}
-lemma [simp]:
- "sin (x + real (Suc m) * pi / 2) =
- cos (x + real (m) * pi / 2)"
-by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto)
-
lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)"
by (auto intro!: DERIV_intros)
@@ -2620,16 +2608,6 @@
apply (subst sin_add, simp)
done
-(*NEEDED??*)
-lemma [simp]:
- "cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)"
-apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto)
-done
-
-(*NEEDED??*)
-lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
-by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto)
-
lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0"
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto)
--- a/src/HOL/ex/Adder.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(* Title: HOL/ex/Adder.thy
- ID: $Id$
- Author: Sergey Tverdyshev (Universitaet des Saarlandes)
-*)
-
-header {* Implementation of carry chain incrementor and adder *}
-
-theory Adder imports Main Word begin
-
-lemma [simp]: "bv_to_nat [b] = bitval b"
- by (simp add: bv_to_nat_helper)
-
-lemma bv_to_nat_helper':
- "bv \<noteq> [] ==> bv_to_nat bv = bitval (hd bv) * 2 ^ (length bv - 1) + bv_to_nat (tl bv)"
- by (cases bv) (simp_all add: bv_to_nat_helper)
-
-definition
- half_adder :: "[bit, bit] => bit list" where
- "half_adder a b = [a bitand b, a bitxor b]"
-
-lemma half_adder_correct: "bv_to_nat (half_adder a b) = bitval a + bitval b"
- apply (simp add: half_adder_def)
- apply (cases a, auto)
- apply (cases b, auto)
- done
-
-lemma [simp]: "length (half_adder a b) = 2"
- by (simp add: half_adder_def)
-
-definition
- full_adder :: "[bit, bit, bit] => bit list" where
- "full_adder a b c =
- (let x = a bitxor b in [a bitand b bitor c bitand x, x bitxor c])"
-
-lemma full_adder_correct:
- "bv_to_nat (full_adder a b c) = bitval a + bitval b + bitval c"
- apply (simp add: full_adder_def Let_def)
- apply (cases a, auto)
- apply (case_tac [!] b, auto)
- apply (case_tac [!] c, auto)
- done
-
-lemma [simp]: "length (full_adder a b c) = 2"
- by (simp add: full_adder_def Let_def)
-
-
-subsection {* Carry chain incrementor *}
-
-consts
- carry_chain_inc :: "[bit list, bit] => bit list"
-primrec
- "carry_chain_inc [] c = [c]"
- "carry_chain_inc (a#as) c =
- (let chain = carry_chain_inc as c
- in half_adder a (hd chain) @ tl chain)"
-
-lemma cci_nonnull: "carry_chain_inc as c \<noteq> []"
- by (cases as) (auto simp add: Let_def half_adder_def)
-
-lemma cci_length [simp]: "length (carry_chain_inc as c) = length as + 1"
- by (induct as) (simp_all add: Let_def)
-
-lemma cci_correct: "bv_to_nat (carry_chain_inc as c) = bv_to_nat as + bitval c"
- apply (induct as)
- apply (cases c, simp_all add: Let_def bv_to_nat_dist_append)
- apply (simp add: half_adder_correct bv_to_nat_helper' [OF cci_nonnull]
- ring_distribs bv_to_nat_helper)
- done
-
-consts
- carry_chain_adder :: "[bit list, bit list, bit] => bit list"
-primrec
- "carry_chain_adder [] bs c = [c]"
- "carry_chain_adder (a # as) bs c =
- (let chain = carry_chain_adder as (tl bs) c
- in full_adder a (hd bs) (hd chain) @ tl chain)"
-
-lemma cca_nonnull: "carry_chain_adder as bs c \<noteq> []"
- by (cases as) (auto simp add: full_adder_def Let_def)
-
-lemma cca_length: "length as = length bs \<Longrightarrow>
- length (carry_chain_adder as bs c) = Suc (length bs)"
- by (induct as arbitrary: bs) (auto simp add: Let_def)
-
-theorem cca_correct:
- "length as = length bs \<Longrightarrow>
- bv_to_nat (carry_chain_adder as bs c) =
- bv_to_nat as + bv_to_nat bs + bitval c"
-proof (induct as arbitrary: bs)
- case Nil
- then show ?case by simp
-next
- case (Cons a as xs)
- note ind = Cons.hyps
- from Cons.prems have len: "Suc (length as) = length xs" by simp
- show ?case
- proof (cases xs)
- case Nil
- with len show ?thesis by simp
- next
- case (Cons b bs)
- with len have len': "length as = length bs" by simp
- then have "bv_to_nat (carry_chain_adder as bs c) = bv_to_nat as + bv_to_nat bs + bitval c"
- by (rule ind)
- with len' and Cons
- show ?thesis
- apply (simp add: Let_def)
- apply (subst bv_to_nat_dist_append)
- apply (simp add: full_adder_correct bv_to_nat_helper' [OF cca_nonnull]
- ring_distribs bv_to_nat_helper cca_length)
- done
- qed
-qed
-
-end
--- a/src/HOL/ex/Codegenerator_Candidates.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/ex/Codegenerator_Candidates.thy Sun May 23 10:38:11 2010 +0100
@@ -13,6 +13,7 @@
Fset
Enum
List_Prefix
+ More_List
Nat_Infinity
Nested_Environment
Option_ord
@@ -23,7 +24,6 @@
RBT
SetsAndFunctions
While_Combinator
- Word
begin
inductive sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
--- a/src/HOL/ex/Execute_Choice.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/ex/Execute_Choice.thy Sun May 23 10:38:11 2010 +0100
@@ -26,7 +26,7 @@
case True then show ?thesis by (simp add: is_empty_def keys_def valuesum_def)
next
case False
- then have l: "\<exists>l. l \<in> dom (Mapping.lookup m)" by (auto simp add: is_empty_def expand_fun_eq mem_def)
+ then have l: "\<exists>l. l \<in> dom (Mapping.lookup m)" by (auto simp add: is_empty_def expand_fun_eq mem_def keys_def)
then have "(let l = SOME l. l \<in> dom (Mapping.lookup m) in
the (Mapping.lookup m l) + (\<Sum>k \<in> dom (Mapping.lookup m) - {l}. the (Mapping.lookup m k))) =
(\<Sum>k \<in> dom (Mapping.lookup m). the (Mapping.lookup m k))"
--- a/src/HOL/ex/MergeSort.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/ex/MergeSort.thy Sun May 23 10:38:11 2010 +0100
@@ -6,7 +6,7 @@
header{*Merge Sort*}
theory MergeSort
-imports Sorting
+imports Multiset
begin
context linorder
@@ -19,23 +19,17 @@
| "merge xs [] = xs"
| "merge [] ys = ys"
-lemma multiset_of_merge[simp]:
- "multiset_of (merge xs ys) = multiset_of xs + multiset_of ys"
-apply(induct xs ys rule: merge.induct)
-apply (auto simp: union_ac)
-done
+lemma multiset_of_merge [simp]:
+ "multiset_of (merge xs ys) = multiset_of xs + multiset_of ys"
+ by (induct xs ys rule: merge.induct) (simp_all add: ac_simps)
-lemma set_merge[simp]: "set (merge xs ys) = set xs \<union> set ys"
-apply(induct xs ys rule: merge.induct)
-apply auto
-done
+lemma set_merge [simp]:
+ "set (merge xs ys) = set xs \<union> set ys"
+ by (induct xs ys rule: merge.induct) auto
-lemma sorted_merge[simp]:
- "sorted (op \<le>) (merge xs ys) = (sorted (op \<le>) xs & sorted (op \<le>) ys)"
-apply(induct xs ys rule: merge.induct)
-apply(simp_all add: ball_Un not_le less_le)
-apply(blast intro: order_trans)
-done
+lemma sorted_merge [simp]:
+ "sorted (merge xs ys) \<longleftrightarrow> sorted xs \<and> sorted ys"
+ by (induct xs ys rule: merge.induct) (auto simp add: ball_Un not_le less_le sorted_Cons)
fun msort :: "'a list \<Rightarrow> 'a list"
where
@@ -44,16 +38,19 @@
| "msort xs = merge (msort (take (size xs div 2) xs))
(msort (drop (size xs div 2) xs))"
-theorem sorted_msort: "sorted (op \<le>) (msort xs)"
-by (induct xs rule: msort.induct) simp_all
+lemma sorted_msort:
+ "sorted (msort xs)"
+ by (induct xs rule: msort.induct) simp_all
-theorem multiset_of_msort: "multiset_of (msort xs) = multiset_of xs"
-apply (induct xs rule: msort.induct)
- apply simp_all
-apply (metis append_take_drop_id drop_Suc_Cons multiset_of.simps(2) multiset_of_append take_Suc_Cons)
-done
+lemma multiset_of_msort:
+ "multiset_of (msort xs) = multiset_of xs"
+ by (induct xs rule: msort.induct)
+ (simp_all, metis append_take_drop_id drop_Suc_Cons multiset_of.simps(2) multiset_of_append take_Suc_Cons)
+
+theorem msort_sort:
+ "sort = msort"
+ by (rule ext, rule properties_for_sort) (fact multiset_of_msort sorted_msort)+
end
-
end
--- a/src/HOL/ex/ROOT.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOL/ex/ROOT.ML Sun May 23 10:38:11 2010 +0100
@@ -7,7 +7,6 @@
"State_Monad",
"Efficient_Nat_examples",
"FuncSet",
- "Word",
"Eval_Examples",
"Codegenerator_Test",
"Codegenerator_Pretty_Test",
@@ -46,7 +45,6 @@
"Unification",
"Primrec",
"Tarski",
- "Adder",
"Classical",
"set",
"Meson_Test",
--- a/src/HOLCF/Cfun.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Cfun.thy Sun May 23 10:38:11 2010 +0100
@@ -100,7 +100,7 @@
subsection {* Continuous function space is pointed *}
lemma UU_CFun: "\<bottom> \<in> CFun"
-by (simp add: CFun_def inst_fun_pcpo cont_const)
+by (simp add: CFun_def inst_fun_pcpo)
instance cfun :: (finite_po, finite_po) finite_po
by (rule typedef_finite_po [OF type_definition_CFun])
@@ -139,9 +139,37 @@
lemma Abs_CFun_inverse2: "cont f \<Longrightarrow> Rep_CFun (Abs_CFun f) = f"
by (simp add: Abs_CFun_inverse CFun_def)
-lemma beta_cfun [simp]: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
+lemma beta_cfun: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
by (simp add: Abs_CFun_inverse2)
+text {* Beta-reduction simproc *}
+
+text {*
+ Given the term @{term "(\<Lambda> x. f x)\<cdot>y"}, the procedure tries to
+ construct the theorem @{term "(\<Lambda> x. f x)\<cdot>y == f y"}. If this
+ theorem cannot be completely solved by the cont2cont rules, then
+ the procedure returns the ordinary conditional @{text beta_cfun}
+ rule.
+
+ The simproc does not solve any more goals that would be solved by
+ using @{text beta_cfun} as a simp rule. The advantage of the
+ simproc is that it can avoid deeply-nested calls to the simplifier
+ that would otherwise be caused by large continuity side conditions.
+*}
+
+simproc_setup beta_cfun_proc ("Abs_CFun f\<cdot>x") = {*
+ fn phi => fn ss => fn ct =>
+ let
+ val dest = Thm.dest_comb;
+ val (f, x) = (apfst (snd o dest o snd o dest) o dest) ct;
+ val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
+ val tr = instantiate' [SOME T, SOME U] [SOME f, SOME x]
+ (mk_meta_eq @{thm beta_cfun});
+ val rules = Cont2ContData.get (Simplifier.the_context ss);
+ val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));
+ in SOME (perhaps (SINGLE (tac 1)) tr) end
+*}
+
text {* Eta-equality for continuous functions *}
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
@@ -301,7 +329,7 @@
text {* cont2cont lemma for @{term Rep_CFun} *}
-lemma cont2cont_Rep_CFun [cont2cont]:
+lemma cont2cont_Rep_CFun [simp, cont2cont]:
assumes f: "cont (\<lambda>x. f x)"
assumes t: "cont (\<lambda>x. t x)"
shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
@@ -341,7 +369,7 @@
has only a single subgoal.
*}
-lemma cont2cont_LAM' [cont2cont]:
+lemma cont2cont_LAM' [simp, cont2cont]:
fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
shows "cont (\<lambda>x. \<Lambda> y. f x y)"
@@ -353,7 +381,7 @@
using f by (rule cont_fst_snd_D1)
qed
-lemma cont2cont_LAM_discrete [cont2cont]:
+lemma cont2cont_LAM_discrete [simp, cont2cont]:
"(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
by (simp add: cont2cont_LAM)
@@ -556,7 +584,7 @@
shows "cont (\<lambda>x. let y = f x in g x y)"
unfolding Let_def using f g2 g1 by (rule cont_apply)
-lemma cont2cont_Let' [cont2cont]:
+lemma cont2cont_Let' [simp, cont2cont]:
assumes f: "cont (\<lambda>x. f x)"
assumes g: "cont (\<lambda>p. g (fst p) (snd p))"
shows "cont (\<lambda>x. let y = f x in g x y)"
--- a/src/HOLCF/Cont.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Cont.thy Sun May 23 10:38:11 2010 +0100
@@ -120,7 +120,7 @@
apply (erule ch2ch_monofun [OF mono])
done
-subsection {* Continuity simproc *}
+subsection {* Collection of continuity rules *}
ML {*
structure Cont2ContData = Named_Thms
@@ -132,34 +132,18 @@
setup Cont2ContData.setup
-text {*
- Given the term @{term "cont f"}, the procedure tries to construct the
- theorem @{term "cont f == True"}. If this theorem cannot be completely
- solved by the introduction rules, then the procedure returns a
- conditional rewrite rule with the unsolved subgoals as premises.
-*}
-
-simproc_setup cont_proc ("cont f") = {*
- fn phi => fn ss => fn ct =>
- let
- val tr = instantiate' [] [SOME ct] @{thm Eq_TrueI};
- val rules = Cont2ContData.get (Simplifier.the_context ss);
- val tac = REPEAT_ALL_NEW (match_tac rules);
- in SINGLE (tac 1) tr end
-*}
-
subsection {* Continuity of basic functions *}
text {* The identity function is continuous *}
-lemma cont_id [cont2cont]: "cont (\<lambda>x. x)"
+lemma cont_id [simp, cont2cont]: "cont (\<lambda>x. x)"
apply (rule contI)
apply (erule cpo_lubI)
done
text {* constant functions are continuous *}
-lemma cont_const [cont2cont]: "cont (\<lambda>x. c)"
+lemma cont_const [simp, cont2cont]: "cont (\<lambda>x. c)"
apply (rule contI)
apply (rule lub_const)
done
@@ -237,7 +221,7 @@
text {* functions with discrete domain *}
-lemma cont_discrete_cpo [cont2cont]: "cont (f::'a::discrete_cpo \<Rightarrow> 'b::cpo)"
+lemma cont_discrete_cpo [simp, cont2cont]: "cont (f::'a::discrete_cpo \<Rightarrow> 'b::cpo)"
apply (rule contI)
apply (drule discrete_chain_const, clarify)
apply (simp add: lub_const)
--- a/src/HOLCF/Fixrec.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Fixrec.thy Sun May 23 10:38:11 2010 +0100
@@ -65,30 +65,6 @@
== "CONST maybe_when\<cdot>t1\<cdot>(\<Lambda> x. t2)\<cdot>m"
-subsubsection {* Monadic bind operator *}
-
-definition
- bind :: "'a maybe \<rightarrow> ('a \<rightarrow> 'b maybe) \<rightarrow> 'b maybe" where
- "bind = (\<Lambda> m f. case m of fail \<Rightarrow> fail | return\<cdot>x \<Rightarrow> f\<cdot>x)"
-
-text {* monad laws *}
-
-lemma bind_strict [simp]: "bind\<cdot>\<bottom>\<cdot>f = \<bottom>"
-by (simp add: bind_def)
-
-lemma bind_fail [simp]: "bind\<cdot>fail\<cdot>f = fail"
-by (simp add: bind_def)
-
-lemma left_unit [simp]: "bind\<cdot>(return\<cdot>a)\<cdot>k = k\<cdot>a"
-by (simp add: bind_def)
-
-lemma right_unit [simp]: "bind\<cdot>m\<cdot>return = m"
-by (rule_tac p=m in maybeE, simp_all)
-
-lemma bind_assoc:
- "bind\<cdot>(bind\<cdot>m\<cdot>k)\<cdot>h = bind\<cdot>m\<cdot>(\<Lambda> a. bind\<cdot>(k\<cdot>a)\<cdot>h)"
-by (rule_tac p=m in maybeE, simp_all)
-
subsubsection {* Run operator *}
definition
@@ -169,7 +145,7 @@
definition
branch :: "('a \<rightarrow> 'b maybe) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c maybe)" where
- "branch p \<equiv> \<Lambda> r x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))"
+ "branch p \<equiv> \<Lambda> r x. maybe_when\<cdot>fail\<cdot>(\<Lambda> y. return\<cdot>(r\<cdot>y))\<cdot>(p\<cdot>x)"
lemma branch_rews:
"p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"
@@ -277,7 +253,7 @@
definition
cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
"cpair_pat p1 p2 = (\<Lambda>(x, y).
- bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. return\<cdot>(a, b))))"
+ maybe_when\<cdot>fail\<cdot>(\<Lambda> a. maybe_when\<cdot>fail\<cdot>(\<Lambda> b. return\<cdot>(a, b))\<cdot>(p2\<cdot>y))\<cdot>(p1\<cdot>x))"
definition
spair_pat ::
@@ -425,7 +401,7 @@
definition
as_pat :: "('a \<rightarrow> 'b maybe) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) maybe" where
- "as_pat p = (\<Lambda> x. bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. return\<cdot>(x, a)))"
+ "as_pat p = (\<Lambda> x. maybe_when\<cdot>fail\<cdot>(\<Lambda> a. return\<cdot>(x, a))\<cdot>(p\<cdot>x))"
definition
lazy_pat :: "('a \<rightarrow> 'b::pcpo maybe) \<Rightarrow> ('a \<rightarrow> 'b maybe)" where
@@ -608,25 +584,6 @@
(@{const_name UU}, @{const_name match_UU}) ]
*}
-hide_const (open) return bind fail run cases
-
-lemmas [fixrec_simp] =
- run_strict run_fail run_return
- mplus_strict mplus_fail mplus_return
- spair_strict_iff
- sinl_defined_iff
- sinr_defined_iff
- up_defined
- ONE_defined
- dist_eq_tr(1,2)
- match_UU_simps
- match_cpair_simps
- match_spair_simps
- match_sinl_simps
- match_sinr_simps
- match_up_simps
- match_ONE_simps
- match_TT_simps
- match_FF_simps
+hide_const (open) return fail run cases
end
--- a/src/HOLCF/IMP/Denotational.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/IMP/Denotational.thy Sun May 23 10:38:11 2010 +0100
@@ -7,6 +7,16 @@
theory Denotational imports HOLCF "../../HOL/IMP/Natural" begin
+text {* Disable conflicting syntax from HOL Map theory. *}
+
+no_syntax
+ "_maplet" :: "['a, 'a] => maplet" ("_ /|->/ _")
+ "_maplets" :: "['a, 'a] => maplet" ("_ /[|->]/ _")
+ "" :: "maplet => maplets" ("_")
+ "_Maplets" :: "[maplet, maplets] => maplets" ("_,/ _")
+ "_MapUpd" :: "['a ~=> 'b, maplets] => 'a ~=> 'b" ("_/'(_')" [900,0]900)
+ "_Map" :: "maplets => 'a ~=> 'b" ("(1[_])")
+
subsection "Definition"
definition
--- a/src/HOLCF/IOA/meta_theory/automaton.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/IOA/meta_theory/automaton.ML Sun May 23 10:38:11 2010 +0100
@@ -485,52 +485,50 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword ["signature", "actions", "inputs",
+val _ = List.app Keyword.keyword ["signature", "actions", "inputs",
"outputs", "internals", "states", "initially", "transitions", "pre",
"post", "transrel", ":=", "compose", "hide_action", "in", "restrict", "to",
"rename"];
-val actionlist = P.list1 P.term;
-val inputslist = P.$$$ "inputs" |-- P.!!! actionlist;
-val outputslist = P.$$$ "outputs" |-- P.!!! actionlist;
-val internalslist = P.$$$ "internals" |-- P.!!! actionlist;
-val stateslist = P.$$$ "states" |-- P.!!! (Scan.repeat1 (P.name --| P.$$$ "::" -- P.typ));
-val initial = P.$$$ "initially" |-- P.!!! P.term;
-val assign_list = P.list1 (P.name --| P.$$$ ":=" -- P.!!! P.term);
-val pre = P.$$$ "pre" |-- P.!!! P.term;
-val post = P.$$$ "post" |-- P.!!! assign_list;
+val actionlist = Parse.list1 Parse.term;
+val inputslist = Parse.$$$ "inputs" |-- Parse.!!! actionlist;
+val outputslist = Parse.$$$ "outputs" |-- Parse.!!! actionlist;
+val internalslist = Parse.$$$ "internals" |-- Parse.!!! actionlist;
+val stateslist =
+ Parse.$$$ "states" |-- Parse.!!! (Scan.repeat1 (Parse.name --| Parse.$$$ "::" -- Parse.typ));
+val initial = Parse.$$$ "initially" |-- Parse.!!! Parse.term;
+val assign_list = Parse.list1 (Parse.name --| Parse.$$$ ":=" -- Parse.!!! Parse.term);
+val pre = Parse.$$$ "pre" |-- Parse.!!! Parse.term;
+val post = Parse.$$$ "post" |-- Parse.!!! assign_list;
val pre1 = (pre -- (Scan.optional post [])) >> mk_trans_of_prepost;
val post1 = ((Scan.optional pre "True") -- post) >> mk_trans_of_prepost;
-val transrel = (P.$$$ "transrel" |-- P.!!! P.term) >> mk_trans_of_rel;
-val transition = P.term -- (transrel || pre1 || post1);
-val translist = P.$$$ "transitions" |-- P.!!! (Scan.repeat1 transition);
+val transrel = (Parse.$$$ "transrel" |-- Parse.!!! Parse.term) >> mk_trans_of_rel;
+val transition = Parse.term -- (transrel || pre1 || post1);
+val translist = Parse.$$$ "transitions" |-- Parse.!!! (Scan.repeat1 transition);
val ioa_decl =
- (P.name -- (P.$$$ "=" |--
- (P.$$$ "signature" |--
- P.!!! (P.$$$ "actions" |--
- (P.typ --
+ (Parse.name -- (Parse.$$$ "=" |--
+ (Parse.$$$ "signature" |--
+ Parse.!!! (Parse.$$$ "actions" |--
+ (Parse.typ --
(Scan.optional inputslist []) --
(Scan.optional outputslist []) --
(Scan.optional internalslist []) --
stateslist --
(Scan.optional initial "True") --
translist))))) >> mk_ioa_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "compose" |-- P.!!! (P.list1 P.name))))
+ (Parse.name -- (Parse.$$$ "=" |-- (Parse.$$$ "compose" |-- Parse.!!! (Parse.list1 Parse.name))))
>> mk_composition_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "hide_action" |--
- P.!!! (P.list1 P.term -- (P.$$$ "in" |-- P.name))))) >> mk_hiding_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "restrict" |--
- P.!!! (P.name -- (P.$$$ "to" |-- P.list1 P.term))))) >> mk_restriction_decl ||
- (P.name -- (P.$$$ "=" |-- (P.$$$ "rename" |-- (P.!!! (P.name -- (P.$$$ "to" |-- P.term))))))
+ (Parse.name -- (Parse.$$$ "=" |-- (Parse.$$$ "hide_action" |--
+ Parse.!!! (Parse.list1 Parse.term -- (Parse.$$$ "in" |-- Parse.name))))) >> mk_hiding_decl ||
+ (Parse.name -- (Parse.$$$ "=" |-- (Parse.$$$ "restrict" |--
+ Parse.!!! (Parse.name -- (Parse.$$$ "to" |-- Parse.list1 Parse.term))))) >> mk_restriction_decl ||
+ (Parse.name -- (Parse.$$$ "=" |--
+ (Parse.$$$ "rename" |-- (Parse.!!! (Parse.name -- (Parse.$$$ "to" |-- Parse.term))))))
>> mk_rename_decl;
val _ =
- OuterSyntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" K.thy_decl
+ Outer_Syntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" Keyword.thy_decl
(ioa_decl >> Toplevel.theory);
end;
-
-end;
--- a/src/HOLCF/IsaMakefile Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/IsaMakefile Sun May 23 10:38:11 2010 +0100
@@ -7,6 +7,7 @@
default: HOLCF
images: HOLCF IOA
test: HOLCF-IMP HOLCF-ex HOLCF-FOCUS \
+ HOLCF-Tutorial \
IOA-ABP IOA-NTP IOA-Modelcheck IOA-Storage IOA-ex
all: images test
@@ -78,6 +79,19 @@
@$(ISABELLE_TOOL) usedir -b -g true -r $(OUT)/HOL HOLCF
+## HOLCF-Tutorial
+
+HOLCF-Tutorial: HOLCF $(LOG)/HOLCF-Tutorial.gz
+
+$(LOG)/HOLCF-Tutorial.gz: $(OUT)/HOLCF \
+ Tutorial/Domain_ex.thy \
+ Tutorial/Fixrec_ex.thy \
+ Tutorial/New_Domain.thy \
+ Tutorial/document/root.tex \
+ Tutorial/ROOT.ML
+ @$(ISABELLE_TOOL) usedir $(OUT)/HOLCF Tutorial
+
+
## HOLCF-IMP
HOLCF-IMP: HOLCF $(LOG)/HOLCF-IMP.gz
@@ -95,15 +109,12 @@
../HOL/Library/Nat_Infinity.thy \
ex/Dagstuhl.thy \
ex/Dnat.thy \
- ex/Domain_ex.thy \
ex/Domain_Proofs.thy \
ex/Fix2.thy \
- ex/Fixrec_ex.thy \
ex/Focus_ex.thy \
ex/Hoare.thy \
ex/Letrec.thy \
ex/Loop.thy \
- ex/New_Domain.thy \
ex/Powerdomain_ex.thy \
ex/Stream.thy \
ex/Strict_Fun.thy \
@@ -201,4 +212,4 @@
$(LOG)/HOLCF-ex.gz $(LOG)/HOLCF-FOCUS.gz $(OUT)/IOA \
$(LOG)/IOA.gz $(LOG)/IOA-ABP.gz $(LOG)/IOA-NTP.gz \
$(LOG)/IOA-Modelcheck.gz $(LOG)/IOA-Storage.gz \
- $(LOG)/IOA-ex.gz
+ $(LOG)/IOA-ex.gz $(LOG)/HOLCF-Tutorial.gz
--- a/src/HOLCF/Product_Cpo.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Product_Cpo.thy Sun May 23 10:38:11 2010 +0100
@@ -221,7 +221,7 @@
apply (erule cpo_lubI [OF ch2ch_snd])
done
-lemma cont2cont_Pair [cont2cont]:
+lemma cont2cont_Pair [simp, cont2cont]:
assumes f: "cont (\<lambda>x. f x)"
assumes g: "cont (\<lambda>x. g x)"
shows "cont (\<lambda>x. (f x, g x))"
@@ -230,9 +230,9 @@
apply (rule cont_const)
done
-lemmas cont2cont_fst [cont2cont] = cont_compose [OF cont_fst]
+lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst]
-lemmas cont2cont_snd [cont2cont] = cont_compose [OF cont_snd]
+lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd]
lemma cont2cont_split:
assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)"
@@ -256,7 +256,7 @@
"cont (\<lambda>p. f (fst p) (snd p)) \<Longrightarrow> cont (\<lambda>y. f x y)"
by (drule cont_compose [OF _ cont_pair2], simp)
-lemma cont2cont_split' [cont2cont]:
+lemma cont2cont_split' [simp, cont2cont]:
assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))"
assumes g: "cont (\<lambda>x. g x)"
shows "cont (\<lambda>x. split (f x) (g x))"
--- a/src/HOLCF/Sum_Cpo.thy Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Sum_Cpo.thy Sun May 23 10:38:11 2010 +0100
@@ -136,8 +136,8 @@
lemma cont_Inr: "cont Inr"
by (intro contI is_lub_Inr cpo_lubI)
-lemmas cont2cont_Inl [cont2cont] = cont_compose [OF cont_Inl]
-lemmas cont2cont_Inr [cont2cont] = cont_compose [OF cont_Inr]
+lemmas cont2cont_Inl [simp, cont2cont] = cont_compose [OF cont_Inl]
+lemmas cont2cont_Inr [simp, cont2cont] = cont_compose [OF cont_Inr]
lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl]
lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr]
@@ -168,7 +168,7 @@
apply (rule cont_sum_case1 [OF f1 g1])
done
-lemma cont2cont_sum_case' [cont2cont]:
+lemma cont2cont_sum_case' [simp, cont2cont]:
assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
assumes g: "cont (\<lambda>p. g (fst p) (snd p))"
assumes h: "cont (\<lambda>x. h x)"
--- a/src/HOLCF/Tools/Domain/domain_constructors.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_constructors.ML Sun May 23 10:38:11 2010 +0100
@@ -42,14 +42,13 @@
(************************** miscellaneous functions ***************************)
-val simple_ss =
- HOL_basic_ss addsimps simp_thms;
+val simple_ss = HOL_basic_ss addsimps simp_thms;
-val beta_ss =
- HOL_basic_ss
- addsimps simp_thms
- addsimps [@{thm beta_cfun}]
- addsimprocs [@{simproc cont_proc}];
+val beta_rules =
+ @{thms beta_cfun cont_id cont_const cont2cont_Rep_CFun cont2cont_LAM'} @
+ @{thms cont2cont_fst cont2cont_snd cont2cont_Pair};
+
+val beta_ss = HOL_basic_ss addsimps (simp_thms @ beta_rules);
fun define_consts
(specs : (binding * term * mixfix) list)
@@ -952,7 +951,7 @@
val (fun1, fun2, taken) = pat_lhs (pat, args);
val defs = @{thm branch_def} :: pat_defs;
val goal = mk_trp (mk_strict fun1);
- val rules = @{thm Fixrec.bind_strict} :: case_rews;
+ val rules = @{thms maybe_when_rews} @ case_rews;
val tacs = [simp_tac (beta_ss addsimps rules) 1];
in prove thy defs goal (K tacs) end;
fun pat_apps (i, (pat, (con, args))) =
@@ -967,7 +966,7 @@
val concl = mk_trp (mk_eq (fun1 ` con_app, rhs));
val goal = Logic.list_implies (assms, concl);
val defs = @{thm branch_def} :: pat_defs;
- val rules = @{thms bind_fail left_unit} @ case_rews;
+ val rules = @{thms maybe_when_rews} @ case_rews;
val tacs = [asm_simp_tac (beta_ss addsimps rules) 1];
in prove thy defs goal (K tacs) end;
in map_index pat_app spec end;
--- a/src/HOLCF/Tools/Domain/domain_extender.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_extender.ML Sun May 23 10:38:11 2010 +0100
@@ -224,27 +224,25 @@
(** outer syntax **)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = OuterKeyword.keyword "lazy";
+val _ = Keyword.keyword "lazy";
val dest_decl : (bool * binding option * string) parser =
- P.$$$ "(" |-- Scan.optional (P.$$$ "lazy" >> K true) false --
- (P.binding >> SOME) -- (P.$$$ "::" |-- P.typ) --| P.$$$ ")" >> P.triple1
- || P.$$$ "(" |-- P.$$$ "lazy" |-- P.typ --| P.$$$ ")"
+ Parse.$$$ "(" |-- Scan.optional (Parse.$$$ "lazy" >> K true) false --
+ (Parse.binding >> SOME) -- (Parse.$$$ "::" |-- Parse.typ) --| Parse.$$$ ")" >> Parse.triple1
+ || Parse.$$$ "(" |-- Parse.$$$ "lazy" |-- Parse.typ --| Parse.$$$ ")"
>> (fn t => (true,NONE,t))
- || P.typ >> (fn t => (false,NONE,t));
+ || Parse.typ >> (fn t => (false,NONE,t));
val cons_decl =
- P.binding -- Scan.repeat dest_decl -- P.opt_mixfix;
+ Parse.binding -- Scan.repeat dest_decl -- Parse.opt_mixfix;
val domain_decl =
- (P.type_args_constrained -- P.binding -- P.opt_mixfix) --
- (P.$$$ "=" |-- P.enum1 "|" cons_decl);
+ (Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix) --
+ (Parse.$$$ "=" |-- Parse.enum1 "|" cons_decl);
val domains_decl =
- Scan.option (P.$$$ "(" |-- P.binding --| P.$$$ ")") --
- P.and_list1 domain_decl;
+ Scan.option (Parse.$$$ "(" |-- Parse.binding --| Parse.$$$ ")") --
+ Parse.and_list1 domain_decl;
fun mk_domain
(definitional : bool)
@@ -267,13 +265,11 @@
end;
val _ =
- OuterSyntax.command "domain" "define recursive domains (HOLCF)"
- K.thy_decl (domains_decl >> (Toplevel.theory o mk_domain false));
+ Outer_Syntax.command "domain" "define recursive domains (HOLCF)"
+ Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain false));
val _ =
- OuterSyntax.command "new_domain" "define recursive domains (HOLCF)"
- K.thy_decl (domains_decl >> (Toplevel.theory o mk_domain true));
+ Outer_Syntax.command "new_domain" "define recursive domains (HOLCF)"
+ Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain true));
end;
-
-end;
--- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML Sun May 23 10:38:11 2010 +0100
@@ -34,11 +34,11 @@
structure Domain_Isomorphism : DOMAIN_ISOMORPHISM =
struct
-val beta_ss =
- HOL_basic_ss
- addsimps simp_thms
- addsimps [@{thm beta_cfun}]
- addsimprocs [@{simproc cont_proc}];
+val beta_rules =
+ @{thms beta_cfun cont_id cont_const cont2cont_Rep_CFun cont2cont_LAM'} @
+ @{thms cont2cont_fst cont2cont_snd cont2cont_Pair cont2cont_split'};
+
+val beta_ss = HOL_basic_ss addsimps (simp_thms @ beta_rules);
val beta_tac = simp_tac beta_ss;
@@ -707,21 +707,20 @@
local
-structure P = OuterParse and K = OuterKeyword
-
val parse_domain_iso :
(string list * binding * mixfix * string * (binding * binding) option)
parser =
- (P.type_args -- P.binding -- P.opt_mixfix -- (P.$$$ "=" |-- P.typ) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding)))
+ (Parse.type_args -- Parse.binding -- Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.typ) --
+ Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding)))
>> (fn ((((vs, t), mx), rhs), morphs) => (vs, t, mx, rhs, morphs));
-val parse_domain_isos = P.and_list1 parse_domain_iso;
+val parse_domain_isos = Parse.and_list1 parse_domain_iso;
in
val _ =
- OuterSyntax.command "domain_isomorphism" "define domain isomorphisms (HOLCF)" K.thy_decl
+ Outer_Syntax.command "domain_isomorphism" "define domain isomorphisms (HOLCF)"
+ Keyword.thy_decl
(parse_domain_isos >> (Toplevel.theory o domain_isomorphism_cmd));
end;
--- a/src/HOLCF/Tools/Domain/domain_take_proofs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_take_proofs.ML Sun May 23 10:38:11 2010 +0100
@@ -101,11 +101,11 @@
take_induct_thms : thm list
};
-val beta_ss =
- HOL_basic_ss
- addsimps simp_thms
- addsimps [@{thm beta_cfun}]
- addsimprocs [@{simproc cont_proc}];
+val beta_rules =
+ @{thms beta_cfun cont_id cont_const cont2cont_Rep_CFun cont2cont_LAM'} @
+ @{thms cont2cont_fst cont2cont_snd cont2cont_Pair};
+
+val beta_ss = HOL_basic_ss addsimps (simp_thms @ beta_rules);
val beta_tac = simp_tac beta_ss;
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Sun May 23 10:38:11 2010 +0100
@@ -198,7 +198,6 @@
fun qualified name = Binding.qualified true name dbind;
val simp = Simplifier.simp_add;
-val fixrec_simp = Fixrec.fixrec_simp_add;
in
thy
@@ -209,7 +208,7 @@
[Rule_Cases.case_names case_ns, Induct.cases_type dname]),
((qualified "when_rews" , when_rews ), [simp]),
((qualified "compacts" , compacts ), [simp]),
- ((qualified "con_rews" , con_rews ), [simp, fixrec_simp]),
+ ((qualified "con_rews" , con_rews ), [simp]),
((qualified "sel_rews" , sel_rews ), [simp]),
((qualified "dis_rews" , dis_rews ), [simp]),
((qualified "pat_rews" , pat_rews ), [simp]),
@@ -218,7 +217,7 @@
((qualified "inverts" , inverts ), [simp]),
((qualified "injects" , injects ), [simp]),
((qualified "take_rews" , take_rews ), [simp]),
- ((qualified "match_rews", mat_rews ), [simp, fixrec_simp])]
+ ((qualified "match_rews", mat_rews ), [simp])]
|> snd
|> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
pat_rews @ dist_les @ dist_eqs)
--- a/src/HOLCF/Tools/cont_consts.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/cont_consts.ML Sun May 23 10:38:11 2010 +0100
@@ -93,12 +93,8 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.command "consts" "declare constants (HOLCF)" K.thy_decl
- (Scan.repeat1 P.const_binding >> (Toplevel.theory o add_consts_cmd));
+ Outer_Syntax.command "consts" "declare constants (HOLCF)" Keyword.thy_decl
+ (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o add_consts_cmd));
end;
-
-end;
--- a/src/HOLCF/Tools/fixrec.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/fixrec.ML Sun May 23 10:38:11 2010 +0100
@@ -13,8 +13,6 @@
val add_fixpat: Thm.binding * term list -> theory -> theory
val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory
val add_matchers: (string * string) list -> theory -> theory
- val fixrec_simp_add: attribute
- val fixrec_simp_del: attribute
val fixrec_simp_tac: Proof.context -> int -> tactic
val setup: theory -> theory
end;
@@ -125,10 +123,23 @@
val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec);
val functional = lambda_tuple lhss (mk_tuple rhss);
val fixpoint = mk_fix (mk_cabs functional);
-
+
val cont_thm =
- Goal.prove lthy [] [] (mk_trp (mk_cont functional))
- (K (simp_tac (simpset_of lthy) 1));
+ let
+ val prop = mk_trp (mk_cont functional);
+ fun err () = error (
+ "Continuity proof failed; please check that cont2cont rules\n" ^
+ "are configured for all non-HOLCF constants.\n" ^
+ "The error occurred for the goal statement:\n" ^
+ Syntax.string_of_term lthy prop);
+ fun check th = case Thm.nprems_of th of 0 => all_tac th | n => err ();
+ val rules = Cont2ContData.get lthy;
+ val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));
+ val slow_tac = simp_tac (simpset_of lthy);
+ val tac = (fast_tac 1 ORELSE slow_tac 1) THEN check;
+ in
+ Goal.prove lthy [] [] prop (K tac)
+ end;
fun one_def (l as Free(n,_)) r =
let val b = Long_Name.base_name n
@@ -294,22 +305,12 @@
(********************** Proving associated theorems **********************)
(*************************************************************************)
-structure FixrecSimpData = Generic_Data
-(
- type T = simpset;
- val empty =
- HOL_basic_ss
- addsimps simp_thms
- addsimps [@{thm beta_cfun}]
- addsimprocs [@{simproc cont_proc}];
- val extend = I;
- val merge = merge_ss;
-);
+fun eta_tac i = CONVERSION Thm.eta_conversion i;
fun fixrec_simp_tac ctxt =
let
val tab = FixrecUnfoldData.get (Context.Proof ctxt);
- val ss = Simplifier.context ctxt (FixrecSimpData.get (Context.Proof ctxt));
+ val ss = Simplifier.simpset_of ctxt;
fun concl t =
if Logic.is_all t then concl (snd (Logic.dest_all t))
else HOLogic.dest_Trueprop (Logic.strip_imp_concl t);
@@ -320,26 +321,18 @@
val unfold_thm = the (Symtab.lookup tab c);
val rule = unfold_thm RS @{thm ssubst_lhs};
in
- CHANGED (rtac rule i THEN asm_simp_tac ss i)
+ CHANGED (rtac rule i THEN eta_tac i THEN asm_simp_tac ss i)
end
in
SUBGOAL (fn ti => the_default no_tac (try tac ti))
end;
-val fixrec_simp_add : attribute =
- Thm.declaration_attribute
- (fn th => FixrecSimpData.map (fn ss => ss addsimps [th]));
-
-val fixrec_simp_del : attribute =
- Thm.declaration_attribute
- (fn th => FixrecSimpData.map (fn ss => ss delsimps [th]));
-
(* proves a block of pattern matching equations as theorems, using unfold *)
fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) =
let
- val ss = Simplifier.context ctxt (FixrecSimpData.get (Context.Proof ctxt));
+ val ss = Simplifier.simpset_of ctxt;
val rule = unfold_thm RS @{thm ssubst_lhs};
- val tac = rtac rule 1 THEN asm_simp_tac ss 1;
+ val tac = rtac rule 1 THEN eta_tac 1 THEN asm_simp_tac ss 1;
fun prove_term t = Goal.prove ctxt [] [] t (K tac);
fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t);
in
@@ -443,23 +436,18 @@
(******************************** Parsers ********************************)
(*************************************************************************)
-local structure P = OuterParse and K = OuterKeyword in
+val _ =
+ Outer_Syntax.local_theory "fixrec" "define recursive functions (HOLCF)" Keyword.thy_decl
+ ((Parse.opt_keyword "permissive" >> not) -- Parse.fixes -- Parse_Spec.where_alt_specs
+ >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs));
-val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl
- ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs
- >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs));
+val _ =
+ Outer_Syntax.command "fixpat" "define rewrites for fixrec functions" Keyword.thy_decl
+ (Parse_Spec.specs >> (Toplevel.theory o add_fixpat_cmd));
-val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
- (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd));
+val setup =
+ Method.setup @{binding fixrec_simp}
+ (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac))
+ "pattern prover for fixrec constants";
end;
-
-val setup =
- Attrib.setup @{binding fixrec_simp}
- (Attrib.add_del fixrec_simp_add fixrec_simp_del)
- "declaration of fixrec simp rule"
- #> Method.setup @{binding fixrec_simp}
- (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac))
- "pattern prover for fixrec constants";
-
-end;
--- a/src/HOLCF/Tools/pcpodef.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/pcpodef.ML Sun May 23 10:38:11 2010 +0100
@@ -355,29 +355,29 @@
(** outer syntax **)
-local structure P = OuterParse and K = OuterKeyword in
-
val typedef_proof_decl =
- Scan.optional (P.$$$ "(" |--
- ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
- --| P.$$$ ")") (true, NONE) --
- (P.type_args_constrained -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
+ Scan.optional (Parse.$$$ "(" |--
+ ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding ||
+ Parse.binding >> (fn s => (true, SOME s)))
+ --| Parse.$$$ ")") (true, NONE) --
+ (Parse.type_args_constrained -- Parse.binding) -- Parse.opt_mixfix --
+ (Parse.$$$ "=" |-- Parse.term) --
+ Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding));
fun mk_pcpodef_proof pcpo ((((((def, opt_name), (args, t)), mx), A), morphs)) =
(if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
((def, the_default t opt_name), (t, args, mx), A, morphs);
val _ =
- OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
+ Outer_Syntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)"
+ Keyword.thy_goal
(typedef_proof_decl >>
(Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true)));
val _ =
- OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
+ Outer_Syntax.command "cpodef" "HOLCF type definition (requires admissibility proof)"
+ Keyword.thy_goal
(typedef_proof_decl >>
(Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false)));
end;
-
-end;
--- a/src/HOLCF/Tools/repdef.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/Tools/repdef.ML Sun May 23 10:38:11 2010 +0100
@@ -168,23 +168,21 @@
(** outer syntax **)
-local structure P = OuterParse and K = OuterKeyword in
-
val repdef_decl =
- Scan.optional (P.$$$ "(" |--
- ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
- --| P.$$$ ")") (true, NONE) --
- (P.type_args_constrained -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
- Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
+ Scan.optional (Parse.$$$ "(" |--
+ ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding ||
+ Parse.binding >> (fn s => (true, SOME s)))
+ --| Parse.$$$ ")") (true, NONE) --
+ (Parse.type_args_constrained -- Parse.binding) --
+ Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) --
+ Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding));
fun mk_repdef ((((((def, opt_name), (args, t)), mx), A), morphs)) =
repdef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs);
val _ =
- OuterSyntax.command "repdef" "HOLCF definition of representable domains" K.thy_decl
+ Outer_Syntax.command "repdef" "HOLCF definition of representable domains" Keyword.thy_decl
(repdef_decl >>
(Toplevel.print oo (Toplevel.theory o mk_repdef)));
end;
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/Domain_ex.thy Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,211 @@
+(* Title: HOLCF/ex/Domain_ex.thy
+ Author: Brian Huffman
+*)
+
+header {* Domain package examples *}
+
+theory Domain_ex
+imports HOLCF
+begin
+
+text {* Domain constructors are strict by default. *}
+
+domain d1 = d1a | d1b "d1" "d1"
+
+lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
+
+text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
+
+domain d2 = d2a | d2b (lazy "d2")
+
+lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
+
+text {* Strict and lazy arguments may be mixed arbitrarily. *}
+
+domain d3 = d3a | d3b (lazy "d2") "d2"
+
+lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
+
+text {* Selectors can be used with strict or lazy constructor arguments. *}
+
+domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
+
+lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
+
+text {* Mixfix declarations can be given for data constructors. *}
+
+domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
+
+lemma "d5a \<noteq> x :#: y :#: z" by simp
+
+text {* Mixfix declarations can also be given for type constructors. *}
+
+domain ('a, 'b) lazypair (infixl ":*:" 25) =
+ lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
+
+lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
+by (rule allI, case_tac p, simp_all)
+
+text {* Non-recursive constructor arguments can have arbitrary types. *}
+
+domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
+
+text {*
+ Indirect recusion is allowed for sums, products, lifting, and the
+ continuous function space. However, the domain package does not
+ generate an induction rule in terms of the constructors.
+*}
+
+domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c (lazy "'a d7 \<rightarrow> 'a")
+ -- "Indirect recursion detected, skipping proofs of (co)induction rules"
+
+text {* Note that @{text d7.induct} is absent. *}
+
+text {*
+ Indirect recursion is also allowed using previously-defined datatypes.
+*}
+
+domain 'a slist = SNil | SCons 'a "'a slist"
+
+domain 'a stree = STip | SBranch "'a stree slist"
+
+text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
+
+domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
+
+text {* Non-regular recursion is not allowed. *}
+(*
+domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
+ -- "illegal direct recursion with different arguments"
+domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
+ -- "illegal direct recursion with different arguments"
+*)
+
+text {*
+ Mutually-recursive datatypes must have all the same type arguments,
+ not necessarily in the same order.
+*}
+
+domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
+ and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
+
+text {* Induction rules for flat datatypes have no admissibility side-condition. *}
+
+domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
+
+lemma "\<lbrakk>P \<bottom>; P Tip; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; P x; P y\<rbrakk> \<Longrightarrow> P (Branch\<cdot>x\<cdot>y)\<rbrakk> \<Longrightarrow> P x"
+by (rule flattree.induct) -- "no admissibility requirement"
+
+text {* Trivial datatypes will produce a warning message. *}
+
+domain triv = Triv triv triv
+ -- "domain @{text Domain_ex.triv} is empty!"
+
+lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
+
+text {* Lazy constructor arguments may have unpointed types. *}
+
+domain natlist = nnil | ncons (lazy "nat discr") natlist
+
+text {* Class constraints may be given for type parameters on the LHS. *}
+
+domain ('a::cpo) box = Box (lazy 'a)
+
+
+subsection {* Generated constants and theorems *}
+
+domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (right :: "'a tree")
+
+lemmas tree_abs_defined_iff =
+ iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
+
+text {* Rules about ismorphism *}
+term tree_rep
+term tree_abs
+thm tree.rep_iso
+thm tree.abs_iso
+thm tree.iso_rews
+
+text {* Rules about constructors *}
+term Leaf
+term Node
+thm Leaf_def Node_def
+thm tree.nchotomy
+thm tree.exhaust
+thm tree.compacts
+thm tree.con_rews
+thm tree.dist_les
+thm tree.dist_eqs
+thm tree.inverts
+thm tree.injects
+
+text {* Rules about case combinator *}
+term tree_when
+thm tree.tree_when_def
+thm tree.when_rews
+
+text {* Rules about selectors *}
+term left
+term right
+thm tree.sel_rews
+
+text {* Rules about discriminators *}
+term is_Leaf
+term is_Node
+thm tree.dis_rews
+
+text {* Rules about pattern match combinators *}
+term Leaf_pat
+term Node_pat
+thm tree.pat_rews
+
+text {* Rules about monadic pattern match combinators *}
+term match_Leaf
+term match_Node
+thm tree.match_rews
+
+text {* Rules about take function *}
+term tree_take
+thm tree.take_def
+thm tree.take_0
+thm tree.take_Suc
+thm tree.take_rews
+thm tree.chain_take
+thm tree.take_take
+thm tree.deflation_take
+thm tree.take_below
+thm tree.take_lemma
+thm tree.lub_take
+thm tree.reach
+thm tree.finite_induct
+
+text {* Rules about finiteness predicate *}
+term tree_finite
+thm tree.finite_def
+thm tree.finite (* only generated for flat datatypes *)
+
+text {* Rules about bisimulation predicate *}
+term tree_bisim
+thm tree.bisim_def
+thm tree.coinduct
+
+text {* Induction rule *}
+thm tree.induct
+
+
+subsection {* Known bugs *}
+
+text {* Declaring a mixfix with spaces causes some strange parse errors. *}
+(*
+domain xx = xx ("x y")
+ -- "Inner syntax error: unexpected end of input"
+*)
+
+text {*
+ Non-cpo type parameters currently do not work.
+*}
+(*
+domain ('a::type) stream = snil | scons (lazy "'a discr") "'a stream"
+*)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/Fixrec_ex.thy Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,246 @@
+(* Title: HOLCF/ex/Fixrec_ex.thy
+ Author: Brian Huffman
+*)
+
+header {* Fixrec package examples *}
+
+theory Fixrec_ex
+imports HOLCF
+begin
+
+subsection {* Basic @{text fixrec} examples *}
+
+text {*
+ Fixrec patterns can mention any constructor defined by the domain
+ package, as well as any of the following built-in constructors:
+ Pair, spair, sinl, sinr, up, ONE, TT, FF.
+*}
+
+text {* Typical usage is with lazy constructors. *}
+
+fixrec down :: "'a u \<rightarrow> 'a"
+where "down\<cdot>(up\<cdot>x) = x"
+
+text {* With strict constructors, rewrite rules may require side conditions. *}
+
+fixrec from_sinl :: "'a \<oplus> 'b \<rightarrow> 'a"
+where "x \<noteq> \<bottom> \<Longrightarrow> from_sinl\<cdot>(sinl\<cdot>x) = x"
+
+text {* Lifting can turn a strict constructor into a lazy one. *}
+
+fixrec from_sinl_up :: "'a u \<oplus> 'b \<rightarrow> 'a"
+where "from_sinl_up\<cdot>(sinl\<cdot>(up\<cdot>x)) = x"
+
+text {* Fixrec also works with the HOL pair constructor. *}
+
+fixrec down2 :: "'a u \<times> 'b u \<rightarrow> 'a \<times> 'b"
+where "down2\<cdot>(up\<cdot>x, up\<cdot>y) = (x, y)"
+
+
+subsection {* Examples using @{text fixrec_simp} *}
+
+text {* A type of lazy lists. *}
+
+domain 'a llist = lNil | lCons (lazy 'a) (lazy "'a llist")
+
+text {* A zip function for lazy lists. *}
+
+text {* Notice that the patterns are not exhaustive. *}
+
+fixrec
+ lzip :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
+where
+ "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+| "lzip\<cdot>lNil\<cdot>lNil = lNil"
+
+text {* @{text fixrec_simp} is useful for producing strictness theorems. *}
+text {* Note that pattern matching is done in left-to-right order. *}
+
+lemma lzip_stricts [simp]:
+ "lzip\<cdot>\<bottom>\<cdot>ys = \<bottom>"
+ "lzip\<cdot>lNil\<cdot>\<bottom> = \<bottom>"
+ "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp+
+
+text {* @{text fixrec_simp} can also produce rules for missing cases. *}
+
+lemma lzip_undefs [simp]:
+ "lzip\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = \<bottom>"
+ "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = \<bottom>"
+by fixrec_simp+
+
+
+subsection {* Pattern matching with bottoms *}
+
+text {*
+ As an alternative to using @{text fixrec_simp}, it is also possible
+ to use bottom as a constructor pattern. When using a bottom
+ pattern, the right-hand-side must also be bottom; otherwise, @{text
+ fixrec} will not be able to prove the equation.
+*}
+
+fixrec
+ from_sinr_up :: "'a \<oplus> 'b\<^sub>\<bottom> \<rightarrow> 'b"
+where
+ "from_sinr_up\<cdot>\<bottom> = \<bottom>"
+| "from_sinr_up\<cdot>(sinr\<cdot>(up\<cdot>x)) = x"
+
+text {*
+ If the function is already strict in that argument, then the bottom
+ pattern does not change the meaning of the function. For example,
+ in the definition of @{term from_sinr_up}, the first equation is
+ actually redundant, and could have been proven separately by
+ @{text fixrec_simp}.
+*}
+
+text {*
+ A bottom pattern can also be used to make a function strict in a
+ certain argument, similar to a bang-pattern in Haskell.
+*}
+
+fixrec
+ seq :: "'a \<rightarrow> 'b \<rightarrow> 'b"
+where
+ "seq\<cdot>\<bottom>\<cdot>y = \<bottom>"
+| "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x\<cdot>y = y"
+
+
+subsection {* Skipping proofs of rewrite rules *}
+
+text {* Another zip function for lazy lists. *}
+
+text {*
+ Notice that this version has overlapping patterns.
+ The second equation cannot be proved as a theorem
+ because it only applies when the first pattern fails.
+*}
+
+fixrec (permissive)
+ lzip2 :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
+where
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+| "lzip2\<cdot>xs\<cdot>ys = lNil"
+
+text {*
+ Usually fixrec tries to prove all equations as theorems.
+ The "permissive" option overrides this behavior, so fixrec
+ does not produce any simp rules.
+*}
+
+text {* Simp rules can be generated later using @{text fixrec_simp}. *}
+
+lemma lzip2_simps [simp]:
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = lNil"
+ "lzip2\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = lNil"
+ "lzip2\<cdot>lNil\<cdot>lNil = lNil"
+by fixrec_simp+
+
+lemma lzip2_stricts [simp]:
+ "lzip2\<cdot>\<bottom>\<cdot>ys = \<bottom>"
+ "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp+
+
+
+subsection {* Mutual recursion with @{text fixrec} *}
+
+text {* Tree and forest types. *}
+
+domain 'a tree = Leaf (lazy 'a) | Branch (lazy "'a forest")
+and 'a forest = Empty | Trees (lazy "'a tree") "'a forest"
+
+text {*
+ To define mutually recursive functions, give multiple type signatures
+ separated by the keyword @{text "and"}.
+*}
+
+fixrec
+ map_tree :: "('a \<rightarrow> 'b) \<rightarrow> ('a tree \<rightarrow> 'b tree)"
+and
+ map_forest :: "('a \<rightarrow> 'b) \<rightarrow> ('a forest \<rightarrow> 'b forest)"
+where
+ "map_tree\<cdot>f\<cdot>(Leaf\<cdot>x) = Leaf\<cdot>(f\<cdot>x)"
+| "map_tree\<cdot>f\<cdot>(Branch\<cdot>ts) = Branch\<cdot>(map_forest\<cdot>f\<cdot>ts)"
+| "map_forest\<cdot>f\<cdot>Empty = Empty"
+| "ts \<noteq> \<bottom> \<Longrightarrow>
+ map_forest\<cdot>f\<cdot>(Trees\<cdot>t\<cdot>ts) = Trees\<cdot>(map_tree\<cdot>f\<cdot>t)\<cdot>(map_forest\<cdot>f\<cdot>ts)"
+
+lemma map_tree_strict [simp]: "map_tree\<cdot>f\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+lemma map_forest_strict [simp]: "map_forest\<cdot>f\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+(*
+ Theorems generated:
+ @{text map_tree_def} @{thm map_tree_def}
+ @{text map_forest_def} @{thm map_forest_def}
+ @{text map_tree.unfold} @{thm map_tree.unfold}
+ @{text map_forest.unfold} @{thm map_forest.unfold}
+ @{text map_tree.simps} @{thm map_tree.simps}
+ @{text map_forest.simps} @{thm map_forest.simps}
+ @{text map_tree_map_forest.induct} @{thm map_tree_map_forest.induct}
+*)
+
+
+subsection {* Looping simp rules *}
+
+text {*
+ The defining equations of a fixrec definition are declared as simp
+ rules by default. In some cases, especially for constants with no
+ arguments or functions with variable patterns, the defining
+ equations may cause the simplifier to loop. In these cases it will
+ be necessary to use a @{text "[simp del]"} declaration.
+*}
+
+fixrec
+ repeat :: "'a \<rightarrow> 'a llist"
+where
+ [simp del]: "repeat\<cdot>x = lCons\<cdot>x\<cdot>(repeat\<cdot>x)"
+
+text {*
+ We can derive other non-looping simp rules for @{const repeat} by
+ using the @{text subst} method with the @{text repeat.simps} rule.
+*}
+
+lemma repeat_simps [simp]:
+ "repeat\<cdot>x \<noteq> \<bottom>"
+ "repeat\<cdot>x \<noteq> lNil"
+ "repeat\<cdot>x = lCons\<cdot>y\<cdot>ys \<longleftrightarrow> x = y \<and> repeat\<cdot>x = ys"
+by (subst repeat.simps, simp)+
+
+lemma llist_when_repeat [simp]:
+ "llist_when\<cdot>z\<cdot>f\<cdot>(repeat\<cdot>x) = f\<cdot>x\<cdot>(repeat\<cdot>x)"
+by (subst repeat.simps, simp)
+
+text {*
+ For mutually-recursive constants, looping might only occur if all
+ equations are in the simpset at the same time. In such cases it may
+ only be necessary to declare @{text "[simp del]"} on one equation.
+*}
+
+fixrec
+ inf_tree :: "'a tree" and inf_forest :: "'a forest"
+where
+ [simp del]: "inf_tree = Branch\<cdot>inf_forest"
+| "inf_forest = Trees\<cdot>inf_tree\<cdot>(Trees\<cdot>inf_tree\<cdot>Empty)"
+
+
+subsection {* Using @{text fixrec} inside locales *}
+
+locale test =
+ fixes foo :: "'a \<rightarrow> 'a"
+ assumes foo_strict: "foo\<cdot>\<bottom> = \<bottom>"
+begin
+
+fixrec
+ bar :: "'a u \<rightarrow> 'a"
+where
+ "bar\<cdot>(up\<cdot>x) = foo\<cdot>x"
+
+lemma bar_strict: "bar\<cdot>\<bottom> = \<bottom>"
+by fixrec_simp
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/New_Domain.thy Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,92 @@
+(* Title: HOLCF/ex/New_Domain.thy
+ Author: Brian Huffman
+*)
+
+header {* Definitional domain package *}
+
+theory New_Domain
+imports HOLCF
+begin
+
+text {*
+ The definitional domain package only works with representable domains,
+ i.e. types in class @{text rep}.
+*}
+
+default_sort rep
+
+text {*
+ Provided that @{text rep} is the default sort, the @{text new_domain}
+ package should work with any type definition supported by the old
+ domain package.
+*}
+
+new_domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
+
+text {*
+ The difference is that the new domain package is completely
+ definitional, and does not generate any axioms. The following type
+ and constant definitions are not produced by the old domain package.
+*}
+
+thm type_definition_llist
+thm llist_abs_def llist_rep_def
+
+text {*
+ The new domain package also adds support for indirect recursion with
+ user-defined datatypes. This definition of a tree datatype uses
+ indirect recursion through the lazy list type constructor.
+*}
+
+new_domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
+
+text {*
+ For indirect-recursive definitions, the domain package is not able to
+ generate a high-level induction rule. (It produces a warning
+ message instead.) The low-level reach lemma (now proved as a
+ theorem, no longer generated as an axiom) can be used to derive
+ other induction rules.
+*}
+
+thm ltree.reach
+
+text {*
+ The definition of the take function uses map functions associated with
+ each type constructor involved in the definition. A map function
+ for the lazy list type has been generated by the new domain package.
+*}
+
+thm ltree.take_rews
+thm llist_map_def
+
+lemma ltree_induct:
+ fixes P :: "'a ltree \<Rightarrow> bool"
+ assumes adm: "adm P"
+ assumes bot: "P \<bottom>"
+ assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
+ assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
+ shows "P x"
+proof -
+ have "P (\<Squnion>i. ltree_take i\<cdot>x)"
+ using adm
+ proof (rule admD)
+ fix i
+ show "P (ltree_take i\<cdot>x)"
+ proof (induct i arbitrary: x)
+ case (0 x)
+ show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
+ next
+ case (Suc n x)
+ show "P (ltree_take (Suc n)\<cdot>x)"
+ apply (cases x)
+ apply (simp add: bot)
+ apply (simp add: Leaf)
+ apply (simp add: Branch Suc)
+ done
+ qed
+ qed (simp add: ltree.chain_take)
+ thus ?thesis
+ by (simp add: ltree.reach)
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/ROOT.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,1 @@
+use_thys ["Domain_ex", "Fixrec_ex", "New_Domain"];
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Tutorial/document/root.tex Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,29 @@
+
+% HOLCF/document/root.tex
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx,isabelle,isabellesym,latexsym}
+\usepackage[only,bigsqcap]{stmaryrd}
+\usepackage[latin1]{inputenc}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+%\isabellestyle{it}
+\pagestyle{myheadings}
+
+\begin{document}
+
+\title{Isabelle/HOLCF Tutorial}
+\maketitle
+
+\tableofcontents
+
+%\newpage
+
+%\renewcommand{\isamarkupheader}[1]%
+%{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOLCF/ex/Domain_ex.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,210 +0,0 @@
-(* Title: HOLCF/ex/Domain_ex.thy
- Author: Brian Huffman
-*)
-
-header {* Domain package examples *}
-
-theory Domain_ex
-imports HOLCF
-begin
-
-text {* Domain constructors are strict by default. *}
-
-domain d1 = d1a | d1b "d1" "d1"
-
-lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
-
-text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
-
-domain d2 = d2a | d2b (lazy "d2")
-
-lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
-
-text {* Strict and lazy arguments may be mixed arbitrarily. *}
-
-domain d3 = d3a | d3b (lazy "d2") "d2"
-
-lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
-
-text {* Selectors can be used with strict or lazy constructor arguments. *}
-
-domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
-
-lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
-
-text {* Mixfix declarations can be given for data constructors. *}
-
-domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
-
-lemma "d5a \<noteq> x :#: y :#: z" by simp
-
-text {* Mixfix declarations can also be given for type constructors. *}
-
-domain ('a, 'b) lazypair (infixl ":*:" 25) =
- lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
-
-lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
-by (rule allI, case_tac p, simp_all)
-
-text {* Non-recursive constructor arguments can have arbitrary types. *}
-
-domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
-
-text {*
- Indirect recusion is allowed for sums, products, lifting, and the
- continuous function space. However, the domain package does not
- generate an induction rule in terms of the constructors.
-*}
-
-domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c (lazy "'a d7 \<rightarrow> 'a")
- -- "Indirect recursion detected, skipping proofs of (co)induction rules"
-(* d7.induct is absent *)
-
-text {*
- Indirect recursion is also allowed using previously-defined datatypes.
-*}
-
-domain 'a slist = SNil | SCons 'a "'a slist"
-
-domain 'a stree = STip | SBranch "'a stree slist"
-
-text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
-
-domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
-
-text {* Non-regular recursion is not allowed. *}
-(*
-domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
- -- "illegal direct recursion with different arguments"
-domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
- -- "illegal direct recursion with different arguments"
-*)
-
-text {*
- Mutually-recursive datatypes must have all the same type arguments,
- not necessarily in the same order.
-*}
-
-domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
- and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
-
-text {* Induction rules for flat datatypes have no admissibility side-condition. *}
-
-domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
-
-lemma "\<lbrakk>P \<bottom>; P Tip; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; P x; P y\<rbrakk> \<Longrightarrow> P (Branch\<cdot>x\<cdot>y)\<rbrakk> \<Longrightarrow> P x"
-by (rule flattree.induct) -- "no admissibility requirement"
-
-text {* Trivial datatypes will produce a warning message. *}
-
-domain triv = Triv triv triv
- -- "domain Domain_ex.triv is empty!"
-
-lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
-
-text {* Lazy constructor arguments may have unpointed types. *}
-
-domain natlist = nnil | ncons (lazy "nat discr") natlist
-
-text {* Class constraints may be given for type parameters on the LHS. *}
-
-domain ('a::cpo) box = Box (lazy 'a)
-
-
-subsection {* Generated constants and theorems *}
-
-domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (right :: "'a tree")
-
-lemmas tree_abs_defined_iff =
- iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
-
-text {* Rules about ismorphism *}
-term tree_rep
-term tree_abs
-thm tree.rep_iso
-thm tree.abs_iso
-thm tree.iso_rews
-
-text {* Rules about constructors *}
-term Leaf
-term Node
-thm Leaf_def Node_def
-thm tree.nchotomy
-thm tree.exhaust
-thm tree.compacts
-thm tree.con_rews
-thm tree.dist_les
-thm tree.dist_eqs
-thm tree.inverts
-thm tree.injects
-
-text {* Rules about case combinator *}
-term tree_when
-thm tree.tree_when_def
-thm tree.when_rews
-
-text {* Rules about selectors *}
-term left
-term right
-thm tree.sel_rews
-
-text {* Rules about discriminators *}
-term is_Leaf
-term is_Node
-thm tree.dis_rews
-
-text {* Rules about pattern match combinators *}
-term Leaf_pat
-term Node_pat
-thm tree.pat_rews
-
-text {* Rules about monadic pattern match combinators *}
-term match_Leaf
-term match_Node
-thm tree.match_rews
-
-text {* Rules about take function *}
-term tree_take
-thm tree.take_def
-thm tree.take_0
-thm tree.take_Suc
-thm tree.take_rews
-thm tree.chain_take
-thm tree.take_take
-thm tree.deflation_take
-thm tree.take_below
-thm tree.take_lemma
-thm tree.lub_take
-thm tree.reach
-thm tree.finite_induct
-
-text {* Rules about finiteness predicate *}
-term tree_finite
-thm tree.finite_def
-thm tree.finite (* only generated for flat datatypes *)
-
-text {* Rules about bisimulation predicate *}
-term tree_bisim
-thm tree.bisim_def
-thm tree.coinduct
-
-text {* Induction rule *}
-thm tree.induct
-
-
-subsection {* Known bugs *}
-
-text {* Declaring a mixfix with spaces causes some strange parse errors. *}
-(*
-domain xx = xx ("x y")
- -- "Inner syntax error: unexpected end of input"
-*)
-
-text {*
- Non-cpo type parameters currently do not work.
-*}
-(*
-domain ('a::type) stream = snil | scons (lazy "'a discr") "'a stream"
-*)
-
-end
--- a/src/HOLCF/ex/Fixrec_ex.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,203 +0,0 @@
-(* Title: HOLCF/ex/Fixrec_ex.thy
- Author: Brian Huffman
-*)
-
-header {* Fixrec package examples *}
-
-theory Fixrec_ex
-imports HOLCF
-begin
-
-subsection {* Basic @{text fixrec} examples *}
-
-text {*
- Fixrec patterns can mention any constructor defined by the domain
- package, as well as any of the following built-in constructors:
- Pair, spair, sinl, sinr, up, ONE, TT, FF.
-*}
-
-text {* Typical usage is with lazy constructors. *}
-
-fixrec down :: "'a u \<rightarrow> 'a"
-where "down\<cdot>(up\<cdot>x) = x"
-
-text {* With strict constructors, rewrite rules may require side conditions. *}
-
-fixrec from_sinl :: "'a \<oplus> 'b \<rightarrow> 'a"
-where "x \<noteq> \<bottom> \<Longrightarrow> from_sinl\<cdot>(sinl\<cdot>x) = x"
-
-text {* Lifting can turn a strict constructor into a lazy one. *}
-
-fixrec from_sinl_up :: "'a u \<oplus> 'b \<rightarrow> 'a"
-where "from_sinl_up\<cdot>(sinl\<cdot>(up\<cdot>x)) = x"
-
-text {* Fixrec also works with the HOL pair constructor. *}
-
-fixrec down2 :: "'a u \<times> 'b u \<rightarrow> 'a \<times> 'b"
-where "down2\<cdot>(up\<cdot>x, up\<cdot>y) = (x, y)"
-
-
-subsection {* Examples using @{text fixrec_simp} *}
-
-text {* A type of lazy lists. *}
-
-domain 'a llist = lNil | lCons (lazy 'a) (lazy "'a llist")
-
-text {* A zip function for lazy lists. *}
-
-text {* Notice that the patterns are not exhaustive. *}
-
-fixrec
- lzip :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
-where
- "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
-| "lzip\<cdot>lNil\<cdot>lNil = lNil"
-
-text {* @{text fixrec_simp} is useful for producing strictness theorems. *}
-text {* Note that pattern matching is done in left-to-right order. *}
-
-lemma lzip_stricts [simp]:
- "lzip\<cdot>\<bottom>\<cdot>ys = \<bottom>"
- "lzip\<cdot>lNil\<cdot>\<bottom> = \<bottom>"
- "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp+
-
-text {* @{text fixrec_simp} can also produce rules for missing cases. *}
-
-lemma lzip_undefs [simp]:
- "lzip\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = \<bottom>"
- "lzip\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = \<bottom>"
-by fixrec_simp+
-
-
-subsection {* Pattern matching with bottoms *}
-
-text {*
- As an alternative to using @{text fixrec_simp}, it is also possible
- to use bottom as a constructor pattern. When using a bottom
- pattern, the right-hand-side must also be bottom; otherwise, @{text
- fixrec} will not be able to prove the equation.
-*}
-
-fixrec
- from_sinr_up :: "'a \<oplus> 'b\<^sub>\<bottom> \<rightarrow> 'b"
-where
- "from_sinr_up\<cdot>\<bottom> = \<bottom>"
-| "from_sinr_up\<cdot>(sinr\<cdot>(up\<cdot>x)) = x"
-
-text {*
- If the function is already strict in that argument, then the bottom
- pattern does not change the meaning of the function. For example,
- in the definition of @{term from_sinr_up}, the first equation is
- actually redundant, and could have been proven separately by
- @{text fixrec_simp}.
-*}
-
-text {*
- A bottom pattern can also be used to make a function strict in a
- certain argument, similar to a bang-pattern in Haskell.
-*}
-
-fixrec
- seq :: "'a \<rightarrow> 'b \<rightarrow> 'b"
-where
- "seq\<cdot>\<bottom>\<cdot>y = \<bottom>"
-| "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x\<cdot>y = y"
-
-
-subsection {* Skipping proofs of rewrite rules *}
-
-text {* Another zip function for lazy lists. *}
-
-text {*
- Notice that this version has overlapping patterns.
- The second equation cannot be proved as a theorem
- because it only applies when the first pattern fails.
-*}
-
-fixrec (permissive)
- lzip2 :: "'a llist \<rightarrow> 'b llist \<rightarrow> ('a \<times> 'b) llist"
-where
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
-| "lzip2\<cdot>xs\<cdot>ys = lNil"
-
-text {*
- Usually fixrec tries to prove all equations as theorems.
- The "permissive" option overrides this behavior, so fixrec
- does not produce any simp rules.
-*}
-
-text {* Simp rules can be generated later using @{text fixrec_simp}. *}
-
-lemma lzip2_simps [simp]:
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>(lCons\<cdot>y\<cdot>ys) = lCons\<cdot>(x, y)\<cdot>(lzip\<cdot>xs\<cdot>ys)"
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>lNil = lNil"
- "lzip2\<cdot>lNil\<cdot>(lCons\<cdot>y\<cdot>ys) = lNil"
- "lzip2\<cdot>lNil\<cdot>lNil = lNil"
-by fixrec_simp+
-
-lemma lzip2_stricts [simp]:
- "lzip2\<cdot>\<bottom>\<cdot>ys = \<bottom>"
- "lzip2\<cdot>(lCons\<cdot>x\<cdot>xs)\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp+
-
-
-subsection {* Mutual recursion with @{text fixrec} *}
-
-text {* Tree and forest types. *}
-
-domain 'a tree = Leaf (lazy 'a) | Branch (lazy "'a forest")
-and 'a forest = Empty | Trees (lazy "'a tree") "'a forest"
-
-text {*
- To define mutually recursive functions, give multiple type signatures
- separated by the keyword @{text "and"}.
-*}
-
-fixrec
- map_tree :: "('a \<rightarrow> 'b) \<rightarrow> ('a tree \<rightarrow> 'b tree)"
-and
- map_forest :: "('a \<rightarrow> 'b) \<rightarrow> ('a forest \<rightarrow> 'b forest)"
-where
- "map_tree\<cdot>f\<cdot>(Leaf\<cdot>x) = Leaf\<cdot>(f\<cdot>x)"
-| "map_tree\<cdot>f\<cdot>(Branch\<cdot>ts) = Branch\<cdot>(map_forest\<cdot>f\<cdot>ts)"
-| "map_forest\<cdot>f\<cdot>Empty = Empty"
-| "ts \<noteq> \<bottom> \<Longrightarrow>
- map_forest\<cdot>f\<cdot>(Trees\<cdot>t\<cdot>ts) = Trees\<cdot>(map_tree\<cdot>f\<cdot>t)\<cdot>(map_forest\<cdot>f\<cdot>ts)"
-
-lemma map_tree_strict [simp]: "map_tree\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-lemma map_forest_strict [simp]: "map_forest\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-text {*
- Theorems generated:
- @{text map_tree_def} @{thm map_tree_def}
- @{text map_forest_def} @{thm map_forest_def}
- @{text map_tree.unfold} @{thm map_tree.unfold}
- @{text map_forest.unfold} @{thm map_forest.unfold}
- @{text map_tree.simps} @{thm map_tree.simps}
- @{text map_forest.simps} @{thm map_forest.simps}
- @{text map_tree_map_forest.induct} @{thm map_tree_map_forest.induct}
-*}
-
-
-subsection {* Using @{text fixrec} inside locales *}
-
-locale test =
- fixes foo :: "'a \<rightarrow> 'a"
- assumes foo_strict: "foo\<cdot>\<bottom> = \<bottom>"
-begin
-
-fixrec
- bar :: "'a u \<rightarrow> 'a"
-where
- "bar\<cdot>(up\<cdot>x) = foo\<cdot>x"
-
-lemma bar_strict: "bar\<cdot>\<bottom> = \<bottom>"
-by fixrec_simp
-
-end
-
-end
--- a/src/HOLCF/ex/New_Domain.thy Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-(* Title: HOLCF/ex/New_Domain.thy
- Author: Brian Huffman
-*)
-
-header {* Definitional domain package *}
-
-theory New_Domain
-imports HOLCF
-begin
-
-text {*
- The definitional domain package only works with representable domains,
- i.e. types in class @{text rep}.
-*}
-
-default_sort rep
-
-text {*
- Provided that @{text rep} is the default sort, the @{text new_domain}
- package should work with any type definition supported by the old
- domain package.
-*}
-
-new_domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
-
-text {*
- The difference is that the new domain package is completely
- definitional, and does not generate any axioms. The following type
- and constant definitions are not produced by the old domain package.
-*}
-
-thm type_definition_llist
-thm llist_abs_def llist_rep_def
-
-text {*
- The new domain package also adds support for indirect recursion with
- user-defined datatypes. This definition of a tree datatype uses
- indirect recursion through the lazy list type constructor.
-*}
-
-new_domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
-
-text {*
- For indirect-recursive definitions, the domain package is not able to
- generate a high-level induction rule. (It produces a warning
- message instead.) The low-level reach lemma (now proved as a
- theorem, no longer generated as an axiom) can be used to derive
- other induction rules.
-*}
-
-thm ltree.reach
-
-text {*
- The definition of the take function uses map functions associated with
- each type constructor involved in the definition. A map function
- for the lazy list type has been generated by the new domain package.
-*}
-
-thm ltree.take_rews
-thm llist_map_def
-
-lemma ltree_induct:
- fixes P :: "'a ltree \<Rightarrow> bool"
- assumes adm: "adm P"
- assumes bot: "P \<bottom>"
- assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
- assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
- shows "P x"
-proof -
- have "P (\<Squnion>i. ltree_take i\<cdot>x)"
- using adm
- proof (rule admD)
- fix i
- show "P (ltree_take i\<cdot>x)"
- proof (induct i arbitrary: x)
- case (0 x)
- show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
- next
- case (Suc n x)
- show "P (ltree_take (Suc n)\<cdot>x)"
- apply (cases x)
- apply (simp add: bot)
- apply (simp add: Leaf)
- apply (simp add: Branch Suc)
- done
- qed
- qed (simp add: ltree.chain_take)
- thus ?thesis
- by (simp add: ltree.reach)
-qed
-
-end
--- a/src/HOLCF/ex/ROOT.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/HOLCF/ex/ROOT.ML Sun May 23 10:38:11 2010 +0100
@@ -4,7 +4,6 @@
*)
use_thys ["Dnat", "Stream", "Dagstuhl", "Focus_ex", "Fix2", "Hoare",
- "Loop", "Fixrec_ex", "Powerdomain_ex", "Domain_ex", "Domain_Proofs",
+ "Loop", "Powerdomain_ex", "Domain_Proofs",
"Letrec",
- "Strict_Fun",
- "New_Domain"];
+ "Strict_Fun"];
--- a/src/Provers/Arith/fast_lin_arith.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML Sun May 23 10:38:11 2010 +0100
@@ -845,8 +845,8 @@
(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
+ val thm1 = Thm.assume ct1
+ val thm2 = Thm.assume ct2
in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]),
ct2, elim_neq neqs (asms', asms@[thm2]))
end
@@ -858,8 +858,8 @@
let
val (thm1, js1) = fwdproof ss tree1 js
val (thm2, js2) = fwdproof ss tree2 js1
- val thm1' = implies_intr ct1 thm1
- val thm2' = implies_intr ct2 thm2
+ val thm1' = Thm.implies_intr ct1 thm1
+ val thm2' = Thm.implies_intr ct2 thm2
in (thm2' COMP (thm1' COMP thm), js2) end;
(* FIXME needs handle THM _ => NONE ? *)
@@ -869,11 +869,11 @@
val thy = ProofContext.theory_of ctxt
val nTconcl = LA_Logic.neg_prop Tconcl
val cnTconcl = cterm_of thy nTconcl
- val nTconclthm = assume cnTconcl
+ val nTconclthm = Thm.assume cnTconcl
val tree = (if split_neq then splitasms ctxt else Tip) (thms @ [nTconclthm])
val (Falsethm, _) = fwdproof ss tree js
val contr = if pos then LA_Logic.ccontr else LA_Logic.notI
- val concl = implies_intr cnTconcl Falsethm COMP contr
+ val concl = Thm.implies_intr cnTconcl Falsethm COMP contr
in SOME (trace_thm ctxt "Proved by lin. arith. prover:" (LA_Logic.mk_Eq concl)) end
(*in case concl contains ?-var, which makes assume fail:*) (* FIXME Variable.import_terms *)
handle THM _ => NONE;
--- a/src/Provers/blast.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Provers/blast.ML Sun May 23 10:38:11 2010 +0100
@@ -1309,7 +1309,7 @@
val setup =
setup_depth_limit #>
Method.setup @{binding blast}
- (Scan.lift (Scan.option OuterParse.nat) --| Method.sections Data.cla_modifiers >>
+ (Scan.lift (Scan.option Parse.nat) --| Method.sections Data.cla_modifiers >>
(fn NONE => Data.cla_meth' blast_tac
| SOME lim => Data.cla_meth' (fn cs => depth_tac cs lim)))
"classical tableau prover";
--- a/src/Provers/clasimp.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Provers/clasimp.ML Sun May 23 10:38:11 2010 +0100
@@ -275,7 +275,7 @@
Method.sections clasimp_modifiers >> K (clasimp_meth' tac);
val auto_method =
- Scan.lift (Scan.option (OuterParse.nat -- OuterParse.nat)) --|
+ Scan.lift (Scan.option (Parse.nat -- Parse.nat)) --|
Method.sections clasimp_modifiers >>
(fn NONE => clasimp_meth (CHANGED_PROP o auto_tac)
| SOME (m, n) => clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n)));
--- a/src/Provers/classical.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Provers/classical.ML Sun May 23 10:38:11 2010 +0100
@@ -1015,8 +1015,8 @@
(** outer syntax **)
val _ =
- OuterSyntax.improper_command "print_claset" "print context of Classical Reasoner"
- OuterKeyword.diag
+ Outer_Syntax.improper_command "print_claset" "print context of Classical Reasoner"
+ Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
Toplevel.keep (print_cs o claset_of o Toplevel.context_of)));
--- a/src/Provers/hypsubst.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Provers/hypsubst.ML Sun May 23 10:38:11 2010 +0100
@@ -143,7 +143,7 @@
[ if inspect_pair bnd false (Data.dest_eq
(Data.dest_Trueprop (#prop (rep_thm th))))
then th RS Data.eq_reflection
- else symmetric(th RS Data.eq_reflection) (*reorient*) ]
+ else Thm.symmetric(th RS Data.eq_reflection) (*reorient*) ]
handle TERM _ => [] | Match => [];
local
--- a/src/Pure/Concurrent/future.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Concurrent/future.ML Sun May 23 10:38:11 2010 +0100
@@ -170,12 +170,14 @@
fun future_job group (e: unit -> 'a) =
let
val result = Single_Assignment.var "future" : 'a result;
+ val pos = Position.thread_data ();
fun job ok =
let
val res =
if ok then
Exn.capture (fn () =>
- Multithreading.with_attributes Multithreading.private_interrupts (fn _ => e ())) ()
+ Multithreading.with_attributes Multithreading.private_interrupts
+ (fn _ => Position.setmp_thread_data pos e ())) ()
else Exn.Exn Exn.Interrupt;
in assign_result group result res end;
in (result, job) end;
--- a/src/Pure/General/position.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/General/position.ML Sun May 23 10:38:11 2010 +0100
@@ -173,9 +173,7 @@
fun thread_data () = the_default none (Thread.getLocal tag);
-fun setmp_thread_data pos f x =
- if ! Output.debugging then f x
- else Library.setmp_thread_data tag (thread_data ()) pos f x;
+fun setmp_thread_data pos = Library.setmp_thread_data tag (thread_data ()) pos;
end;
--- a/src/Pure/General/scan.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/General/scan.scala Sun May 23 10:38:11 2010 +0100
@@ -258,47 +258,44 @@
/* outer syntax tokens */
- def token(symbols: Symbol.Interpretation, is_command: String => Boolean):
- Parser[Outer_Lex.Token] =
+ def token(symbols: Symbol.Interpretation, is_command: String => Boolean): Parser[Token] =
{
- import Outer_Lex.Token_Kind, Outer_Lex.Token
-
val id = one(symbols.is_letter) ~ many(symbols.is_letdig) ^^ { case x ~ y => x + y }
val nat = many1(symbols.is_digit)
val id_nat = id ~ opt("." ~ nat) ^^ { case x ~ Some(y ~ z) => x + y + z case x ~ None => x }
val ident = id ~ rep("." ~> id) ^^
- { case x ~ Nil => Token(Token_Kind.IDENT, x)
- case x ~ ys => Token(Token_Kind.LONG_IDENT, (x :: ys).mkString(".")) }
+ { case x ~ Nil => Token(Token.Kind.IDENT, x)
+ case x ~ ys => Token(Token.Kind.LONG_IDENT, (x :: ys).mkString(".")) }
- val var_ = "?" ~ id_nat ^^ { case x ~ y => Token(Token_Kind.VAR, x + y) }
- val type_ident = "'" ~ id ^^ { case x ~ y => Token(Token_Kind.TYPE_IDENT, x + y) }
- val type_var = "?'" ~ id_nat ^^ { case x ~ y => Token(Token_Kind.TYPE_VAR, x + y) }
- val nat_ = nat ^^ (x => Token(Token_Kind.NAT, x))
+ val var_ = "?" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.VAR, x + y) }
+ val type_ident = "'" ~ id ^^ { case x ~ y => Token(Token.Kind.TYPE_IDENT, x + y) }
+ val type_var = "?'" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.TYPE_VAR, x + y) }
+ val nat_ = nat ^^ (x => Token(Token.Kind.NAT, x))
val sym_ident =
(many1(symbols.is_symbolic_char) |
one(sym => symbols.is_symbolic(sym) & Symbol.is_wellformed(sym))) ^^
- (x => Token(Token_Kind.SYM_IDENT, x))
+ (x => Token(Token.Kind.SYM_IDENT, x))
- val space = many1(symbols.is_blank) ^^ (x => Token(Token_Kind.SPACE, x))
+ val space = many1(symbols.is_blank) ^^ (x => Token(Token.Kind.SPACE, x))
- val string = quoted("\"") ^^ (x => Token(Token_Kind.STRING, x))
- val alt_string = quoted("`") ^^ (x => Token(Token_Kind.ALT_STRING, x))
+ val string = quoted("\"") ^^ (x => Token(Token.Kind.STRING, x))
+ val alt_string = quoted("`") ^^ (x => Token(Token.Kind.ALT_STRING, x))
val junk = many1(sym => !(symbols.is_blank(sym)))
val bad_delimiter =
- ("\"" | "`" | "{*" | "(*") ~ junk ^^ { case x ~ y => Token(Token_Kind.BAD_INPUT, x + y) }
- val bad = junk ^^ (x => Token(Token_Kind.BAD_INPUT, x))
+ ("\"" | "`" | "{*" | "(*") ~ junk ^^ { case x ~ y => Token(Token.Kind.BAD_INPUT, x + y) }
+ val bad = junk ^^ (x => Token(Token.Kind.BAD_INPUT, x))
/* tokens */
- (space | (string | (alt_string | (verbatim ^^ (x => Token(Token_Kind.VERBATIM, x)) |
- comment ^^ (x => Token(Token_Kind.COMMENT, x)))))) |
+ (space | (string | (alt_string | (verbatim ^^ (x => Token(Token.Kind.VERBATIM, x)) |
+ comment ^^ (x => Token(Token.Kind.COMMENT, x)))))) |
bad_delimiter |
((ident | (var_ | (type_ident | (type_var | (nat_ | sym_ident))))) |||
- keyword ^^ (x => Token(if (is_command(x)) Token_Kind.COMMAND else Token_Kind.KEYWORD, x))) |
+ keyword ^^ (x => Token(if (is_command(x)) Token.Kind.COMMAND else Token.Kind.KEYWORD, x))) |
bad
}
}
--- a/src/Pure/General/symbol_pos.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/General/symbol_pos.ML Sun May 23 10:38:11 2010 +0100
@@ -4,17 +4,12 @@
Symbols with explicit position information.
*)
-signature BASIC_SYMBOL_POS =
+signature SYMBOL_POS =
sig
type T = Symbol.symbol * Position.T
val symbol: T -> Symbol.symbol
val $$$ : Symbol.symbol -> T list -> T list * T list
val ~$$$ : Symbol.symbol -> T list -> T list * T list
-end
-
-signature SYMBOL_POS =
-sig
- include BASIC_SYMBOL_POS
val content: T list -> string
val untabify_content: T list -> string
val is_eof: T -> bool
@@ -185,5 +180,10 @@
end;
-structure Basic_Symbol_Pos: BASIC_SYMBOL_POS = Symbol_Pos; (*not open by default*)
+structure Basic_Symbol_Pos = (*not open by default*)
+struct
+ val symbol = Symbol_Pos.symbol;
+ val $$$ = Symbol_Pos.$$$;
+ val ~$$$ = Symbol_Pos.~$$$;
+end;
--- a/src/Pure/IsaMakefile Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/IsaMakefile Sun May 23 10:38:11 2010 +0100
@@ -12,8 +12,6 @@
## global settings
-SHELL = /bin/bash
-
SRC = $(ISABELLE_HOME)/src
OUT = $(ISABELLE_OUTPUT)
LOG = $(OUT)/log
@@ -65,16 +63,16 @@
Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML \
Isar/constdefs.ML Isar/context_rules.ML Isar/element.ML \
Isar/expression.ML Isar/isar_cmd.ML Isar/isar_document.ML \
- Isar/isar_syn.ML Isar/local_defs.ML Isar/local_syntax.ML \
- Isar/local_theory.ML Isar/locale.ML Isar/method.ML \
- Isar/object_logic.ML Isar/obtain.ML Isar/outer_keyword.ML \
- Isar/outer_lex.ML Isar/outer_parse.ML Isar/outer_syntax.ML \
- Isar/overloading.ML Isar/proof.ML Isar/proof_context.ML \
- Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML \
- Isar/rule_insts.ML Isar/runtime.ML Isar/skip_proof.ML \
- Isar/spec_parse.ML Isar/spec_rules.ML Isar/specification.ML \
- Isar/theory_target.ML Isar/toplevel.ML Isar/typedecl.ML \
- Isar/value_parse.ML ML/ml_antiquote.ML ML/ml_compiler.ML \
+ Isar/isar_syn.ML Isar/keyword.ML Isar/local_defs.ML \
+ Isar/local_syntax.ML Isar/local_theory.ML Isar/locale.ML \
+ Isar/method.ML Isar/object_logic.ML Isar/obtain.ML \
+ Isar/outer_syntax.ML Isar/overloading.ML Isar/parse.ML \
+ Isar/parse_spec.ML Isar/parse_value.ML Isar/proof.ML \
+ Isar/proof_context.ML Isar/proof_display.ML Isar/proof_node.ML \
+ Isar/rule_cases.ML Isar/rule_insts.ML Isar/runtime.ML \
+ Isar/skip_proof.ML Isar/spec_rules.ML Isar/specification.ML \
+ Isar/theory_target.ML Isar/token.ML Isar/toplevel.ML \
+ Isar/typedecl.ML ML/ml_antiquote.ML ML/ml_compiler.ML \
ML/ml_compiler_polyml-5.3.ML ML/ml_context.ML ML/ml_env.ML \
ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML ML/ml_thms.ML \
ML-Systems/install_pp_polyml.ML ML-Systems/install_pp_polyml-5.3.ML \
--- a/src/Pure/Isar/args.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/args.ML Sun May 23 10:38:11 2010 +0100
@@ -7,10 +7,9 @@
signature ARGS =
sig
- type token = OuterLex.token
type src
- val src: (string * token list) * Position.T -> src
- val dest_src: src -> (string * token list) * Position.T
+ val src: (string * Token.T list) * Position.T -> src
+ val dest_src: src -> (string * Token.T list) * Position.T
val pretty_src: Proof.context -> src -> Pretty.T
val map_name: (string -> string) -> src -> src
val morph_values: morphism -> src -> src
@@ -57,8 +56,8 @@
val const: bool -> string context_parser
val const_proper: bool -> string context_parser
val goal_spec: ((int -> tactic) -> tactic) context_parser
- val parse: token list parser
- val parse1: (string -> bool) -> token list parser
+ val parse: Token.T list parser
+ val parse1: (string -> bool) -> Token.T list parser
val attribs: (string -> string) -> src list parser
val opt_attribs: (string -> string) -> src list parser
val thm_name: (string -> string) -> string -> (binding * src list) parser
@@ -70,16 +69,9 @@
structure Args: ARGS =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-
-type token = T.token;
-
-
-
(** datatype src **)
-datatype src = Src of (string * token list) * Position.T;
+datatype src = Src of (string * Token.T list) * Position.T;
val src = Src;
fun dest_src (Src src) = src;
@@ -88,12 +80,12 @@
let
val prt_thm = Pretty.backquote o Display.pretty_thm ctxt;
fun prt arg =
- (case T.get_value arg of
- SOME (T.Text s) => Pretty.str (quote s)
- | SOME (T.Typ T) => Syntax.pretty_typ ctxt T
- | SOME (T.Term t) => Syntax.pretty_term ctxt t
- | SOME (T.Fact ths) => Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
- | _ => Pretty.str (T.unparse arg));
+ (case Token.get_value arg of
+ SOME (Token.Text s) => Pretty.str (quote s)
+ | SOME (Token.Typ T) => Syntax.pretty_typ ctxt T
+ | SOME (Token.Term t) => Syntax.pretty_term ctxt t
+ | SOME (Token.Fact ths) => Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
+ | _ => Pretty.str (Token.unparse arg));
val (s, args) = #1 (dest_src src);
in Pretty.block (Pretty.breaks (Pretty.str s :: map prt args)) end;
@@ -103,15 +95,15 @@
(* values *)
-fun morph_values phi = map_args (T.map_value
- (fn T.Text s => T.Text s
- | T.Typ T => T.Typ (Morphism.typ phi T)
- | T.Term t => T.Term (Morphism.term phi t)
- | T.Fact ths => T.Fact (Morphism.fact phi ths)
- | T.Attribute att => T.Attribute (Morphism.transform phi att)));
+fun morph_values phi = map_args (Token.map_value
+ (fn Token.Text s => Token.Text s
+ | Token.Typ T => Token.Typ (Morphism.typ phi T)
+ | Token.Term t => Token.Term (Morphism.term phi t)
+ | Token.Fact ths => Token.Fact (Morphism.fact phi ths)
+ | Token.Attribute att => Token.Attribute (Morphism.transform phi att)));
-val assignable = map_args T.assignable;
-val closure = map_args T.closure;
+val assignable = map_args Token.assignable;
+val closure = map_args Token.closure;
@@ -125,17 +117,17 @@
(* basic *)
-fun token atom = Scan.ahead P.not_eof --| atom;
+fun token atom = Scan.ahead Parse.not_eof --| atom;
val ident = token
- (P.short_ident || P.long_ident || P.sym_ident || P.term_var ||
- P.type_ident || P.type_var || P.number);
+ (Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var ||
+ Parse.type_ident || Parse.type_var || Parse.number);
-val string = token (P.string || P.verbatim);
-val alt_string = token P.alt_string;
-val symbolic = token P.keyword_ident_or_symbolic;
+val string = token (Parse.string || Parse.verbatim);
+val alt_string = token Parse.alt_string;
+val symbolic = token Parse.keyword_ident_or_symbolic;
-fun $$$ x = (ident >> T.content_of || P.keyword)
+fun $$$ x = (ident >> Token.content_of || Parse.keyword)
:|-- (fn y => if x = y then Scan.succeed x else Scan.fail);
@@ -154,39 +146,40 @@
fun mode s = Scan.lift (Scan.optional (parens ($$$ s) >> K true) false);
fun maybe scan = $$$ "_" >> K NONE || scan >> SOME;
-val name_source = named >> T.source_of;
-val name_source_position = named >> T.source_position_of;
+val name_source = named >> Token.source_of;
+val name_source_position = named >> Token.source_position_of;
-val name = named >> T.content_of;
-val binding = P.position name >> Binding.make;
-val alt_name = alt_string >> T.content_of;
-val symbol = symbolic >> T.content_of;
+val name = named >> Token.content_of;
+val binding = Parse.position name >> Binding.make;
+val alt_name = alt_string >> Token.content_of;
+val symbol = symbolic >> Token.content_of;
val liberal_name = symbol || name;
-val var = (ident >> T.content_of) :|-- (fn x =>
+val var = (ident >> Token.content_of) :|-- (fn x =>
(case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail));
(* values *)
fun value dest = Scan.some (fn arg =>
- (case T.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE));
+ (case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE));
fun evaluate mk eval arg =
- let val x = eval arg in (T.assign (SOME (mk x)) arg; x) end;
+ let val x = eval arg in (Token.assign (SOME (mk x)) arg; x) end;
+
+val internal_text = value (fn Token.Text s => s);
+val internal_typ = value (fn Token.Typ T => T);
+val internal_term = value (fn Token.Term t => t);
+val internal_fact = value (fn Token.Fact ths => ths);
+val internal_attribute = value (fn Token.Attribute att => att);
-val internal_text = value (fn T.Text s => s);
-val internal_typ = value (fn T.Typ T => T);
-val internal_term = value (fn T.Term t => t);
-val internal_fact = value (fn T.Fact ths => ths);
-val internal_attribute = value (fn T.Attribute att => att);
-
-fun named_text intern = internal_text || named >> evaluate T.Text (intern o T.content_of);
-fun named_typ readT = internal_typ || named >> evaluate T.Typ (readT o T.source_of);
-fun named_term read = internal_term || named >> evaluate T.Term (read o T.source_of);
-fun named_fact get = internal_fact || named >> evaluate T.Fact (get o T.content_of) ||
- alt_string >> evaluate T.Fact (get o T.source_of);
-fun named_attribute att = internal_attribute || named >> evaluate T.Attribute (att o T.content_of);
+fun named_text intern = internal_text || named >> evaluate Token.Text (intern o Token.content_of);
+fun named_typ readT = internal_typ || named >> evaluate Token.Typ (readT o Token.source_of);
+fun named_term read = internal_term || named >> evaluate Token.Term (read o Token.source_of);
+fun named_fact get = internal_fact || named >> evaluate Token.Fact (get o Token.content_of) ||
+ alt_string >> evaluate Token.Fact (get o Token.source_of);
+fun named_attribute att =
+ internal_attribute || named >> evaluate Token.Attribute (att o Token.content_of);
(* terms and types *)
@@ -216,12 +209,12 @@
(* improper method arguments *)
val from_to =
- P.nat -- ($$$ "-" |-- P.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) ||
- P.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) ||
- P.nat >> (fn i => fn tac => tac i) ||
+ Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) ||
+ Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) ||
+ Parse.nat >> (fn i => fn tac => tac i) ||
$$$ "!" >> K ALLGOALS;
-val goal = $$$ "[" |-- P.!!! (from_to --| $$$ "]");
+val goal = $$$ "[" |-- Parse.!!! (from_to --| $$$ "]");
fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x;
@@ -229,10 +222,10 @@
fun parse_args is_symid =
let
- val keyword_symid = token (P.keyword_with is_symid);
- fun atom blk = P.group "argument"
+ val keyword_symid = token (Parse.keyword_with is_symid);
+ fun atom blk = Parse.group "argument"
(ident || keyword_symid || string || alt_string ||
- (if blk then token (P.$$$ ",") else Scan.fail));
+ (if blk then token (Parse.$$$ ",") else Scan.fail));
fun args blk x = Scan.optional (args1 blk) [] x
and args1 blk x =
@@ -240,10 +233,11 @@
(Scan.repeat1 (atom blk) ||
argsp "(" ")" ||
argsp "[" "]")) >> flat) x
- and argsp l r x = (token (P.$$$ l) ::: P.!!! (args true @@@ (token (P.$$$ r) >> single))) x;
+ and argsp l r x =
+ (token (Parse.$$$ l) ::: Parse.!!! (args true @@@ (token (Parse.$$$ r) >> single))) x;
in (args, args1) end;
-val parse = #1 (parse_args T.ident_or_symbolic) false;
+val parse = #1 (parse_args Token.ident_or_symbolic) false;
fun parse1 is_symid = #2 (parse_args is_symid) false;
@@ -252,9 +246,9 @@
fun attribs intern =
let
val attrib_name = internal_text || (symbolic || named)
- >> evaluate T.Text (intern o T.content_of);
- val attrib = P.position (attrib_name -- P.!!! parse) >> src;
- in $$$ "[" |-- P.!!! (P.list attrib --| $$$ "]") end;
+ >> evaluate Token.Text (intern o Token.content_of);
+ val attrib = Parse.position (attrib_name -- Parse.!!! parse) >> src;
+ in $$$ "[" |-- Parse.!!! (Parse.list attrib --| $$$ "]") end;
fun opt_attribs intern = Scan.optional (attribs intern) [];
@@ -273,11 +267,11 @@
(** syntax wrapper **)
fun syntax kind scan (Src ((s, args), pos)) st =
- (case Scan.error (Scan.finite' T.stopper (Scan.option scan)) (st, args) of
+ (case Scan.error (Scan.finite' Token.stopper (Scan.option scan)) (st, args) of
(SOME x, (st', [])) => (x, st')
| (_, (_, args')) =>
error (kind ^ " " ^ quote s ^ Position.str_of pos ^ ": bad arguments\n " ^
- space_implode " " (map T.unparse args')));
+ space_implode " " (map Token.unparse args')));
fun context_syntax kind scan src = apsnd Context.the_proof o syntax kind scan src o Context.Proof;
--- a/src/Pure/Isar/attrib.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/attrib.ML Sun May 23 10:38:11 2010 +0100
@@ -47,10 +47,6 @@
structure Attrib: ATTRIB =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-
-
(* source and bindings *)
type src = Args.src;
@@ -168,10 +164,10 @@
(** parsing attributed theorems **)
-val thm_sel = P.$$$ "(" |-- P.list1
- (P.nat --| P.minus -- P.nat >> Facts.FromTo ||
- P.nat --| P.minus >> Facts.From ||
- P.nat >> Facts.Single) --| P.$$$ ")";
+val thm_sel = Parse.$$$ "(" |-- Parse.list1
+ (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo ||
+ Parse.nat --| Parse.minus >> Facts.From ||
+ Parse.nat >> Facts.Single) --| Parse.$$$ ")";
local
@@ -184,7 +180,7 @@
val get_fact = get o Facts.Fact;
fun get_named pos name = get (Facts.Named ((name, pos), NONE));
in
- P.$$$ "[" |-- Args.attribs (intern thy) --| P.$$$ "]" >> (fn srcs =>
+ Parse.$$$ "[" |-- Args.attribs (intern thy) --| Parse.$$$ "]" >> (fn srcs =>
let
val atts = map (attribute_i thy) srcs;
val (context', th') = Library.apply atts (context, Drule.dummy_thm);
@@ -192,7 +188,7 @@
||
(Scan.ahead Args.alt_name -- Args.named_fact get_fact
>> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
- Scan.ahead (P.position fact_name) :|-- (fn (name, pos) =>
+ Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) =>
Args.named_fact (get_named pos) -- Scan.option thm_sel
>> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))
-- Args.opt_attribs (intern thy) >> (fn ((name, thmref, fact), srcs) =>
@@ -217,17 +213,17 @@
(* internal *)
-fun internal att = Args.src (("Pure.attribute", [T.mk_attribute att]), Position.none);
+fun internal att = Args.src (("Pure.attribute", [Token.mk_attribute att]), Position.none);
(* rule composition *)
val COMP_att =
- Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm
+ Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B)));
val THEN_att =
- Scan.lift (Scan.optional (Args.bracks P.nat) 1) -- thm
+ Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)));
val OF_att =
@@ -267,7 +263,7 @@
val eta_long =
Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion));
-val rotated = Scan.optional P.int 1 >> (fn n => Thm.rule_attribute (K (rotate_prems n)));
+val rotated = Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute (K (rotate_prems n)));
(* theory setup *)
@@ -285,9 +281,9 @@
"rename bound variables in abstractions" #>
setup (Binding.name "unfolded") unfolded "unfolded definitions" #>
setup (Binding.name "folded") folded "folded definitions" #>
- setup (Binding.name "consumes") (Scan.lift (Scan.optional P.nat 1) >> Rule_Cases.consumes)
+ setup (Binding.name "consumes") (Scan.lift (Scan.optional Parse.nat 1) >> Rule_Cases.consumes)
"number of consumed facts" #>
- setup (Binding.name "constraints") (Scan.lift P.nat >> Rule_Cases.constraints)
+ setup (Binding.name "constraints") (Scan.lift Parse.nat >> Rule_Cases.constraints)
"number of equality constraints" #>
setup (Binding.name "case_names") (Scan.lift (Scan.repeat1 Args.name) >> Rule_Cases.case_names)
"named rule cases" #>
@@ -295,7 +291,7 @@
(Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion)
"named conclusion of rule cases" #>
setup (Binding.name "params")
- (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params)
+ (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params)
"named rule parameters" #>
setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.export_without_context)))
"result put into standard form (legacy)" #>
@@ -345,13 +341,13 @@
local
-val equals = P.$$$ "=";
+val equals = Parse.$$$ "=";
fun scan_value (Config.Bool _) =
equals -- Args.$$$ "false" >> K (Config.Bool false) ||
equals -- Args.$$$ "true" >> K (Config.Bool true) ||
Scan.succeed (Config.Bool true)
- | scan_value (Config.Int _) = equals |-- P.int >> Config.Int
+ | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int
| scan_value (Config.String _) = equals |-- Args.name >> Config.String;
fun scan_config thy config =
--- a/src/Pure/Isar/calculation.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/calculation.ML Sun May 23 10:38:11 2010 +0100
@@ -14,7 +14,8 @@
val sym_del: attribute
val symmetric: attribute
val also: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
- val also_cmd: (Facts.ref * Attrib.src list) list option -> bool -> Proof.state -> Proof.state Seq.seq
+ val also_cmd: (Facts.ref * Attrib.src list) list option ->
+ bool -> Proof.state -> Proof.state Seq.seq
val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
val finally_cmd: (Facts.ref * Attrib.src list) list option -> bool ->
Proof.state -> Proof.state Seq.seq
@@ -27,7 +28,7 @@
(** calculation data **)
-structure CalculationData = Generic_Data
+structure Data = Generic_Data
(
type T = (thm Item_Net.T * thm list) * (thm list * int) option;
val empty = ((Thm.elim_rules, []), NONE);
@@ -37,7 +38,7 @@
);
fun print_rules ctxt =
- let val ((trans, sym), _) = CalculationData.get (Context.Proof ctxt) in
+ let val ((trans, sym), _) = Data.get (Context.Proof ctxt) in
[Pretty.big_list "transitivity rules:"
(map (Display.pretty_thm ctxt) (Item_Net.content trans)),
Pretty.big_list "symmetry rules:" (map (Display.pretty_thm ctxt) sym)]
@@ -48,7 +49,7 @@
(* access calculation *)
fun get_calculation state =
- (case #2 (CalculationData.get (Context.Proof (Proof.context_of state))) of
+ (case #2 (Data.get (Context.Proof (Proof.context_of state))) of
NONE => NONE
| SOME (thms, lev) => if lev = Proof.level state then SOME thms else NONE);
@@ -56,7 +57,7 @@
fun put_calculation calc =
`Proof.level #-> (fn lev => Proof.map_context (Context.proof_map
- (CalculationData.map (apsnd (K (Option.map (rpair lev) calc))))))
+ (Data.map (apsnd (K (Option.map (rpair lev) calc))))))
#> Proof.put_thms false (calculationN, calc);
@@ -65,22 +66,22 @@
(* add/del rules *)
-val trans_add = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.update);
-val trans_del = Thm.declaration_attribute (CalculationData.map o apfst o apfst o Item_Net.remove);
+val trans_add = Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.update);
+val trans_del = Thm.declaration_attribute (Data.map o apfst o apfst o Item_Net.remove);
val sym_add =
- Thm.declaration_attribute (CalculationData.map o apfst o apsnd o Thm.add_thm)
+ Thm.declaration_attribute (Data.map o apfst o apsnd o Thm.add_thm)
#> Context_Rules.elim_query NONE;
val sym_del =
- Thm.declaration_attribute (CalculationData.map o apfst o apsnd o Thm.del_thm)
+ Thm.declaration_attribute (Data.map o apfst o apsnd o Thm.del_thm)
#> Context_Rules.rule_del;
(* symmetric *)
val symmetric = Thm.rule_attribute (fn x => fn th =>
- (case Seq.chop 2 (Drule.multi_resolves [th] (#2 (#1 (CalculationData.get x)))) of
+ (case Seq.chop 2 (Drule.multi_resolves [th] (#2 (#1 (Data.get x)))) of
([th'], _) => Drule.zero_var_indexes th'
| ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
| _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
@@ -103,28 +104,31 @@
(** proof commands **)
-fun err_if b msg = if b then error msg else ();
-
fun assert_sane final =
if final then Proof.assert_forward else Proof.assert_forward_or_chain;
-fun maintain_calculation false calc = put_calculation (SOME calc)
- | maintain_calculation true calc = put_calculation NONE #> Proof.chain_facts calc;
-
-fun print_calculation false _ _ = ()
- | print_calculation true ctxt calc = Pretty.writeln
- (ProofContext.pretty_fact ctxt (ProofContext.full_name ctxt (Binding.name calculationN), calc));
+fun maintain_calculation int final calc state =
+ let
+ val state' = put_calculation (SOME calc) state;
+ val ctxt' = Proof.context_of state';
+ val _ =
+ if int then
+ Pretty.writeln
+ (ProofContext.pretty_fact ctxt'
+ (ProofContext.full_name ctxt' (Binding.name calculationN), calc))
+ else ();
+ in state' |> final ? (put_calculation NONE #> Proof.chain_facts calc) end;
(* also and finally *)
-val get_rules = #1 o CalculationData.get o Context.Proof o Proof.context_of;
+val get_rules = #1 o Data.get o Context.Proof o Proof.context_of;
fun calculate prep_rules final raw_rules int state =
let
val strip_assums_concl = Logic.strip_assums_concl o Thm.prop_of;
val eq_prop = op aconv o pairself (Envir.beta_eta_contract o strip_assums_concl);
- fun projection ths th = Library.exists (Library.curry eq_prop th) ths;
+ fun projection ths th = exists (curry eq_prop th) ths;
val opt_rules = Option.map (prep_rules state) raw_rules;
fun combine ths =
@@ -141,11 +145,12 @@
(case get_calculation state of
NONE => (true, Seq.single facts)
| SOME calc => (false, Seq.map single (combine (calc @ facts))));
+
+ val _ = initial andalso final andalso error "No calculation yet";
+ val _ = initial andalso is_some opt_rules andalso
+ error "Initial calculation -- no rules to be given";
in
- err_if (initial andalso final) "No calculation yet";
- err_if (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given";
- calculations |> Seq.map (fn calc => (print_calculation int (Proof.context_of state) calc;
- state |> maintain_calculation final calc))
+ calculations |> Seq.map (fn calc => maintain_calculation int final calc state)
end;
val also = calculate (K I) false;
@@ -164,11 +169,8 @@
NONE => (true, [])
| SOME thms => (false, thms));
val calc = thms @ facts;
- in
- err_if (initial andalso final) "No calculation yet";
- print_calculation int (Proof.context_of state) calc;
- state |> maintain_calculation final calc
- end;
+ val _ = initial andalso final andalso error "No calculation yet";
+ in maintain_calculation int final calc state end;
val moreover = collect false;
val ultimately = collect true;
--- a/src/Pure/Isar/class.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/class.ML Sun May 23 10:38:11 2010 +0100
@@ -293,7 +293,8 @@
|-> (fn (param_map, params, assm_axiom) =>
`(fn thy => calculate thy class sups base_sort param_map assm_axiom)
#-> (fn (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) =>
- Locale.add_registration (class, base_morph)
+ Locale.add_registration (class, case eq_morph of SOME eq_morph => base_morph $> eq_morph | NONE => base_morph)
+ (*FIXME should avoid modification of base morphism, but then problem with sublocale linorder < distrib_lattice*)
(Option.map (rpair true) eq_morph) export_morph
#> register class sups params base_sort base_morph export_morph axiom assm_intro of_class))
|> Theory_Target.init (SOME class)
--- a/src/Pure/Isar/context_rules.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/context_rules.ML Sun May 23 10:38:11 2010 +0100
@@ -203,7 +203,7 @@
fun add a b c x =
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) --
- Scan.option OuterParse.nat) >> (fn (f, n) => f n)) x;
+ Scan.option Parse.nat) >> (fn (f, n) => f n)) x;
val _ = Context.>> (Context.map_theory
(Attrib.setup (Binding.name "intro") (add intro_bang intro intro_query)
--- a/src/Pure/Isar/isar_cmd.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/isar_cmd.ML Sun May 23 10:38:11 2010 +0100
@@ -507,7 +507,7 @@
fun check_text (txt, pos) state =
(Position.report Markup.doc_source pos;
- ignore (ThyOutput.eval_antiquote (#1 (OuterKeyword.get_lexicons ())) state (txt, pos)));
+ ignore (ThyOutput.eval_antiquote (#1 (Keyword.get_lexicons ())) state (txt, pos)));
fun header_markup txt = Toplevel.keep (fn state =>
if Toplevel.is_toplevel state then check_text txt state
--- a/src/Pure/Isar/isar_document.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/isar_document.ML Sun May 23 10:38:11 2010 +0100
@@ -275,32 +275,29 @@
(** concrete syntax **)
-local structure P = OuterParse structure V = ValueParse in
+val _ =
+ Outer_Syntax.internal_command "Isar.define_command"
+ (Parse.string -- Parse.string >> (fn (id, text) =>
+ Toplevel.position (Position.id_only id) o
+ Toplevel.imperative (fn () =>
+ define_command id (Outer_Syntax.prepare_command (Position.id id) text))));
val _ =
- OuterSyntax.internal_command "Isar.define_command"
- (P.string -- P.string >> (fn (id, text) =>
- Toplevel.position (Position.id_only id) o
- Toplevel.imperative (fn () =>
- define_command id (OuterSyntax.prepare_command (Position.id id) text))));
-
-val _ =
- OuterSyntax.internal_command "Isar.begin_document"
- (P.string -- P.string >> (fn (id, path) =>
+ Outer_Syntax.internal_command "Isar.begin_document"
+ (Parse.string -- Parse.string >> (fn (id, path) =>
Toplevel.imperative (fn () => begin_document id (Path.explode path))));
val _ =
- OuterSyntax.internal_command "Isar.end_document"
- (P.string >> (fn id => Toplevel.imperative (fn () => end_document id)));
+ Outer_Syntax.internal_command "Isar.end_document"
+ (Parse.string >> (fn id => Toplevel.imperative (fn () => end_document id)));
val _ =
- OuterSyntax.internal_command "Isar.edit_document"
- (P.string -- P.string -- V.list (P.string -- (P.string >> SOME) || P.string >> rpair NONE)
+ Outer_Syntax.internal_command "Isar.edit_document"
+ (Parse.string -- Parse.string --
+ Parse_Value.list (Parse.string -- (Parse.string >> SOME) || Parse.string >> rpair NONE)
>> (fn ((id, new_id), edits) =>
Toplevel.position (Position.id_only new_id) o
Toplevel.imperative (fn () => edit_document id new_id edits)));
end;
-end;
-
--- a/src/Pure/Isar/isar_syn.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/isar_syn.ML Sun May 23 10:38:11 2010 +0100
@@ -7,15 +7,12 @@
structure IsarSyn: sig end =
struct
-structure P = OuterParse and K = OuterKeyword;
-
-
(** keywords **)
(*keep keywords consistent with the parsers, otherwise be prepared for
unexpected errors*)
-val _ = List.app OuterKeyword.keyword
+val _ = List.app Keyword.keyword
["!!", "!", "%", "(", ")", "+", ",", "--", ":", "::", ";", "<", "<=",
"=", "==", "=>", "?", "[", "\\<equiv>", "\\<leftharpoondown>",
"\\<rightharpoonup>", "\\<rightleftharpoons>", "\\<subseteq>", "]",
@@ -30,54 +27,54 @@
(** init and exit **)
val _ =
- OuterSyntax.command "theory" "begin theory" (K.tag_theory K.thy_begin)
+ Outer_Syntax.command "theory" "begin theory" (Keyword.tag_theory Keyword.thy_begin)
(Thy_Header.args >> (Toplevel.print oo IsarCmd.init_theory));
val _ =
- OuterSyntax.command "end" "end (local) theory" (K.tag_theory K.thy_end)
+ Outer_Syntax.command "end" "end (local) theory" (Keyword.tag_theory Keyword.thy_end)
(Scan.succeed (Toplevel.exit o Toplevel.end_local_theory));
(** markup commands **)
-val _ = OuterSyntax.markup_command ThyOutput.Markup "header" "theory header" K.diag
- (P.doc_source >> IsarCmd.header_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "header" "theory header" Keyword.diag
+ (Parse.doc_source >> IsarCmd.header_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "chapter" "chapter heading"
- K.thy_heading (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "chapter" "chapter heading"
+ Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "section" "section heading"
- K.thy_heading (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "section" "section heading"
+ Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsection" "subsection heading"
- K.thy_heading (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsection" "subsection heading"
+ Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
val _ =
- OuterSyntax.markup_command ThyOutput.Markup "subsubsection" "subsubsection heading"
- K.thy_heading (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+ Outer_Syntax.markup_command ThyOutput.Markup "subsubsection" "subsubsection heading"
+ Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
-val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "text" "formal comment (theory)"
- K.thy_decl (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.MarkupEnv "text" "formal comment (theory)"
+ Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Verbatim "text_raw" "raw document preparation text"
- K.thy_decl (P.opt_target -- P.doc_source >> IsarCmd.local_theory_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Verbatim "text_raw" "raw document preparation text"
+ Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "sect" "formal comment (proof)"
- (K.tag_proof K.prf_heading) (P.doc_source >> IsarCmd.proof_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "sect" "formal comment (proof)"
+ (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsect" "formal comment (proof)"
- (K.tag_proof K.prf_heading) (P.doc_source >> IsarCmd.proof_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsect" "formal comment (proof)"
+ (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsubsect" "formal comment (proof)"
- (K.tag_proof K.prf_heading) (P.doc_source >> IsarCmd.proof_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsubsect" "formal comment (proof)"
+ (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
-val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "txt" "formal comment (proof)"
- (K.tag_proof K.prf_decl) (P.doc_source >> IsarCmd.proof_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.MarkupEnv "txt" "formal comment (proof)"
+ (Keyword.tag_proof Keyword.prf_decl) (Parse.doc_source >> IsarCmd.proof_markup);
-val _ = OuterSyntax.markup_command ThyOutput.Verbatim "txt_raw"
- "raw document preparation text (proof)" (K.tag_proof K.prf_decl)
- (P.doc_source >> IsarCmd.proof_markup);
+val _ = Outer_Syntax.markup_command ThyOutput.Verbatim "txt_raw"
+ "raw document preparation text (proof)" (Keyword.tag_proof Keyword.prf_decl)
+ (Parse.doc_source >> IsarCmd.proof_markup);
@@ -86,207 +83,217 @@
(* classes and sorts *)
val _ =
- OuterSyntax.command "classes" "declare type classes" K.thy_decl
- (Scan.repeat1 (P.binding -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
- P.!!! (P.list1 P.xname)) []) >> (Toplevel.theory o fold AxClass.axiomatize_class_cmd));
+ Outer_Syntax.command "classes" "declare type classes" Keyword.thy_decl
+ (Scan.repeat1 (Parse.binding -- Scan.optional ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |--
+ Parse.!!! (Parse.list1 Parse.xname)) [])
+ >> (Toplevel.theory o fold AxClass.axiomatize_class_cmd));
val _ =
- OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" K.thy_decl
- (P.and_list1 (P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname))
+ Outer_Syntax.command "classrel" "state inclusion of type classes (axiomatic!)" Keyword.thy_decl
+ (Parse.and_list1 (Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<")
+ |-- Parse.!!! Parse.xname))
>> (Toplevel.theory o AxClass.axiomatize_classrel_cmd));
val _ =
- OuterSyntax.local_theory "default_sort" "declare default sort for explicit type variables"
- K.thy_decl
- (P.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));
+ Outer_Syntax.local_theory "default_sort" "declare default sort for explicit type variables"
+ Keyword.thy_decl
+ (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));
(* types *)
val _ =
- OuterSyntax.local_theory "typedecl" "type declaration" K.thy_decl
- (P.type_args -- P.binding -- P.opt_mixfix
+ Outer_Syntax.local_theory "typedecl" "type declaration" Keyword.thy_decl
+ (Parse.type_args -- Parse.binding -- Parse.opt_mixfix
>> (fn ((args, a), mx) => Typedecl.typedecl (a, map (rpair dummyS) args, mx) #> snd));
val _ =
- OuterSyntax.local_theory "types" "declare type abbreviations" K.thy_decl
+ Outer_Syntax.local_theory "types" "declare type abbreviations" Keyword.thy_decl
(Scan.repeat1
- (P.type_args -- P.binding -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_mixfix')))
+ (Parse.type_args -- Parse.binding --
+ (Parse.$$$ "=" |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')))
>> (fold (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)));
val _ =
- OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"
- K.thy_decl (Scan.repeat1 P.binding >> (Toplevel.theory o Sign.add_nonterminals));
+ Outer_Syntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"
+ Keyword.thy_decl (Scan.repeat1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals));
val _ =
- OuterSyntax.command "arities" "state type arities (axiomatic!)" K.thy_decl
- (Scan.repeat1 P.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd));
+ Outer_Syntax.command "arities" "state type arities (axiomatic!)" Keyword.thy_decl
+ (Scan.repeat1 Parse.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd));
(* consts and syntax *)
val _ =
- OuterSyntax.command "judgment" "declare object-logic judgment" K.thy_decl
- (P.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));
+ Outer_Syntax.command "judgment" "declare object-logic judgment" Keyword.thy_decl
+ (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));
val _ =
- OuterSyntax.command "consts" "declare constants" K.thy_decl
- (Scan.repeat1 P.const_binding >> (Toplevel.theory o Sign.add_consts));
+ Outer_Syntax.command "consts" "declare constants" Keyword.thy_decl
+ (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts));
-val opt_overloaded = P.opt_keyword "overloaded";
+val opt_overloaded = Parse.opt_keyword "overloaded";
val _ =
- OuterSyntax.command "finalconsts" "declare constants as final" K.thy_decl
- (opt_overloaded -- Scan.repeat1 P.term >> (uncurry (Toplevel.theory oo Theory.add_finals)));
+ Outer_Syntax.command "finalconsts" "declare constants as final" Keyword.thy_decl
+ (opt_overloaded -- Scan.repeat1 Parse.term >> (uncurry (Toplevel.theory oo Theory.add_finals)));
val mode_spec =
- (P.$$$ "output" >> K ("", false)) || P.name -- Scan.optional (P.$$$ "output" >> K false) true;
+ (Parse.$$$ "output" >> K ("", false)) ||
+ Parse.name -- Scan.optional (Parse.$$$ "output" >> K false) true;
val opt_mode =
- Scan.optional (P.$$$ "(" |-- P.!!! (mode_spec --| P.$$$ ")")) Syntax.mode_default;
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (mode_spec --| Parse.$$$ ")")) Syntax.mode_default;
val _ =
- OuterSyntax.command "syntax" "declare syntactic constants" K.thy_decl
- (opt_mode -- Scan.repeat1 P.const >> (Toplevel.theory o uncurry Sign.add_modesyntax));
+ Outer_Syntax.command "syntax" "declare syntactic constants" Keyword.thy_decl
+ (opt_mode -- Scan.repeat1 Parse.const >> (Toplevel.theory o uncurry Sign.add_modesyntax));
val _ =
- OuterSyntax.command "no_syntax" "delete syntax declarations" K.thy_decl
- (opt_mode -- Scan.repeat1 P.const >> (Toplevel.theory o uncurry Sign.del_modesyntax));
+ Outer_Syntax.command "no_syntax" "delete syntax declarations" Keyword.thy_decl
+ (opt_mode -- Scan.repeat1 Parse.const >> (Toplevel.theory o uncurry Sign.del_modesyntax));
(* translations *)
val trans_pat =
- Scan.optional (P.$$$ "(" |-- P.!!! (P.xname --| P.$$$ ")")) "logic" -- P.string;
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Parse.xname --| Parse.$$$ ")")) "logic"
+ -- Parse.string;
fun trans_arrow toks =
- ((P.$$$ "\\<rightharpoonup>" || P.$$$ "=>") >> K Syntax.ParseRule ||
- (P.$$$ "\\<leftharpoondown>" || P.$$$ "<=") >> K Syntax.PrintRule ||
- (P.$$$ "\\<rightleftharpoons>" || P.$$$ "==") >> K Syntax.ParsePrintRule) toks;
+ ((Parse.$$$ "\\<rightharpoonup>" || Parse.$$$ "=>") >> K Syntax.ParseRule ||
+ (Parse.$$$ "\\<leftharpoondown>" || Parse.$$$ "<=") >> K Syntax.PrintRule ||
+ (Parse.$$$ "\\<rightleftharpoons>" || Parse.$$$ "==") >> K Syntax.ParsePrintRule) toks;
val trans_line =
- trans_pat -- P.!!! (trans_arrow -- trans_pat)
+ trans_pat -- Parse.!!! (trans_arrow -- trans_pat)
>> (fn (left, (arr, right)) => arr (left, right));
val _ =
- OuterSyntax.command "translations" "declare syntax translation rules" K.thy_decl
+ Outer_Syntax.command "translations" "declare syntax translation rules" Keyword.thy_decl
(Scan.repeat1 trans_line >> (Toplevel.theory o Sign.add_trrules));
val _ =
- OuterSyntax.command "no_translations" "remove syntax translation rules" K.thy_decl
+ Outer_Syntax.command "no_translations" "remove syntax translation rules" Keyword.thy_decl
(Scan.repeat1 trans_line >> (Toplevel.theory o Sign.del_trrules));
(* axioms and definitions *)
val _ =
- OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" K.thy_decl
- (Scan.repeat1 SpecParse.spec >>
+ Outer_Syntax.command "axioms" "state arbitrary propositions (axiomatic!)" Keyword.thy_decl
+ (Scan.repeat1 Parse_Spec.spec >>
(Toplevel.theory o
(IsarCmd.add_axioms o
tap (fn _ => legacy_feature "Old 'axioms' command -- use 'axiomatization' instead"))));
val opt_unchecked_overloaded =
- Scan.optional (P.$$$ "(" |-- P.!!!
- (((P.$$$ "unchecked" >> K true) -- Scan.optional (P.$$$ "overloaded" >> K true) false ||
- P.$$$ "overloaded" >> K (false, true)) --| P.$$$ ")")) (false, false);
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!!
+ (((Parse.$$$ "unchecked" >> K true) -- Scan.optional (Parse.$$$ "overloaded" >> K true) false ||
+ Parse.$$$ "overloaded" >> K (false, true)) --| Parse.$$$ ")")) (false, false);
val _ =
- OuterSyntax.command "defs" "define constants" K.thy_decl
+ Outer_Syntax.command "defs" "define constants" Keyword.thy_decl
(opt_unchecked_overloaded --
- Scan.repeat1 (SpecParse.thm_name ":" -- P.prop >> (fn ((x, y), z) => ((x, z), y)))
+ Scan.repeat1 (Parse_Spec.thm_name ":" -- Parse.prop >> (fn ((x, y), z) => ((x, z), y)))
>> (Toplevel.theory o IsarCmd.add_defs));
(* old constdefs *)
val old_constdecl =
- P.binding --| P.where_ >> (fn x => (x, NONE, NoSyn)) ||
- P.binding -- (P.$$$ "::" |-- P.!!! P.typ >> SOME) -- P.opt_mixfix'
- --| Scan.option P.where_ >> P.triple1 ||
- P.binding -- (P.mixfix >> pair NONE) --| Scan.option P.where_ >> P.triple2;
+ Parse.binding --| Parse.where_ >> (fn x => (x, NONE, NoSyn)) ||
+ Parse.binding -- (Parse.$$$ "::" |-- Parse.!!! Parse.typ >> SOME) -- Parse.opt_mixfix'
+ --| Scan.option Parse.where_ >> Parse.triple1 ||
+ Parse.binding -- (Parse.mixfix >> pair NONE) --| Scan.option Parse.where_ >> Parse.triple2;
-val old_constdef = Scan.option old_constdecl -- (SpecParse.opt_thm_name ":" -- P.prop);
+val old_constdef = Scan.option old_constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop);
val structs =
- Scan.optional ((P.$$$ "(" -- P.$$$ "structure") |-- P.!!! (P.simple_fixes --| P.$$$ ")")) [];
+ Scan.optional ((Parse.$$$ "(" -- Parse.$$$ "structure")
+ |-- Parse.!!! (Parse.simple_fixes --| Parse.$$$ ")")) [];
val _ =
- OuterSyntax.command "constdefs" "old-style constant definition" K.thy_decl
+ Outer_Syntax.command "constdefs" "old-style constant definition" Keyword.thy_decl
(structs -- Scan.repeat1 old_constdef >> (Toplevel.theory o Constdefs.add_constdefs));
(* constant definitions and abbreviations *)
val _ =
- OuterSyntax.local_theory "definition" "constant definition" K.thy_decl
- (SpecParse.constdef >> (fn args => #2 o Specification.definition_cmd args));
+ Outer_Syntax.local_theory "definition" "constant definition" Keyword.thy_decl
+ (Parse_Spec.constdef >> (fn args => #2 o Specification.definition_cmd args));
val _ =
- OuterSyntax.local_theory "abbreviation" "constant abbreviation" K.thy_decl
- (opt_mode -- (Scan.option SpecParse.constdecl -- P.prop)
- >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
+ Outer_Syntax.local_theory "abbreviation" "constant abbreviation" Keyword.thy_decl
+ (opt_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop)
+ >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
val _ =
- OuterSyntax.local_theory "type_notation" "add concrete syntax for type constructors" K.thy_decl
- (opt_mode -- P.and_list1 (P.xname -- P.mixfix)
- >> (fn (mode, args) => Specification.type_notation_cmd true mode args));
+ Outer_Syntax.local_theory "type_notation" "add concrete syntax for type constructors"
+ Keyword.thy_decl
+ (opt_mode -- Parse.and_list1 (Parse.xname -- Parse.mixfix)
+ >> (fn (mode, args) => Specification.type_notation_cmd true mode args));
val _ =
- OuterSyntax.local_theory "no_type_notation" "delete concrete syntax for type constructors" K.thy_decl
- (opt_mode -- P.and_list1 (P.xname -- P.mixfix)
- >> (fn (mode, args) => Specification.type_notation_cmd false mode args));
+ Outer_Syntax.local_theory "no_type_notation" "delete concrete syntax for type constructors"
+ Keyword.thy_decl
+ (opt_mode -- Parse.and_list1 (Parse.xname -- Parse.mixfix)
+ >> (fn (mode, args) => Specification.type_notation_cmd false mode args));
val _ =
- OuterSyntax.local_theory "notation" "add concrete syntax for constants / fixed variables" K.thy_decl
- (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
- >> (fn (mode, args) => Specification.notation_cmd true mode args));
+ Outer_Syntax.local_theory "notation" "add concrete syntax for constants / fixed variables"
+ Keyword.thy_decl
+ (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
+ >> (fn (mode, args) => Specification.notation_cmd true mode args));
val _ =
- OuterSyntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables" K.thy_decl
- (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
- >> (fn (mode, args) => Specification.notation_cmd false mode args));
+ Outer_Syntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables"
+ Keyword.thy_decl
+ (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
+ >> (fn (mode, args) => Specification.notation_cmd false mode args));
(* constant specifications *)
val _ =
- OuterSyntax.command "axiomatization" "axiomatic constant specification" K.thy_decl
- (Scan.optional P.fixes [] --
- Scan.optional (P.where_ |-- P.!!! (P.and_list1 SpecParse.specs)) []
- >> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));
+ Outer_Syntax.command "axiomatization" "axiomatic constant specification" Keyword.thy_decl
+ (Scan.optional Parse.fixes [] --
+ Scan.optional (Parse.where_ |-- Parse.!!! (Parse.and_list1 Parse_Spec.specs)) []
+ >> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));
(* theorems *)
fun theorems kind =
- SpecParse.name_facts >> (fn args => #2 o Specification.theorems_cmd kind args);
+ Parse_Spec.name_facts >> (fn args => #2 o Specification.theorems_cmd kind args);
val _ =
- OuterSyntax.local_theory "theorems" "define theorems" K.thy_decl (theorems Thm.theoremK);
+ Outer_Syntax.local_theory "theorems" "define theorems" Keyword.thy_decl (theorems Thm.theoremK);
val _ =
- OuterSyntax.local_theory "lemmas" "define lemmas" K.thy_decl (theorems Thm.lemmaK);
+ Outer_Syntax.local_theory "lemmas" "define lemmas" Keyword.thy_decl (theorems Thm.lemmaK);
val _ =
- OuterSyntax.local_theory "declare" "declare theorems" K.thy_decl
- (P.and_list1 SpecParse.xthms1
+ Outer_Syntax.local_theory "declare" "declare theorems" Keyword.thy_decl
+ (Parse.and_list1 Parse_Spec.xthms1
>> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.empty_binding, flat args)]));
(* name space entry path *)
val _ =
- OuterSyntax.command "global" "disable prefixing of theory name" K.thy_decl
+ Outer_Syntax.command "global" "disable prefixing of theory name" Keyword.thy_decl
(Scan.succeed (Toplevel.theory Sign.root_path));
val _ =
- OuterSyntax.command "local" "enable prefixing of theory name" K.thy_decl
+ Outer_Syntax.command "local" "enable prefixing of theory name" Keyword.thy_decl
(Scan.succeed (Toplevel.theory Sign.local_path));
fun hide_names name hide what =
- OuterSyntax.command name ("hide " ^ what ^ " from name space") K.thy_decl
- ((P.opt_keyword "open" >> not) -- Scan.repeat1 P.xname >>
+ Outer_Syntax.command name ("hide " ^ what ^ " from name space") Keyword.thy_decl
+ ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 Parse.xname >>
(Toplevel.theory o uncurry hide));
val _ = hide_names "hide_class" IsarCmd.hide_class "classes";
@@ -305,180 +312,187 @@
(Context.proof_map (ML_Env.inherit (Context.Proof (Proof.context_of prf)))) prf;
val _ =
- OuterSyntax.command "use" "ML text from file" (K.tag_ml K.thy_decl)
- (P.path >> (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> propagate_env)));
+ Outer_Syntax.command "use" "ML text from file" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.path >>
+ (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> propagate_env)));
val _ =
- OuterSyntax.command "ML" "ML text within theory or local theory" (K.tag_ml K.thy_decl)
- (P.ML_source >> (fn (txt, pos) =>
+ Outer_Syntax.command "ML" "ML text within theory or local theory"
+ (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.ML_source >> (fn (txt, pos) =>
Toplevel.generic_theory
(ML_Context.exec (fn () => ML_Context.eval true pos txt) #> propagate_env)));
val _ =
- OuterSyntax.command "ML_prf" "ML text within proof" (K.tag_proof K.prf_decl)
- (P.ML_source >> (fn (txt, pos) =>
+ Outer_Syntax.command "ML_prf" "ML text within proof" (Keyword.tag_proof Keyword.prf_decl)
+ (Parse.ML_source >> (fn (txt, pos) =>
Toplevel.proof (Proof.map_context (Context.proof_map
(ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf)));
val _ =
- OuterSyntax.command "ML_val" "diagnostic ML text" (K.tag_ml K.diag)
- (P.ML_source >> IsarCmd.ml_diag true);
+ Outer_Syntax.command "ML_val" "diagnostic ML text" (Keyword.tag_ml Keyword.diag)
+ (Parse.ML_source >> IsarCmd.ml_diag true);
val _ =
- OuterSyntax.command "ML_command" "diagnostic ML text (silent)" (K.tag_ml K.diag)
- (P.ML_source >> (Toplevel.no_timing oo IsarCmd.ml_diag false));
+ Outer_Syntax.command "ML_command" "diagnostic ML text (silent)" (Keyword.tag_ml Keyword.diag)
+ (Parse.ML_source >> (Toplevel.no_timing oo IsarCmd.ml_diag false));
val _ =
- OuterSyntax.command "setup" "ML theory setup" (K.tag_ml K.thy_decl)
- (P.ML_source >> (Toplevel.theory o IsarCmd.global_setup));
+ Outer_Syntax.command "setup" "ML theory setup" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.ML_source >> (Toplevel.theory o IsarCmd.global_setup));
val _ =
- OuterSyntax.local_theory "local_setup" "ML local theory setup" (K.tag_ml K.thy_decl)
- (P.ML_source >> IsarCmd.local_setup);
+ Outer_Syntax.local_theory "local_setup" "ML local theory setup" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.ML_source >> IsarCmd.local_setup);
val _ =
- OuterSyntax.command "attribute_setup" "define attribute in ML" (K.tag_ml K.thy_decl)
- (P.position P.name -- P.!!! (P.$$$ "=" |-- P.ML_source -- P.text)
- >> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));
+ Outer_Syntax.command "attribute_setup" "define attribute in ML" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.position Parse.name -- Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Parse.text)
+ >> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));
val _ =
- OuterSyntax.command "method_setup" "define proof method in ML" (K.tag_ml K.thy_decl)
- (P.position P.name -- P.!!! (P.$$$ "=" |-- P.ML_source -- P.text)
- >> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));
+ Outer_Syntax.command "method_setup" "define proof method in ML" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.position Parse.name -- Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Parse.text)
+ >> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));
val _ =
- OuterSyntax.local_theory "declaration" "generic ML declaration" (K.tag_ml K.thy_decl)
- (P.opt_keyword "pervasive" -- P.ML_source >> uncurry IsarCmd.declaration);
+ Outer_Syntax.local_theory "declaration" "generic ML declaration" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.opt_keyword "pervasive" -- Parse.ML_source >> uncurry IsarCmd.declaration);
val _ =
- OuterSyntax.local_theory "simproc_setup" "define simproc in ML" (K.tag_ml K.thy_decl)
- (P.name -- (P.$$$ "(" |-- P.enum1 "|" P.term --| P.$$$ ")" --| P.$$$ "=") --
- P.ML_source -- Scan.optional (P.$$$ "identifier" |-- Scan.repeat1 P.xname) []
+ Outer_Syntax.local_theory "simproc_setup" "define simproc in ML" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.name --
+ (Parse.$$$ "(" |-- Parse.enum1 "|" Parse.term --| Parse.$$$ ")" --| Parse.$$$ "=") --
+ Parse.ML_source -- Scan.optional (Parse.$$$ "identifier" |-- Scan.repeat1 Parse.xname) []
>> (fn (((a, b), c), d) => IsarCmd.simproc_setup a b c d));
(* translation functions *)
-val trfun = P.opt_keyword "advanced" -- P.ML_source;
+val trfun = Parse.opt_keyword "advanced" -- Parse.ML_source;
val _ =
- OuterSyntax.command "parse_ast_translation" "install parse ast translation functions"
- (K.tag_ml K.thy_decl)
+ Outer_Syntax.command "parse_ast_translation" "install parse ast translation functions"
+ (Keyword.tag_ml Keyword.thy_decl)
(trfun >> (Toplevel.theory o IsarCmd.parse_ast_translation));
val _ =
- OuterSyntax.command "parse_translation" "install parse translation functions"
- (K.tag_ml K.thy_decl)
+ Outer_Syntax.command "parse_translation" "install parse translation functions"
+ (Keyword.tag_ml Keyword.thy_decl)
(trfun >> (Toplevel.theory o IsarCmd.parse_translation));
val _ =
- OuterSyntax.command "print_translation" "install print translation functions"
- (K.tag_ml K.thy_decl)
+ Outer_Syntax.command "print_translation" "install print translation functions"
+ (Keyword.tag_ml Keyword.thy_decl)
(trfun >> (Toplevel.theory o IsarCmd.print_translation));
val _ =
- OuterSyntax.command "typed_print_translation" "install typed print translation functions"
- (K.tag_ml K.thy_decl)
+ Outer_Syntax.command "typed_print_translation" "install typed print translation functions"
+ (Keyword.tag_ml Keyword.thy_decl)
(trfun >> (Toplevel.theory o IsarCmd.typed_print_translation));
val _ =
- OuterSyntax.command "print_ast_translation" "install print ast translation functions"
- (K.tag_ml K.thy_decl)
+ Outer_Syntax.command "print_ast_translation" "install print ast translation functions"
+ (Keyword.tag_ml Keyword.thy_decl)
(trfun >> (Toplevel.theory o IsarCmd.print_ast_translation));
(* oracles *)
val _ =
- OuterSyntax.command "oracle" "declare oracle" (K.tag_ml K.thy_decl)
- (P.position P.name -- (P.$$$ "=" |-- P.ML_source) >>
+ Outer_Syntax.command "oracle" "declare oracle" (Keyword.tag_ml Keyword.thy_decl)
+ (Parse.position Parse.name -- (Parse.$$$ "=" |-- Parse.ML_source) >>
(fn (x, y) => Toplevel.theory (IsarCmd.oracle x y)));
(* local theories *)
val _ =
- OuterSyntax.command "context" "enter local theory context" K.thy_decl
- (P.name --| P.begin >> (fn name =>
+ Outer_Syntax.command "context" "enter local theory context" Keyword.thy_decl
+ (Parse.name --| Parse.begin >> (fn name =>
Toplevel.print o Toplevel.begin_local_theory true (Theory_Target.context_cmd name)));
(* locales *)
val locale_val =
- SpecParse.locale_expression false --
- Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
- Scan.repeat1 SpecParse.context_element >> pair ([], []);
+ Parse_Spec.locale_expression false --
+ Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
+ Scan.repeat1 Parse_Spec.context_element >> pair ([], []);
val _ =
- OuterSyntax.command "locale" "define named proof context" K.thy_decl
- (P.binding -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (([], []), []) -- P.opt_begin
+ Outer_Syntax.command "locale" "define named proof context" Keyword.thy_decl
+ (Parse.binding --
+ Scan.optional (Parse.$$$ "=" |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
>> (fn ((name, (expr, elems)), begin) =>
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin
(Expression.add_locale_cmd name Binding.empty expr elems #> snd)));
val _ =
- OuterSyntax.command "sublocale"
- "prove sublocale relation between a locale and a locale expression" K.thy_goal
- (P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! (SpecParse.locale_expression false)
+ Outer_Syntax.command "sublocale"
+ "prove sublocale relation between a locale and a locale expression" Keyword.thy_goal
+ (Parse.xname --| (Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") --
+ Parse.!!! (Parse_Spec.locale_expression false)
>> (fn (loc, expr) =>
Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd loc expr)));
val _ =
- OuterSyntax.command "interpretation"
- "prove interpretation of locale expression in theory" K.thy_goal
- (P.!!! (SpecParse.locale_expression true) --
- Scan.optional (P.where_ |-- P.and_list1 (SpecParse.opt_thm_name ":" -- P.prop)) []
+ Outer_Syntax.command "interpretation"
+ "prove interpretation of locale expression in theory" Keyword.thy_goal
+ (Parse.!!! (Parse_Spec.locale_expression true) --
+ Scan.optional (Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []
>> (fn (expr, equations) => Toplevel.print o
Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations)));
val _ =
- OuterSyntax.command "interpret"
+ Outer_Syntax.command "interpret"
"prove interpretation of locale expression in proof context"
- (K.tag_proof K.prf_goal)
- (P.!!! (SpecParse.locale_expression true)
+ (Keyword.tag_proof Keyword.prf_goal)
+ (Parse.!!! (Parse_Spec.locale_expression true)
>> (fn expr => Toplevel.print o Toplevel.proof' (Expression.interpret_cmd expr)));
(* classes *)
val class_val =
- SpecParse.class_expr --
- Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
- Scan.repeat1 SpecParse.context_element >> pair [];
+ Parse_Spec.class_expr --
+ Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
+ Scan.repeat1 Parse_Spec.context_element >> pair [];
val _ =
- OuterSyntax.command "class" "define type class" K.thy_decl
- (P.binding -- Scan.optional (P.$$$ "=" |-- class_val) ([], []) -- P.opt_begin
+ Outer_Syntax.command "class" "define type class" Keyword.thy_decl
+ (Parse.binding -- Scan.optional (Parse.$$$ "=" |-- class_val) ([], []) -- Parse.opt_begin
>> (fn ((name, (supclasses, elems)), begin) =>
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin
(Class.class_cmd name supclasses elems #> snd)));
val _ =
- OuterSyntax.local_theory_to_proof "subclass" "prove a subclass relation" K.thy_goal
- (P.xname >> Class.subclass_cmd);
+ Outer_Syntax.local_theory_to_proof "subclass" "prove a subclass relation" Keyword.thy_goal
+ (Parse.xname >> Class.subclass_cmd);
val _ =
- OuterSyntax.command "instantiation" "instantiate and prove type arity" K.thy_decl
- (P.multi_arity --| P.begin
+ Outer_Syntax.command "instantiation" "instantiate and prove type arity" Keyword.thy_decl
+ (Parse.multi_arity --| Parse.begin
>> (fn arities => Toplevel.print o
Toplevel.begin_local_theory true (Theory_Target.instantiation_cmd arities)));
val _ =
- OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal
- ((P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> Class.classrel_cmd ||
- P.multi_arity >> Class.instance_arity_cmd)
- >> (Toplevel.print oo Toplevel.theory_to_proof)
- || Scan.succeed (Toplevel.print o Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));
+ Outer_Syntax.command "instance" "prove type arity or subclass relation" Keyword.thy_goal
+ ((Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |-- Parse.!!! Parse.xname)
+ >> Class.classrel_cmd ||
+ Parse.multi_arity >> Class.instance_arity_cmd)
+ >> (Toplevel.print oo Toplevel.theory_to_proof) ||
+ Scan.succeed
+ (Toplevel.print o Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));
(* arbitrary overloading *)
val _ =
- OuterSyntax.command "overloading" "overloaded definitions" K.thy_decl
- (Scan.repeat1 (P.name --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term --
- Scan.optional (P.$$$ "(" |-- (P.$$$ "unchecked" >> K false) --| P.$$$ ")") true
- >> P.triple1) --| P.begin
+ Outer_Syntax.command "overloading" "overloaded definitions" Keyword.thy_decl
+ (Scan.repeat1 (Parse.name --| (Parse.$$$ "\\<equiv>" || Parse.$$$ "==") -- Parse.term --
+ Scan.optional (Parse.$$$ "(" |-- (Parse.$$$ "unchecked" >> K false) --| Parse.$$$ ")") true
+ >> Parse.triple1) --| Parse.begin
>> (fn operations => Toplevel.print o
Toplevel.begin_local_theory true (Theory_Target.overloading_cmd operations)));
@@ -486,8 +500,8 @@
(* code generation *)
val _ =
- OuterSyntax.command "code_datatype" "define set of code datatype constructors" K.thy_decl
- (Scan.repeat1 P.term >> (Toplevel.theory o Code.add_datatype_cmd));
+ Outer_Syntax.command "code_datatype" "define set of code datatype constructors" Keyword.thy_decl
+ (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd));
@@ -496,13 +510,13 @@
(* statements *)
fun gen_theorem schematic kind =
- OuterSyntax.local_theory_to_proof'
+ Outer_Syntax.local_theory_to_proof'
(if schematic then "schematic_" ^ kind else kind)
("state " ^ (if schematic then "schematic " ^ kind else kind))
- (if schematic then K.thy_schematic_goal else K.thy_goal)
- (Scan.optional (SpecParse.opt_thm_name ":" --|
- Scan.ahead (SpecParse.locale_keyword || SpecParse.statement_keyword)) Attrib.empty_binding --
- SpecParse.general_statement >> (fn (a, (elems, concl)) =>
+ (if schematic then Keyword.thy_schematic_goal else Keyword.thy_goal)
+ (Scan.optional (Parse_Spec.opt_thm_name ":" --|
+ Scan.ahead (Parse_Spec.locale_keyword || Parse_Spec.statement_keyword)) Attrib.empty_binding --
+ Parse_Spec.general_statement >> (fn (a, (elems, concl)) =>
((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd)
kind NONE (K I) a elems concl)));
@@ -514,207 +528,209 @@
val _ = gen_theorem true Thm.corollaryK;
val _ =
- OuterSyntax.local_theory_to_proof "example_proof"
- "example proof body, without any result" K.thy_schematic_goal
+ Outer_Syntax.local_theory_to_proof "example_proof"
+ "example proof body, without any result" Keyword.thy_schematic_goal
(Scan.succeed
(Specification.schematic_theorem_cmd "" NONE (K I)
Attrib.empty_binding [] (Element.Shows []) false #> Proof.enter_forward));
val _ =
- OuterSyntax.command "have" "state local goal"
- (K.tag_proof K.prf_goal)
- (SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.have));
+ Outer_Syntax.command "have" "state local goal"
+ (Keyword.tag_proof Keyword.prf_goal)
+ (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.have));
val _ =
- OuterSyntax.command "hence" "abbreviates \"then have\""
- (K.tag_proof K.prf_goal)
- (SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.hence));
+ Outer_Syntax.command "hence" "abbreviates \"then have\""
+ (Keyword.tag_proof Keyword.prf_goal)
+ (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.hence));
val _ =
- OuterSyntax.command "show" "state local goal, solving current obligation"
- (K.tag_proof K.prf_asm_goal)
- (SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.show));
+ Outer_Syntax.command "show" "state local goal, solving current obligation"
+ (Keyword.tag_proof Keyword.prf_asm_goal)
+ (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.show));
val _ =
- OuterSyntax.command "thus" "abbreviates \"then show\""
- (K.tag_proof K.prf_asm_goal)
- (SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.thus));
+ Outer_Syntax.command "thus" "abbreviates \"then show\""
+ (Keyword.tag_proof Keyword.prf_asm_goal)
+ (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.thus));
(* facts *)
-val facts = P.and_list1 SpecParse.xthms1;
+val facts = Parse.and_list1 Parse_Spec.xthms1;
val _ =
- OuterSyntax.command "then" "forward chaining"
- (K.tag_proof K.prf_chain)
+ Outer_Syntax.command "then" "forward chaining"
+ (Keyword.tag_proof Keyword.prf_chain)
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.chain));
val _ =
- OuterSyntax.command "from" "forward chaining from given facts"
- (K.tag_proof K.prf_chain)
+ Outer_Syntax.command "from" "forward chaining from given facts"
+ (Keyword.tag_proof Keyword.prf_chain)
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd)));
val _ =
- OuterSyntax.command "with" "forward chaining from given and current facts"
- (K.tag_proof K.prf_chain)
+ Outer_Syntax.command "with" "forward chaining from given and current facts"
+ (Keyword.tag_proof Keyword.prf_chain)
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd)));
val _ =
- OuterSyntax.command "note" "define facts"
- (K.tag_proof K.prf_decl)
- (SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));
+ Outer_Syntax.command "note" "define facts"
+ (Keyword.tag_proof Keyword.prf_decl)
+ (Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));
val _ =
- OuterSyntax.command "using" "augment goal facts"
- (K.tag_proof K.prf_decl)
+ Outer_Syntax.command "using" "augment goal facts"
+ (Keyword.tag_proof Keyword.prf_decl)
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd)));
val _ =
- OuterSyntax.command "unfolding" "unfold definitions in goal and facts"
- (K.tag_proof K.prf_decl)
+ Outer_Syntax.command "unfolding" "unfold definitions in goal and facts"
+ (Keyword.tag_proof Keyword.prf_decl)
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd)));
(* proof context *)
val _ =
- OuterSyntax.command "fix" "fix local variables (Skolem constants)"
- (K.tag_proof K.prf_asm)
- (P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));
+ Outer_Syntax.command "fix" "fix local variables (Skolem constants)"
+ (Keyword.tag_proof Keyword.prf_asm)
+ (Parse.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));
val _ =
- OuterSyntax.command "assume" "assume propositions"
- (K.tag_proof K.prf_asm)
- (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));
+ Outer_Syntax.command "assume" "assume propositions"
+ (Keyword.tag_proof Keyword.prf_asm)
+ (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));
val _ =
- OuterSyntax.command "presume" "assume propositions, to be established later"
- (K.tag_proof K.prf_asm)
- (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));
+ Outer_Syntax.command "presume" "assume propositions, to be established later"
+ (Keyword.tag_proof Keyword.prf_asm)
+ (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));
val _ =
- OuterSyntax.command "def" "local definition"
- (K.tag_proof K.prf_asm)
- (P.and_list1
- (SpecParse.opt_thm_name ":" --
- ((P.binding -- P.opt_mixfix) -- ((P.$$$ "\\<equiv>" || P.$$$ "==") |-- P.!!! P.termp)))
+ Outer_Syntax.command "def" "local definition"
+ (Keyword.tag_proof Keyword.prf_asm)
+ (Parse.and_list1
+ (Parse_Spec.opt_thm_name ":" --
+ ((Parse.binding -- Parse.opt_mixfix) --
+ ((Parse.$$$ "\\<equiv>" || Parse.$$$ "==") |-- Parse.!!! Parse.termp)))
>> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd)));
val _ =
- OuterSyntax.command "obtain" "generalized existence"
- (K.tag_proof K.prf_asm_goal)
- (P.parname -- Scan.optional (P.fixes --| P.where_) [] -- SpecParse.statement
+ Outer_Syntax.command "obtain" "generalized existence"
+ (Keyword.tag_proof Keyword.prf_asm_goal)
+ (Parse.parname -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement
>> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain_cmd x y z)));
val _ =
- OuterSyntax.command "guess" "wild guessing (unstructured)"
- (K.tag_proof K.prf_asm_goal)
- (Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));
+ Outer_Syntax.command "guess" "wild guessing (unstructured)"
+ (Keyword.tag_proof Keyword.prf_asm_goal)
+ (Scan.optional Parse.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));
val _ =
- OuterSyntax.command "let" "bind text variables"
- (K.tag_proof K.prf_decl)
- (P.and_list1 (P.and_list1 P.term -- (P.$$$ "=" |-- P.term))
+ Outer_Syntax.command "let" "bind text variables"
+ (Keyword.tag_proof Keyword.prf_decl)
+ (Parse.and_list1 (Parse.and_list1 Parse.term -- (Parse.$$$ "=" |-- Parse.term))
>> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));
val _ =
- OuterSyntax.command "write" "add concrete syntax for constants / fixed variables"
- (K.tag_proof K.prf_decl)
- (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
+ Outer_Syntax.command "write" "add concrete syntax for constants / fixed variables"
+ (Keyword.tag_proof Keyword.prf_decl)
+ (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
>> (fn (mode, args) => Toplevel.print o Toplevel.proof (Proof.write_cmd mode args)));
val case_spec =
- (P.$$$ "(" |-- P.!!! (P.xname -- Scan.repeat1 (P.maybe P.name) --| P.$$$ ")") ||
- P.xname >> rpair []) -- SpecParse.opt_attribs >> P.triple1;
+ (Parse.$$$ "(" |--
+ Parse.!!! (Parse.xname -- Scan.repeat1 (Parse.maybe Parse.name) --| Parse.$$$ ")") ||
+ Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> Parse.triple1;
val _ =
- OuterSyntax.command "case" "invoke local context"
- (K.tag_proof K.prf_asm)
+ Outer_Syntax.command "case" "invoke local context"
+ (Keyword.tag_proof Keyword.prf_asm)
(case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd)));
(* proof structure *)
val _ =
- OuterSyntax.command "{" "begin explicit proof block"
- (K.tag_proof K.prf_open)
+ Outer_Syntax.command "{" "begin explicit proof block"
+ (Keyword.tag_proof Keyword.prf_open)
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.begin_block));
val _ =
- OuterSyntax.command "}" "end explicit proof block"
- (K.tag_proof K.prf_close)
+ Outer_Syntax.command "}" "end explicit proof block"
+ (Keyword.tag_proof Keyword.prf_close)
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.end_block));
val _ =
- OuterSyntax.command "next" "enter next proof block"
- (K.tag_proof K.prf_block)
+ Outer_Syntax.command "next" "enter next proof block"
+ (Keyword.tag_proof Keyword.prf_block)
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.next_block));
(* end proof *)
val _ =
- OuterSyntax.command "qed" "conclude (sub-)proof"
- (K.tag_proof K.qed_block)
+ Outer_Syntax.command "qed" "conclude (sub-)proof"
+ (Keyword.tag_proof Keyword.qed_block)
(Scan.option Method.parse >> IsarCmd.qed);
val _ =
- OuterSyntax.command "by" "terminal backward proof"
- (K.tag_proof K.qed)
+ Outer_Syntax.command "by" "terminal backward proof"
+ (Keyword.tag_proof Keyword.qed)
(Method.parse -- Scan.option Method.parse >> IsarCmd.terminal_proof);
val _ =
- OuterSyntax.command ".." "default proof"
- (K.tag_proof K.qed)
+ Outer_Syntax.command ".." "default proof"
+ (Keyword.tag_proof Keyword.qed)
(Scan.succeed IsarCmd.default_proof);
val _ =
- OuterSyntax.command "." "immediate proof"
- (K.tag_proof K.qed)
+ Outer_Syntax.command "." "immediate proof"
+ (Keyword.tag_proof Keyword.qed)
(Scan.succeed IsarCmd.immediate_proof);
val _ =
- OuterSyntax.command "done" "done proof"
- (K.tag_proof K.qed)
+ Outer_Syntax.command "done" "done proof"
+ (Keyword.tag_proof Keyword.qed)
(Scan.succeed IsarCmd.done_proof);
val _ =
- OuterSyntax.command "sorry" "skip proof (quick-and-dirty mode only!)"
- (K.tag_proof K.qed)
+ Outer_Syntax.command "sorry" "skip proof (quick-and-dirty mode only!)"
+ (Keyword.tag_proof Keyword.qed)
(Scan.succeed IsarCmd.skip_proof);
val _ =
- OuterSyntax.command "oops" "forget proof"
- (K.tag_proof K.qed_global)
+ Outer_Syntax.command "oops" "forget proof"
+ (Keyword.tag_proof Keyword.qed_global)
(Scan.succeed Toplevel.forget_proof);
(* proof steps *)
val _ =
- OuterSyntax.command "defer" "shuffle internal proof state"
- (K.tag_proof K.prf_script)
- (Scan.option P.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer)));
+ Outer_Syntax.command "defer" "shuffle internal proof state"
+ (Keyword.tag_proof Keyword.prf_script)
+ (Scan.option Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer)));
val _ =
- OuterSyntax.command "prefer" "shuffle internal proof state"
- (K.tag_proof K.prf_script)
- (P.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer)));
+ Outer_Syntax.command "prefer" "shuffle internal proof state"
+ (Keyword.tag_proof Keyword.prf_script)
+ (Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer)));
val _ =
- OuterSyntax.command "apply" "initial refinement step (unstructured)"
- (K.tag_proof K.prf_script)
+ Outer_Syntax.command "apply" "initial refinement step (unstructured)"
+ (Keyword.tag_proof Keyword.prf_script)
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply)));
val _ =
- OuterSyntax.command "apply_end" "terminal refinement (unstructured)"
- (K.tag_proof K.prf_script)
+ Outer_Syntax.command "apply_end" "terminal refinement (unstructured)"
+ (Keyword.tag_proof Keyword.prf_script)
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply_end)));
val _ =
- OuterSyntax.command "proof" "backward proof"
- (K.tag_proof K.prf_block)
+ Outer_Syntax.command "proof" "backward proof"
+ (Keyword.tag_proof Keyword.prf_block)
(Scan.option Method.parse >> (fn m => Toplevel.print o
Toplevel.actual_proof (Proof_Node.applys (Proof.proof m)) o
Toplevel.skip_proof (fn i => i + 1)));
@@ -723,47 +739,48 @@
(* calculational proof commands *)
val calc_args =
- Scan.option (P.$$$ "(" |-- P.!!! ((SpecParse.xthms1 --| P.$$$ ")")));
+ Scan.option (Parse.$$$ "(" |-- Parse.!!! ((Parse_Spec.xthms1 --| Parse.$$$ ")")));
val _ =
- OuterSyntax.command "also" "combine calculation and current facts"
- (K.tag_proof K.prf_decl)
+ Outer_Syntax.command "also" "combine calculation and current facts"
+ (Keyword.tag_proof Keyword.prf_decl)
(calc_args >> (Toplevel.proofs' o Calculation.also_cmd));
val _ =
- OuterSyntax.command "finally" "combine calculation and current facts, exhibit result"
- (K.tag_proof K.prf_chain)
+ Outer_Syntax.command "finally" "combine calculation and current facts, exhibit result"
+ (Keyword.tag_proof Keyword.prf_chain)
(calc_args >> (Toplevel.proofs' o Calculation.finally_cmd));
val _ =
- OuterSyntax.command "moreover" "augment calculation by current facts"
- (K.tag_proof K.prf_decl)
+ Outer_Syntax.command "moreover" "augment calculation by current facts"
+ (Keyword.tag_proof Keyword.prf_decl)
(Scan.succeed (Toplevel.proof' Calculation.moreover));
val _ =
- OuterSyntax.command "ultimately" "augment calculation by current facts, exhibit result"
- (K.tag_proof K.prf_chain)
+ Outer_Syntax.command "ultimately" "augment calculation by current facts, exhibit result"
+ (Keyword.tag_proof Keyword.prf_chain)
(Scan.succeed (Toplevel.proof' Calculation.ultimately));
(* proof navigation *)
val _ =
- OuterSyntax.command "back" "backtracking of proof command"
- (K.tag_proof K.prf_script)
+ Outer_Syntax.command "back" "backtracking of proof command"
+ (Keyword.tag_proof Keyword.prf_script)
(Scan.succeed (Toplevel.print o Toplevel.actual_proof Proof_Node.back o Toplevel.skip_proof I));
(* nested commands *)
val props_text =
- Scan.optional ValueParse.properties [] -- P.position P.string >> (fn (props, (str, pos)) =>
- (Position.of_properties (Position.default_properties pos props), str));
+ Scan.optional Parse_Value.properties [] -- Parse.position Parse.string
+ >> (fn (props, (str, pos)) =>
+ (Position.of_properties (Position.default_properties pos props), str));
val _ =
- OuterSyntax.improper_command "Isabelle.command" "nested Isabelle command" K.control
+ Outer_Syntax.improper_command "Isabelle.command" "nested Isabelle command" Keyword.control
(props_text :|-- (fn (pos, str) =>
- (case OuterSyntax.parse pos str of
+ (case Outer_Syntax.parse pos str of
[tr] => Scan.succeed (K tr)
| _ => Scan.fail_with (K "exactly one command expected"))
handle ERROR msg => Scan.fail_with (K msg)));
@@ -772,151 +789,153 @@
(** diagnostic commands (for interactive mode only) **)
-val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
-val opt_bang = Scan.optional (P.$$$ "!" >> K true) false;
+val opt_modes =
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
+
+val opt_bang = Scan.optional (Parse.$$$ "!" >> K true) false;
val _ =
- OuterSyntax.improper_command "pretty_setmargin" "change default margin for pretty printing"
- K.diag (P.nat >> (Toplevel.no_timing oo IsarCmd.pretty_setmargin));
+ Outer_Syntax.improper_command "pretty_setmargin" "change default margin for pretty printing"
+ Keyword.diag (Parse.nat >> (Toplevel.no_timing oo IsarCmd.pretty_setmargin));
val _ =
- OuterSyntax.improper_command "help" "print outer syntax commands" K.diag
- (Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));
+ Outer_Syntax.improper_command "help" "print outer syntax commands" Keyword.diag
+ (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));
val _ =
- OuterSyntax.improper_command "print_commands" "print outer syntax commands" K.diag
- (Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));
+ Outer_Syntax.improper_command "print_commands" "print outer syntax commands" Keyword.diag
+ (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));
val _ =
- OuterSyntax.improper_command "print_configs" "print configuration options" K.diag
+ Outer_Syntax.improper_command "print_configs" "print configuration options" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_configs));
val _ =
- OuterSyntax.improper_command "print_context" "print theory context name" K.diag
+ Outer_Syntax.improper_command "print_context" "print theory context name" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_context));
val _ =
- OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)" K.diag
- (opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theory));
+ Outer_Syntax.improper_command "print_theory" "print logical theory contents (verbose!)"
+ Keyword.diag (opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theory));
val _ =
- OuterSyntax.improper_command "print_syntax" "print inner syntax of context (verbose!)" K.diag
- (Scan.succeed (Toplevel.no_timing o IsarCmd.print_syntax));
+ Outer_Syntax.improper_command "print_syntax" "print inner syntax of context (verbose!)"
+ Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.print_syntax));
val _ =
- OuterSyntax.improper_command "print_abbrevs" "print constant abbreviation of context" K.diag
- (Scan.succeed (Toplevel.no_timing o IsarCmd.print_abbrevs));
+ Outer_Syntax.improper_command "print_abbrevs" "print constant abbreviation of context"
+ Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.print_abbrevs));
val _ =
- OuterSyntax.improper_command "print_theorems"
- "print theorems of local theory or proof context" K.diag
+ Outer_Syntax.improper_command "print_theorems"
+ "print theorems of local theory or proof context" Keyword.diag
(opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theorems));
val _ =
- OuterSyntax.improper_command "print_locales" "print locales of this theory" K.diag
+ Outer_Syntax.improper_command "print_locales" "print locales of this theory" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_locales));
val _ =
- OuterSyntax.improper_command "print_classes" "print classes of this theory" K.diag
+ Outer_Syntax.improper_command "print_classes" "print classes of this theory" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory
o Toplevel.keep (Class.print_classes o Toplevel.theory_of)));
val _ =
- OuterSyntax.improper_command "print_locale" "print locale of this theory" K.diag
- (opt_bang -- P.xname >> (Toplevel.no_timing oo IsarCmd.print_locale));
+ Outer_Syntax.improper_command "print_locale" "print locale of this theory" Keyword.diag
+ (opt_bang -- Parse.xname >> (Toplevel.no_timing oo IsarCmd.print_locale));
val _ =
- OuterSyntax.improper_command "print_interps"
- "print interpretations of locale for this theory" K.diag
- (P.xname >> (Toplevel.no_timing oo IsarCmd.print_registrations));
+ Outer_Syntax.improper_command "print_interps"
+ "print interpretations of locale for this theory" Keyword.diag
+ (Parse.xname >> (Toplevel.no_timing oo IsarCmd.print_registrations));
val _ =
- OuterSyntax.improper_command "print_attributes" "print attributes of this theory" K.diag
+ Outer_Syntax.improper_command "print_attributes" "print attributes of this theory" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_attributes));
val _ =
- OuterSyntax.improper_command "print_simpset" "print context of Simplifier" K.diag
+ Outer_Syntax.improper_command "print_simpset" "print context of Simplifier" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_simpset));
val _ =
- OuterSyntax.improper_command "print_rules" "print intro/elim rules" K.diag
+ Outer_Syntax.improper_command "print_rules" "print intro/elim rules" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_rules));
val _ =
- OuterSyntax.improper_command "print_trans_rules" "print transitivity rules" K.diag
+ Outer_Syntax.improper_command "print_trans_rules" "print transitivity rules" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_trans_rules));
val _ =
- OuterSyntax.improper_command "print_methods" "print methods of this theory" K.diag
+ Outer_Syntax.improper_command "print_methods" "print methods of this theory" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_methods));
val _ =
- OuterSyntax.improper_command "print_antiquotations" "print antiquotations (global)" K.diag
+ Outer_Syntax.improper_command "print_antiquotations" "print antiquotations (global)" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_antiquotations));
val _ =
- OuterSyntax.improper_command "thy_deps" "visualize theory dependencies"
- K.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.thy_deps));
+ Outer_Syntax.improper_command "thy_deps" "visualize theory dependencies"
+ Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.thy_deps));
val _ =
- OuterSyntax.improper_command "class_deps" "visualize class dependencies"
- K.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.class_deps));
+ Outer_Syntax.improper_command "class_deps" "visualize class dependencies"
+ Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.class_deps));
val _ =
- OuterSyntax.improper_command "thm_deps" "visualize theorem dependencies"
- K.diag (SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.thm_deps));
+ Outer_Syntax.improper_command "thm_deps" "visualize theorem dependencies"
+ Keyword.diag (Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.thm_deps));
val _ =
- OuterSyntax.improper_command "print_binds" "print term bindings of proof context" K.diag
+ Outer_Syntax.improper_command "print_binds" "print term bindings of proof context" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_binds));
val _ =
- OuterSyntax.improper_command "print_facts" "print facts of proof context" K.diag
+ Outer_Syntax.improper_command "print_facts" "print facts of proof context" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_facts));
val _ =
- OuterSyntax.improper_command "print_cases" "print cases of proof context" K.diag
+ Outer_Syntax.improper_command "print_cases" "print cases of proof context" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_cases));
val _ =
- OuterSyntax.improper_command "print_statement" "print theorems as long statements" K.diag
- (opt_modes -- SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_stmts));
+ Outer_Syntax.improper_command "print_statement" "print theorems as long statements" Keyword.diag
+ (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_stmts));
val _ =
- OuterSyntax.improper_command "thm" "print theorems" K.diag
- (opt_modes -- SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_thms));
+ Outer_Syntax.improper_command "thm" "print theorems" Keyword.diag
+ (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_thms));
val _ =
- OuterSyntax.improper_command "prf" "print proof terms of theorems" K.diag
- (opt_modes -- Scan.option SpecParse.xthms1
+ Outer_Syntax.improper_command "prf" "print proof terms of theorems" Keyword.diag
+ (opt_modes -- Scan.option Parse_Spec.xthms1
>> (Toplevel.no_timing oo IsarCmd.print_prfs false));
val _ =
- OuterSyntax.improper_command "full_prf" "print full proof terms of theorems" K.diag
- (opt_modes -- Scan.option SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_prfs true));
+ Outer_Syntax.improper_command "full_prf" "print full proof terms of theorems" Keyword.diag
+ (opt_modes -- Scan.option Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_prfs true));
val _ =
- OuterSyntax.improper_command "prop" "read and print proposition" K.diag
- (opt_modes -- P.term >> (Toplevel.no_timing oo IsarCmd.print_prop));
+ Outer_Syntax.improper_command "prop" "read and print proposition" Keyword.diag
+ (opt_modes -- Parse.term >> (Toplevel.no_timing oo IsarCmd.print_prop));
val _ =
- OuterSyntax.improper_command "term" "read and print term" K.diag
- (opt_modes -- P.term >> (Toplevel.no_timing oo IsarCmd.print_term));
+ Outer_Syntax.improper_command "term" "read and print term" Keyword.diag
+ (opt_modes -- Parse.term >> (Toplevel.no_timing oo IsarCmd.print_term));
val _ =
- OuterSyntax.improper_command "typ" "read and print type" K.diag
- (opt_modes -- P.typ >> (Toplevel.no_timing oo IsarCmd.print_type));
+ Outer_Syntax.improper_command "typ" "read and print type" Keyword.diag
+ (opt_modes -- Parse.typ >> (Toplevel.no_timing oo IsarCmd.print_type));
val _ =
- OuterSyntax.improper_command "print_codesetup" "print code generator setup" K.diag
+ Outer_Syntax.improper_command "print_codesetup" "print code generator setup" Keyword.diag
(Scan.succeed
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
(Code.print_codesetup o Toplevel.theory_of)));
val _ =
- OuterSyntax.improper_command "unused_thms" "find unused theorems" K.diag
- (Scan.option ((Scan.repeat1 (Scan.unless P.minus P.name) --| P.minus) --
- Scan.option (Scan.repeat1 (Scan.unless P.minus P.name))) >>
+ Outer_Syntax.improper_command "unused_thms" "find unused theorems" Keyword.diag
+ (Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --
+ Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >>
(Toplevel.no_timing oo IsarCmd.unused_thms));
@@ -924,66 +943,66 @@
(** system commands (for interactive mode only) **)
val _ =
- OuterSyntax.improper_command "cd" "change current working directory" K.diag
- (P.path >> (Toplevel.no_timing oo IsarCmd.cd));
+ Outer_Syntax.improper_command "cd" "change current working directory" Keyword.diag
+ (Parse.path >> (Toplevel.no_timing oo IsarCmd.cd));
val _ =
- OuterSyntax.improper_command "pwd" "print current working directory" K.diag
+ Outer_Syntax.improper_command "pwd" "print current working directory" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.pwd));
val _ =
- OuterSyntax.improper_command "use_thy" "use theory file" K.diag
- (P.name >> (fn name =>
+ Outer_Syntax.improper_command "use_thy" "use theory file" Keyword.diag
+ (Parse.name >> (fn name =>
Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.use_thy name)));
val _ =
- OuterSyntax.improper_command "touch_thy" "outdate theory, including descendants" K.diag
- (P.name >> (fn name =>
+ Outer_Syntax.improper_command "touch_thy" "outdate theory, including descendants" Keyword.diag
+ (Parse.name >> (fn name =>
Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.touch_thy name)));
val _ =
- OuterSyntax.improper_command "remove_thy" "remove theory from loader database" K.diag
- (P.name >> (fn name =>
+ Outer_Syntax.improper_command "remove_thy" "remove theory from loader database" Keyword.diag
+ (Parse.name >> (fn name =>
Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.remove_thy name)));
val _ =
- OuterSyntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"
- K.diag (P.name >> (fn name =>
+ Outer_Syntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"
+ Keyword.diag (Parse.name >> (fn name =>
Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.kill_thy name)));
val _ =
- OuterSyntax.improper_command "display_drafts" "display raw source files with symbols"
- K.diag (Scan.repeat1 P.path >> (Toplevel.no_timing oo IsarCmd.display_drafts));
+ Outer_Syntax.improper_command "display_drafts" "display raw source files with symbols"
+ Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo IsarCmd.display_drafts));
val _ =
- OuterSyntax.improper_command "print_drafts" "print raw source files with symbols"
- K.diag (Scan.repeat1 P.path >> (Toplevel.no_timing oo IsarCmd.print_drafts));
+ Outer_Syntax.improper_command "print_drafts" "print raw source files with symbols"
+ Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo IsarCmd.print_drafts));
val opt_limits =
- Scan.option P.nat -- Scan.option (P.$$$ "," |-- P.!!! P.nat);
+ Scan.option Parse.nat -- Scan.option (Parse.$$$ "," |-- Parse.!!! Parse.nat);
val _ =
- OuterSyntax.improper_command "pr" "print current proof state (if present)" K.diag
+ Outer_Syntax.improper_command "pr" "print current proof state (if present)" Keyword.diag
(opt_modes -- opt_limits >> (Toplevel.no_timing oo IsarCmd.pr));
val _ =
- OuterSyntax.improper_command "disable_pr" "disable printing of toplevel state" K.diag
+ Outer_Syntax.improper_command "disable_pr" "disable printing of toplevel state" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.disable_pr));
val _ =
- OuterSyntax.improper_command "enable_pr" "enable printing of toplevel state" K.diag
+ Outer_Syntax.improper_command "enable_pr" "enable printing of toplevel state" Keyword.diag
(Scan.succeed (Toplevel.no_timing o IsarCmd.enable_pr));
val _ =
- OuterSyntax.improper_command "commit" "commit current session to ML database" K.diag
- (P.opt_unit >> K (Toplevel.no_timing o Toplevel.imperative Secure.commit));
+ Outer_Syntax.improper_command "commit" "commit current session to ML database" Keyword.diag
+ (Parse.opt_unit >> K (Toplevel.no_timing o Toplevel.imperative Secure.commit));
val _ =
- OuterSyntax.improper_command "quit" "quit Isabelle" K.control
- (P.opt_unit >> (Toplevel.no_timing oo K IsarCmd.quit));
+ Outer_Syntax.improper_command "quit" "quit Isabelle" Keyword.control
+ (Parse.opt_unit >> (Toplevel.no_timing oo K IsarCmd.quit));
val _ =
- OuterSyntax.improper_command "exit" "exit Isar loop" K.control
+ Outer_Syntax.improper_command "exit" "exit Isar loop" Keyword.control
(Scan.succeed (Toplevel.no_timing o IsarCmd.exit));
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/keyword.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,213 @@
+(* Title: Pure/Isar/keyword.ML
+ Author: Makarius
+
+Isar command keyword classification and global keyword tables.
+*)
+
+signature KEYWORD =
+sig
+ type T
+ val kind_of: T -> string
+ val control: T
+ val diag: T
+ val thy_begin: T
+ val thy_switch: T
+ val thy_end: T
+ val thy_heading: T
+ val thy_decl: T
+ val thy_script: T
+ val thy_goal: T
+ val thy_schematic_goal: T
+ val qed: T
+ val qed_block: T
+ val qed_global: T
+ val prf_heading: T
+ val prf_goal: T
+ val prf_block: T
+ val prf_open: T
+ val prf_close: T
+ val prf_chain: T
+ val prf_decl: T
+ val prf_asm: T
+ val prf_asm_goal: T
+ val prf_script: T
+ val kinds: string list
+ val update_tags: string -> string list -> string list
+ val tag: string -> T -> T
+ val tags_of: T -> string list
+ val tag_theory: T -> T
+ val tag_proof: T -> T
+ val tag_ml: T -> T
+ val get_lexicons: unit -> Scan.lexicon * Scan.lexicon
+ val is_keyword: string -> bool
+ val dest_keywords: unit -> string list
+ val dest_commands: unit -> string list
+ val command_keyword: string -> T option
+ val command_tags: string -> string list
+ val keyword_status_reportN: string
+ val report: unit -> unit
+ val keyword: string -> unit
+ val command: string -> T -> unit
+ val is_diag: string -> bool
+ val is_control: string -> bool
+ val is_regular: string -> bool
+ val is_theory_begin: string -> bool
+ val is_theory: string -> bool
+ val is_proof: string -> bool
+ val is_theory_goal: string -> bool
+ val is_proof_goal: string -> bool
+ val is_schematic_goal: string -> bool
+ val is_qed: string -> bool
+ val is_qed_global: string -> bool
+end;
+
+structure Keyword: KEYWORD =
+struct
+
+(** keyword classification **)
+
+datatype T = Keyword of string * string list;
+
+fun kind s = Keyword (s, []);
+fun kind_of (Keyword (s, _)) = s;
+
+
+(* kinds *)
+
+val control = kind "control";
+val diag = kind "diag";
+val thy_begin = kind "theory-begin";
+val thy_switch = kind "theory-switch";
+val thy_end = kind "theory-end";
+val thy_heading = kind "theory-heading";
+val thy_decl = kind "theory-decl";
+val thy_script = kind "theory-script";
+val thy_goal = kind "theory-goal";
+val thy_schematic_goal = kind "theory-schematic-goal";
+val qed = kind "qed";
+val qed_block = kind "qed-block";
+val qed_global = kind "qed-global";
+val prf_heading = kind "proof-heading";
+val prf_goal = kind "proof-goal";
+val prf_block = kind "proof-block";
+val prf_open = kind "proof-open";
+val prf_close = kind "proof-close";
+val prf_chain = kind "proof-chain";
+val prf_decl = kind "proof-decl";
+val prf_asm = kind "proof-asm";
+val prf_asm_goal = kind "proof-asm-goal";
+val prf_script = kind "proof-script";
+
+val kinds =
+ [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal,
+ thy_schematic_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open,
+ prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]
+ |> map kind_of;
+
+
+(* tags *)
+
+fun update_tags t ts = t :: remove (op = : string * string -> bool) t ts;
+
+fun tag t (Keyword (s, ts)) = Keyword (s, update_tags t ts);
+fun tags_of (Keyword (_, ts)) = ts;
+
+val tag_theory = tag "theory";
+val tag_proof = tag "proof";
+val tag_ml = tag "ML";
+
+
+
+(** global keyword tables **)
+
+local
+
+val global_commands = Unsynchronized.ref (Symtab.empty: T Symtab.table);
+val global_lexicons = Unsynchronized.ref (Scan.empty_lexicon, Scan.empty_lexicon);
+
+in
+
+fun get_commands () = ! global_commands;
+fun get_lexicons () = ! global_lexicons;
+
+fun change_commands f = CRITICAL (fn () => Unsynchronized.change global_commands f);
+fun change_lexicons f = CRITICAL (fn () => Unsynchronized.change global_lexicons f);
+
+end;
+
+
+(* lookup *)
+
+fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s);
+fun dest_keywords () = sort_strings (Scan.dest_lexicon (#1 (get_lexicons ())));
+
+fun dest_commands () = sort_strings (Symtab.keys (get_commands ()));
+fun command_keyword name = Symtab.lookup (get_commands ()) name;
+fun command_tags name = these (Option.map tags_of (command_keyword name));
+
+
+(* reports *)
+
+val keyword_status_reportN = "keyword_status_report";
+
+fun report_message s =
+ (if print_mode_active keyword_status_reportN then Output.status else writeln) s;
+
+fun report_keyword name =
+ report_message (Markup.markup (Markup.keyword_decl name)
+ ("Outer syntax keyword: " ^ quote name));
+
+fun report_command (name, kind) =
+ report_message (Markup.markup (Markup.command_decl name (kind_of kind))
+ ("Outer syntax command: " ^ quote name ^ " (" ^ kind_of kind ^ ")"));
+
+fun report () =
+ let val (keywords, commands) = CRITICAL (fn () =>
+ (dest_keywords (), sort_wrt #1 (Symtab.dest (get_commands ()))))
+ in
+ List.app report_keyword keywords;
+ List.app report_command commands
+ end;
+
+
+(* augment tables *)
+
+fun keyword name =
+ (change_lexicons (apfst (Scan.extend_lexicon (Symbol.explode name)));
+ report_keyword name);
+
+fun command name kind =
+ (change_lexicons (apsnd (Scan.extend_lexicon (Symbol.explode name)));
+ change_commands (Symtab.update (name, kind));
+ report_command (name, kind));
+
+
+(* command categories *)
+
+fun command_category ks name =
+ (case command_keyword name of
+ NONE => false
+ | SOME k => member (op = o pairself kind_of) ks k);
+
+val is_diag = command_category [diag];
+val is_control = command_category [control];
+fun is_regular name = not (is_diag name orelse is_control name);
+
+val is_theory_begin = command_category [thy_begin];
+
+val is_theory = command_category
+ [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal, thy_schematic_goal];
+
+val is_proof = command_category
+ [qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
+ prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
+
+val is_theory_goal = command_category [thy_goal, thy_schematic_goal];
+val is_proof_goal = command_category [prf_goal, prf_asm_goal];
+val is_schematic_goal = command_category [thy_schematic_goal];
+val is_qed = command_category [qed, qed_block];
+val is_qed_global = command_category [qed_global];
+
+end;
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/keyword.scala Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,72 @@
+/* Title: Pure/Isar/keyword.scala
+ Author: Makarius
+
+Isar command keyword classification and keyword tables.
+*/
+
+package isabelle
+
+
+object Keyword
+{
+ /* kinds */
+
+ val MINOR = "minor"
+ val CONTROL = "control"
+ val DIAG = "diag"
+ val THY_BEGIN = "theory-begin"
+ val THY_SWITCH = "theory-switch"
+ val THY_END = "theory-end"
+ val THY_HEADING = "theory-heading"
+ val THY_DECL = "theory-decl"
+ val THY_SCRIPT = "theory-script"
+ val THY_GOAL = "theory-goal"
+ val QED = "qed"
+ val QED_BLOCK = "qed-block"
+ val QED_GLOBAL = "qed-global"
+ val PRF_HEADING = "proof-heading"
+ val PRF_GOAL = "proof-goal"
+ val PRF_BLOCK = "proof-block"
+ val PRF_OPEN = "proof-open"
+ val PRF_CLOSE = "proof-close"
+ val PRF_CHAIN = "proof-chain"
+ val PRF_DECL = "proof-decl"
+ val PRF_ASM = "proof-asm"
+ val PRF_ASM_GOAL = "proof-asm-goal"
+ val PRF_SCRIPT = "proof-script"
+
+
+ /* categories */
+
+ val minor = Set(MINOR)
+ val control = Set(CONTROL)
+ val diag = Set(DIAG)
+ val heading = Set(THY_HEADING, PRF_HEADING)
+ val theory1 = Set(THY_BEGIN, THY_SWITCH, THY_END)
+ val theory2 = Set(THY_DECL, THY_GOAL)
+ val proof1 =
+ Set(QED, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL)
+ val proof2 = Set(PRF_ASM, PRF_ASM_GOAL)
+ val improper = Set(THY_SCRIPT, PRF_SCRIPT)
+
+
+ /* reports */
+
+ object Keyword_Decl {
+ def unapply(msg: XML.Tree): Option[String] =
+ msg match {
+ case XML.Elem(Markup.KEYWORD_DECL, List((Markup.NAME, name)), _) => Some(name)
+ case _ => None
+ }
+ }
+
+ object Command_Decl {
+ def unapply(msg: XML.Tree): Option[(String, String)] =
+ msg match {
+ case XML.Elem(Markup.COMMAND_DECL, List((Markup.NAME, name), (Markup.KIND, kind)), _) =>
+ Some((name, kind))
+ case _ => None
+ }
+ }
+}
+
--- a/src/Pure/Isar/locale.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/locale.ML Sun May 23 10:38:11 2010 +0100
@@ -384,10 +384,9 @@
|-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) =>
(name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs);
-fun all_registrations thy = Registrations.get thy (* FIXME clone *)
- (* with inherited mixins *)
- |-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) =>
- (name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs);
+fun all_registrations thy = Registrations.get thy
+ |-> (fn regs => fn mixins => map (reg_morph mixins) regs);
+ (* without inherited mixins *)
fun activate_notes' activ_elem transfer thy export (name, morph) input =
let
@@ -463,21 +462,19 @@
fun add_registration (name, base_morph) mixin export thy =
let
val base = instance_of thy name base_morph;
- val mix = case mixin of NONE => Morphism.identity
- | SOME (mix, _) => mix;
in
if member (ident_eq thy) (get_idents (Context.Theory thy)) (name, base)
then thy
else
(get_idents (Context.Theory thy), thy)
(* add new registrations with inherited mixins *)
- |> roundup thy (add_reg export) export (name, base_morph $> mix) (* FIXME *)
+ |> roundup thy (add_reg export) export (name, base_morph)
|> snd
(* add mixin *)
|> (case mixin of NONE => I
- | SOME mixin => amend_registration (name, base_morph $> mix) mixin export)
+ | SOME mixin => amend_registration (name, base_morph) mixin export)
(* activate import hierarchy as far as not already active *)
- |> Context.theory_map (activate_facts' export (name, base_morph $> mix))
+ |> Context.theory_map (activate_facts' export (name, base_morph))
end;
--- a/src/Pure/Isar/method.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/method.ML Sun May 23 10:38:11 2010 +0100
@@ -382,9 +382,6 @@
(** concrete syntax **)
-structure P = OuterParse;
-
-
(* sections *)
type modifier = (Proof.context -> Proof.context) * attribute;
@@ -407,30 +404,30 @@
(* extra rule methods *)
fun xrule_meth m =
- Scan.lift (Scan.optional (Args.parens OuterParse.nat) 0) -- Attrib.thms >>
+ Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >>
(fn (n, ths) => K (m n ths));
(* outer parser *)
fun is_symid_meth s =
- s <> "|" andalso s <> "?" andalso s <> "+" andalso OuterLex.ident_or_symbolic s;
+ s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s;
local
fun meth4 x =
- (P.position (P.xname >> rpair []) >> (Source o Args.src) ||
- P.$$$ "(" |-- P.!!! (meth0 --| P.$$$ ")")) x
+ (Parse.position (Parse.xname >> rpair []) >> (Source o Args.src) ||
+ Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
and meth3 x =
- (meth4 --| P.$$$ "?" >> Try ||
- meth4 --| P.$$$ "+" >> Repeat1 ||
- meth4 -- (P.$$$ "[" |-- Scan.optional P.nat 1 --| P.$$$ "]") >> (SelectGoals o swap) ||
+ (meth4 --| Parse.$$$ "?" >> Try ||
+ meth4 --| Parse.$$$ "+" >> Repeat1 ||
+ meth4 -- (Parse.$$$ "[" |-- Scan.optional Parse.nat 1 --| Parse.$$$ "]") >> (SelectGoals o swap) ||
meth4) x
and meth2 x =
- (P.position (P.xname -- Args.parse1 is_symid_meth) >> (Source o Args.src) ||
+ (Parse.position (Parse.xname -- Args.parse1 is_symid_meth) >> (Source o Args.src) ||
meth3) x
-and meth1 x = (P.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x
-and meth0 x = (P.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x;
+and meth1 x = (Parse.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x
+and meth0 x = (Parse.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x;
in val parse = meth3 end;
@@ -463,12 +460,12 @@
setup (Binding.name "rename_tac") (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >>
(fn (quant, xs) => K (SIMPLE_METHOD'' quant (Tactic.rename_tac xs))))
"rename parameters of goal" #>
- setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional P.int 1) >>
+ setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
(fn (quant, i) => K (SIMPLE_METHOD'' quant (Tactic.rotate_tac i))))
"rotate assumptions of goal" #>
- setup (Binding.name "tactic") (Scan.lift (P.position Args.name) >> tactic)
+ setup (Binding.name "tactic") (Scan.lift (Parse.position Args.name) >> tactic)
"ML tactic as proof method" #>
- setup (Binding.name "raw_tactic") (Scan.lift (P.position Args.name) >> raw_tactic)
+ setup (Binding.name "raw_tactic") (Scan.lift (Parse.position Args.name) >> raw_tactic)
"ML tactic as raw proof method"));
--- a/src/Pure/Isar/outer_keyword.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,211 +0,0 @@
-(* Title: Pure/Isar/outer_keyword.ML
- Author: Makarius
-
-Isar command keyword classification and global keyword tables.
-*)
-
-signature OUTER_KEYWORD =
-sig
- type T
- val kind_of: T -> string
- val control: T
- val diag: T
- val thy_begin: T
- val thy_switch: T
- val thy_end: T
- val thy_heading: T
- val thy_decl: T
- val thy_script: T
- val thy_goal: T
- val thy_schematic_goal: T
- val qed: T
- val qed_block: T
- val qed_global: T
- val prf_heading: T
- val prf_goal: T
- val prf_block: T
- val prf_open: T
- val prf_close: T
- val prf_chain: T
- val prf_decl: T
- val prf_asm: T
- val prf_asm_goal: T
- val prf_script: T
- val kinds: string list
- val update_tags: string -> string list -> string list
- val tag: string -> T -> T
- val tags_of: T -> string list
- val tag_theory: T -> T
- val tag_proof: T -> T
- val tag_ml: T -> T
- val get_lexicons: unit -> Scan.lexicon * Scan.lexicon
- val is_keyword: string -> bool
- val dest_keywords: unit -> string list
- val dest_commands: unit -> string list
- val command_keyword: string -> T option
- val command_tags: string -> string list
- val keyword_status_reportN: string
- val report: unit -> unit
- val keyword: string -> unit
- val command: string -> T -> unit
- val is_diag: string -> bool
- val is_control: string -> bool
- val is_regular: string -> bool
- val is_theory_begin: string -> bool
- val is_theory: string -> bool
- val is_proof: string -> bool
- val is_theory_goal: string -> bool
- val is_proof_goal: string -> bool
- val is_schematic_goal: string -> bool
- val is_qed: string -> bool
- val is_qed_global: string -> bool
-end;
-
-structure OuterKeyword: OUTER_KEYWORD =
-struct
-
-(** keyword classification **)
-
-datatype T = Keyword of string * string list;
-
-fun kind s = Keyword (s, []);
-fun kind_of (Keyword (s, _)) = s;
-
-
-(* kinds *)
-
-val control = kind "control";
-val diag = kind "diag";
-val thy_begin = kind "theory-begin";
-val thy_switch = kind "theory-switch";
-val thy_end = kind "theory-end";
-val thy_heading = kind "theory-heading";
-val thy_decl = kind "theory-decl";
-val thy_script = kind "theory-script";
-val thy_goal = kind "theory-goal";
-val thy_schematic_goal = kind "theory-schematic-goal";
-val qed = kind "qed";
-val qed_block = kind "qed-block";
-val qed_global = kind "qed-global";
-val prf_heading = kind "proof-heading";
-val prf_goal = kind "proof-goal";
-val prf_block = kind "proof-block";
-val prf_open = kind "proof-open";
-val prf_close = kind "proof-close";
-val prf_chain = kind "proof-chain";
-val prf_decl = kind "proof-decl";
-val prf_asm = kind "proof-asm";
-val prf_asm_goal = kind "proof-asm-goal";
-val prf_script = kind "proof-script";
-
-val kinds =
- [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal,
- thy_schematic_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open,
- prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]
- |> map kind_of;
-
-
-(* tags *)
-
-fun update_tags t ts = t :: remove (op = : string * string -> bool) t ts;
-
-fun tag t (Keyword (s, ts)) = Keyword (s, update_tags t ts);
-fun tags_of (Keyword (_, ts)) = ts;
-
-val tag_theory = tag "theory";
-val tag_proof = tag "proof";
-val tag_ml = tag "ML";
-
-
-
-(** global keyword tables **)
-
-local
-
-val global_commands = Unsynchronized.ref (Symtab.empty: T Symtab.table);
-val global_lexicons = Unsynchronized.ref (Scan.empty_lexicon, Scan.empty_lexicon);
-
-in
-
-fun get_commands () = ! global_commands;
-fun get_lexicons () = ! global_lexicons;
-
-fun change_commands f = CRITICAL (fn () => Unsynchronized.change global_commands f);
-fun change_lexicons f = CRITICAL (fn () => Unsynchronized.change global_lexicons f);
-
-end;
-
-
-(* lookup *)
-
-fun is_keyword s = Scan.is_literal (#1 (get_lexicons ())) (Symbol.explode s);
-fun dest_keywords () = sort_strings (Scan.dest_lexicon (#1 (get_lexicons ())));
-
-fun dest_commands () = sort_strings (Symtab.keys (get_commands ()));
-fun command_keyword name = Symtab.lookup (get_commands ()) name;
-fun command_tags name = these (Option.map tags_of (command_keyword name));
-
-
-(* reports *)
-
-val keyword_status_reportN = "keyword_status_report";
-
-fun report_message s =
- (if print_mode_active keyword_status_reportN then Output.status else writeln) s;
-
-fun report_keyword name =
- report_message (Markup.markup (Markup.keyword_decl name)
- ("Outer syntax keyword: " ^ quote name));
-
-fun report_command (name, kind) =
- report_message (Markup.markup (Markup.command_decl name (kind_of kind))
- ("Outer syntax command: " ^ quote name ^ " (" ^ kind_of kind ^ ")"));
-
-fun report () =
- let val (keywords, commands) = CRITICAL (fn () =>
- (dest_keywords (), sort_wrt #1 (Symtab.dest (get_commands ()))))
- in
- List.app report_keyword keywords;
- List.app report_command commands
- end;
-
-
-(* augment tables *)
-
-fun keyword name =
- (change_lexicons (apfst (Scan.extend_lexicon (Symbol.explode name)));
- report_keyword name);
-
-fun command name kind =
- (change_lexicons (apsnd (Scan.extend_lexicon (Symbol.explode name)));
- change_commands (Symtab.update (name, kind));
- report_command (name, kind));
-
-
-(* command categories *)
-
-fun command_category ks name =
- (case command_keyword name of
- NONE => false
- | SOME k => member (op = o pairself kind_of) ks k);
-
-val is_diag = command_category [diag];
-val is_control = command_category [control];
-fun is_regular name = not (is_diag name orelse is_control name);
-
-val is_theory_begin = command_category [thy_begin];
-
-val is_theory = command_category
- [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal, thy_schematic_goal];
-
-val is_proof = command_category
- [qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
- prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
-
-val is_theory_goal = command_category [thy_goal, thy_schematic_goal];
-val is_proof_goal = command_category [prf_goal, prf_asm_goal];
-val is_schematic_goal = command_category [thy_schematic_goal];
-val is_qed = command_category [qed, qed_block];
-val is_qed_global = command_category [qed_global];
-
-end;
--- a/src/Pure/Isar/outer_keyword.scala Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,72 +0,0 @@
-/* Title: Pure/Isar/outer_keyword.scala
- Author: Makarius
-
-Isar command keyword classification and keyword tables.
-*/
-
-package isabelle
-
-
-object Outer_Keyword
-{
- /* kinds */
-
- val MINOR = "minor"
- val CONTROL = "control"
- val DIAG = "diag"
- val THY_BEGIN = "theory-begin"
- val THY_SWITCH = "theory-switch"
- val THY_END = "theory-end"
- val THY_HEADING = "theory-heading"
- val THY_DECL = "theory-decl"
- val THY_SCRIPT = "theory-script"
- val THY_GOAL = "theory-goal"
- val QED = "qed"
- val QED_BLOCK = "qed-block"
- val QED_GLOBAL = "qed-global"
- val PRF_HEADING = "proof-heading"
- val PRF_GOAL = "proof-goal"
- val PRF_BLOCK = "proof-block"
- val PRF_OPEN = "proof-open"
- val PRF_CLOSE = "proof-close"
- val PRF_CHAIN = "proof-chain"
- val PRF_DECL = "proof-decl"
- val PRF_ASM = "proof-asm"
- val PRF_ASM_GOAL = "proof-asm-goal"
- val PRF_SCRIPT = "proof-script"
-
-
- /* categories */
-
- val minor = Set(MINOR)
- val control = Set(CONTROL)
- val diag = Set(DIAG)
- val heading = Set(THY_HEADING, PRF_HEADING)
- val theory1 = Set(THY_BEGIN, THY_SWITCH, THY_END)
- val theory2 = Set(THY_DECL, THY_GOAL)
- val proof1 =
- Set(QED, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL)
- val proof2 = Set(PRF_ASM, PRF_ASM_GOAL)
- val improper = Set(THY_SCRIPT, PRF_SCRIPT)
-
-
- /* reports */
-
- object Keyword_Decl {
- def unapply(msg: XML.Tree): Option[String] =
- msg match {
- case XML.Elem(Markup.KEYWORD_DECL, List((Markup.NAME, name)), _) => Some(name)
- case _ => None
- }
- }
-
- object Command_Decl {
- def unapply(msg: XML.Tree): Option[(String, String)] =
- msg match {
- case XML.Elem(Markup.COMMAND_DECL, List((Markup.NAME, name), (Markup.KIND, kind)), _) =>
- Some((name, kind))
- case _ => None
- }
- }
-}
-
--- a/src/Pure/Isar/outer_lex.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,390 +0,0 @@
-(* Title: Pure/Isar/outer_lex.ML
- Author: Markus Wenzel, TU Muenchen
-
-Outer lexical syntax for Isabelle/Isar.
-*)
-
-signature OUTER_LEX =
-sig
- datatype token_kind =
- Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
- Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
- Malformed | Error of string | Sync | EOF
- datatype value =
- Text of string | Typ of typ | Term of term | Fact of thm list |
- Attribute of morphism -> attribute
- type token
- val str_of_kind: token_kind -> string
- val position_of: token -> Position.T
- val end_position_of: token -> Position.T
- val pos_of: token -> string
- val eof: token
- val is_eof: token -> bool
- val not_eof: token -> bool
- val not_sync: token -> bool
- val stopper: token Scan.stopper
- val kind_of: token -> token_kind
- val is_kind: token_kind -> token -> bool
- val keyword_with: (string -> bool) -> token -> bool
- val ident_with: (string -> bool) -> token -> bool
- val is_proper: token -> bool
- val is_semicolon: token -> bool
- val is_comment: token -> bool
- val is_begin_ignore: token -> bool
- val is_end_ignore: token -> bool
- val is_blank: token -> bool
- val is_newline: token -> bool
- val source_of: token -> string
- val source_position_of: token -> Symbol_Pos.text * Position.T
- val content_of: token -> string
- val unparse: token -> string
- val text_of: token -> string * string
- val get_value: token -> value option
- val map_value: (value -> value) -> token -> token
- val mk_text: string -> token
- val mk_typ: typ -> token
- val mk_term: term -> token
- val mk_fact: thm list -> token
- val mk_attribute: (morphism -> attribute) -> token
- val assignable: token -> token
- val assign: value option -> token -> unit
- val closure: token -> token
- val ident_or_symbolic: string -> bool
- val !!! : string -> (Symbol_Pos.T list -> 'a) -> Symbol_Pos.T list -> 'a
- val source_proper: (token, 'a) Source.source -> (token, (token, 'a) Source.source) Source.source
- val source': {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
- (Symbol_Pos.T, 'a) Source.source -> (token, (Symbol_Pos.T, 'a) Source.source) Source.source
- val source: {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
- Position.T -> (Symbol.symbol, 'a) Source.source -> (token,
- (Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source
- val read_antiq: Scan.lexicon -> (token list -> 'a * token list) ->
- Symbol_Pos.T list * Position.T -> 'a
-end;
-
-structure OuterLex: OUTER_LEX =
-struct
-
-(** tokens **)
-
-(* token values *)
-
-(*The value slot assigns an (optional) internal value to a token,
- usually as a side-effect of special scanner setup (see also
- args.ML). Note that an assignable ref designates an intermediate
- state of internalization -- it is NOT meant to persist.*)
-
-datatype value =
- Text of string |
- Typ of typ |
- Term of term |
- Fact of thm list |
- Attribute of morphism -> attribute;
-
-datatype slot =
- Slot |
- Value of value option |
- Assignable of value option Unsynchronized.ref;
-
-
-(* datatype token *)
-
-datatype token_kind =
- Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
- Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
- Malformed | Error of string | Sync | EOF;
-
-datatype token = Token of (Symbol_Pos.text * Position.range) * (token_kind * string) * slot;
-
-val str_of_kind =
- fn Command => "command"
- | Keyword => "keyword"
- | Ident => "identifier"
- | LongIdent => "long identifier"
- | SymIdent => "symbolic identifier"
- | Var => "schematic variable"
- | TypeIdent => "type variable"
- | TypeVar => "schematic type variable"
- | Nat => "number"
- | String => "string"
- | AltString => "back-quoted string"
- | Verbatim => "verbatim text"
- | Space => "white space"
- | Comment => "comment text"
- | InternalValue => "internal value"
- | Malformed => "malformed symbolic character"
- | Error _ => "bad input"
- | Sync => "sync marker"
- | EOF => "end-of-file";
-
-
-(* position *)
-
-fun position_of (Token ((_, (pos, _)), _, _)) = pos;
-fun end_position_of (Token ((_, (_, pos)), _, _)) = pos;
-
-val pos_of = Position.str_of o position_of;
-
-
-(* control tokens *)
-
-fun mk_eof pos = Token (("", (pos, Position.none)), (EOF, ""), Slot);
-val eof = mk_eof Position.none;
-
-fun is_eof (Token (_, (EOF, _), _)) = true
- | is_eof _ = false;
-
-val not_eof = not o is_eof;
-
-fun not_sync (Token (_, (Sync, _), _)) = false
- | not_sync _ = true;
-
-val stopper =
- Scan.stopper (fn [] => eof | toks => mk_eof (end_position_of (List.last toks))) is_eof;
-
-
-(* kind of token *)
-
-fun kind_of (Token (_, (k, _), _)) = k;
-fun is_kind k (Token (_, (k', _), _)) = k = k';
-
-fun keyword_with pred (Token (_, (Keyword, x), _)) = pred x
- | keyword_with _ _ = false;
-
-fun ident_with pred (Token (_, (Ident, x), _)) = pred x
- | ident_with _ _ = false;
-
-fun is_proper (Token (_, (Space, _), _)) = false
- | is_proper (Token (_, (Comment, _), _)) = false
- | is_proper _ = true;
-
-fun is_semicolon (Token (_, (Keyword, ";"), _)) = true
- | is_semicolon _ = false;
-
-fun is_comment (Token (_, (Comment, _), _)) = true
- | is_comment _ = false;
-
-fun is_begin_ignore (Token (_, (Comment, "<"), _)) = true
- | is_begin_ignore _ = false;
-
-fun is_end_ignore (Token (_, (Comment, ">"), _)) = true
- | is_end_ignore _ = false;
-
-
-(* blanks and newlines -- space tokens obey lines *)
-
-fun is_blank (Token (_, (Space, x), _)) = not (String.isSuffix "\n" x)
- | is_blank _ = false;
-
-fun is_newline (Token (_, (Space, x), _)) = String.isSuffix "\n" x
- | is_newline _ = false;
-
-
-(* token content *)
-
-fun source_of (Token ((source, (pos, _)), _, _)) =
- YXML.string_of (XML.Elem (Markup.tokenN, Position.properties_of pos, [XML.Text source]));
-
-fun source_position_of (Token ((source, (pos, _)), _, _)) = (source, pos);
-
-fun content_of (Token (_, (_, x), _)) = x;
-
-
-(* unparse *)
-
-fun escape q =
- implode o map (fn s => if s = q orelse s = "\\" then "\\" ^ s else s) o Symbol.explode;
-
-fun unparse (Token (_, (kind, x), _)) =
- (case kind of
- String => x |> quote o escape "\""
- | AltString => x |> enclose "`" "`" o escape "`"
- | Verbatim => x |> enclose "{*" "*}"
- | Comment => x |> enclose "(*" "*)"
- | Malformed => space_implode " " (explode x)
- | Sync => ""
- | EOF => ""
- | _ => x);
-
-fun text_of tok =
- if is_semicolon tok then ("terminator", "")
- else
- let
- val k = str_of_kind (kind_of tok);
- val s = unparse tok
- handle ERROR _ => Symbol.separate_chars (content_of tok);
- in
- if s = "" then (k, "")
- else if size s < 40 andalso not (exists_string (fn c => c = "\n") s) then (k ^ " " ^ s, "")
- else (k, s)
- end;
-
-
-
-(** associated values **)
-
-(* access values *)
-
-fun get_value (Token (_, _, Value v)) = v
- | get_value _ = NONE;
-
-fun map_value f (Token (x, y, Value (SOME v))) = Token (x, y, Value (SOME (f v)))
- | map_value _ tok = tok;
-
-
-(* make values *)
-
-fun mk_value k v = Token ((k, Position.no_range), (InternalValue, k), Value (SOME v));
-
-val mk_text = mk_value "<text>" o Text;
-val mk_typ = mk_value "<typ>" o Typ;
-val mk_term = mk_value "<term>" o Term;
-val mk_fact = mk_value "<fact>" o Fact;
-val mk_attribute = mk_value "<attribute>" o Attribute;
-
-
-(* static binding *)
-
-(*1st stage: make empty slots assignable*)
-fun assignable (Token (x, y, Slot)) = Token (x, y, Assignable (Unsynchronized.ref NONE))
- | assignable tok = tok;
-
-(*2nd stage: assign values as side-effect of scanning*)
-fun assign v (Token (_, _, Assignable r)) = r := v
- | assign _ _ = ();
-
-(*3rd stage: static closure of final values*)
-fun closure (Token (x, y, Assignable (Unsynchronized.ref v))) = Token (x, y, Value v)
- | closure tok = tok;
-
-
-
-(** scanners **)
-
-open Basic_Symbol_Pos;
-
-fun !!! msg = Symbol_Pos.!!! ("Outer lexical error: " ^ msg);
-
-
-(* scan symbolic idents *)
-
-val is_sym_char = member (op =) (explode "!#$%&*+-/<=>?@^_|~");
-
-val scan_symid =
- Scan.many1 (is_sym_char o symbol) ||
- Scan.one (Symbol.is_symbolic o symbol) >> single;
-
-fun is_symid str =
- (case try Symbol.explode str of
- SOME [s] => Symbol.is_symbolic s orelse is_sym_char s
- | SOME ss => forall is_sym_char ss
- | _ => false);
-
-fun ident_or_symbolic "begin" = false
- | ident_or_symbolic ":" = true
- | ident_or_symbolic "::" = true
- | ident_or_symbolic s = Syntax.is_identifier s orelse is_symid s;
-
-
-(* scan verbatim text *)
-
-val scan_verb =
- $$$ "*" --| Scan.ahead (~$$$ "}") ||
- Scan.one (fn (s, _) => s <> "*" andalso Symbol.is_regular s) >> single;
-
-val scan_verbatim =
- (Symbol_Pos.scan_pos --| $$$ "{" --| $$$ "*") -- !!! "missing end of verbatim text"
- (Symbol_Pos.change_prompt
- ((Scan.repeat scan_verb >> flat) -- ($$$ "*" |-- $$$ "}" |-- Symbol_Pos.scan_pos)));
-
-
-(* scan space *)
-
-fun is_space s = Symbol.is_blank s andalso s <> "\n";
-
-val scan_space =
- Scan.many1 (is_space o symbol) @@@ Scan.optional ($$$ "\n") [] ||
- Scan.many (is_space o symbol) @@@ $$$ "\n";
-
-
-(* scan comment *)
-
-val scan_comment =
- Symbol_Pos.scan_pos -- (Symbol_Pos.scan_comment_body !!! -- Symbol_Pos.scan_pos);
-
-
-(* scan malformed symbols *)
-
-val scan_malformed =
- $$$ Symbol.malformed |--
- Symbol_Pos.change_prompt (Scan.many (Symbol.is_regular o symbol))
- --| Scan.option ($$$ Symbol.end_malformed);
-
-
-
-(** token sources **)
-
-fun source_proper src = src |> Source.filter is_proper;
-
-local
-
-fun token_leq ((_, syms1), (_, syms2)) = length syms1 <= length syms2;
-
-fun token k ss =
- Token ((Symbol_Pos.implode ss, Symbol_Pos.range ss), (k, Symbol_Pos.untabify_content ss), Slot);
-
-fun token_range k (pos1, (ss, pos2)) =
- Token (Symbol_Pos.implode_range pos1 pos2 ss, (k, Symbol_Pos.untabify_content ss), Slot);
-
-fun scan (lex1, lex2) = !!! "bad input"
- (Symbol_Pos.scan_string >> token_range String ||
- Symbol_Pos.scan_alt_string >> token_range AltString ||
- scan_verbatim >> token_range Verbatim ||
- scan_comment >> token_range Comment ||
- scan_space >> token Space ||
- scan_malformed >> token Malformed ||
- Scan.one (Symbol.is_sync o symbol) >> (token Sync o single) ||
- (Scan.max token_leq
- (Scan.max token_leq
- (Scan.literal lex2 >> pair Command)
- (Scan.literal lex1 >> pair Keyword))
- (Syntax.scan_longid >> pair LongIdent ||
- Syntax.scan_id >> pair Ident ||
- Syntax.scan_var >> pair Var ||
- Syntax.scan_tid >> pair TypeIdent ||
- Syntax.scan_tvar >> pair TypeVar ||
- Syntax.scan_nat >> pair Nat ||
- scan_symid >> pair SymIdent) >> uncurry token));
-
-fun recover msg =
- Scan.many ((Symbol.is_regular andf (not o Symbol.is_blank)) o symbol)
- >> (single o token (Error msg));
-
-in
-
-fun source' {do_recover} get_lex =
- Source.source Symbol_Pos.stopper (Scan.bulk (fn xs => scan (get_lex ()) xs))
- (Option.map (rpair recover) do_recover);
-
-fun source do_recover get_lex pos src =
- Symbol_Pos.source pos src
- |> source' do_recover get_lex;
-
-end;
-
-
-(* read_antiq *)
-
-fun read_antiq lex scan (syms, pos) =
- let
- fun err msg = cat_error msg ("Malformed antiquotation" ^ Position.str_of pos ^ ":\n" ^
- "@{" ^ Symbol_Pos.content syms ^ "}");
-
- val res =
- Source.of_list syms
- |> source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
- |> source_proper
- |> Source.source stopper (Scan.error (Scan.bulk scan)) NONE
- |> Source.exhaust;
- in (case res of [x] => x | _ => err "") handle ERROR msg => err msg end;
-
-end;
--- a/src/Pure/Isar/outer_lex.scala Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,94 +0,0 @@
-/* Title: Pure/Isar/outer_lex.scala
- Author: Makarius
-
-Outer lexical syntax for Isabelle/Isar.
-*/
-
-package isabelle
-
-
-object Outer_Lex
-{
- /* tokens */
-
- object Token_Kind extends Enumeration
- {
- val COMMAND = Value("command")
- val KEYWORD = Value("keyword")
- val IDENT = Value("identifier")
- val LONG_IDENT = Value("long identifier")
- val SYM_IDENT = Value("symbolic identifier")
- val VAR = Value("schematic variable")
- val TYPE_IDENT = Value("type variable")
- val TYPE_VAR = Value("schematic type variable")
- val NAT = Value("number")
- val STRING = Value("string")
- val ALT_STRING = Value("back-quoted string")
- val VERBATIM = Value("verbatim text")
- val SPACE = Value("white space")
- val COMMENT = Value("comment text")
- val BAD_INPUT = Value("bad input")
- val UNPARSED = Value("unparsed input")
- }
-
- sealed case class Token(val kind: Token_Kind.Value, val source: String)
- {
- def is_command: Boolean = kind == Token_Kind.COMMAND
- def is_delimited: Boolean =
- kind == Token_Kind.STRING ||
- kind == Token_Kind.ALT_STRING ||
- kind == Token_Kind.VERBATIM ||
- kind == Token_Kind.COMMENT
- def is_name: Boolean =
- kind == Token_Kind.IDENT ||
- kind == Token_Kind.SYM_IDENT ||
- kind == Token_Kind.STRING ||
- kind == Token_Kind.NAT
- def is_xname: Boolean = is_name || kind == Token_Kind.LONG_IDENT
- def is_text: Boolean = is_xname || kind == Token_Kind.VERBATIM
- def is_space: Boolean = kind == Token_Kind.SPACE
- def is_comment: Boolean = kind == Token_Kind.COMMENT
- def is_ignored: Boolean = is_space || is_comment
- def is_unparsed: Boolean = kind == Token_Kind.UNPARSED
-
- def content: String =
- if (kind == Token_Kind.STRING) Scan.Lexicon.empty.quoted_content("\"", source)
- else if (kind == Token_Kind.ALT_STRING) Scan.Lexicon.empty.quoted_content("`", source)
- else if (kind == Token_Kind.VERBATIM) Scan.Lexicon.empty.verbatim_content(source)
- else if (kind == Token_Kind.COMMENT) Scan.Lexicon.empty.comment_content(source)
- else source
-
- def text: (String, String) =
- if (kind == Token_Kind.COMMAND && source == ";") ("terminator", "")
- else (kind.toString, source)
- }
-
-
- /* token reader */
-
- class Line_Position(val line: Int) extends scala.util.parsing.input.Position
- {
- def column = 0
- def lineContents = ""
- override def toString = line.toString
-
- def advance(token: Token): Line_Position =
- {
- var n = 0
- for (c <- token.content if c == '\n') n += 1
- if (n == 0) this else new Line_Position(line + n)
- }
- }
-
- abstract class Reader extends scala.util.parsing.input.Reader[Token]
-
- private class Token_Reader(tokens: List[Token], val pos: Line_Position) extends Reader
- {
- def first = tokens.head
- def rest = new Token_Reader(tokens.tail, pos.advance(first))
- def atEnd = tokens.isEmpty
- }
-
- def reader(tokens: List[Token]): Reader = new Token_Reader(tokens, new Line_Position(1))
-}
-
--- a/src/Pure/Isar/outer_parse.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,343 +0,0 @@
-(* Title: Pure/Isar/outer_parse.ML
- Author: Markus Wenzel, TU Muenchen
-
-Generic parsers for Isabelle/Isar outer syntax.
-*)
-
-signature OUTER_PARSE =
-sig
- type token = OuterLex.token
- type 'a parser = token list -> 'a * token list
- type 'a context_parser = Context.generic * token list -> 'a * (Context.generic * token list)
- val group: string -> (token list -> 'a) -> token list -> 'a
- val !!! : (token list -> 'a) -> token list -> 'a
- val !!!! : (token list -> 'a) -> token list -> 'a
- val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
- val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
- val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
- val not_eof: token parser
- val position: (token list -> 'a * 'b) -> token list -> ('a * Position.T) * 'b
- val command: string parser
- val keyword: string parser
- val short_ident: string parser
- val long_ident: string parser
- val sym_ident: string parser
- val minus: string parser
- val term_var: string parser
- val type_ident: string parser
- val type_var: string parser
- val number: string parser
- val string: string parser
- val alt_string: string parser
- val verbatim: string parser
- val sync: string parser
- val eof: string parser
- val keyword_with: (string -> bool) -> string parser
- val keyword_ident_or_symbolic: string parser
- val $$$ : string -> string parser
- val reserved: string -> string parser
- val semicolon: string parser
- val underscore: string parser
- val maybe: 'a parser -> 'a option parser
- val tag_name: string parser
- val tags: string list parser
- val opt_unit: unit parser
- val opt_keyword: string -> bool parser
- val begin: string parser
- val opt_begin: bool parser
- val nat: int parser
- val int: int parser
- val enum: string -> 'a parser -> 'a list parser
- val enum1: string -> 'a parser -> 'a list parser
- val and_list: 'a parser -> 'a list parser
- val and_list1: 'a parser -> 'a list parser
- val enum': string -> 'a context_parser -> 'a list context_parser
- val enum1': string -> 'a context_parser -> 'a list context_parser
- val and_list': 'a context_parser -> 'a list context_parser
- val and_list1': 'a context_parser -> 'a list context_parser
- val list: 'a parser -> 'a list parser
- val list1: 'a parser -> 'a list parser
- val name: bstring parser
- val binding: binding parser
- val xname: xstring parser
- val text: string parser
- val path: Path.T parser
- val parname: string parser
- val parbinding: binding parser
- val sort: string parser
- val arity: (string * string list * string) parser
- val multi_arity: (string list * string list * string) parser
- val type_args: string list parser
- val type_args_constrained: (string * string option) list parser
- val typ_group: string parser
- val typ: string parser
- val mixfix: mixfix parser
- val mixfix': mixfix parser
- val opt_mixfix: mixfix parser
- val opt_mixfix': mixfix parser
- val where_: string parser
- val const: (string * string * mixfix) parser
- val const_binding: (binding * string * mixfix) parser
- val params: (binding * string option) list parser
- val simple_fixes: (binding * string option) list parser
- val fixes: (binding * string option * mixfix) list parser
- val for_fixes: (binding * string option * mixfix) list parser
- val for_simple_fixes: (binding * string option) list parser
- val ML_source: (Symbol_Pos.text * Position.T) parser
- val doc_source: (Symbol_Pos.text * Position.T) parser
- val term_group: string parser
- val prop_group: string parser
- val term: string parser
- val prop: string parser
- val propp: (string * string list) parser
- val termp: (string * string list) parser
- val target: xstring parser
- val opt_target: xstring option parser
-end;
-
-structure OuterParse: OUTER_PARSE =
-struct
-
-structure T = OuterLex;
-type token = T.token;
-
-type 'a parser = token list -> 'a * token list;
-type 'a context_parser = Context.generic * token list -> 'a * (Context.generic * token list);
-
-
-(** error handling **)
-
-(* group atomic parsers (no cuts!) *)
-
-fun fail_with s = Scan.fail_with
- (fn [] => s ^ " expected (past end-of-file!)"
- | (tok :: _) =>
- (case T.text_of tok of
- (txt, "") => s ^ " expected,\nbut " ^ txt ^ T.pos_of tok ^ " was found"
- | (txt1, txt2) => s ^ " expected,\nbut " ^ txt1 ^ T.pos_of tok ^ " was found:\n" ^ txt2));
-
-fun group s scan = scan || fail_with s;
-
-
-(* cut *)
-
-fun cut kind scan =
- let
- fun get_pos [] = " (past end-of-file!)"
- | get_pos (tok :: _) = T.pos_of tok;
-
- fun err (toks, NONE) = kind ^ get_pos toks
- | err (toks, SOME msg) =
- if String.isPrefix kind msg then msg
- else kind ^ get_pos toks ^ ": " ^ msg;
- in Scan.!! err scan end;
-
-fun !!! scan = cut "Outer syntax error" scan;
-fun !!!! scan = cut "Corrupted outer syntax in presentation" scan;
-
-
-
-(** basic parsers **)
-
-(* utils *)
-
-fun triple1 ((x, y), z) = (x, y, z);
-fun triple2 (x, (y, z)) = (x, y, z);
-fun triple_swap ((x, y), z) = ((x, z), y);
-
-
-(* tokens *)
-
-fun RESET_VALUE atom = (*required for all primitive parsers*)
- Scan.ahead (Scan.one (K true)) -- atom >> (fn (arg, x) => (T.assign NONE arg; x));
-
-
-val not_eof = RESET_VALUE (Scan.one T.not_eof);
-
-fun position scan = (Scan.ahead not_eof >> T.position_of) -- scan >> Library.swap;
-fun source_position atom = Scan.ahead atom |-- not_eof >> T.source_position_of;
-fun inner_syntax atom = Scan.ahead atom |-- not_eof >> T.source_of;
-
-fun kind k =
- group (T.str_of_kind k) (RESET_VALUE (Scan.one (T.is_kind k) >> T.content_of));
-
-val command = kind T.Command;
-val keyword = kind T.Keyword;
-val short_ident = kind T.Ident;
-val long_ident = kind T.LongIdent;
-val sym_ident = kind T.SymIdent;
-val term_var = kind T.Var;
-val type_ident = kind T.TypeIdent;
-val type_var = kind T.TypeVar;
-val number = kind T.Nat;
-val string = kind T.String;
-val alt_string = kind T.AltString;
-val verbatim = kind T.Verbatim;
-val sync = kind T.Sync;
-val eof = kind T.EOF;
-
-fun keyword_with pred = RESET_VALUE (Scan.one (T.keyword_with pred) >> T.content_of);
-val keyword_ident_or_symbolic = keyword_with T.ident_or_symbolic;
-
-fun $$$ x =
- group (T.str_of_kind T.Keyword ^ " " ^ quote x) (keyword_with (fn y => x = y));
-
-fun reserved x =
- group ("reserved identifier " ^ quote x)
- (RESET_VALUE (Scan.one (T.ident_with (fn y => x = y)) >> T.content_of));
-
-val semicolon = $$$ ";";
-
-val minus = sym_ident :-- (fn "-" => Scan.succeed () | _ => Scan.fail) >> #1;
-val underscore = sym_ident :-- (fn "_" => Scan.succeed () | _ => Scan.fail) >> #1;
-fun maybe scan = underscore >> K NONE || scan >> SOME;
-
-val nat = number >> (#1 o Library.read_int o Symbol.explode);
-val int = Scan.optional (minus >> K ~1) 1 -- nat >> op *;
-
-val tag_name = group "tag name" (short_ident || string);
-val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);
-
-val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) ();
-fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false;
-
-val begin = $$$ "begin";
-val opt_begin = Scan.optional (begin >> K true) false;
-
-
-(* enumerations *)
-
-fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan);
-fun enum sep scan = enum1 sep scan || Scan.succeed [];
-
-fun enum1' sep scan = scan ::: Scan.repeat (Scan.lift ($$$ sep) |-- scan);
-fun enum' sep scan = enum1' sep scan || Scan.succeed [];
-
-fun and_list1 scan = enum1 "and" scan;
-fun and_list scan = enum "and" scan;
-
-fun and_list1' scan = enum1' "and" scan;
-fun and_list' scan = enum' "and" scan;
-
-fun list1 scan = enum1 "," scan;
-fun list scan = enum "," scan;
-
-
-(* names and text *)
-
-val name = group "name declaration" (short_ident || sym_ident || string || number);
-val binding = position name >> Binding.make;
-val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number);
-val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);
-val path = group "file name/path specification" name >> Path.explode;
-
-val parname = Scan.optional ($$$ "(" |-- name --| $$$ ")") "";
-val parbinding = Scan.optional ($$$ "(" |-- binding --| $$$ ")") Binding.empty;
-
-
-(* sorts and arities *)
-
-val sort = group "sort" (inner_syntax xname);
-
-val arity = xname -- ($$$ "::" |-- !!!
- (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
-
-val multi_arity = and_list1 xname -- ($$$ "::" |-- !!!
- (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
-
-
-(* types *)
-
-val typ_group = group "type"
- (short_ident || long_ident || sym_ident || type_ident || type_var || string || number);
-
-val typ = inner_syntax typ_group;
-
-fun type_arguments arg =
- arg >> single ||
- $$$ "(" |-- !!! (list1 arg --| $$$ ")") ||
- Scan.succeed [];
-
-val type_args = type_arguments type_ident;
-val type_args_constrained = type_arguments (type_ident -- Scan.option ($$$ "::" |-- !!! sort));
-
-
-(* mixfix annotations *)
-
-val mfix = string --
- !!! (Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) [] --
- Scan.optional nat Syntax.max_pri) >> (Mixfix o triple2);
-
-val infx = $$$ "infix" |-- !!! (string -- nat >> Infix);
-val infxl = $$$ "infixl" |-- !!! (string -- nat >> Infixl);
-val infxr = $$$ "infixr" |-- !!! (string -- nat >> Infixr);
-
-val binder = $$$ "binder" |--
- !!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n))))
- >> (Binder o triple2);
-
-fun annotation guard fix = $$$ "(" |-- guard (fix --| $$$ ")");
-fun opt_annotation guard fix = Scan.optional (annotation guard fix) NoSyn;
-
-val mixfix = annotation !!! (mfix || binder || infxl || infxr || infx);
-val mixfix' = annotation I (mfix || binder || infxl || infxr || infx);
-val opt_mixfix = opt_annotation !!! (mfix || binder || infxl || infxr || infx);
-val opt_mixfix' = opt_annotation I (mfix || binder || infxl || infxr || infx);
-
-
-(* fixes *)
-
-val where_ = $$$ "where";
-
-val const = name -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;
-val const_binding = binding -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;
-
-val params = Scan.repeat1 binding -- Scan.option ($$$ "::" |-- !!! typ)
- >> (fn (xs, T) => map (rpair T) xs);
-
-val simple_fixes = and_list1 params >> flat;
-
-val fixes =
- and_list1 (binding -- Scan.option ($$$ "::" |-- typ) -- mixfix >> (single o triple1) ||
- params >> map Syntax.no_syn) >> flat;
-
-val for_fixes = Scan.optional ($$$ "for" |-- !!! fixes) [];
-val for_simple_fixes = Scan.optional ($$$ "for" |-- !!! simple_fixes) [];
-
-
-(* embedded source text *)
-
-val ML_source = source_position (group "ML source" text);
-val doc_source = source_position (group "document source" text);
-
-
-(* terms *)
-
-val trm = short_ident || long_ident || sym_ident || term_var || number || string;
-
-val term_group = group "term" trm;
-val prop_group = group "proposition" trm;
-
-val term = inner_syntax term_group;
-val prop = inner_syntax prop_group;
-
-
-(* patterns *)
-
-val is_terms = Scan.repeat1 ($$$ "is" |-- term);
-val is_props = Scan.repeat1 ($$$ "is" |-- prop);
-
-val propp = prop -- Scan.optional ($$$ "(" |-- !!! (is_props --| $$$ ")")) [];
-val termp = term -- Scan.optional ($$$ "(" |-- !!! (is_terms --| $$$ ")")) [];
-
-
-(* targets *)
-
-val target = ($$$ "(" -- $$$ "in") |-- !!! (xname --| $$$ ")");
-val opt_target = Scan.option target;
-
-end;
-
-type 'a parser = 'a OuterParse.parser;
-type 'a context_parser = 'a OuterParse.context_parser;
-
--- a/src/Pure/Isar/outer_parse.scala Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,77 +0,0 @@
-/* Title: Pure/Isar/outer_parse.scala
- Author: Makarius
-
-Generic parsers for Isabelle/Isar outer syntax.
-*/
-
-package isabelle
-
-import scala.util.parsing.combinator.Parsers
-
-
-object Outer_Parse
-{
- /* parsing tokens */
-
- trait Parser extends Parsers
- {
- type Elem = Outer_Lex.Token
-
- def filter_proper = true
-
- private def proper(in: Input): Input =
- if (in.atEnd || !in.first.is_ignored || !filter_proper) in
- else proper(in.rest)
-
- def token(s: String, pred: Elem => Boolean): Parser[Elem] = new Parser[Elem]
- {
- def apply(raw_input: Input) =
- {
- val in = proper(raw_input)
- if (in.atEnd) Failure(s + " expected (past end-of-file!)", in)
- else {
- val token = in.first
- if (pred(token)) Success(token, proper(in.rest))
- else
- token.text match {
- case (txt, "") =>
- Failure(s + " expected,\nbut " + txt + " was found", in)
- case (txt1, txt2) =>
- Failure(s + " expected,\nbut " + txt1 + " was found:\n" + txt2, in)
- }
- }
- }
- }
-
- def not_eof: Parser[Elem] = token("input token", _ => true)
- def eof: Parser[Unit] = not(not_eof)
-
- def atom(s: String, pred: Elem => Boolean): Parser[String] =
- token(s, pred) ^^ (_.content)
-
- def keyword(name: String): Parser[String] =
- atom(Outer_Lex.Token_Kind.KEYWORD.toString + " \"" + name + "\"",
- tok => tok.kind == Outer_Lex.Token_Kind.KEYWORD && tok.content == name)
-
- def name: Parser[String] = atom("name declaration", _.is_name)
- def xname: Parser[String] = atom("name reference", _.is_xname)
- def text: Parser[String] = atom("text", _.is_text)
- def ML_source: Parser[String] = atom("ML source", _.is_text)
- def doc_source: Parser[String] = atom("document source", _.is_text)
- def path: Parser[String] = atom("file name/path specification", _.is_name)
-
- private def tag_name: Parser[String] =
- atom("tag name", tok =>
- tok.kind == Outer_Lex.Token_Kind.IDENT ||
- tok.kind == Outer_Lex.Token_Kind.STRING)
-
- def tags: Parser[List[String]] = rep(keyword("%") ~> tag_name)
-
-
- /* wrappers */
-
- def parse[T](p: Parser[T], in: Outer_Lex.Reader): ParseResult[T] = p(in)
- def parse_all[T](p: Parser[T], in: Outer_Lex.Reader): ParseResult[T] = parse(phrase(p), in)
- }
-}
-
--- a/src/Pure/Isar/outer_syntax.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/outer_syntax.ML Sun May 23 10:38:11 2010 +0100
@@ -9,23 +9,23 @@
signature OUTER_SYNTAX =
sig
- val command: string -> string -> OuterKeyword.T ->
+ val command: string -> string -> Keyword.T ->
(Toplevel.transition -> Toplevel.transition) parser -> unit
- val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T ->
+ val markup_command: ThyOutput.markup -> string -> string -> Keyword.T ->
(Toplevel.transition -> Toplevel.transition) parser -> unit
- val improper_command: string -> string -> OuterKeyword.T ->
+ val improper_command: string -> string -> Keyword.T ->
(Toplevel.transition -> Toplevel.transition) parser -> unit
val internal_command: string -> (Toplevel.transition -> Toplevel.transition) parser -> unit
- val local_theory': string -> string -> OuterKeyword.T ->
+ val local_theory': string -> string -> Keyword.T ->
(bool -> local_theory -> local_theory) parser -> unit
- val local_theory: string -> string -> OuterKeyword.T ->
+ val local_theory: string -> string -> Keyword.T ->
(local_theory -> local_theory) parser -> unit
- val local_theory_to_proof': string -> string -> OuterKeyword.T ->
+ val local_theory_to_proof': string -> string -> Keyword.T ->
(bool -> local_theory -> Proof.state) parser -> unit
- val local_theory_to_proof: string -> string -> OuterKeyword.T ->
+ val local_theory_to_proof: string -> string -> Keyword.T ->
(local_theory -> Proof.state) parser -> unit
val print_outer_syntax: unit -> unit
- val scan: Position.T -> string -> OuterLex.token list
+ val scan: Position.T -> string -> Token.T list
val parse: Position.T -> string -> Toplevel.transition list
val process_file: Path.T -> theory -> theory
type isar
@@ -34,15 +34,9 @@
val load_thy: string -> Position.T -> string list -> bool -> unit -> unit
end;
-structure OuterSyntax: OUTER_SYNTAX =
+structure Outer_Syntax: OUTER_SYNTAX =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-type 'a parser = 'a P.parser;
-
-
-
(** outer syntax **)
(* command parsers *)
@@ -62,20 +56,20 @@
local
fun terminate false = Scan.succeed ()
- | terminate true = P.group "end of input" (Scan.option P.sync -- P.semicolon >> K ());
+ | terminate true = Parse.group "end of input" (Scan.option Parse.sync -- Parse.semicolon >> K ());
fun body cmd (name, _) =
(case cmd name of
SOME (Command {int_only, parse, ...}) =>
- P.!!! (Scan.prompt (name ^ "# ") (P.tags |-- parse >> pair int_only))
+ Parse.!!! (Scan.prompt (name ^ "# ") (Parse.tags |-- parse >> pair int_only))
| NONE => sys_error ("no parser for outer syntax command " ^ quote name));
in
fun parse_command do_terminate cmd =
- P.semicolon >> K NONE ||
- P.sync >> K NONE ||
- (P.position P.command :-- body cmd) --| terminate do_terminate
+ Parse.semicolon >> K NONE ||
+ Parse.sync >> K NONE ||
+ (Parse.position Parse.command :-- body cmd) --| terminate do_terminate
>> (fn ((name, pos), (int_only, f)) =>
SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
Toplevel.interactive int_only |> f));
@@ -105,7 +99,7 @@
fun get_markups () = ! global_markups;
fun get_command () = Symtab.lookup (get_commands ());
-fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_command ()));
+fun get_syntax () = CRITICAL (fn () => (Keyword.get_lexicons (), get_command ()));
fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;
@@ -113,7 +107,7 @@
(* augment syntax *)
fun add_command name kind cmd = CRITICAL (fn () =>
- (OuterKeyword.command name kind;
+ (Keyword.command name kind;
if not (Symtab.defined (get_commands ()) name) then ()
else warning ("Redefining outer syntax command " ^ quote name);
change_commands (Symtab.update (name, cmd))));
@@ -130,13 +124,13 @@
end;
fun internal_command name parse =
- command name "(internal)" OuterKeyword.control (parse >> (fn tr => Toplevel.no_timing o tr));
+ command name "(internal)" Keyword.control (parse >> (fn tr => Toplevel.no_timing o tr));
(* local_theory commands *)
fun local_theory_command do_print trans name comment kind parse =
- command name comment kind (P.opt_target -- parse
+ command name comment kind (Parse.opt_target -- parse
>> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));
val local_theory' = local_theory_command false Toplevel.local_theory';
@@ -157,7 +151,7 @@
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
val (int_cmds, cmds) = List.partition #3 (dest_commands ());
in
- [Pretty.strs ("syntax keywords:" :: map quote (OuterKeyword.dest_keywords ())),
+ [Pretty.strs ("syntax keywords:" :: map quote (Keyword.dest_keywords ())),
Pretty.big_list "commands:" (map pretty_cmd cmds),
Pretty.big_list "interactive-only commands:" (map pretty_cmd int_cmds)]
|> Pretty.chunks |> Pretty.writeln
@@ -172,18 +166,18 @@
fun toplevel_source term do_recover cmd src =
let
val no_terminator =
- Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof));
+ Scan.unless Parse.semicolon (Scan.one (Token.not_sync andf Token.not_eof));
fun recover int =
(int, fn _ => Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]);
in
src
- |> T.source_proper
- |> Source.source T.stopper
- (Scan.bulk (P.$$$ "--" -- P.!!! P.doc_source >> K NONE || P.not_eof >> SOME))
+ |> Token.source_proper
+ |> Source.source Token.stopper
+ (Scan.bulk (Parse.$$$ "--" -- Parse.!!! Parse.doc_source >> K NONE || Parse.not_eof >> SOME))
(Option.map recover do_recover)
|> Source.map_filter I
- |> Source.source T.stopper
- (Scan.bulk (fn xs => P.!!! (parse_command term (cmd ())) xs))
+ |> Source.source Token.stopper
+ (Scan.bulk (fn xs => Parse.!!! (parse_command term (cmd ())) xs))
(Option.map recover do_recover)
|> Source.map_filter I
end;
@@ -194,13 +188,13 @@
fun scan pos str =
Source.of_string str
|> Symbol.source {do_recover = false}
- |> T.source {do_recover = SOME false} OuterKeyword.get_lexicons pos
+ |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
|> Source.exhaust;
fun parse pos str =
Source.of_string str
|> Symbol.source {do_recover = false}
- |> T.source {do_recover = SOME false} OuterKeyword.get_lexicons pos
+ |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
|> toplevel_source false NONE get_command
|> Source.exhaust;
@@ -223,7 +217,7 @@
type isar =
(Toplevel.transition, (Toplevel.transition option,
- (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token,
+ (Token.T, (Token.T option, (Token.T, (Token.T,
(Symbol_Pos.T, Position.T * (Symbol.symbol, (string, unit) Source.source)
Source.source) Source.source) Source.source) Source.source)
Source.source) Source.source) Source.source) Source.source;
@@ -231,7 +225,7 @@
fun isar term : isar =
Source.tty
|> Symbol.source {do_recover = true}
- |> T.source {do_recover = SOME true} OuterKeyword.get_lexicons Position.none
+ |> Token.source {do_recover = SOME true} Keyword.get_lexicons Position.none
|> toplevel_source term (SOME true) get_command;
@@ -291,9 +285,10 @@
val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
fun after_load () =
- ThyOutput.present_thy (#1 lexs) OuterKeyword.command_tags is_markup (Lazy.force results) toks
+ ThyOutput.present_thy (#1 lexs) Keyword.command_tags is_markup (Lazy.force results) toks
|> Buffer.content
|> Present.theory_output name;
in after_load end;
end;
+
--- a/src/Pure/Isar/outer_syntax.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/outer_syntax.scala Sun May 23 10:38:11 2010 +0100
@@ -12,7 +12,7 @@
class Outer_Syntax(symbols: Symbol.Interpretation)
{
- protected val keywords: Map[String, String] = Map((";" -> Outer_Keyword.DIAG))
+ protected val keywords: Map[String, String] = Map((";" -> Keyword.DIAG))
protected val lexicon: Scan.Lexicon = Scan.Lexicon.empty
lazy val completion: Completion = new Completion + symbols // FIXME !?
@@ -28,18 +28,18 @@
}
}
- def + (name: String): Outer_Syntax = this + (name, Outer_Keyword.MINOR)
+ def + (name: String): Outer_Syntax = this + (name, Keyword.MINOR)
def is_command(name: String): Boolean =
keywords.get(name) match {
- case Some(kind) => kind != Outer_Keyword.MINOR
+ case Some(kind) => kind != Keyword.MINOR
case None => false
}
/* tokenize */
- def scan(input: Reader[Char]): List[Outer_Lex.Token] =
+ def scan(input: Reader[Char]): List[Token] =
{
import lexicon._
@@ -49,6 +49,6 @@
}
}
- def scan(input: CharSequence): List[Outer_Lex.Token] =
+ def scan(input: CharSequence): List[Token] =
scan(new CharSequenceReader(input))
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/parse.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,341 @@
+(* Title: Pure/Isar/parse.ML
+ Author: Markus Wenzel, TU Muenchen
+
+Generic parsers for Isabelle/Isar outer syntax.
+*)
+
+signature PARSE =
+sig
+ type 'a parser = Token.T list -> 'a * Token.T list
+ type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list)
+ val group: string -> (Token.T list -> 'a) -> Token.T list -> 'a
+ val !!! : (Token.T list -> 'a) -> Token.T list -> 'a
+ val !!!! : (Token.T list -> 'a) -> Token.T list -> 'a
+ val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
+ val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
+ val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
+ val not_eof: Token.T parser
+ val position: (Token.T list -> 'a * 'b) -> Token.T list -> ('a * Position.T) * 'b
+ val command: string parser
+ val keyword: string parser
+ val short_ident: string parser
+ val long_ident: string parser
+ val sym_ident: string parser
+ val minus: string parser
+ val term_var: string parser
+ val type_ident: string parser
+ val type_var: string parser
+ val number: string parser
+ val string: string parser
+ val alt_string: string parser
+ val verbatim: string parser
+ val sync: string parser
+ val eof: string parser
+ val keyword_with: (string -> bool) -> string parser
+ val keyword_ident_or_symbolic: string parser
+ val $$$ : string -> string parser
+ val reserved: string -> string parser
+ val semicolon: string parser
+ val underscore: string parser
+ val maybe: 'a parser -> 'a option parser
+ val tag_name: string parser
+ val tags: string list parser
+ val opt_unit: unit parser
+ val opt_keyword: string -> bool parser
+ val begin: string parser
+ val opt_begin: bool parser
+ val nat: int parser
+ val int: int parser
+ val enum: string -> 'a parser -> 'a list parser
+ val enum1: string -> 'a parser -> 'a list parser
+ val and_list: 'a parser -> 'a list parser
+ val and_list1: 'a parser -> 'a list parser
+ val enum': string -> 'a context_parser -> 'a list context_parser
+ val enum1': string -> 'a context_parser -> 'a list context_parser
+ val and_list': 'a context_parser -> 'a list context_parser
+ val and_list1': 'a context_parser -> 'a list context_parser
+ val list: 'a parser -> 'a list parser
+ val list1: 'a parser -> 'a list parser
+ val name: bstring parser
+ val binding: binding parser
+ val xname: xstring parser
+ val text: string parser
+ val path: Path.T parser
+ val parname: string parser
+ val parbinding: binding parser
+ val sort: string parser
+ val arity: (string * string list * string) parser
+ val multi_arity: (string list * string list * string) parser
+ val type_args: string list parser
+ val type_args_constrained: (string * string option) list parser
+ val typ_group: string parser
+ val typ: string parser
+ val mixfix: mixfix parser
+ val mixfix': mixfix parser
+ val opt_mixfix: mixfix parser
+ val opt_mixfix': mixfix parser
+ val where_: string parser
+ val const: (string * string * mixfix) parser
+ val const_binding: (binding * string * mixfix) parser
+ val params: (binding * string option) list parser
+ val simple_fixes: (binding * string option) list parser
+ val fixes: (binding * string option * mixfix) list parser
+ val for_fixes: (binding * string option * mixfix) list parser
+ val for_simple_fixes: (binding * string option) list parser
+ val ML_source: (Symbol_Pos.text * Position.T) parser
+ val doc_source: (Symbol_Pos.text * Position.T) parser
+ val term_group: string parser
+ val prop_group: string parser
+ val term: string parser
+ val prop: string parser
+ val propp: (string * string list) parser
+ val termp: (string * string list) parser
+ val target: xstring parser
+ val opt_target: xstring option parser
+end;
+
+structure Parse: PARSE =
+struct
+
+type 'a parser = Token.T list -> 'a * Token.T list;
+type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list);
+
+
+(** error handling **)
+
+(* group atomic parsers (no cuts!) *)
+
+fun fail_with s = Scan.fail_with
+ (fn [] => s ^ " expected (past end-of-file!)"
+ | (tok :: _) =>
+ (case Token.text_of tok of
+ (txt, "") =>
+ s ^ " expected,\nbut " ^ txt ^ Token.pos_of tok ^ " was found"
+ | (txt1, txt2) =>
+ s ^ " expected,\nbut " ^ txt1 ^ Token.pos_of tok ^ " was found:\n" ^ txt2));
+
+fun group s scan = scan || fail_with s;
+
+
+(* cut *)
+
+fun cut kind scan =
+ let
+ fun get_pos [] = " (past end-of-file!)"
+ | get_pos (tok :: _) = Token.pos_of tok;
+
+ fun err (toks, NONE) = kind ^ get_pos toks
+ | err (toks, SOME msg) =
+ if String.isPrefix kind msg then msg
+ else kind ^ get_pos toks ^ ": " ^ msg;
+ in Scan.!! err scan end;
+
+fun !!! scan = cut "Outer syntax error" scan;
+fun !!!! scan = cut "Corrupted outer syntax in presentation" scan;
+
+
+
+(** basic parsers **)
+
+(* utils *)
+
+fun triple1 ((x, y), z) = (x, y, z);
+fun triple2 (x, (y, z)) = (x, y, z);
+fun triple_swap ((x, y), z) = ((x, z), y);
+
+
+(* tokens *)
+
+fun RESET_VALUE atom = (*required for all primitive parsers*)
+ Scan.ahead (Scan.one (K true)) -- atom >> (fn (arg, x) => (Token.assign NONE arg; x));
+
+
+val not_eof = RESET_VALUE (Scan.one Token.not_eof);
+
+fun position scan = (Scan.ahead not_eof >> Token.position_of) -- scan >> Library.swap;
+fun source_position atom = Scan.ahead atom |-- not_eof >> Token.source_position_of;
+fun inner_syntax atom = Scan.ahead atom |-- not_eof >> Token.source_of;
+
+fun kind k =
+ group (Token.str_of_kind k) (RESET_VALUE (Scan.one (Token.is_kind k) >> Token.content_of));
+
+val command = kind Token.Command;
+val keyword = kind Token.Keyword;
+val short_ident = kind Token.Ident;
+val long_ident = kind Token.LongIdent;
+val sym_ident = kind Token.SymIdent;
+val term_var = kind Token.Var;
+val type_ident = kind Token.TypeIdent;
+val type_var = kind Token.TypeVar;
+val number = kind Token.Nat;
+val string = kind Token.String;
+val alt_string = kind Token.AltString;
+val verbatim = kind Token.Verbatim;
+val sync = kind Token.Sync;
+val eof = kind Token.EOF;
+
+fun keyword_with pred = RESET_VALUE (Scan.one (Token.keyword_with pred) >> Token.content_of);
+val keyword_ident_or_symbolic = keyword_with Token.ident_or_symbolic;
+
+fun $$$ x =
+ group (Token.str_of_kind Token.Keyword ^ " " ^ quote x) (keyword_with (fn y => x = y));
+
+fun reserved x =
+ group ("reserved identifier " ^ quote x)
+ (RESET_VALUE (Scan.one (Token.ident_with (fn y => x = y)) >> Token.content_of));
+
+val semicolon = $$$ ";";
+
+val minus = sym_ident :-- (fn "-" => Scan.succeed () | _ => Scan.fail) >> #1;
+val underscore = sym_ident :-- (fn "_" => Scan.succeed () | _ => Scan.fail) >> #1;
+fun maybe scan = underscore >> K NONE || scan >> SOME;
+
+val nat = number >> (#1 o Library.read_int o Symbol.explode);
+val int = Scan.optional (minus >> K ~1) 1 -- nat >> op *;
+
+val tag_name = group "tag name" (short_ident || string);
+val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);
+
+val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) ();
+fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false;
+
+val begin = $$$ "begin";
+val opt_begin = Scan.optional (begin >> K true) false;
+
+
+(* enumerations *)
+
+fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan);
+fun enum sep scan = enum1 sep scan || Scan.succeed [];
+
+fun enum1' sep scan = scan ::: Scan.repeat (Scan.lift ($$$ sep) |-- scan);
+fun enum' sep scan = enum1' sep scan || Scan.succeed [];
+
+fun and_list1 scan = enum1 "and" scan;
+fun and_list scan = enum "and" scan;
+
+fun and_list1' scan = enum1' "and" scan;
+fun and_list' scan = enum' "and" scan;
+
+fun list1 scan = enum1 "," scan;
+fun list scan = enum "," scan;
+
+
+(* names and text *)
+
+val name = group "name declaration" (short_ident || sym_ident || string || number);
+val binding = position name >> Binding.make;
+val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number);
+val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);
+val path = group "file name/path specification" name >> Path.explode;
+
+val parname = Scan.optional ($$$ "(" |-- name --| $$$ ")") "";
+val parbinding = Scan.optional ($$$ "(" |-- binding --| $$$ ")") Binding.empty;
+
+
+(* sorts and arities *)
+
+val sort = group "sort" (inner_syntax xname);
+
+val arity = xname -- ($$$ "::" |-- !!!
+ (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
+
+val multi_arity = and_list1 xname -- ($$$ "::" |-- !!!
+ (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
+
+
+(* types *)
+
+val typ_group = group "type"
+ (short_ident || long_ident || sym_ident || type_ident || type_var || string || number);
+
+val typ = inner_syntax typ_group;
+
+fun type_arguments arg =
+ arg >> single ||
+ $$$ "(" |-- !!! (list1 arg --| $$$ ")") ||
+ Scan.succeed [];
+
+val type_args = type_arguments type_ident;
+val type_args_constrained = type_arguments (type_ident -- Scan.option ($$$ "::" |-- !!! sort));
+
+
+(* mixfix annotations *)
+
+val mfix = string --
+ !!! (Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) [] --
+ Scan.optional nat Syntax.max_pri) >> (Mixfix o triple2);
+
+val infx = $$$ "infix" |-- !!! (string -- nat >> Infix);
+val infxl = $$$ "infixl" |-- !!! (string -- nat >> Infixl);
+val infxr = $$$ "infixr" |-- !!! (string -- nat >> Infixr);
+
+val binder = $$$ "binder" |--
+ !!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n))))
+ >> (Binder o triple2);
+
+fun annotation guard fix = $$$ "(" |-- guard (fix --| $$$ ")");
+fun opt_annotation guard fix = Scan.optional (annotation guard fix) NoSyn;
+
+val mixfix = annotation !!! (mfix || binder || infxl || infxr || infx);
+val mixfix' = annotation I (mfix || binder || infxl || infxr || infx);
+val opt_mixfix = opt_annotation !!! (mfix || binder || infxl || infxr || infx);
+val opt_mixfix' = opt_annotation I (mfix || binder || infxl || infxr || infx);
+
+
+(* fixes *)
+
+val where_ = $$$ "where";
+
+val const = name -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;
+val const_binding = binding -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;
+
+val params = Scan.repeat1 binding -- Scan.option ($$$ "::" |-- !!! typ)
+ >> (fn (xs, T) => map (rpair T) xs);
+
+val simple_fixes = and_list1 params >> flat;
+
+val fixes =
+ and_list1 (binding -- Scan.option ($$$ "::" |-- typ) -- mixfix >> (single o triple1) ||
+ params >> map Syntax.no_syn) >> flat;
+
+val for_fixes = Scan.optional ($$$ "for" |-- !!! fixes) [];
+val for_simple_fixes = Scan.optional ($$$ "for" |-- !!! simple_fixes) [];
+
+
+(* embedded source text *)
+
+val ML_source = source_position (group "ML source" text);
+val doc_source = source_position (group "document source" text);
+
+
+(* terms *)
+
+val trm = short_ident || long_ident || sym_ident || term_var || number || string;
+
+val term_group = group "term" trm;
+val prop_group = group "proposition" trm;
+
+val term = inner_syntax term_group;
+val prop = inner_syntax prop_group;
+
+
+(* patterns *)
+
+val is_terms = Scan.repeat1 ($$$ "is" |-- term);
+val is_props = Scan.repeat1 ($$$ "is" |-- prop);
+
+val propp = prop -- Scan.optional ($$$ "(" |-- !!! (is_props --| $$$ ")")) [];
+val termp = term -- Scan.optional ($$$ "(" |-- !!! (is_terms --| $$$ ")")) [];
+
+
+(* targets *)
+
+val target = ($$$ "(" -- $$$ "in") |-- !!! (xname --| $$$ ")");
+val opt_target = Scan.option target;
+
+end;
+
+type 'a parser = 'a Parse.parser;
+type 'a context_parser = 'a Parse.context_parser;
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/parse.scala Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,77 @@
+/* Title: Pure/Isar/outer_parse.scala
+ Author: Makarius
+
+Generic parsers for Isabelle/Isar outer syntax.
+*/
+
+package isabelle
+
+import scala.util.parsing.combinator.Parsers
+
+
+object Parse
+{
+ /* parsing tokens */
+
+ trait Parser extends Parsers
+ {
+ type Elem = Token
+
+ def filter_proper = true
+
+ private def proper(in: Input): Input =
+ if (in.atEnd || !in.first.is_ignored || !filter_proper) in
+ else proper(in.rest)
+
+ def token(s: String, pred: Elem => Boolean): Parser[Elem] = new Parser[Elem]
+ {
+ def apply(raw_input: Input) =
+ {
+ val in = proper(raw_input)
+ if (in.atEnd) Failure(s + " expected (past end-of-file!)", in)
+ else {
+ val token = in.first
+ if (pred(token)) Success(token, proper(in.rest))
+ else
+ token.text match {
+ case (txt, "") =>
+ Failure(s + " expected,\nbut " + txt + " was found", in)
+ case (txt1, txt2) =>
+ Failure(s + " expected,\nbut " + txt1 + " was found:\n" + txt2, in)
+ }
+ }
+ }
+ }
+
+ def not_eof: Parser[Elem] = token("input token", _ => true)
+ def eof: Parser[Unit] = not(not_eof)
+
+ def atom(s: String, pred: Elem => Boolean): Parser[String] =
+ token(s, pred) ^^ (_.content)
+
+ def keyword(name: String): Parser[String] =
+ atom(Token.Kind.KEYWORD.toString + " \"" + name + "\"",
+ tok => tok.kind == Token.Kind.KEYWORD && tok.content == name)
+
+ def name: Parser[String] = atom("name declaration", _.is_name)
+ def xname: Parser[String] = atom("name reference", _.is_xname)
+ def text: Parser[String] = atom("text", _.is_text)
+ def ML_source: Parser[String] = atom("ML source", _.is_text)
+ def doc_source: Parser[String] = atom("document source", _.is_text)
+ def path: Parser[String] = atom("file name/path specification", _.is_name)
+
+ private def tag_name: Parser[String] =
+ atom("tag name", tok =>
+ tok.kind == Token.Kind.IDENT ||
+ tok.kind == Token.Kind.STRING)
+
+ def tags: Parser[List[String]] = rep(keyword("%") ~> tag_name)
+
+
+ /* wrappers */
+
+ def parse[T](p: Parser[T], in: Token.Reader): ParseResult[T] = p(in)
+ def parse_all[T](p: Parser[T], in: Token.Reader): ParseResult[T] = parse(phrase(p), in)
+ }
+}
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/parse_spec.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,164 @@
+(* Title: Pure/Isar/parse_spec.ML
+ Author: Makarius
+
+Parsers for complex specifications.
+*)
+
+signature PARSE_SPEC =
+sig
+ val attrib: Attrib.src parser
+ val attribs: Attrib.src list parser
+ val opt_attribs: Attrib.src list parser
+ val thm_name: string -> Attrib.binding parser
+ val opt_thm_name: string -> Attrib.binding parser
+ val spec: (Attrib.binding * string) parser
+ val specs: (Attrib.binding * string list) parser
+ val alt_specs: (Attrib.binding * string) list parser
+ val where_alt_specs: (Attrib.binding * string) list parser
+ val xthm: (Facts.ref * Attrib.src list) parser
+ val xthms1: (Facts.ref * Attrib.src list) list parser
+ val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
+ val constdecl: (binding * string option * mixfix) parser
+ val constdef: ((binding * string option * mixfix) option * (Attrib.binding * string)) parser
+ val locale_mixfix: mixfix parser
+ val locale_fixes: (binding * string option * mixfix) list parser
+ val locale_insts: (string option list * (Attrib.binding * string) list) parser
+ val class_expr: string list parser
+ val locale_expression: bool -> Expression.expression parser
+ val locale_keyword: string parser
+ val context_element: Element.context parser
+ val statement: (Attrib.binding * (string * string list) list) list parser
+ val general_statement: (Element.context list * Element.statement) parser
+ val statement_keyword: string parser
+end;
+
+structure Parse_Spec: PARSE_SPEC =
+struct
+
+(* theorem specifications *)
+
+val attrib =
+ Parse.position ((Parse.keyword_ident_or_symbolic || Parse.xname) -- Parse.!!! Args.parse)
+ >> Args.src;
+
+val attribs = Parse.$$$ "[" |-- Parse.list attrib --| Parse.$$$ "]";
+val opt_attribs = Scan.optional attribs [];
+
+fun thm_name s = Parse.binding -- opt_attribs --| Parse.$$$ s;
+
+fun opt_thm_name s =
+ Scan.optional ((Parse.binding -- opt_attribs || attribs >> pair Binding.empty) --| Parse.$$$ s)
+ Attrib.empty_binding;
+
+val spec = opt_thm_name ":" -- Parse.prop;
+val specs = opt_thm_name ":" -- Scan.repeat1 Parse.prop;
+
+val alt_specs =
+ Parse.enum1 "|"
+ (spec --| Scan.option (Scan.ahead (Parse.name || Parse.$$$ "[") -- Parse.!!! (Parse.$$$ "|")));
+
+val where_alt_specs = Parse.where_ |-- Parse.!!! alt_specs;
+
+val xthm =
+ Parse.$$$ "[" |-- attribs --| Parse.$$$ "]" >> pair (Facts.named "") ||
+ (Parse.alt_string >> Facts.Fact ||
+ Parse.position Parse.xname -- Scan.option Attrib.thm_sel >> Facts.Named) -- opt_attribs;
+
+val xthms1 = Scan.repeat1 xthm;
+
+val name_facts = Parse.and_list1 (opt_thm_name "=" -- xthms1);
+
+
+(* basic constant specifications *)
+
+val constdecl =
+ Parse.binding --
+ (Parse.where_ >> K (NONE, NoSyn) ||
+ Parse.$$$ "::" |-- Parse.!!! ((Parse.typ >> SOME) -- Parse.opt_mixfix' --| Parse.where_) ||
+ Scan.ahead (Parse.$$$ "(") |-- Parse.!!! (Parse.mixfix' --| Parse.where_ >> pair NONE))
+ >> Parse.triple2;
+
+val constdef = Scan.option constdecl -- (opt_thm_name ":" -- Parse.prop);
+
+
+(* locale and context elements *)
+
+val locale_mixfix =
+ Parse.$$$ "(" -- Parse.$$$ "structure" -- Parse.!!! (Parse.$$$ ")") >> K Structure ||
+ Parse.mixfix;
+
+val locale_fixes =
+ Parse.and_list1 (Parse.binding -- Scan.option (Parse.$$$ "::" |-- Parse.typ) -- locale_mixfix
+ >> (single o Parse.triple1) ||
+ Parse.params >> map Syntax.no_syn) >> flat;
+
+val locale_insts =
+ Scan.optional
+ (Parse.$$$ "[" |-- Parse.!!! (Scan.repeat1 (Parse.maybe Parse.term) --| Parse.$$$ "]")) [] --
+ Scan.optional (Parse.where_ |-- Parse.and_list1 (opt_thm_name ":" -- Parse.prop)) [];
+
+local
+
+val loc_element =
+ Parse.$$$ "fixes" |-- Parse.!!! locale_fixes >> Element.Fixes ||
+ Parse.$$$ "constrains" |--
+ Parse.!!! (Parse.and_list1 (Parse.name -- (Parse.$$$ "::" |-- Parse.typ)))
+ >> Element.Constrains ||
+ Parse.$$$ "assumes" |-- Parse.!!! (Parse.and_list1 (opt_thm_name ":" -- Scan.repeat1 Parse.propp))
+ >> Element.Assumes ||
+ Parse.$$$ "defines" |-- Parse.!!! (Parse.and_list1 (opt_thm_name ":" -- Parse.propp))
+ >> Element.Defines ||
+ Parse.$$$ "notes" |-- Parse.!!! (Parse.and_list1 (opt_thm_name "=" -- xthms1))
+ >> (curry Element.Notes "");
+
+fun plus1_unless test scan =
+ scan ::: Scan.repeat (Parse.$$$ "+" |-- Scan.unless test (Parse.!!! scan));
+
+fun prefix mandatory =
+ Parse.name --
+ (Parse.$$$ "!" >> K true || Parse.$$$ "?" >> K false || Scan.succeed mandatory) --|
+ Parse.$$$ ":";
+
+val instance = Parse.where_ |--
+ Parse.and_list1 (Parse.name -- (Parse.$$$ "=" |-- Parse.term)) >> Expression.Named ||
+ Scan.repeat1 (Parse.maybe Parse.term) >> Expression.Positional;
+
+in
+
+val locale_keyword =
+ Parse.$$$ "fixes" || Parse.$$$ "constrains" || Parse.$$$ "assumes" ||
+ Parse.$$$ "defines" || Parse.$$$ "notes";
+
+val class_expr = plus1_unless locale_keyword Parse.xname;
+
+fun locale_expression mandatory =
+ let
+ val expr2 = Parse.xname;
+ val expr1 = Scan.optional (prefix mandatory) ("", false) -- expr2 --
+ Scan.optional instance (Expression.Named []) >> (fn ((p, l), i) => (l, (p, i)));
+ val expr0 = plus1_unless locale_keyword expr1;
+ in expr0 -- Scan.optional (Parse.$$$ "for" |-- Parse.!!! locale_fixes) [] end;
+
+val context_element = Parse.group "context element" loc_element;
+
+end;
+
+
+(* statements *)
+
+val statement = Parse.and_list1 (opt_thm_name ":" -- Scan.repeat1 Parse.propp);
+
+val obtain_case =
+ Parse.parbinding -- (Scan.optional (Parse.simple_fixes --| Parse.where_) [] --
+ (Parse.and_list1 (Scan.repeat1 Parse.prop) >> flat));
+
+val general_statement =
+ statement >> (fn x => ([], Element.Shows x)) ||
+ Scan.repeat context_element --
+ (Parse.$$$ "obtains" |-- Parse.!!! (Parse.enum1 "|" obtain_case) >> Element.Obtains ||
+ Parse.$$$ "shows" |-- Parse.!!! statement >> Element.Shows);
+
+val statement_keyword = Parse.$$$ "obtains" || Parse.$$$ "shows";
+
+end;
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/parse_value.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,42 @@
+(* Title: Pure/Isar/parse_value.ML
+ Author: Makarius
+
+Outer syntax parsers for basic ML values.
+*)
+
+signature PARSE_VALUE =
+sig
+ val comma: 'a parser -> 'a parser
+ val equal: 'a parser -> 'a parser
+ val parens: 'a parser -> 'a parser
+ val unit: unit parser
+ val pair: 'a parser -> 'b parser -> ('a * 'b) parser
+ val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
+ val list: 'a parser -> 'a list parser
+ val properties: Properties.T parser
+end;
+
+structure Parse_Value: PARSE_VALUE =
+struct
+
+(* syntax utilities *)
+
+fun comma p = Parse.$$$ "," |-- Parse.!!! p;
+fun equal p = Parse.$$$ "=" |-- Parse.!!! p;
+fun parens p = Parse.$$$ "(" |-- Parse.!!! (p --| Parse.$$$ ")");
+
+
+(* tuples *)
+
+val unit = parens (Scan.succeed ());
+fun pair p1 p2 = parens (p1 -- comma p2);
+fun triple p1 p2 p3 = parens (p1 -- comma p2 -- comma p3) >> Parse.triple1;
+
+
+(* lists *)
+
+fun list p = parens (Parse.enum "," p);
+val properties = list (Parse.string -- equal Parse.string);
+
+end;
+
--- a/src/Pure/Isar/rule_insts.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/rule_insts.ML Sun May 23 10:38:11 2010 +0100
@@ -29,10 +29,6 @@
structure RuleInsts: RULE_INSTS =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-
-
(** reading instantiations **)
local
@@ -101,11 +97,11 @@
val (type_insts, term_insts) = List.partition (is_tvar o fst) mixed_insts;
val internal_insts = term_insts |> map_filter
- (fn (xi, T.Term t) => SOME (xi, t)
- | (_, T.Text _) => NONE
+ (fn (xi, Token.Term t) => SOME (xi, t)
+ | (_, Token.Text _) => NONE
| (xi, _) => error_var "Term argument expected for " xi);
val external_insts = term_insts |> map_filter
- (fn (xi, T.Text s) => SOME (xi, s) | _ => NONE);
+ (fn (xi, Token.Text s) => SOME (xi, s) | _ => NONE);
(* mixed type instantiations *)
@@ -115,8 +111,8 @@
val S = the_sort tvars xi;
val T =
(case arg of
- T.Text s => Syntax.read_typ ctxt s
- | T.Typ T => T
+ Token.Text s => Syntax.read_typ ctxt s
+ | Token.Typ T => T
| _ => error_var "Type argument expected for " xi);
in
if Sign.of_sort thy (T, S) then ((xi, S), T)
@@ -173,9 +169,9 @@
val _ = (*assign internalized values*)
mixed_insts |> List.app (fn (arg, (xi, _)) =>
if is_tvar xi then
- T.assign (SOME (T.Typ (the (AList.lookup (op =) type_insts xi)))) arg
+ Token.assign (SOME (Token.Typ (the (AList.lookup (op =) type_insts xi)))) arg
else
- T.assign (SOME (T.Term (the (AList.lookup (op =) term_insts xi)))) arg);
+ Token.assign (SOME (Token.Term (the (AList.lookup (op =) term_insts xi)))) arg);
in
Drule.instantiate insts thm |> Rule_Cases.save thm
end;
@@ -198,7 +194,7 @@
fun read_instantiate ctxt args thm =
read_instantiate_mixed (ctxt |> ProofContext.set_mode ProofContext.mode_schematic) (* FIXME !? *)
- (map (fn (x, y) => (T.eof, (x, T.Text y))) args) thm;
+ (map (fn (x, y) => (Token.eof, (x, Token.Text y))) args) thm;
fun instantiate_tac ctxt args = PRIMITIVE (read_instantiate ctxt args);
@@ -211,18 +207,18 @@
local
val value =
- Args.internal_typ >> T.Typ ||
- Args.internal_term >> T.Term ||
- Args.name_source >> T.Text;
+ Args.internal_typ >> Token.Typ ||
+ Args.internal_term >> Token.Term ||
+ Args.name_source >> Token.Text;
-val inst = Args.var -- (Args.$$$ "=" |-- Scan.ahead P.not_eof -- value)
+val inst = Args.var -- (Args.$$$ "=" |-- Scan.ahead Parse.not_eof -- value)
>> (fn (xi, (a, v)) => (a, (xi, v)));
in
val _ = Context.>> (Context.map_theory
(Attrib.setup (Binding.name "where")
- (Scan.lift (P.and_list inst) >> (fn args =>
+ (Scan.lift (Parse.and_list inst) >> (fn args =>
Thm.rule_attribute (fn context => read_instantiate_mixed (Context.proof_of context) args)))
"named instantiation of theorem"));
@@ -234,10 +230,10 @@
local
val value =
- Args.internal_term >> T.Term ||
- Args.name_source >> T.Text;
+ Args.internal_term >> Token.Term ||
+ Args.name_source >> Token.Text;
-val inst = Scan.ahead P.not_eof -- Args.maybe value;
+val inst = Scan.ahead Parse.not_eof -- Args.maybe value;
val concl = Args.$$$ "concl" -- Args.colon;
val insts =
@@ -387,7 +383,8 @@
fun method inst_tac tac =
Args.goal_spec --
Scan.optional (Scan.lift
- (P.and_list1 (Args.var -- (Args.$$$ "=" |-- P.!!! Args.name_source)) --| Args.$$$ "in")) [] --
+ (Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.name_source)) --|
+ Args.$$$ "in")) [] --
Attrib.thms >>
(fn ((quant, insts), thms) => fn ctxt => METHOD (fn facts =>
if null insts then quant (Method.insert_tac facts THEN' tac ctxt thms)
@@ -425,5 +422,5 @@
end;
-structure BasicRuleInsts: BASIC_RULE_INSTS = RuleInsts;
-open BasicRuleInsts;
+structure Basic_Rule_Insts: BASIC_RULE_INSTS = RuleInsts;
+open Basic_Rule_Insts;
--- a/src/Pure/Isar/spec_parse.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,155 +0,0 @@
-(* Title: Pure/Isar/spec_parse.ML
- Author: Makarius
-
-Parsers for complex specifications.
-*)
-
-signature SPEC_PARSE =
-sig
- val attrib: Attrib.src parser
- val attribs: Attrib.src list parser
- val opt_attribs: Attrib.src list parser
- val thm_name: string -> Attrib.binding parser
- val opt_thm_name: string -> Attrib.binding parser
- val spec: (Attrib.binding * string) parser
- val specs: (Attrib.binding * string list) parser
- val alt_specs: (Attrib.binding * string) list parser
- val where_alt_specs: (Attrib.binding * string) list parser
- val xthm: (Facts.ref * Attrib.src list) parser
- val xthms1: (Facts.ref * Attrib.src list) list parser
- val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
- val constdecl: (binding * string option * mixfix) parser
- val constdef: ((binding * string option * mixfix) option * (Attrib.binding * string)) parser
- val locale_mixfix: mixfix parser
- val locale_fixes: (binding * string option * mixfix) list parser
- val locale_insts: (string option list * (Attrib.binding * string) list) parser
- val class_expr: string list parser
- val locale_expression: bool -> Expression.expression parser
- val locale_keyword: string parser
- val context_element: Element.context parser
- val statement: (Attrib.binding * (string * string list) list) list parser
- val general_statement: (Element.context list * Element.statement) parser
- val statement_keyword: string parser
-end;
-
-structure SpecParse: SPEC_PARSE =
-struct
-
-structure P = OuterParse;
-
-
-(* theorem specifications *)
-
-val attrib = P.position ((P.keyword_ident_or_symbolic || P.xname) -- P.!!! Args.parse) >> Args.src;
-val attribs = P.$$$ "[" |-- P.list attrib --| P.$$$ "]";
-val opt_attribs = Scan.optional attribs [];
-
-fun thm_name s = P.binding -- opt_attribs --| P.$$$ s;
-
-fun opt_thm_name s =
- Scan.optional ((P.binding -- opt_attribs || attribs >> pair Binding.empty) --| P.$$$ s)
- Attrib.empty_binding;
-
-val spec = opt_thm_name ":" -- P.prop;
-val specs = opt_thm_name ":" -- Scan.repeat1 P.prop;
-
-val alt_specs =
- P.enum1 "|" (spec --| Scan.option (Scan.ahead (P.name || P.$$$ "[") -- P.!!! (P.$$$ "|")));
-
-val where_alt_specs = P.where_ |-- P.!!! alt_specs;
-
-val xthm =
- P.$$$ "[" |-- attribs --| P.$$$ "]" >> pair (Facts.named "") ||
- (P.alt_string >> Facts.Fact ||
- P.position P.xname -- Scan.option Attrib.thm_sel >> Facts.Named) -- opt_attribs;
-
-val xthms1 = Scan.repeat1 xthm;
-
-val name_facts = P.and_list1 (opt_thm_name "=" -- xthms1);
-
-
-(* basic constant specifications *)
-
-val constdecl =
- P.binding --
- (P.where_ >> K (NONE, NoSyn) ||
- P.$$$ "::" |-- P.!!! ((P.typ >> SOME) -- P.opt_mixfix' --| P.where_) ||
- Scan.ahead (P.$$$ "(") |-- P.!!! (P.mixfix' --| P.where_ >> pair NONE))
- >> P.triple2;
-
-val constdef = Scan.option constdecl -- (opt_thm_name ":" -- P.prop);
-
-
-(* locale and context elements *)
-
-val locale_mixfix = P.$$$ "(" -- P.$$$ "structure" -- P.!!! (P.$$$ ")") >> K Structure || P.mixfix;
-
-val locale_fixes =
- P.and_list1 (P.binding -- Scan.option (P.$$$ "::" |-- P.typ) -- locale_mixfix
- >> (single o P.triple1) ||
- P.params >> map Syntax.no_syn) >> flat;
-
-val locale_insts =
- Scan.optional (P.$$$ "[" |-- P.!!! (Scan.repeat1 (P.maybe P.term) --| P.$$$ "]")) []
- -- Scan.optional (P.where_ |-- P.and_list1 (opt_thm_name ":" -- P.prop)) [];
-
-local
-
-val loc_element =
- P.$$$ "fixes" |-- P.!!! locale_fixes >> Element.Fixes ||
- P.$$$ "constrains" |-- P.!!! (P.and_list1 (P.name -- (P.$$$ "::" |-- P.typ)))
- >> Element.Constrains ||
- P.$$$ "assumes" |-- P.!!! (P.and_list1 (opt_thm_name ":" -- Scan.repeat1 P.propp))
- >> Element.Assumes ||
- P.$$$ "defines" |-- P.!!! (P.and_list1 (opt_thm_name ":" -- P.propp))
- >> Element.Defines ||
- P.$$$ "notes" |-- P.!!! (P.and_list1 (opt_thm_name "=" -- xthms1))
- >> (curry Element.Notes "");
-
-fun plus1_unless test scan =
- scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
-
-fun prefix mandatory =
- P.name -- (P.$$$ "!" >> K true || P.$$$ "?" >> K false || Scan.succeed mandatory) --| P.$$$ ":";
-
-val instance = P.where_ |--
- P.and_list1 (P.name -- (P.$$$ "=" |-- P.term)) >> Expression.Named ||
- Scan.repeat1 (P.maybe P.term) >> Expression.Positional;
-
-in
-
-val locale_keyword = P.$$$ "fixes" || P.$$$ "constrains" || P.$$$ "assumes" ||
- P.$$$ "defines" || P.$$$ "notes";
-
-val class_expr = plus1_unless locale_keyword P.xname;
-
-fun locale_expression mandatory =
- let
- val expr2 = P.xname;
- val expr1 = Scan.optional (prefix mandatory) ("", false) -- expr2 --
- Scan.optional instance (Expression.Named []) >> (fn ((p, l), i) => (l, (p, i)));
- val expr0 = plus1_unless locale_keyword expr1;
- in expr0 -- Scan.optional (P.$$$ "for" |-- P.!!! locale_fixes) [] end;
-
-val context_element = P.group "context element" loc_element;
-
-end;
-
-
-(* statements *)
-
-val statement = P.and_list1 (opt_thm_name ":" -- Scan.repeat1 P.propp);
-
-val obtain_case =
- P.parbinding -- (Scan.optional (P.simple_fixes --| P.where_) [] --
- (P.and_list1 (Scan.repeat1 P.prop) >> flat));
-
-val general_statement =
- statement >> (fn x => ([], Element.Shows x)) ||
- Scan.repeat context_element --
- (P.$$$ "obtains" |-- P.!!! (P.enum1 "|" obtain_case) >> Element.Obtains ||
- P.$$$ "shows" |-- P.!!! statement >> Element.Shows);
-
-val statement_keyword = P.$$$ "obtains" || P.$$$ "shows";
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/token.ML Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,389 @@
+(* Title: Pure/Isar/token.ML
+ Author: Markus Wenzel, TU Muenchen
+
+Outer token syntax for Isabelle/Isar.
+*)
+
+signature TOKEN =
+sig
+ datatype kind =
+ Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
+ Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
+ Malformed | Error of string | Sync | EOF
+ datatype value =
+ Text of string | Typ of typ | Term of term | Fact of thm list |
+ Attribute of morphism -> attribute
+ type T
+ val str_of_kind: kind -> string
+ val position_of: T -> Position.T
+ val end_position_of: T -> Position.T
+ val pos_of: T -> string
+ val eof: T
+ val is_eof: T -> bool
+ val not_eof: T -> bool
+ val not_sync: T -> bool
+ val stopper: T Scan.stopper
+ val kind_of: T -> kind
+ val is_kind: kind -> T -> bool
+ val keyword_with: (string -> bool) -> T -> bool
+ val ident_with: (string -> bool) -> T -> bool
+ val is_proper: T -> bool
+ val is_semicolon: T -> bool
+ val is_comment: T -> bool
+ val is_begin_ignore: T -> bool
+ val is_end_ignore: T -> bool
+ val is_blank: T -> bool
+ val is_newline: T -> bool
+ val source_of: T -> string
+ val source_position_of: T -> Symbol_Pos.text * Position.T
+ val content_of: T -> string
+ val unparse: T -> string
+ val text_of: T -> string * string
+ val get_value: T -> value option
+ val map_value: (value -> value) -> T -> T
+ val mk_text: string -> T
+ val mk_typ: typ -> T
+ val mk_term: term -> T
+ val mk_fact: thm list -> T
+ val mk_attribute: (morphism -> attribute) -> T
+ val assignable: T -> T
+ val assign: value option -> T -> unit
+ val closure: T -> T
+ val ident_or_symbolic: string -> bool
+ val !!! : string -> (Symbol_Pos.T list -> 'a) -> Symbol_Pos.T list -> 'a
+ val source_proper: (T, 'a) Source.source -> (T, (T, 'a) Source.source) Source.source
+ val source': {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
+ (Symbol_Pos.T, 'a) Source.source -> (T, (Symbol_Pos.T, 'a) Source.source) Source.source
+ val source: {do_recover: bool Option.option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
+ Position.T -> (Symbol.symbol, 'a) Source.source -> (T,
+ (Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source
+ val read_antiq: Scan.lexicon -> (T list -> 'a * T list) -> Symbol_Pos.T list * Position.T -> 'a
+end;
+
+structure Token: TOKEN =
+struct
+
+(** tokens **)
+
+(* token values *)
+
+(*The value slot assigns an (optional) internal value to a token,
+ usually as a side-effect of special scanner setup (see also
+ args.ML). Note that an assignable ref designates an intermediate
+ state of internalization -- it is NOT meant to persist.*)
+
+datatype value =
+ Text of string |
+ Typ of typ |
+ Term of term |
+ Fact of thm list |
+ Attribute of morphism -> attribute;
+
+datatype slot =
+ Slot |
+ Value of value option |
+ Assignable of value option Unsynchronized.ref;
+
+
+(* datatype token *)
+
+datatype kind =
+ Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
+ Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
+ Malformed | Error of string | Sync | EOF;
+
+datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot;
+
+val str_of_kind =
+ fn Command => "command"
+ | Keyword => "keyword"
+ | Ident => "identifier"
+ | LongIdent => "long identifier"
+ | SymIdent => "symbolic identifier"
+ | Var => "schematic variable"
+ | TypeIdent => "type variable"
+ | TypeVar => "schematic type variable"
+ | Nat => "number"
+ | String => "string"
+ | AltString => "back-quoted string"
+ | Verbatim => "verbatim text"
+ | Space => "white space"
+ | Comment => "comment text"
+ | InternalValue => "internal value"
+ | Malformed => "malformed symbolic character"
+ | Error _ => "bad input"
+ | Sync => "sync marker"
+ | EOF => "end-of-file";
+
+
+(* position *)
+
+fun position_of (Token ((_, (pos, _)), _, _)) = pos;
+fun end_position_of (Token ((_, (_, pos)), _, _)) = pos;
+
+val pos_of = Position.str_of o position_of;
+
+
+(* control tokens *)
+
+fun mk_eof pos = Token (("", (pos, Position.none)), (EOF, ""), Slot);
+val eof = mk_eof Position.none;
+
+fun is_eof (Token (_, (EOF, _), _)) = true
+ | is_eof _ = false;
+
+val not_eof = not o is_eof;
+
+fun not_sync (Token (_, (Sync, _), _)) = false
+ | not_sync _ = true;
+
+val stopper =
+ Scan.stopper (fn [] => eof | toks => mk_eof (end_position_of (List.last toks))) is_eof;
+
+
+(* kind of token *)
+
+fun kind_of (Token (_, (k, _), _)) = k;
+fun is_kind k (Token (_, (k', _), _)) = k = k';
+
+fun keyword_with pred (Token (_, (Keyword, x), _)) = pred x
+ | keyword_with _ _ = false;
+
+fun ident_with pred (Token (_, (Ident, x), _)) = pred x
+ | ident_with _ _ = false;
+
+fun is_proper (Token (_, (Space, _), _)) = false
+ | is_proper (Token (_, (Comment, _), _)) = false
+ | is_proper _ = true;
+
+fun is_semicolon (Token (_, (Keyword, ";"), _)) = true
+ | is_semicolon _ = false;
+
+fun is_comment (Token (_, (Comment, _), _)) = true
+ | is_comment _ = false;
+
+fun is_begin_ignore (Token (_, (Comment, "<"), _)) = true
+ | is_begin_ignore _ = false;
+
+fun is_end_ignore (Token (_, (Comment, ">"), _)) = true
+ | is_end_ignore _ = false;
+
+
+(* blanks and newlines -- space tokens obey lines *)
+
+fun is_blank (Token (_, (Space, x), _)) = not (String.isSuffix "\n" x)
+ | is_blank _ = false;
+
+fun is_newline (Token (_, (Space, x), _)) = String.isSuffix "\n" x
+ | is_newline _ = false;
+
+
+(* token content *)
+
+fun source_of (Token ((source, (pos, _)), _, _)) =
+ YXML.string_of (XML.Elem (Markup.tokenN, Position.properties_of pos, [XML.Text source]));
+
+fun source_position_of (Token ((source, (pos, _)), _, _)) = (source, pos);
+
+fun content_of (Token (_, (_, x), _)) = x;
+
+
+(* unparse *)
+
+fun escape q =
+ implode o map (fn s => if s = q orelse s = "\\" then "\\" ^ s else s) o Symbol.explode;
+
+fun unparse (Token (_, (kind, x), _)) =
+ (case kind of
+ String => x |> quote o escape "\""
+ | AltString => x |> enclose "`" "`" o escape "`"
+ | Verbatim => x |> enclose "{*" "*}"
+ | Comment => x |> enclose "(*" "*)"
+ | Malformed => space_implode " " (explode x)
+ | Sync => ""
+ | EOF => ""
+ | _ => x);
+
+fun text_of tok =
+ if is_semicolon tok then ("terminator", "")
+ else
+ let
+ val k = str_of_kind (kind_of tok);
+ val s = unparse tok
+ handle ERROR _ => Symbol.separate_chars (content_of tok);
+ in
+ if s = "" then (k, "")
+ else if size s < 40 andalso not (exists_string (fn c => c = "\n") s) then (k ^ " " ^ s, "")
+ else (k, s)
+ end;
+
+
+
+(** associated values **)
+
+(* access values *)
+
+fun get_value (Token (_, _, Value v)) = v
+ | get_value _ = NONE;
+
+fun map_value f (Token (x, y, Value (SOME v))) = Token (x, y, Value (SOME (f v)))
+ | map_value _ tok = tok;
+
+
+(* make values *)
+
+fun mk_value k v = Token ((k, Position.no_range), (InternalValue, k), Value (SOME v));
+
+val mk_text = mk_value "<text>" o Text;
+val mk_typ = mk_value "<typ>" o Typ;
+val mk_term = mk_value "<term>" o Term;
+val mk_fact = mk_value "<fact>" o Fact;
+val mk_attribute = mk_value "<attribute>" o Attribute;
+
+
+(* static binding *)
+
+(*1st stage: make empty slots assignable*)
+fun assignable (Token (x, y, Slot)) = Token (x, y, Assignable (Unsynchronized.ref NONE))
+ | assignable tok = tok;
+
+(*2nd stage: assign values as side-effect of scanning*)
+fun assign v (Token (_, _, Assignable r)) = r := v
+ | assign _ _ = ();
+
+(*3rd stage: static closure of final values*)
+fun closure (Token (x, y, Assignable (Unsynchronized.ref v))) = Token (x, y, Value v)
+ | closure tok = tok;
+
+
+
+(** scanners **)
+
+open Basic_Symbol_Pos;
+
+fun !!! msg = Symbol_Pos.!!! ("Outer lexical error: " ^ msg);
+
+
+(* scan symbolic idents *)
+
+val is_sym_char = member (op =) (explode "!#$%&*+-/<=>?@^_|~");
+
+val scan_symid =
+ Scan.many1 (is_sym_char o symbol) ||
+ Scan.one (Symbol.is_symbolic o symbol) >> single;
+
+fun is_symid str =
+ (case try Symbol.explode str of
+ SOME [s] => Symbol.is_symbolic s orelse is_sym_char s
+ | SOME ss => forall is_sym_char ss
+ | _ => false);
+
+fun ident_or_symbolic "begin" = false
+ | ident_or_symbolic ":" = true
+ | ident_or_symbolic "::" = true
+ | ident_or_symbolic s = Syntax.is_identifier s orelse is_symid s;
+
+
+(* scan verbatim text *)
+
+val scan_verb =
+ $$$ "*" --| Scan.ahead (~$$$ "}") ||
+ Scan.one (fn (s, _) => s <> "*" andalso Symbol.is_regular s) >> single;
+
+val scan_verbatim =
+ (Symbol_Pos.scan_pos --| $$$ "{" --| $$$ "*") -- !!! "missing end of verbatim text"
+ (Symbol_Pos.change_prompt
+ ((Scan.repeat scan_verb >> flat) -- ($$$ "*" |-- $$$ "}" |-- Symbol_Pos.scan_pos)));
+
+
+(* scan space *)
+
+fun is_space s = Symbol.is_blank s andalso s <> "\n";
+
+val scan_space =
+ Scan.many1 (is_space o symbol) @@@ Scan.optional ($$$ "\n") [] ||
+ Scan.many (is_space o symbol) @@@ $$$ "\n";
+
+
+(* scan comment *)
+
+val scan_comment =
+ Symbol_Pos.scan_pos -- (Symbol_Pos.scan_comment_body !!! -- Symbol_Pos.scan_pos);
+
+
+(* scan malformed symbols *)
+
+val scan_malformed =
+ $$$ Symbol.malformed |--
+ Symbol_Pos.change_prompt (Scan.many (Symbol.is_regular o symbol))
+ --| Scan.option ($$$ Symbol.end_malformed);
+
+
+
+(** token sources **)
+
+fun source_proper src = src |> Source.filter is_proper;
+
+local
+
+fun token_leq ((_, syms1), (_, syms2)) = length syms1 <= length syms2;
+
+fun token k ss =
+ Token ((Symbol_Pos.implode ss, Symbol_Pos.range ss), (k, Symbol_Pos.untabify_content ss), Slot);
+
+fun token_range k (pos1, (ss, pos2)) =
+ Token (Symbol_Pos.implode_range pos1 pos2 ss, (k, Symbol_Pos.untabify_content ss), Slot);
+
+fun scan (lex1, lex2) = !!! "bad input"
+ (Symbol_Pos.scan_string >> token_range String ||
+ Symbol_Pos.scan_alt_string >> token_range AltString ||
+ scan_verbatim >> token_range Verbatim ||
+ scan_comment >> token_range Comment ||
+ scan_space >> token Space ||
+ scan_malformed >> token Malformed ||
+ Scan.one (Symbol.is_sync o symbol) >> (token Sync o single) ||
+ (Scan.max token_leq
+ (Scan.max token_leq
+ (Scan.literal lex2 >> pair Command)
+ (Scan.literal lex1 >> pair Keyword))
+ (Syntax.scan_longid >> pair LongIdent ||
+ Syntax.scan_id >> pair Ident ||
+ Syntax.scan_var >> pair Var ||
+ Syntax.scan_tid >> pair TypeIdent ||
+ Syntax.scan_tvar >> pair TypeVar ||
+ Syntax.scan_nat >> pair Nat ||
+ scan_symid >> pair SymIdent) >> uncurry token));
+
+fun recover msg =
+ Scan.many ((Symbol.is_regular andf (not o Symbol.is_blank)) o symbol)
+ >> (single o token (Error msg));
+
+in
+
+fun source' {do_recover} get_lex =
+ Source.source Symbol_Pos.stopper (Scan.bulk (fn xs => scan (get_lex ()) xs))
+ (Option.map (rpair recover) do_recover);
+
+fun source do_recover get_lex pos src =
+ Symbol_Pos.source pos src
+ |> source' do_recover get_lex;
+
+end;
+
+
+(* read_antiq *)
+
+fun read_antiq lex scan (syms, pos) =
+ let
+ fun err msg = cat_error msg ("Malformed antiquotation" ^ Position.str_of pos ^ ":\n" ^
+ "@{" ^ Symbol_Pos.content syms ^ "}");
+
+ val res =
+ Source.of_list syms
+ |> source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
+ |> source_proper
+ |> Source.source stopper (Scan.error (Scan.bulk scan)) NONE
+ |> Source.exhaust;
+ in (case res of [x] => x | _ => err "") handle ERROR msg => err msg end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/token.scala Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,95 @@
+/* Title: Pure/Isar/token.scala
+ Author: Makarius
+
+Outer token syntax for Isabelle/Isar.
+*/
+
+package isabelle
+
+
+object Token
+{
+ /* tokens */
+
+ object Kind extends Enumeration
+ {
+ val COMMAND = Value("command")
+ val KEYWORD = Value("keyword")
+ val IDENT = Value("identifier")
+ val LONG_IDENT = Value("long identifier")
+ val SYM_IDENT = Value("symbolic identifier")
+ val VAR = Value("schematic variable")
+ val TYPE_IDENT = Value("type variable")
+ val TYPE_VAR = Value("schematic type variable")
+ val NAT = Value("number")
+ val STRING = Value("string")
+ val ALT_STRING = Value("back-quoted string")
+ val VERBATIM = Value("verbatim text")
+ val SPACE = Value("white space")
+ val COMMENT = Value("comment text")
+ val BAD_INPUT = Value("bad input")
+ val UNPARSED = Value("unparsed input")
+ }
+
+
+ /* token reader */
+
+ class Line_Position(val line: Int) extends scala.util.parsing.input.Position
+ {
+ def column = 0
+ def lineContents = ""
+ override def toString = line.toString
+
+ def advance(token: Token): Line_Position =
+ {
+ var n = 0
+ for (c <- token.content if c == '\n') n += 1
+ if (n == 0) this else new Line_Position(line + n)
+ }
+ }
+
+ abstract class Reader extends scala.util.parsing.input.Reader[Token]
+
+ private class Token_Reader(tokens: List[Token], val pos: Line_Position) extends Reader
+ {
+ def first = tokens.head
+ def rest = new Token_Reader(tokens.tail, pos.advance(first))
+ def atEnd = tokens.isEmpty
+ }
+
+ def reader(tokens: List[Token]): Reader = new Token_Reader(tokens, new Line_Position(1))
+}
+
+
+sealed case class Token(val kind: Token.Kind.Value, val source: String)
+{
+ def is_command: Boolean = kind == Token.Kind.COMMAND
+ def is_delimited: Boolean =
+ kind == Token.Kind.STRING ||
+ kind == Token.Kind.ALT_STRING ||
+ kind == Token.Kind.VERBATIM ||
+ kind == Token.Kind.COMMENT
+ def is_name: Boolean =
+ kind == Token.Kind.IDENT ||
+ kind == Token.Kind.SYM_IDENT ||
+ kind == Token.Kind.STRING ||
+ kind == Token.Kind.NAT
+ def is_xname: Boolean = is_name || kind == Token.Kind.LONG_IDENT
+ def is_text: Boolean = is_xname || kind == Token.Kind.VERBATIM
+ def is_space: Boolean = kind == Token.Kind.SPACE
+ def is_comment: Boolean = kind == Token.Kind.COMMENT
+ def is_ignored: Boolean = is_space || is_comment
+ def is_unparsed: Boolean = kind == Token.Kind.UNPARSED
+
+ def content: String =
+ if (kind == Token.Kind.STRING) Scan.Lexicon.empty.quoted_content("\"", source)
+ else if (kind == Token.Kind.ALT_STRING) Scan.Lexicon.empty.quoted_content("`", source)
+ else if (kind == Token.Kind.VERBATIM) Scan.Lexicon.empty.verbatim_content(source)
+ else if (kind == Token.Kind.COMMENT) Scan.Lexicon.empty.comment_content(source)
+ else source
+
+ def text: (String, String) =
+ if (kind == Token.Kind.COMMAND && source == ";") ("terminator", "")
+ else (kind.toString, source)
+}
+
--- a/src/Pure/Isar/toplevel.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Isar/toplevel.ML Sun May 23 10:38:11 2010 +0100
@@ -660,8 +660,8 @@
let val st' = command tr st in
if immediate orelse
null proof_trs orelse
- OuterKeyword.is_schematic_goal (name_of tr) orelse
- exists (OuterKeyword.is_qed_global o name_of) proof_trs orelse
+ Keyword.is_schematic_goal (name_of tr) orelse
+ exists (Keyword.is_qed_global o name_of) proof_trs orelse
not (can proof_of st') orelse
Proof.is_relevant (proof_of st')
then
--- a/src/Pure/Isar/value_parse.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-(* Title: Pure/Isar/value_parse.ML
- Author: Makarius
-
-Outer syntax parsers for basic ML values.
-*)
-
-signature VALUE_PARSE =
-sig
- val comma: 'a parser -> 'a parser
- val equal: 'a parser -> 'a parser
- val parens: 'a parser -> 'a parser
- val unit: unit parser
- val pair: 'a parser -> 'b parser -> ('a * 'b) parser
- val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
- val list: 'a parser -> 'a list parser
- val properties: Properties.T parser
-end;
-
-structure ValueParse: VALUE_PARSE =
-struct
-
-structure P = OuterParse;
-
-
-(* syntax utilities *)
-
-fun comma p = P.$$$ "," |-- P.!!! p;
-fun equal p = P.$$$ "=" |-- P.!!! p;
-fun parens p = P.$$$ "(" |-- P.!!! (p --| P.$$$ ")");
-
-
-(* tuples *)
-
-val unit = parens (Scan.succeed ());
-fun pair p1 p2 = parens (p1 -- comma p2);
-fun triple p1 p2 p3 = parens (p1 -- comma p2 -- comma p3) >> P.triple1;
-
-
-(* lists *)
-
-fun list p = parens (P.enum "," p);
-val properties = list (P.string -- equal P.string);
-
-end;
-
--- a/src/Pure/ML-Systems/polyml_common.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ML-Systems/polyml_common.ML Sun May 23 10:38:11 2010 +0100
@@ -6,7 +6,11 @@
exception Interrupt = SML90.Interrupt;
use "General/exn.ML";
-use "ML-Systems/single_assignment_polyml.ML";
+
+if List.exists (fn s => s = "SingleAssignment") (PolyML.Compiler.structureNames ())
+then ()
+else use "ML-Systems/single_assignment_polyml.ML";
+
use "ML-Systems/multithreading.ML";
use "ML-Systems/time_limit.ML";
use "ML-Systems/timing.ML";
--- a/src/Pure/ML/ml_antiquote.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ML/ml_antiquote.ML Sun May 23 10:38:11 2010 +0100
@@ -57,12 +57,11 @@
(** misc antiquotations **)
-structure P = OuterParse;
-
val _ = inline "make_string" (Scan.succeed ml_make_string);
val _ = value "binding"
- (Scan.lift (P.position Args.name) >> (fn name => ML_Syntax.atomic (ML_Syntax.make_binding name)));
+ (Scan.lift (Parse.position Args.name)
+ >> (fn name => ML_Syntax.atomic (ML_Syntax.make_binding name)));
val _ = value "theory"
(Scan.lift Args.name >> (fn name => "ThyInfo.get_theory " ^ ML_Syntax.print_string name)
@@ -80,11 +79,12 @@
val _ = inline "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term));
val _ = macro "let" (Args.context --
- Scan.lift (P.and_list1 (P.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source)))
+ Scan.lift
+ (Parse.and_list1 (Parse.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source)))
>> (fn (ctxt, args) => #2 (ProofContext.match_bind true args ctxt)));
val _ = macro "note" (Args.context :|-- (fn ctxt =>
- P.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms >> (fn ((a, srcs), ths) =>
+ Parse.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms >> (fn ((a, srcs), ths) =>
((a, map (Attrib.attribute (ProofContext.theory_of ctxt)) srcs), [(ths, [])])))
>> (fn args => #2 (ProofContext.note_thmss "" args ctxt))));
@@ -118,7 +118,7 @@
(* type constructors *)
-fun type_name kind check = Args.context -- Scan.lift (OuterParse.position Args.name_source)
+fun type_name kind check = Args.context -- Scan.lift (Parse.position Args.name_source)
>> (fn (ctxt, (s, pos)) =>
let
val Type (c, _) = ProofContext.read_type_name_proper ctxt false s;
@@ -137,7 +137,7 @@
(* constants *)
-fun const_name check = Args.context -- Scan.lift (OuterParse.position Args.name_source)
+fun const_name check = Args.context -- Scan.lift (Parse.position Args.name_source)
>> (fn (ctxt, (s, pos)) =>
let
val Const (c, _) = ProofContext.read_const_proper ctxt false s;
@@ -151,13 +151,13 @@
val _ = inline "syntax_const"
- (Args.context -- Scan.lift (OuterParse.position Args.name) >> (fn (ctxt, (c, pos)) =>
+ (Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (c, pos)) =>
if Syntax.is_const (ProofContext.syn_of ctxt) c then ML_Syntax.print_string c
else error ("Unknown syntax const: " ^ quote c ^ Position.str_of pos)));
val _ = inline "const"
(Args.context -- Scan.lift Args.name_source -- Scan.optional
- (Scan.lift (Args.$$$ "(") |-- OuterParse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
+ (Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
>> (fn ((ctxt, raw_c), Ts) =>
let
val Const (c, _) = ProofContext.read_const_proper ctxt true raw_c;
--- a/src/Pure/ML/ml_context.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ML/ml_context.ML Sun May 23 10:38:11 2010 +0100
@@ -112,11 +112,8 @@
local
-structure P = OuterParse;
-structure T = OuterLex;
-
val antiq =
- P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof)
+ Parse.!!! (Parse.position Parse.xname -- Args.parse --| Scan.ahead Parse.eof)
>> (fn ((x, pos), y) => Args.src ((x, y), pos));
val begin_env = ML_Lex.tokenize "structure Isabelle =\nstruct\n";
@@ -138,14 +135,15 @@
NONE => error ("No context -- cannot expand ML antiquotations" ^ Position.str_of pos)
| SOME ctxt => Context.proof_of ctxt);
- val lex = #1 (OuterKeyword.get_lexicons ());
+ val lex = #1 (Keyword.get_lexicons ());
fun no_decl _ = ([], []);
fun expand (Antiquote.Text tok) state = (K ([], [tok]), state)
| expand (Antiquote.Antiq (ss, range)) (scope, background) =
let
val context = Stack.top scope;
- val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context;
+ val (f, context') =
+ antiquotation (Token.read_antiq lex antiq (ss, #1 range)) context;
val (decl, background') = f background;
val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range));
in (decl', (Stack.map_top (K context') scope, background')) end
--- a/src/Pure/ML/ml_thms.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ML/ml_thms.ML Sun May 23 10:38:11 2010 +0100
@@ -54,7 +54,7 @@
val _ = ML_Context.add_antiq "lemma"
(fn _ => Args.context -- Args.mode "open" --
- Scan.lift (OuterParse.and_list1 (Scan.repeat1 goal) --
+ Scan.lift (Parse.and_list1 (Scan.repeat1 goal) --
(by |-- Method.parse -- Scan.option Method.parse)) >>
(fn ((ctxt, is_open), (raw_propss, methods)) => fn background =>
let
--- a/src/Pure/PIDE/command.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/PIDE/command.scala Sun May 23 10:38:11 2010 +0100
@@ -20,6 +20,7 @@
val UNPROCESSED = Value("UNPROCESSED")
val FINISHED = Value("FINISHED")
val FAILED = Value("FAILED")
+ val UNDEFINED = Value("UNDEFINED")
}
case class HighlightInfo(highlight: String) { override def toString = highlight }
--- a/src/Pure/PIDE/document.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/PIDE/document.scala Sun May 23 10:38:11 2010 +0100
@@ -46,7 +46,7 @@
/* unparsed dummy commands */
def unparsed(source: String) =
- new Command(null, List(Outer_Lex.Token(Outer_Lex.Token_Kind.UNPARSED, source)))
+ new Command(null, List(Token(Token.Kind.UNPARSED, source)))
def is_unparsed(command: Command) = command.id == null
@@ -175,9 +175,12 @@
System.err.println("assign_states: " + (System.currentTimeMillis - time0) + " ms elapsed time")
}
- def current_state(cmd: Command): Option[State] =
+ def current_state(cmd: Command): State =
{
require(assignment.is_finished)
- (assignment.join).get(cmd).map(_.current_state)
+ (assignment.join).get(cmd) match {
+ case Some(cmd_state) => cmd_state.current_state
+ case None => new State(cmd, Command.Status.UNDEFINED, Nil, cmd.empty_markup)
+ }
}
}
--- a/src/Pure/PIDE/state.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/PIDE/state.scala Sun May 23 10:38:11 2010 +0100
@@ -74,7 +74,7 @@
{
val changed: State =
message match {
- case XML.Elem(Markup.MESSAGE, (Markup.CLASS, Markup.STATUS) :: _, elems) =>
+ case XML.Elem(Markup.STATUS, _, elems) =>
(this /: elems)((state, elem) =>
elem match {
case XML.Elem(Markup.UNPROCESSED, _, _) => state.set_status(Command.Status.UNPROCESSED)
--- a/src/Pure/Proof/extraction.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Proof/extraction.ML Sun May 23 10:38:11 2010 +0100
@@ -750,31 +750,29 @@
(**** interface ****)
-structure P = OuterParse and K = OuterKeyword;
-
-val parse_vars = Scan.optional (P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") [];
+val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") [];
val _ =
- OuterSyntax.command "realizers"
+ Outer_Syntax.command "realizers"
"specify realizers for primitive axioms / theorems, together with correctness proof"
- K.thy_decl
- (Scan.repeat1 (P.xname -- parse_vars --| P.$$$ ":" -- P.string -- P.string) >>
+ Keyword.thy_decl
+ (Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >>
(fn xs => Toplevel.theory (fn thy => add_realizers
(map (fn (((a, vs), s1), s2) => (PureThy.get_thm thy a, (vs, s1, s2))) xs) thy)));
val _ =
- OuterSyntax.command "realizability"
- "add equations characterizing realizability" K.thy_decl
- (Scan.repeat1 P.string >> (Toplevel.theory o add_realizes_eqns));
+ Outer_Syntax.command "realizability"
+ "add equations characterizing realizability" Keyword.thy_decl
+ (Scan.repeat1 Parse.string >> (Toplevel.theory o add_realizes_eqns));
val _ =
- OuterSyntax.command "extract_type"
- "add equations characterizing type of extracted program" K.thy_decl
- (Scan.repeat1 P.string >> (Toplevel.theory o add_typeof_eqns));
+ Outer_Syntax.command "extract_type"
+ "add equations characterizing type of extracted program" Keyword.thy_decl
+ (Scan.repeat1 Parse.string >> (Toplevel.theory o add_typeof_eqns));
val _ =
- OuterSyntax.command "extract" "extract terms from proofs" K.thy_decl
- (Scan.repeat1 (P.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>
+ Outer_Syntax.command "extract" "extract terms from proofs" Keyword.thy_decl
+ (Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>
extract (map (apfst (PureThy.get_thm thy)) xs) thy)));
val etype_of = etype_of o add_syntax;
--- a/src/Pure/ProofGeneral/pgip_parser.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ProofGeneral/pgip_parser.ML Sun May 23 10:38:11 2010 +0100
@@ -12,7 +12,6 @@
structure PgipParser: PGIP_PARSER =
struct
-structure K = OuterKeyword;
structure D = PgipMarkup;
structure I = PgipIsabelle;
@@ -51,42 +50,42 @@
fun closegoal text = [D.Closegoal {text = text}, D.Closeblock {}];
-fun command k f = Symtab.update_new (OuterKeyword.kind_of k, f);
+fun command k f = Symtab.update_new (Keyword.kind_of k, f);
val command_keywords = Symtab.empty
- |> command K.control badcmd
- |> command K.diag (fn text => [D.Spuriouscmd {text = text}])
- |> command K.thy_begin thy_begin
- |> command K.thy_switch badcmd
- |> command K.thy_end (fn text => [D.Closeblock {}, D.Closetheory {text = text}])
- |> command K.thy_heading thy_heading
- |> command K.thy_decl thy_decl
- |> command K.thy_script thy_decl
- |> command K.thy_goal goal
- |> command K.thy_schematic_goal goal
- |> command K.qed closegoal
- |> command K.qed_block closegoal
- |> command K.qed_global (fn text => [D.Giveupgoal {text = text}])
- |> command K.prf_heading (fn text => [D.Doccomment {text = text}])
- |> command K.prf_goal goal
- |> command K.prf_block prf_block
- |> command K.prf_open prf_open
- |> command K.prf_close (fn text => [D.Proofstep {text = text}, D.Closeblock {}])
- |> command K.prf_chain proofstep
- |> command K.prf_decl proofstep
- |> command K.prf_asm proofstep
- |> command K.prf_asm_goal goal
- |> command K.prf_script proofstep;
+ |> command Keyword.control badcmd
+ |> command Keyword.diag (fn text => [D.Spuriouscmd {text = text}])
+ |> command Keyword.thy_begin thy_begin
+ |> command Keyword.thy_switch badcmd
+ |> command Keyword.thy_end (fn text => [D.Closeblock {}, D.Closetheory {text = text}])
+ |> command Keyword.thy_heading thy_heading
+ |> command Keyword.thy_decl thy_decl
+ |> command Keyword.thy_script thy_decl
+ |> command Keyword.thy_goal goal
+ |> command Keyword.thy_schematic_goal goal
+ |> command Keyword.qed closegoal
+ |> command Keyword.qed_block closegoal
+ |> command Keyword.qed_global (fn text => [D.Giveupgoal {text = text}])
+ |> command Keyword.prf_heading (fn text => [D.Doccomment {text = text}])
+ |> command Keyword.prf_goal goal
+ |> command Keyword.prf_block prf_block
+ |> command Keyword.prf_open prf_open
+ |> command Keyword.prf_close (fn text => [D.Proofstep {text = text}, D.Closeblock {}])
+ |> command Keyword.prf_chain proofstep
+ |> command Keyword.prf_decl proofstep
+ |> command Keyword.prf_asm proofstep
+ |> command Keyword.prf_asm_goal goal
+ |> command Keyword.prf_script proofstep;
-val _ = subset (op =) (OuterKeyword.kinds, Symtab.keys command_keywords)
+val _ = subset (op =) (Keyword.kinds, Symtab.keys command_keywords)
orelse sys_error "Incomplete coverage of command keywords";
fun parse_command "sorry" text = [D.Postponegoal {text = text}, D.Closeblock {}]
| parse_command name text =
- (case OuterKeyword.command_keyword name of
+ (case Keyword.command_keyword name of
NONE => [D.Unparseable {text = text}]
| SOME k =>
- (case Symtab.lookup command_keywords (OuterKeyword.kind_of k) of
+ (case Symtab.lookup command_keywords (Keyword.kind_of k) of
NONE => [D.Unparseable {text = text}]
| SOME f => f text));
@@ -104,6 +103,6 @@
fun pgip_parser pos str =
- maps parse_span (ThySyntax.parse_spans (OuterKeyword.get_lexicons ()) pos str);
+ maps parse_span (ThySyntax.parse_spans (Keyword.get_lexicons ()) pos str);
end;
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML Sun May 23 10:38:11 2010 +0100
@@ -188,48 +188,44 @@
(* additional outer syntax for Isar *)
-local structure P = OuterParse and K = OuterKeyword in
-
fun prP () =
- OuterSyntax.improper_command "ProofGeneral.pr" "print state (internal)" K.diag
+ Outer_Syntax.improper_command "ProofGeneral.pr" "print state (internal)" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
if Toplevel.is_toplevel state orelse Toplevel.is_theory state then tell_clear_goals ()
else (Toplevel.quiet := false; Toplevel.print_state true state))));
fun undoP () = (*undo without output -- historical*)
- OuterSyntax.improper_command "ProofGeneral.undo" "(internal)" K.control
+ Outer_Syntax.improper_command "ProofGeneral.undo" "(internal)" Keyword.control
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1)));
fun restartP () =
- OuterSyntax.improper_command "ProofGeneral.restart" "(internal)" K.control
- (P.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart)));
+ Outer_Syntax.improper_command "ProofGeneral.restart" "(internal)" Keyword.control
+ (Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart)));
fun kill_proofP () =
- OuterSyntax.improper_command "ProofGeneral.kill_proof" "(internal)" K.control
+ Outer_Syntax.improper_command "ProofGeneral.kill_proof" "(internal)" Keyword.control
(Scan.succeed (Toplevel.no_timing o
Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));
fun inform_file_processedP () =
- OuterSyntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" K.control
- (P.name >> (fn file =>
+ Outer_Syntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" Keyword.control
+ (Parse.name >> (fn file =>
Toplevel.no_timing o Toplevel.imperative (fn () => inform_file_processed file)));
fun inform_file_retractedP () =
- OuterSyntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" K.control
- (P.name >> (Toplevel.no_timing oo
+ Outer_Syntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" Keyword.control
+ (Parse.name >> (Toplevel.no_timing oo
(fn file => Toplevel.imperative (fn () => inform_file_retracted file))));
fun process_pgipP () =
- OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" K.control
- (P.text >> (Toplevel.no_timing oo
+ Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
+ (Parse.text >> (Toplevel.no_timing oo
(fn txt => Toplevel.imperative (fn () => ProofGeneralPgip.process_pgip txt))));
fun init_outer_syntax () = List.app (fn f => f ())
[prP, undoP, restartP, kill_proofP, inform_file_processedP,
inform_file_retractedP, process_pgipP];
-end;
-
(* init *)
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML Sun May 23 10:38:11 2010 +0100
@@ -312,8 +312,8 @@
(** lexicalstructure element with keywords (PGIP version of elisp keywords file) **)
fun lexicalstructure_keywords () =
- let val keywords = OuterKeyword.dest_keywords ()
- val commands = OuterKeyword.dest_commands ()
+ let val keywords = Keyword.dest_keywords ()
+ val commands = Keyword.dest_commands ()
fun keyword_elt kind keyword =
XML.Elem("keyword", [("word", keyword), ("category", kind)], [])
in
@@ -382,7 +382,7 @@
(* Sending commands to Isar *)
-fun isarcmd s = Isar.>>> (OuterSyntax.parse Position.none s);
+fun isarcmd s = Isar.>>> (Outer_Syntax.parse Position.none s);
(* TODO:
- apply a command given a transition function;
@@ -1013,8 +1013,8 @@
(* Extra command for embedding prover-control inside document (obscure/debug usage). *)
fun init_outer_syntax () =
- OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" OuterKeyword.control
- (OuterParse.text >> (Toplevel.no_timing oo
+ Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
+ (Parse.text >> (Toplevel.no_timing oo
(fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));
--- a/src/Pure/ROOT.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/ROOT.ML Sun May 23 10:38:11 2010 +0100
@@ -167,10 +167,10 @@
use "Proof/proofchecker.ML";
(*outer syntax*)
-use "Isar/outer_lex.ML";
-use "Isar/outer_keyword.ML";
-use "Isar/outer_parse.ML";
-use "Isar/value_parse.ML";
+use "Isar/token.ML";
+use "Isar/keyword.ML";
+use "Isar/parse.ML";
+use "Isar/parse_value.ML";
use "Isar/args.ML";
(*ML support*)
@@ -220,7 +220,7 @@
use "Isar/code.ML";
(*specifications*)
-use "Isar/spec_parse.ML";
+use "Isar/parse_spec.ML";
use "Isar/spec_rules.ML";
use "Isar/specification.ML";
use "Isar/typedecl.ML";
@@ -292,3 +292,23 @@
use "pure_setup.ML";
+
+(* legacy aliases *)
+
+structure Legacy =
+struct
+
+structure OuterKeyword = Keyword;
+
+structure OuterLex =
+struct
+ open Token;
+ type token = T;
+end;
+
+structure OuterParse = Parse;
+structure OuterSyntax = Outer_Syntax;
+structure SpecParse = Parse_Spec;
+
+end;
+
--- a/src/Pure/System/gui_setup.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/gui_setup.scala Sun May 23 10:38:11 2010 +0100
@@ -6,8 +6,8 @@
package isabelle
-import scala.swing._
-import scala.swing.event._
+import scala.swing.{Button, FlowPanel, BorderPanel, MainFrame, TextArea, SwingApplication}
+import scala.swing.event.ButtonClicked
object GUI_Setup extends SwingApplication
@@ -27,16 +27,14 @@
editable = false
columns = 80
rows = 20
- xLayoutAlignment = 0.5
}
- val ok = new Button {
- text = "OK"
- xLayoutAlignment = 0.5
- }
- contents = new BoxPanel(Orientation.Vertical) {
- contents += text
- contents += ok
- }
+ val ok = new Button { text = "OK" }
+ val ok_panel = new FlowPanel(FlowPanel.Alignment.Center)(ok)
+
+ val panel = new BorderPanel
+ panel.layout(text) = BorderPanel.Position.Center
+ panel.layout(ok_panel) = BorderPanel.Position.South
+ contents = panel
// values
if (Platform.is_windows)
--- a/src/Pure/System/isabelle_process.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/isabelle_process.ML Sun May 23 10:38:11 2010 +0100
@@ -89,9 +89,9 @@
fun init out =
(Unsynchronized.change print_mode
- (fold (update op =) [isabelle_processN, OuterKeyword.keyword_status_reportN, Pretty.symbolicN]);
+ (fold (update op =) [isabelle_processN, Keyword.keyword_status_reportN, Pretty.symbolicN]);
setup_channels out |> init_message;
- OuterKeyword.report ();
+ Keyword.report ();
Output.status (Markup.markup Markup.ready "");
Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true});
--- a/src/Pure/System/isabelle_process.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/isabelle_process.scala Sun May 23 10:38:11 2010 +0100
@@ -84,7 +84,7 @@
class Result(val kind: Kind.Value, val props: List[(String, String)], val body: List[XML.Tree])
{
- def message = XML.Elem(Markup.MESSAGE, (Markup.CLASS, Kind.markup(kind)) :: props, body)
+ def message = XML.Elem(Kind.markup(kind), props, body)
override def toString: String =
{
--- a/src/Pure/System/isabelle_system.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/isabelle_system.scala Sun May 23 10:38:11 2010 +0100
@@ -18,10 +18,20 @@
import scala.collection.mutable
-class Isabelle_System extends Standard_System
+class Isabelle_System(this_isabelle_home: String) extends Standard_System
{
+ def this() = this(null)
+
+
/** Isabelle environment **/
+ /*
+ isabelle_home precedence:
+ (1) this_isabelle_home as explicit argument
+ (2) ISABELLE_HOME process environment variable (e.g. inherited from running isabelle tool)
+ (3) isabelle.home system property (e.g. via JVM application boot process)
+ */
+
private val environment: Map[String, String] =
{
import scala.collection.JavaConversions._
@@ -30,13 +40,15 @@
("THIS_JAVA" -> this_java())
val isabelle_home =
- env0.get("ISABELLE_HOME") match {
- case None | Some("") =>
- val path = java.lang.System.getProperty("isabelle.home")
- if (path == null || path == "") error("Unknown Isabelle home directory")
- else path
- case Some(path) => path
- }
+ if (this_isabelle_home != null) this_isabelle_home
+ else
+ env0.get("ISABELLE_HOME") match {
+ case None | Some("") =>
+ val path = java.lang.System.getProperty("isabelle.home")
+ if (path == null || path == "") error("Unknown Isabelle home directory")
+ else path
+ case Some(path) => path
+ }
Standard_System.with_tmp_file("settings") { dump =>
val shell_prefix =
@@ -150,6 +162,15 @@
def platform_file(path: String) = new File(platform_path(path))
+ /* try_read */
+
+ def try_read(path: String): String =
+ {
+ val file = platform_file(path)
+ if (file.isFile) Source.fromFile(file).mkString else ""
+ }
+
+
/* source files */
private def try_file(file: File) = if (file.isFile) Some(file) else None
@@ -304,11 +325,8 @@
/* symbols */
private def read_symbols(path: String): List[String] =
- {
- val file = platform_file(path)
- if (file.isFile) Source.fromFile(file).getLines("\n").toList
- else Nil
- }
+ Library.chunks(try_read(path)).map(_.toString).toList
+
val symbols = new Symbol.Interpretation(
read_symbols("$ISABELLE_HOME/etc/symbols") :::
read_symbols("$ISABELLE_HOME_USER/etc/symbols"))
--- a/src/Pure/System/isar.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/isar.ML Sun May 23 10:38:11 2010 +0100
@@ -80,14 +80,14 @@
fun linear_undo n = edit_history n (K (find_and_undo (K true)));
fun undo n = edit_history n (fn st => fn hist =>
- find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist);
+ find_and_undo (if Toplevel.is_proof st then K true else Keyword.is_theory) hist);
fun kill () = edit_history 1 (fn st => fn hist =>
find_and_undo
- (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist);
+ (if Toplevel.is_proof st then Keyword.is_theory else Keyword.is_theory_begin) hist);
fun kill_proof () = edit_history 1 (fn st => fn hist =>
- if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist
+ if Toplevel.is_proof st then find_and_undo Keyword.is_theory hist
else raise Toplevel.UNDEF);
end;
@@ -102,9 +102,9 @@
| SOME (st', NONE) =>
let
val name = Toplevel.name_of tr;
- val _ = if OuterKeyword.is_theory_begin name then init () else ();
+ val _ = if Keyword.is_theory_begin name then init () else ();
val _ =
- if OuterKeyword.is_regular name
+ if Keyword.is_regular name
then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else ();
in true end);
@@ -141,7 +141,7 @@
Secure.open_unsynchronized ();
if do_init then init () else ();
if welcome then writeln (Session.welcome ()) else ();
- uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
+ uninterruptible (fn _ => fn () => raw_loop secure (Outer_Syntax.isar sync)) ());
end;
@@ -157,7 +157,6 @@
local
-structure P = OuterParse and K = OuterKeyword;
val op >> = Scan.>>;
in
@@ -165,33 +164,35 @@
(* global history *)
val _ =
- OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" K.control
+ Outer_Syntax.improper_command "init_toplevel" "init toplevel point-of-interest" Keyword.control
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative init));
val _ =
- OuterSyntax.improper_command "linear_undo" "undo commands" K.control
- (Scan.optional P.nat 1 >>
+ Outer_Syntax.improper_command "linear_undo" "undo commands" Keyword.control
+ (Scan.optional Parse.nat 1 >>
(fn n => Toplevel.no_timing o Toplevel.imperative (fn () => linear_undo n)));
val _ =
- OuterSyntax.improper_command "undo" "undo commands (skipping closed proofs)" K.control
- (Scan.optional P.nat 1 >>
+ Outer_Syntax.improper_command "undo" "undo commands (skipping closed proofs)" Keyword.control
+ (Scan.optional Parse.nat 1 >>
(fn n => Toplevel.no_timing o Toplevel.imperative (fn () => undo n)));
val _ =
- OuterSyntax.improper_command "undos_proof" "undo commands (skipping closed proofs)" K.control
- (Scan.optional P.nat 1 >> (fn n => Toplevel.no_timing o
+ Outer_Syntax.improper_command "undos_proof" "undo commands (skipping closed proofs)"
+ Keyword.control
+ (Scan.optional Parse.nat 1 >> (fn n => Toplevel.no_timing o
Toplevel.keep (fn state =>
if Toplevel.is_proof state then (undo n; print ()) else raise Toplevel.UNDEF)));
val _ =
- OuterSyntax.improper_command "cannot_undo" "partial undo -- Proof General legacy" K.control
- (P.name >>
+ Outer_Syntax.improper_command "cannot_undo" "partial undo -- Proof General legacy"
+ Keyword.control
+ (Parse.name >>
(fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => undo 1)
| txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt))));
val _ =
- OuterSyntax.improper_command "kill" "kill partial proof or theory development" K.control
+ Outer_Syntax.improper_command "kill" "kill partial proof or theory development" Keyword.control
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
end;
--- a/src/Pure/System/session.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/session.ML Sun May 23 10:38:11 2010 +0100
@@ -48,7 +48,7 @@
(if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else "");
val _ =
- OuterSyntax.improper_command "welcome" "print welcome message" OuterKeyword.diag
+ Outer_Syntax.improper_command "welcome" "print welcome message" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o welcome)));
--- a/src/Pure/System/session.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/System/session.scala Sun May 23 10:38:11 2010 +0100
@@ -104,7 +104,7 @@
def bad_result(result: Isabelle_Process.Result)
{
- System.err.println("Ignoring prover result: " + result)
+ System.err.println("Ignoring prover result: " + result.message.toString)
}
def handle_result(result: Isabelle_Process.Result)
@@ -140,8 +140,8 @@
}
// keyword declarations
- case List(Outer_Keyword.Command_Decl(name, kind)) => syntax += (name, kind)
- case List(Outer_Keyword.Keyword_Decl(name)) => syntax += name
+ case List(Keyword.Command_Decl(name, kind)) => syntax += (name, kind)
+ case List(Keyword.Keyword_Decl(name)) => syntax += name
case _ => if (!result.is_ready) bad_result(result)
}
--- a/src/Pure/Thy/latex.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/latex.ML Sun May 23 10:38:11 2010 +0100
@@ -9,7 +9,7 @@
val output_known_symbols: (string -> bool) * (string -> bool) ->
Symbol.symbol list -> string
val output_symbols: Symbol.symbol list -> string
- val output_basic: OuterLex.token -> string
+ val output_basic: Token.T -> string
val output_markup: string -> string -> string
val output_markup_env: string -> string -> string
val output_verbatim: string -> string
@@ -99,24 +99,22 @@
(* token output *)
-structure T = OuterLex;
-
-val invisible_token = T.keyword_with (fn s => s = ";") orf T.is_kind T.Comment;
+val invisible_token = Token.keyword_with (fn s => s = ";") orf Token.is_kind Token.Comment;
fun output_basic tok =
- let val s = T.content_of tok in
+ let val s = Token.content_of tok in
if invisible_token tok then ""
- else if T.is_kind T.Command tok then
+ else if Token.is_kind Token.Command tok then
"\\isacommand{" ^ output_syms s ^ "}"
- else if T.is_kind T.Keyword tok andalso Syntax.is_ascii_identifier s then
+ else if Token.is_kind Token.Keyword tok andalso Syntax.is_ascii_identifier s then
"\\isakeyword{" ^ output_syms s ^ "}"
- else if T.is_kind T.String tok then
+ else if Token.is_kind Token.String tok then
enclose "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}" (output_syms s)
- else if T.is_kind T.AltString tok then
+ else if Token.is_kind Token.AltString tok then
enclose "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}" (output_syms s)
- else if T.is_kind T.Verbatim tok then
+ else if Token.is_kind Token.Verbatim tok then
let
- val (txt, pos) = T.source_position_of tok;
+ val (txt, pos) = Token.source_position_of tok;
val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);
val out = implode (map output_syms_antiq ants);
in enclose "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" out end
--- a/src/Pure/Thy/term_style.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/term_style.ML Sun May 23 10:38:11 2010 +0100
@@ -43,7 +43,7 @@
(* style parsing *)
-fun parse_single ctxt = OuterParse.position (OuterParse.xname -- Args.parse)
+fun parse_single ctxt = Parse.position (Parse.xname -- Args.parse)
>> (fn x as ((name, _), _) => fst (Args.context_syntax "style"
(Scan.lift (the_style (ProofContext.theory_of ctxt) name))
(Args.src x) ctxt |>> (fn f => f ctxt)));
--- a/src/Pure/Thy/thy_header.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_header.ML Sun May 23 10:38:11 2010 +0100
@@ -6,18 +6,13 @@
signature THY_HEADER =
sig
- val args: OuterLex.token list ->
- (string * string list * (string * bool) list) * OuterLex.token list
+ val args: Token.T list -> (string * string list * (string * bool) list) * Token.T list
val read: Position.T -> (string, 'a) Source.source -> string * string list * (string * bool) list
end;
structure Thy_Header: THY_HEADER =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-
-
(* keywords *)
val headerN = "header";
@@ -32,31 +27,36 @@
(* header args *)
-val file_name = P.group "file name" P.name;
-val theory_name = P.group "theory name" P.name;
+val file_name = Parse.group "file name" Parse.name;
+val theory_name = Parse.group "theory name" Parse.name;
-val file = P.$$$ "(" |-- P.!!! (file_name --| P.$$$ ")") >> rpair false || file_name >> rpair true;
-val uses = Scan.optional (P.$$$ usesN |-- P.!!! (Scan.repeat1 file)) [];
+val file =
+ Parse.$$$ "(" |-- Parse.!!! (file_name --| Parse.$$$ ")") >> rpair false ||
+ file_name >> rpair true;
+
+val uses = Scan.optional (Parse.$$$ usesN |-- Parse.!!! (Scan.repeat1 file)) [];
val args =
- theory_name -- (P.$$$ importsN |-- P.!!! (Scan.repeat1 theory_name -- uses --| P.$$$ beginN))
- >> P.triple2;
+ theory_name -- (Parse.$$$ importsN |--
+ Parse.!!! (Scan.repeat1 theory_name -- uses --| Parse.$$$ beginN))
+ >> Parse.triple2;
(* read header *)
val header =
- (P.$$$ headerN -- P.tags) |-- (P.!!! (P.doc_source -- Scan.repeat P.semicolon --
- (P.$$$ theoryN -- P.tags) |-- args)) ||
- (P.$$$ theoryN -- P.tags) |-- P.!!! args;
+ (Parse.$$$ headerN -- Parse.tags) |--
+ (Parse.!!! (Parse.doc_source -- Scan.repeat Parse.semicolon --
+ (Parse.$$$ theoryN -- Parse.tags) |-- args)) ||
+ (Parse.$$$ theoryN -- Parse.tags) |-- Parse.!!! args;
fun read pos src =
let val res =
src
|> Symbol.source {do_recover = false}
- |> T.source {do_recover = NONE} (fn () => (header_lexicon, Scan.empty_lexicon)) pos
- |> T.source_proper
- |> Source.source T.stopper (Scan.single (Scan.error (P.!!! header))) NONE
+ |> Token.source {do_recover = NONE} (fn () => (header_lexicon, Scan.empty_lexicon)) pos
+ |> Token.source_proper
+ |> Source.source Token.stopper (Scan.single (Scan.error (Parse.!!! header))) NONE
|> Source.get_single;
in
(case res of SOME (x, _) => x
--- a/src/Pure/Thy/thy_header.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_header.scala Sun May 23 10:38:11 2010 +0100
@@ -27,7 +27,7 @@
}
-class Thy_Header(symbols: Symbol.Interpretation) extends Outer_Parse.Parser
+class Thy_Header(symbols: Symbol.Interpretation) extends Parse.Parser
{
import Thy_Header._
@@ -59,14 +59,14 @@
def read(file: File): Header =
{
val token = lexicon.token(symbols, _ => false)
- val toks = new mutable.ListBuffer[Outer_Lex.Token]
+ val toks = new mutable.ListBuffer[Token]
def scan_to_begin(in: Reader[Char])
{
token(in) match {
case lexicon.Success(tok, rest) =>
toks += tok
- if (!(tok.kind == Outer_Lex.Token_Kind.KEYWORD && tok.content == BEGIN))
+ if (!(tok.kind == Token.Kind.KEYWORD && tok.content == BEGIN))
scan_to_begin(rest)
case _ =>
}
@@ -74,7 +74,7 @@
val reader = Scan.byte_reader(file)
try { scan_to_begin(reader) } finally { reader.close }
- parse(commit(header), Outer_Lex.reader(toks.toList)) match {
+ parse(commit(header), Token.reader(toks.toList)) match {
case Success(result, _) => result
case bad => error(bad.toString)
}
--- a/src/Pure/Thy/thy_info.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_info.ML Sun May 23 10:38:11 2010 +0100
@@ -346,7 +346,7 @@
val _ = CRITICAL (fn () =>
change_deps name (Option.map (fn {master, text, parents, files, ...} =>
make_deps upd_time master text parents files)));
- val after_load = OuterSyntax.load_thy name pos text (time orelse ! Output.timing);
+ val after_load = Outer_Syntax.load_thy name pos text (time orelse ! Output.timing);
val _ =
CRITICAL (fn () =>
(change_deps name
--- a/src/Pure/Thy/thy_output.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_output.ML Sun May 23 10:38:11 2010 +0100
@@ -24,7 +24,7 @@
val modes: string list Unsynchronized.ref
val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string
val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->
- (Toplevel.transition * Toplevel.state) list -> (OuterLex.token, 'a) Source.source -> Buffer.T
+ (Toplevel.transition * Toplevel.state) list -> (Token.T, 'a) Source.source -> Buffer.T
val pretty_text: string -> Pretty.T
val pretty_term: Proof.context -> term -> Pretty.T
val pretty_thm: Proof.context -> thm -> Pretty.T
@@ -36,10 +36,6 @@
structure ThyOutput: THY_OUTPUT =
struct
-structure T = OuterLex;
-structure P = OuterParse;
-
-
(** global options **)
val display = Unsynchronized.ref false;
@@ -132,12 +128,16 @@
local
-val property = P.xname -- Scan.optional (P.$$$ "=" |-- P.!!! P.xname) "";
-val properties = Scan.optional (P.$$$ "[" |-- P.!!! (P.enum "," property --| P.$$$ "]")) [];
+val property =
+ Parse.xname -- Scan.optional (Parse.$$$ "=" |-- Parse.!!! Parse.xname) "";
+
+val properties =
+ Scan.optional (Parse.$$$ "[" |-- Parse.!!! (Parse.enum "," property --| Parse.$$$ "]")) [];
in
-val antiq = P.!!! (P.position P.xname -- properties -- Args.parse --| Scan.ahead P.eof)
+val antiq =
+ Parse.!!! (Parse.position Parse.xname -- properties -- Args.parse --| Scan.ahead Parse.eof)
>> (fn (((x, pos), y), z) => (y, Args.src ((x, z), pos)));
end;
@@ -151,7 +151,7 @@
let
fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
| expand (Antiquote.Antiq (ss, (pos, _))) =
- let val (opts, src) = T.read_antiq lex antiq (ss, pos) in
+ let val (opts, src) = Token.read_antiq lex antiq (ss, pos) in
options opts (fn () => command src state) (); (*preview errors!*)
PrintMode.with_modes (! modes @ Latex.modes)
(Output.no_warnings_CRITICAL (options opts (fn () => command src state))) ()
@@ -175,7 +175,7 @@
datatype token =
NoToken
- | BasicToken of T.token
+ | BasicToken of Token.T
| MarkupToken of string * (string * Position.T)
| MarkupEnvToken of string * (string * Position.T)
| VerbatimToken of string * Position.T;
@@ -192,10 +192,10 @@
fun basic_token pred (BasicToken tok) = pred tok
| basic_token _ _ = false;
-val improper_token = basic_token (not o T.is_proper);
-val comment_token = basic_token T.is_comment;
-val blank_token = basic_token T.is_blank;
-val newline_token = basic_token T.is_newline;
+val improper_token = basic_token (not o Token.is_proper);
+val comment_token = basic_token Token.is_comment;
+val blank_token = basic_token Token.is_blank;
+val newline_token = basic_token Token.is_newline;
(* command spans *)
@@ -249,7 +249,7 @@
val Span ((cmd_name, cmd_pos, cmd_tags), srcs, span_newline) = span;
val (tag, tags) = tag_stack;
- val tag' = try hd (fold OuterKeyword.update_tags cmd_tags (the_list tag));
+ val tag' = try hd (fold Keyword.update_tags cmd_tags (the_list tag));
val active_tag' =
if is_some tag' then tag'
@@ -300,27 +300,27 @@
local
val space_proper =
- Scan.one T.is_blank -- Scan.many T.is_comment -- Scan.one T.is_proper;
+ Scan.one Token.is_blank -- Scan.many Token.is_comment -- Scan.one Token.is_proper;
-val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore);
+val is_improper = not o (Token.is_proper orf Token.is_begin_ignore orf Token.is_end_ignore);
val improper = Scan.many is_improper;
val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
-val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one T.is_blank));
+val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
-val opt_newline = Scan.option (Scan.one T.is_newline);
+val opt_newline = Scan.option (Scan.one Token.is_newline);
val ignore =
- Scan.depend (fn d => opt_newline |-- Scan.one T.is_begin_ignore
+ Scan.depend (fn d => opt_newline |-- Scan.one Token.is_begin_ignore
>> pair (d + 1)) ||
- Scan.depend (fn d => Scan.one T.is_end_ignore --|
+ Scan.depend (fn d => Scan.one Token.is_end_ignore --|
(if d = 0 then Scan.fail_with (K "Bad nesting of meta-comments") else opt_newline)
>> pair (d - 1));
-val tag = (improper -- P.$$$ "%" -- improper) |-- P.!!! (P.tag_name --| blank_end);
+val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);
val locale =
- Scan.option ((P.$$$ "(" -- improper -- P.$$$ "in") |--
- P.!!! (improper |-- P.xname --| (improper -- P.$$$ ")")));
+ Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
+ Parse.!!! (improper |-- Parse.xname --| (improper -- Parse.$$$ ")")));
in
@@ -332,26 +332,28 @@
>> (fn d => (NONE, (NoToken, ("", d))));
fun markup mark mk flag = Scan.peek (fn d =>
- improper |-- P.position (Scan.one (T.is_kind T.Command andf is_markup mark o T.content_of)) --
+ improper |--
+ Parse.position
+ (Scan.one (Token.is_kind Token.Command andf is_markup mark o Token.content_of)) --
Scan.repeat tag --
- P.!!!! ((improper -- locale -- improper) |-- P.doc_source --| improper_end)
+ Parse.!!!! ((improper -- locale -- improper) |-- Parse.doc_source --| improper_end)
>> (fn (((tok, pos), tags), txt) =>
- let val name = T.content_of tok
+ let val name = Token.content_of tok
in (SOME (name, pos, tags), (mk (name, txt), (flag, d))) end));
val command = Scan.peek (fn d =>
- P.position (Scan.one (T.is_kind T.Command)) --
+ Parse.position (Scan.one (Token.is_kind Token.Command)) --
Scan.repeat tag
>> (fn ((tok, pos), tags) =>
- let val name = T.content_of tok
+ let val name = Token.content_of tok
in (SOME (name, pos, tags), (BasicToken tok, (Latex.markup_false, d))) end));
val cmt = Scan.peek (fn d =>
- P.$$$ "--" |-- P.!!!! (improper |-- P.doc_source)
+ Parse.$$$ "--" |-- Parse.!!!! (improper |-- Parse.doc_source)
>> (fn txt => (NONE, (MarkupToken ("cmt", txt), ("", d)))));
val other = Scan.peek (fn d =>
- P.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
+ Parse.not_eof >> (fn tok => (NONE, (BasicToken tok, ("", d)))));
val token =
ignored ||
@@ -363,8 +365,8 @@
(* spans *)
- val is_eof = fn (_, (BasicToken x, _)) => T.is_eof x | _ => false;
- val stopper = Scan.stopper (K (NONE, (BasicToken T.eof, ("", 0)))) is_eof;
+ val is_eof = fn (_, (BasicToken x, _)) => Token.is_eof x | _ => false;
+ val stopper = Scan.stopper (K (NONE, (BasicToken Token.eof, ("", 0)))) is_eof;
val cmd = Scan.one (is_some o fst);
val non_cmd = Scan.one (is_none o fst andf not o is_eof) >> #2;
@@ -386,8 +388,8 @@
val spans =
src
- |> Source.filter (not o T.is_semicolon)
- |> Source.source' 0 T.stopper (Scan.error (Scan.bulk token)) NONE
+ |> Source.filter (not o Token.is_semicolon)
+ |> Source.source' 0 Token.stopper (Scan.error (Scan.bulk token)) NONE
|> Source.source stopper (Scan.error (Scan.bulk span)) NONE
|> Source.exhaust;
@@ -486,7 +488,7 @@
(* default output *)
-val str_of_source = space_implode " " o map T.unparse o #2 o #1 o Args.dest_src;
+val str_of_source = space_implode " " o map Token.unparse o #2 o #1 o Args.dest_src;
fun maybe_pretty_source pretty src xs =
map pretty xs (*always pretty in order to exhibit errors!*)
@@ -565,7 +567,7 @@
(* embedded lemma *)
-val _ = OuterKeyword.keyword "by";
+val _ = Keyword.keyword "by";
val _ = antiquotation "lemma"
(Args.prop -- Scan.lift (Args.$$$ "by" |-- Method.parse -- Scan.option Method.parse))
--- a/src/Pure/Thy/thy_syntax.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_syntax.ML Sun May 23 10:38:11 2010 +0100
@@ -7,18 +7,17 @@
signature THY_SYNTAX =
sig
val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
- (OuterLex.token, (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, 'a)
+ (Token.T, (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, 'a)
Source.source) Source.source) Source.source) Source.source
- val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> OuterLex.token list
- val present_token: OuterLex.token -> output
- val report_token: OuterLex.token -> unit
+ val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list
+ val present_token: Token.T -> output
+ val report_token: Token.T -> unit
datatype span_kind = Command of string | Ignored | Malformed
type span
val span_kind: span -> span_kind
- val span_content: span -> OuterLex.token list
+ val span_content: span -> Token.T list
val span_range: span -> Position.range
- val span_source: (OuterLex.token, 'a) Source.source ->
- (span, (OuterLex.token, 'a) Source.source) Source.source
+ val span_source: (Token.T, 'a) Source.source -> (span, (Token.T, 'a) Source.source) Source.source
val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
val present_span: span -> output
val report_span: span -> unit
@@ -29,18 +28,13 @@
structure ThySyntax: THY_SYNTAX =
struct
-structure K = OuterKeyword;
-structure T = OuterLex;
-structure P = OuterParse;
-
-
(** tokens **)
(* parse *)
fun token_source lexs pos src =
Symbol.source {do_recover = true} src
- |> T.source {do_recover = SOME false} (K lexs) pos;
+ |> Token.source {do_recover = SOME false} (K lexs) pos;
fun parse_tokens lexs pos str =
Source.of_string str
@@ -53,33 +47,33 @@
local
val token_kind_markup =
- fn T.Command => (Markup.commandN, [])
- | T.Keyword => (Markup.keywordN, [])
- | T.Ident => Markup.ident
- | T.LongIdent => Markup.ident
- | T.SymIdent => Markup.ident
- | T.Var => Markup.var
- | T.TypeIdent => Markup.tfree
- | T.TypeVar => Markup.tvar
- | T.Nat => Markup.ident
- | T.String => Markup.string
- | T.AltString => Markup.altstring
- | T.Verbatim => Markup.verbatim
- | T.Space => Markup.none
- | T.Comment => Markup.comment
- | T.InternalValue => Markup.none
- | T.Malformed => Markup.malformed
- | T.Error _ => Markup.malformed
- | T.Sync => Markup.control
- | T.EOF => Markup.control;
+ fn Token.Command => (Markup.commandN, [])
+ | Token.Keyword => (Markup.keywordN, [])
+ | Token.Ident => Markup.ident
+ | Token.LongIdent => Markup.ident
+ | Token.SymIdent => Markup.ident
+ | Token.Var => Markup.var
+ | Token.TypeIdent => Markup.tfree
+ | Token.TypeVar => Markup.tvar
+ | Token.Nat => Markup.ident
+ | Token.String => Markup.string
+ | Token.AltString => Markup.altstring
+ | Token.Verbatim => Markup.verbatim
+ | Token.Space => Markup.none
+ | Token.Comment => Markup.comment
+ | Token.InternalValue => Markup.none
+ | Token.Malformed => Markup.malformed
+ | Token.Error _ => Markup.malformed
+ | Token.Sync => Markup.control
+ | Token.EOF => Markup.control;
in
fun present_token tok =
- Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
+ Markup.enclose (token_kind_markup (Token.kind_of tok)) (Output.output (Token.unparse tok));
fun report_token tok =
- Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
+ Position.report (token_kind_markup (Token.kind_of tok)) (Token.position_of tok);
end;
@@ -90,7 +84,7 @@
(* type span *)
datatype span_kind = Command of string | Ignored | Malformed;
-datatype span = Span of span_kind * OuterLex.token list;
+datatype span = Span of span_kind * Token.T list;
fun span_kind (Span (k, _)) = k;
fun span_content (Span (_, toks)) = toks;
@@ -100,8 +94,8 @@
[] => (Position.none, Position.none)
| toks =>
let
- val start_pos = T.position_of (hd toks);
- val end_pos = T.end_position_of (List.last toks);
+ val start_pos = Token.position_of (hd toks);
+ val end_pos = Token.end_position_of (List.last toks);
in (start_pos, end_pos) end);
@@ -109,19 +103,20 @@
local
-val is_whitespace = T.is_kind T.Space orf T.is_kind T.Comment;
+val is_whitespace = Token.is_kind Token.Space orf Token.is_kind Token.Comment;
-val body = Scan.unless (Scan.many is_whitespace -- Scan.ahead (P.command || P.eof)) P.not_eof;
+val body =
+ Scan.unless (Scan.many is_whitespace -- Scan.ahead (Parse.command || Parse.eof)) Parse.not_eof;
val span =
- Scan.ahead P.command -- P.not_eof -- Scan.repeat body
+ Scan.ahead Parse.command -- Parse.not_eof -- Scan.repeat body
>> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
in
-fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
+fun span_source src = Source.source Token.stopper (Scan.bulk span) NONE src;
end;
@@ -175,13 +170,13 @@
val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
if d <= 0 then Scan.fail
else
- command_with K.is_qed_global >> pair ~1 ||
- command_with K.is_proof_goal >> pair (d + 1) ||
- (if d = 0 then Scan.fail else command_with K.is_qed >> pair (d - 1)) ||
- Scan.unless (command_with K.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
+ command_with Keyword.is_qed_global >> pair ~1 ||
+ command_with Keyword.is_proof_goal >> pair (d + 1) ||
+ (if d = 0 then Scan.fail else command_with Keyword.is_qed >> pair (d - 1)) ||
+ Scan.unless (command_with Keyword.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
val unit =
- command_with K.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
+ command_with Keyword.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
Scan.one not_eof >> (fn a => (a, [], true));
in
--- a/src/Pure/Thy/thy_syntax.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Thy/thy_syntax.scala Sun May 23 10:38:11 2010 +0100
@@ -9,7 +9,7 @@
object Thy_Syntax
{
- private val parser = new Outer_Parse.Parser
+ private val parser = new Parse.Parser
{
override def filter_proper = false
@@ -22,13 +22,13 @@
}
}
- type Span = List[Outer_Lex.Token]
+ type Span = List[Token]
- def parse_spans(toks: List[Outer_Lex.Token]): List[Span] =
+ def parse_spans(toks: List[Token]): List[Span] =
{
import parser._
- parse(rep(command_span), Outer_Lex.reader(toks)) match {
+ parse(rep(command_span), Token.reader(toks)) match {
case Success(spans, rest) if rest.atEnd => spans
case bad => error(bad.toString)
}
--- a/src/Pure/Tools/find_consts.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Tools/find_consts.ML Sun May 23 10:38:11 2010 +0100
@@ -148,23 +148,15 @@
Toplevel.unknown_theory o Toplevel.keep (fn state =>
find_consts (Proof.context_of (Toplevel.enter_proof_body state)) spec);
-local
-
-structure P = OuterParse and K = OuterKeyword;
-
val criterion =
- P.reserved "strict" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Strict ||
- P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Name ||
- P.xname >> Loose;
-
-in
+ Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Strict ||
+ Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name ||
+ Parse.xname >> Loose;
val _ =
- OuterSyntax.improper_command "find_consts" "search constants by type pattern" K.diag
- (Scan.repeat ((Scan.option P.minus >> is_none) -- criterion)
+ Outer_Syntax.improper_command "find_consts" "search constants by type pattern" Keyword.diag
+ (Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion)
>> (Toplevel.no_timing oo find_consts_cmd));
end;
-end;
-
--- a/src/Pure/Tools/find_theorems.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/Tools/find_theorems.ML Sun May 23 10:38:11 2010 +0100
@@ -465,28 +465,27 @@
local
-structure P = OuterParse and K = OuterKeyword;
-
val criterion =
- P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Name ||
- P.reserved "intro" >> K Intro ||
- P.reserved "introiff" >> K IntroIff ||
- P.reserved "elim" >> K Elim ||
- P.reserved "dest" >> K Dest ||
- P.reserved "solves" >> K Solves ||
- P.reserved "simp" |-- P.!!! (P.$$$ ":" |-- P.term) >> Simp ||
- P.term >> Pattern;
+ Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Name ||
+ Parse.reserved "intro" >> K Intro ||
+ Parse.reserved "introiff" >> K IntroIff ||
+ Parse.reserved "elim" >> K Elim ||
+ Parse.reserved "dest" >> K Dest ||
+ Parse.reserved "solves" >> K Solves ||
+ Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp ||
+ Parse.term >> Pattern;
val options =
Scan.optional
- (P.$$$ "(" |--
- P.!!! (Scan.option P.nat -- Scan.optional (P.reserved "with_dups" >> K false) true
- --| P.$$$ ")")) (NONE, true);
+ (Parse.$$$ "(" |--
+ Parse.!!! (Scan.option Parse.nat -- Scan.optional (Parse.reserved "with_dups" >> K false) true
+ --| Parse.$$$ ")")) (NONE, true);
in
val _ =
- OuterSyntax.improper_command "find_theorems" "print theorems meeting specified criteria" K.diag
- (options -- Scan.repeat (((Scan.option P.minus >> is_none) -- criterion))
+ Outer_Syntax.improper_command "find_theorems" "print theorems meeting specified criteria"
+ Keyword.diag
+ (options -- Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion))
>> (Toplevel.no_timing oo find_theorems_cmd));
end;
--- a/src/Pure/axclass.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/axclass.ML Sun May 23 10:38:11 2010 +0100
@@ -196,9 +196,9 @@
fun the_classrel thy (c1, c2) =
(case Symreltab.lookup (proven_classrels_of thy) (c1, c2) of
- SOME thm => Thm.transfer thy thm
+ SOME thm => thm
| NONE => error ("Unproven class relation " ^
- Syntax.string_of_classrel (ProofContext.init_global thy) [c1, c2]));
+ Syntax.string_of_classrel (ProofContext.init_global thy) [c1, c2])); (* FIXME stale thy (!?) *)
fun put_trancl_classrel ((c1, c2), th) thy =
let
@@ -206,7 +206,7 @@
val classrels = proven_classrels_of thy;
fun reflcl_classrel (c1', c2') =
- if c1' = c2' then NONE else SOME (the_classrel thy (c1', c2'));
+ if c1' = c2' then NONE else SOME (Thm.transfer thy (the_classrel thy (c1', c2')));
fun gen_classrel (c1_pred, c2_succ) =
let
val th' =
@@ -243,9 +243,9 @@
fun the_arity thy (a, Ss, c) =
(case AList.lookup (op =) (Symtab.lookup_list (proven_arities_of thy) a) (c, Ss) of
- SOME (thm, _) => Thm.transfer thy thm
+ SOME (thm, _) => thm
| NONE => error ("Unproven type arity " ^
- Syntax.string_of_arity (ProofContext.init_global thy) (a, Ss, [c])));
+ Syntax.string_of_arity (ProofContext.init_global thy) (a, Ss, [c]))); (* FIXME stale thy (!?) *)
fun thynames_of_arity thy (c, a) =
Symtab.lookup_list (proven_arities_of thy) a
@@ -267,7 +267,7 @@
val completions = super_class_completions |> map (fn c1 =>
let
val th1 =
- (th RS the_classrel thy (c, c1))
+ (th RS Thm.transfer thy (the_classrel thy (c, c1)))
|> Drule.instantiate' std_vars []
|> Thm.close_derivation;
in ((th1, thy_name), c1) end);
@@ -412,8 +412,6 @@
(* primitive rules *)
-val shyps_topped = forall null o #shyps o Thm.rep_thm;
-
fun add_classrel raw_th thy =
let
val th = Thm.strip_shyps (Thm.transfer thy raw_th);
@@ -422,9 +420,8 @@
val rel = Logic.dest_classrel prop handle TERM _ => err ();
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
val th' = th
- |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [c1])))] []
- |> Thm.unconstrainT;
- val _ = shyps_topped th' orelse raise Fail "add_classrel: nontop shyps after unconstrain";
+ |> Thm.unconstrainT
+ |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [])))] [];
in
thy
|> Sign.primitive_classrel (c1, c2)
@@ -441,16 +438,15 @@
val args = Name.names Name.context Name.aT Ss;
val T = Type (t, map TFree args);
- val std_vars = map (fn (a, S) => SOME (ctyp_of thy (TVar ((a, 0), S)))) args;
+ val std_vars = map (fn (a, S) => SOME (ctyp_of thy (TVar ((a, 0), [])))) args;
val missing_params = Sign.complete_sort thy [c]
|> maps (these o Option.map #params o try (get_info thy))
|> filter_out (fn (const, _) => can (get_inst_param thy) (const, t))
|> (map o apsnd o map_atyps) (K T);
val th' = th
- |> Drule.instantiate' std_vars []
- |> Thm.unconstrainT;
- val _ = shyps_topped th' orelse raise Fail "add_arity: nontop shyps after unconstrain";
+ |> Thm.unconstrainT
+ |> Drule.instantiate' std_vars [];
in
thy
|> fold (#2 oo declare_overloaded) missing_params
--- a/src/Pure/build-jars Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/build-jars Sun May 23 10:38:11 2010 +0100
@@ -33,10 +33,10 @@
General/xml.scala
General/yxml.scala
Isar/isar_document.scala
- Isar/outer_keyword.scala
- Isar/outer_lex.scala
- Isar/outer_parse.scala
+ Isar/keyword.scala
Isar/outer_syntax.scala
+ Isar/parse.scala
+ Isar/token.scala
PIDE/change.scala
PIDE/command.scala
PIDE/document.scala
--- a/src/Pure/codegen.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/codegen.ML Sun May 23 10:38:11 2010 +0100
@@ -959,44 +959,42 @@
| _ => error ("Malformed annotation: " ^ quote s));
-structure P = OuterParse and K = OuterKeyword;
-
-val _ = List.app OuterKeyword.keyword ["attach", "file", "contains"];
+val _ = List.app Keyword.keyword ["attach", "file", "contains"];
fun strip_whitespace s = implode (fst (take_suffix (fn c => c = "\n" orelse c = " ")
(snd (take_prefix (fn c => c = "\n" orelse c = " ") (explode s))))) ^ "\n";
-val parse_attach = Scan.repeat (P.$$$ "attach" |--
- Scan.optional (P.$$$ "(" |-- P.xname --| P.$$$ ")") "" --
- (P.verbatim >> strip_whitespace));
+val parse_attach = Scan.repeat (Parse.$$$ "attach" |--
+ Scan.optional (Parse.$$$ "(" |-- Parse.xname --| Parse.$$$ ")") "" --
+ (Parse.verbatim >> strip_whitespace));
val _ =
- OuterSyntax.command "types_code"
- "associate types with target language types" K.thy_decl
- (Scan.repeat1 (P.xname --| P.$$$ "(" -- P.string --| P.$$$ ")" -- parse_attach) >>
+ Outer_Syntax.command "types_code"
+ "associate types with target language types" Keyword.thy_decl
+ (Scan.repeat1 (Parse.xname --| Parse.$$$ "(" -- Parse.string --| Parse.$$$ ")" -- parse_attach) >>
(fn xs => Toplevel.theory (fn thy => fold (assoc_type o
(fn ((name, mfx), aux) => (name, (parse_mixfix
(Syntax.read_typ_global thy) mfx, aux)))) xs thy)));
val _ =
- OuterSyntax.command "consts_code"
- "associate constants with target language code" K.thy_decl
+ Outer_Syntax.command "consts_code"
+ "associate constants with target language code" Keyword.thy_decl
(Scan.repeat1
- (P.term --|
- P.$$$ "(" -- P.string --| P.$$$ ")" -- parse_attach) >>
+ (Parse.term --|
+ Parse.$$$ "(" -- Parse.string --| Parse.$$$ ")" -- parse_attach) >>
(fn xs => Toplevel.theory (fn thy => fold (assoc_const o
(fn ((const, mfx), aux) =>
(const, (parse_mixfix (Syntax.read_term_global thy) mfx, aux)))) xs thy)));
fun parse_code lib =
- Scan.optional (P.$$$ "(" |-- P.enum "," P.xname --| P.$$$ ")") (!mode) --
- (if lib then Scan.optional P.name "" else P.name) --
- Scan.option (P.$$$ "file" |-- P.name) --
+ Scan.optional (Parse.$$$ "(" |-- Parse.enum "," Parse.xname --| Parse.$$$ ")") (!mode) --
+ (if lib then Scan.optional Parse.name "" else Parse.name) --
+ Scan.option (Parse.$$$ "file" |-- Parse.name) --
(if lib then Scan.succeed []
- else Scan.optional (P.$$$ "imports" |-- Scan.repeat1 P.name) []) --|
- P.$$$ "contains" --
- ( Scan.repeat1 (P.name --| P.$$$ "=" -- P.term)
- || Scan.repeat1 (P.term >> pair "")) >>
+ else Scan.optional (Parse.$$$ "imports" |-- Scan.repeat1 Parse.name) []) --|
+ Parse.$$$ "contains" --
+ ( Scan.repeat1 (Parse.name --| Parse.$$$ "=" -- Parse.term)
+ || Scan.repeat1 (Parse.term >> pair "")) >>
(fn ((((mode', module), opt_fname), modules), xs) => Toplevel.theory (fn thy =>
let
val mode'' = (if lib then insert (op =) "library" else I) (remove (op =) "library" mode');
@@ -1020,13 +1018,13 @@
#> add_preprocessor unfold_preprocessor;
val _ =
- OuterSyntax.command "code_library"
- "generates code for terms (one structure for each theory)" K.thy_decl
+ Outer_Syntax.command "code_library"
+ "generates code for terms (one structure for each theory)" Keyword.thy_decl
(parse_code true);
val _ =
- OuterSyntax.command "code_module"
- "generates code for terms (single structure, incremental)" K.thy_decl
+ Outer_Syntax.command "code_module"
+ "generates code for terms (single structure, incremental)" Keyword.thy_decl
(parse_code false);
end;
--- a/src/Pure/conv.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/conv.ML Sun May 23 10:38:11 2010 +0100
@@ -1,5 +1,6 @@
(* Title: Pure/conv.ML
Author: Amine Chaieb, TU Muenchen
+ Author: Sascha Boehme, TU Muenchen
Author: Makarius
Conversions: primitive equality reasoning.
@@ -32,10 +33,16 @@
val arg1_conv: conv -> conv
val fun2_conv: conv -> conv
val binop_conv: conv -> conv
+ val binder_conv: (cterm * Proof.context -> conv) -> Proof.context -> conv
val forall_conv: (cterm * Proof.context -> conv) -> Proof.context -> conv
val implies_conv: conv -> conv -> conv
val implies_concl_conv: conv -> conv
val rewr_conv: thm -> conv
+ val rewrs_conv: thm list -> conv
+ val sub_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val bottom_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val top_conv: (Proof.context -> conv) -> Proof.context -> conv
+ val top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv
val params_conv: int -> (Proof.context -> conv) -> Proof.context -> conv
val prems_conv: int -> conv -> conv
val concl_conv: int -> conv -> conv
@@ -105,6 +112,29 @@
fun binop_conv cv = combination_conv (arg_conv cv) cv;
+fun binder_conv cv ctxt = arg_conv (abs_conv cv ctxt);
+
+
+(* subterm structure *)
+
+(*cf. SUB_CONV in HOL*)
+fun sub_conv conv ctxt =
+ comb_conv (conv ctxt) else_conv
+ abs_conv (conv o snd) ctxt else_conv
+ all_conv;
+
+(*cf. BOTTOM_CONV in HOL*)
+fun bottom_conv conv ctxt ct =
+ (sub_conv (bottom_conv conv) ctxt then_conv conv ctxt) ct;
+
+(*cf. TOP_CONV in HOL*)
+fun top_conv conv ctxt ct =
+ (conv ctxt then_conv sub_conv (top_conv conv) ctxt) ct;
+
+(*cf. TOP_SWEEP_CONV in HOL*)
+fun top_sweep_conv conv ctxt ct =
+ (conv ctxt else_conv sub_conv (top_sweep_conv conv) ctxt) ct;
+
(* primitive logic *)
@@ -136,6 +166,8 @@
handle Pattern.MATCH => raise CTERM ("rewr_conv", [lhs, ct])
end;
+fun rewrs_conv rules = first_conv (map rewr_conv rules);
+
(* conversions on HHF rules *)
--- a/src/Pure/drule.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/drule.ML Sun May 23 10:38:11 2010 +0100
@@ -204,11 +204,11 @@
(** Standardization of rules **)
(*Generalization over a list of variables*)
-val forall_intr_list = fold_rev forall_intr;
+val forall_intr_list = fold_rev Thm.forall_intr;
(*Generalization over Vars -- canonical order*)
fun forall_intr_vars th =
- fold forall_intr
+ fold Thm.forall_intr
(map (Thm.cterm_of (Thm.theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th [])) th;
fun outer_params t =
@@ -245,10 +245,10 @@
fun generalize names th = Thm.generalize names (Thm.maxidx_of th + 1) th;
(*specialization over a list of cterms*)
-val forall_elim_list = fold forall_elim;
+val forall_elim_list = fold Thm.forall_elim;
(*maps A1,...,An |- B to [| A1;...;An |] ==> B*)
-val implies_intr_list = fold_rev implies_intr;
+val implies_intr_list = fold_rev Thm.implies_intr;
(*maps [| A1;...;An |] ==> B and [A1,...,An] to B*)
fun implies_elim_list impth ths = fold Thm.elim_implies ths impth;
@@ -278,7 +278,7 @@
This step can lose information.*)
fun flexflex_unique th =
if null (tpairs_of th) then th else
- case distinct Thm.eq_thm (Seq.list_of (flexflex_rule th)) of
+ case distinct Thm.eq_thm (Seq.list_of (Thm.flexflex_rule th)) of
[th] => th
| [] => raise THM("flexflex_unique: impossible constraints", 0, [th])
| _ => raise THM("flexflex_unique: multiple unifiers", 0, [th]);
@@ -464,8 +464,8 @@
fun extensional eq =
let val eq' =
- abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (cprop_of eq)))) eq
- in equal_elim (eta_conversion (cprop_of eq')) eq' end;
+ Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (cprop_of eq)))) eq
+ in Thm.equal_elim (Thm.eta_conversion (cprop_of eq')) eq' end;
val equals_cong =
store_standard_thm_open (Binding.name "equals_cong")
@@ -478,13 +478,13 @@
val AC = read_prop "A ==> C"
val A = read_prop "A"
in
- store_standard_thm_open (Binding.name "imp_cong") (implies_intr ABC (equal_intr
- (implies_intr AB (implies_intr A
- (equal_elim (implies_elim (assume ABC) (assume A))
- (implies_elim (assume AB) (assume A)))))
- (implies_intr AC (implies_intr A
- (equal_elim (symmetric (implies_elim (assume ABC) (assume A)))
- (implies_elim (assume AC) (assume A)))))))
+ store_standard_thm_open (Binding.name "imp_cong") (Thm.implies_intr ABC (Thm.equal_intr
+ (Thm.implies_intr AB (Thm.implies_intr A
+ (Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))
+ (Thm.implies_elim (Thm.assume AB) (Thm.assume A)))))
+ (Thm.implies_intr AC (Thm.implies_intr A
+ (Thm.equal_elim (Thm.symmetric (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)))
+ (Thm.implies_elim (Thm.assume AC) (Thm.assume A)))))))
end;
val swap_prems_eq =
@@ -495,11 +495,11 @@
val B = read_prop "B"
in
store_standard_thm_open (Binding.name "swap_prems_eq")
- (equal_intr
- (implies_intr ABC (implies_intr B (implies_intr A
- (implies_elim (implies_elim (assume ABC) (assume A)) (assume B)))))
- (implies_intr BAC (implies_intr A (implies_intr B
- (implies_elim (implies_elim (assume BAC) (assume B)) (assume A))))))
+ (Thm.equal_intr
+ (Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A
+ (Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B)))))
+ (Thm.implies_intr BAC (Thm.implies_intr A (Thm.implies_intr B
+ (Thm.implies_elim (Thm.implies_elim (Thm.assume BAC) (Thm.assume B)) (Thm.assume A))))))
end;
val imp_cong_rule = Thm.combination o Thm.combination (Thm.reflexive implies);
@@ -513,16 +513,18 @@
val rhs_of = snd o dest_eq
in
fun beta_eta_conversion t =
- let val thm = beta_conversion true t
- in transitive thm (eta_conversion (rhs_of thm)) end
+ let val thm = Thm.beta_conversion true t
+ in Thm.transitive thm (Thm.eta_conversion (rhs_of thm)) end
end;
-fun eta_long_conversion ct = transitive (beta_eta_conversion ct)
- (symmetric (beta_eta_conversion (cterm_fun (Pattern.eta_long []) ct)));
+fun eta_long_conversion ct =
+ Thm.transitive
+ (beta_eta_conversion ct)
+ (Thm.symmetric (beta_eta_conversion (cterm_fun (Pattern.eta_long []) ct)));
(*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*)
fun eta_contraction_rule th =
- equal_elim (eta_conversion (cprop_of th)) th;
+ Thm.equal_elim (Thm.eta_conversion (cprop_of th)) th;
(* abs_def *)
@@ -578,7 +580,7 @@
val VW = read_prop "V ==> W";
in
store_standard_thm_open (Binding.name "revcut_rl")
- (implies_intr V (implies_intr VW (implies_elim (assume VW) (assume V))))
+ (Thm.implies_intr V (Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V))))
end;
(*for deleting an unwanted assumption*)
@@ -586,7 +588,7 @@
let
val V = read_prop "V";
val W = read_prop "W";
- val thm = implies_intr V (implies_intr W (assume W));
+ val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W));
in store_standard_thm_open (Binding.name "thin_rl") thm end;
(* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*)
@@ -597,8 +599,8 @@
val x = certify (Free ("x", Term.aT []));
in
store_standard_thm_open (Binding.name "triv_forall_equality")
- (equal_intr (implies_intr QV (forall_elim x (assume QV)))
- (implies_intr V (forall_intr x (assume V))))
+ (Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV)))
+ (Thm.implies_intr V (Thm.forall_intr x (Thm.assume V))))
end;
(* (PROP ?Phi ==> PROP ?Phi ==> PROP ?Psi) ==>
@@ -610,7 +612,7 @@
val A = read_prop "Phi";
in
store_standard_thm_open (Binding.name "distinct_prems_rl")
- (implies_intr_list [AAB, A] (implies_elim_list (assume AAB) [assume A, assume A]))
+ (implies_intr_list [AAB, A] (implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A]))
end;
(* (PROP ?PhiA ==> PROP ?PhiB ==> PROP ?Psi) ==>
@@ -620,15 +622,15 @@
val swap_prems_rl =
let
val cmajor = read_prop "PhiA ==> PhiB ==> Psi";
- val major = assume cmajor;
+ val major = Thm.assume cmajor;
val cminor1 = read_prop "PhiA";
- val minor1 = assume cminor1;
+ val minor1 = Thm.assume cminor1;
val cminor2 = read_prop "PhiB";
- val minor2 = assume cminor2;
+ val minor2 = Thm.assume cminor2;
in
store_standard_thm_open (Binding.name "swap_prems_rl")
- (implies_intr cmajor (implies_intr cminor2 (implies_intr cminor1
- (implies_elim (implies_elim major minor1) minor2))))
+ (Thm.implies_intr cmajor (Thm.implies_intr cminor2 (Thm.implies_intr cminor1
+ (Thm.implies_elim (Thm.implies_elim major minor1) minor2))))
end;
(* [| PROP ?phi ==> PROP ?psi; PROP ?psi ==> PROP ?phi |]
@@ -641,7 +643,7 @@
val QP = read_prop "psi ==> phi";
in
store_standard_thm_open (Binding.name "equal_intr_rule")
- (implies_intr PQ (implies_intr QP (equal_intr (assume PQ) (assume QP))))
+ (Thm.implies_intr PQ (Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP))))
end;
(* PROP ?phi == PROP ?psi ==> PROP ?phi ==> PROP ?psi *)
@@ -651,7 +653,7 @@
val P = read_prop "phi";
in
store_standard_thm_open (Binding.name "equal_elim_rule1")
- (Thm.equal_elim (assume eq) (assume P) |> implies_intr_list [eq, P])
+ (Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P])
end;
(* PROP ?psi == PROP ?phi ==> PROP ?phi ==> PROP ?psi *)
@@ -917,7 +919,7 @@
fun ren (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, ren t)
| ren (t $ u) = ren t $ ren u
| ren t = t;
- in equal_elim (reflexive (cert (ren (Thm.prop_of thm)))) thm end;
+ in Thm.equal_elim (Thm.reflexive (cert (ren (Thm.prop_of thm)))) thm end;
(* renaming in left-to-right order *)
@@ -937,7 +939,7 @@
in (xs'', t' $ u') end
| rename xs t = (xs, t);
in case rename xs prop of
- ([], prop') => equal_elim (reflexive (cert prop')) thm
+ ([], prop') => Thm.equal_elim (Thm.reflexive (cert prop')) thm
| _ => error "More names than abstractions in theorem"
end;
--- a/src/Pure/library.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/library.scala Sun May 23 10:38:11 2010 +0100
@@ -11,8 +11,20 @@
import javax.swing.JOptionPane
+import scala.swing.ComboBox
+import scala.swing.event.SelectionChanged
+
+
object Library
{
+ /* partial functions */
+
+ def undefined[A, B] = new PartialFunction[A, B] {
+ def apply(x: A): B = throw new NoSuchElementException("undefined")
+ def isDefinedAt(x: A) = false
+ }
+
+
/* separate */
def separate[A](s: A, list: List[A]): List[A] =
@@ -88,6 +100,31 @@
def error_dialog = simple_dialog(JOptionPane.ERROR_MESSAGE, "Error") _
+ /* zoom box */
+
+ class Zoom_Box(apply_factor: Int => Unit) extends ComboBox[String](
+ List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%"))
+ {
+ val Factor = "([0-9]+)%?"r
+ def parse(text: String): Int =
+ text match {
+ case Factor(s) =>
+ val i = Integer.parseInt(s)
+ if (10 <= i && i <= 1000) i else 100
+ case _ => 100
+ }
+ def print(i: Int): String = i.toString + "%"
+
+ makeEditable()(c => new ComboBox.BuiltInEditor(c)(text => print(parse(text)), x => x))
+ reactions += {
+ case SelectionChanged(_) => apply_factor(parse(selection.item))
+ }
+ listenTo(selection)
+ selection.index = 3
+ prototypeDisplayValue = Some("00000%")
+ }
+
+
/* timing */
def timeit[A](message: String)(e: => A) =
--- a/src/Pure/meta_simplifier.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/meta_simplifier.ML Sun May 23 10:38:11 2010 +0100
@@ -832,9 +832,9 @@
fun check_conv msg ss thm thm' =
let
- val thm'' = transitive thm thm' handle THM _ =>
- transitive thm (transitive
- (symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm')
+ val thm'' = Thm.transitive thm thm' handle THM _ =>
+ Thm.transitive thm (Thm.transitive
+ (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm')
in if msg then trace_thm (fn () => "SUCCEEDED") ss thm' else (); SOME thm'' end
handle THM _ =>
let
@@ -972,8 +972,8 @@
| some => some)))
else proc_rews ps;
in case eta_t of
- Abs _ $ _ => SOME (transitive eta_thm
- (beta_conversion false eta_t'), skel0)
+ Abs _ $ _ => SOME (Thm.transitive eta_thm
+ (Thm.beta_conversion false eta_t'), skel0)
| _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of
NONE => proc_rews (Net.match_term procs eta_t)
| some => some)
@@ -1006,7 +1006,7 @@
fun transitive1 NONE NONE = NONE
| transitive1 (SOME thm1) NONE = SOME thm1
| transitive1 NONE (SOME thm2) = SOME thm2
- | transitive1 (SOME thm1) (SOME thm2) = SOME (transitive thm1 thm2)
+ | transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2)
fun transitive2 thm = transitive1 (SOME thm);
fun transitive3 thm = transitive1 thm o SOME;
@@ -1020,7 +1020,7 @@
some as SOME thm1 =>
(case rewritec (prover, thy, maxidx) ss (Thm.rhs_of thm1) of
SOME (thm2, skel2) =>
- transitive2 (transitive thm1 thm2)
+ transitive2 (Thm.transitive thm1 thm2)
(botc skel2 ss (Thm.rhs_of thm2))
| NONE => some)
| NONE =>
@@ -1031,7 +1031,7 @@
and try_botc ss t =
(case botc skel0 ss t of
- SOME trec1 => trec1 | NONE => (reflexive t))
+ SOME trec1 => trec1 | NONE => (Thm.reflexive t))
and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 =
(case term_of t0 of
@@ -1048,16 +1048,16 @@
val ss' = add_bound ((b', T), a) ss;
val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0;
in case botc skel' ss' t' of
- SOME thm => SOME (abstract_rule a v thm)
+ SOME thm => SOME (Thm.abstract_rule a v thm)
| NONE => NONE
end
| t $ _ => (case t of
Const ("==>", _) $ _ => impc t0 ss
| Abs _ =>
- let val thm = beta_conversion false t0
+ let val thm = Thm.beta_conversion false t0
in case subc skel0 ss (Thm.rhs_of thm) of
NONE => SOME thm
- | SOME thm' => SOME (transitive thm thm')
+ | SOME thm' => SOME (Thm.transitive thm thm')
end
| _ =>
let fun appc () =
@@ -1070,11 +1070,11 @@
(case botc tskel ss ct of
SOME thm1 =>
(case botc uskel ss cu of
- SOME thm2 => SOME (combination thm1 thm2)
- | NONE => SOME (combination thm1 (reflexive cu)))
+ SOME thm2 => SOME (Thm.combination thm1 thm2)
+ | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu)))
| NONE =>
(case botc uskel ss cu of
- SOME thm1 => SOME (combination (reflexive ct) thm1)
+ SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1)
| NONE => NONE))
end
val (h, ts) = strip_comb t
@@ -1095,7 +1095,7 @@
in case botc skel ss cl of
NONE => thm
| SOME thm' => transitive3 thm
- (combination thm' (reflexive cr))
+ (Thm.combination thm' (Thm.reflexive cr))
end handle Pattern.MATCH => appc ()))
| _ => appc ()
end)
@@ -1110,7 +1110,7 @@
(fn () => "Cannot add premise as rewrite rule because it contains (type) unknowns:")
ss prem; ([], NONE))
else
- let val asm = assume prem
+ let val asm = Thm.assume prem
in (extract_safe_rrules (ss, asm), SOME asm) end
and add_rrules (rrss, asms) ss =
@@ -1119,14 +1119,14 @@
and disch r prem eq =
let
val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq);
- val eq' = implies_elim (Thm.instantiate
+ val eq' = Thm.implies_elim (Thm.instantiate
([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong)
- (implies_intr prem eq)
+ (Thm.implies_intr prem eq)
in if not r then eq' else
let
val (prem', concl) = Thm.dest_implies lhs;
val (prem'', _) = Thm.dest_implies rhs
- in transitive (transitive
+ in Thm.transitive (Thm.transitive
(Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)])
Drule.swap_prems_eq) eq')
(Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)])
@@ -1182,7 +1182,7 @@
in mut_impc prems concl rrss asms (prem' :: prems')
(rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true)
(take i prems)
- (Drule.imp_cong_rule eqn (reflexive (Drule.list_implies
+ (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies
(drop i prems, concl))))) :: eqns)
ss (length prems') ~1
end
@@ -1197,13 +1197,13 @@
in (case botc skel0 ss1 conc of
NONE => (case thm1 of
NONE => NONE
- | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (reflexive conc)))
+ | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc)))
| SOME thm2 =>
let val thm2' = disch false prem1 thm2
in (case thm1 of
NONE => SOME thm2'
| SOME thm1' =>
- SOME (transitive (Drule.imp_cong_rule thm1' (reflexive conc)) thm2'))
+ SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2'))
end)
end
@@ -1321,7 +1321,7 @@
val keys = sort_distinct (rev_order o int_ord) (map #2 keylist)
in map (AList.find (op =) keylist) keys end;
-val rev_defs = sort_lhs_depths o map symmetric;
+val rev_defs = sort_lhs_depths o map Thm.symmetric;
fun fold_rule defs = fold rewrite_rule (rev_defs defs);
fun fold_goals_tac defs = EVERY (map rewrite_goals_tac (rev_defs defs));
--- a/src/Pure/old_goals.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/old_goals.ML Sun May 23 10:38:11 2010 +0100
@@ -134,7 +134,7 @@
val maxidx = Thm.maxidx_of state
val cterm = Thm.cterm_of (Thm.theory_of_thm state)
val chyps = map cterm hyps
- val hypths = map assume chyps
+ val hypths = map Thm.assume chyps
val subprems = map (Thm.forall_elim_vars 0) hypths
val fparams = map Free params
val cparams = map cterm fparams
@@ -165,7 +165,7 @@
val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
val emBs = map (cterm o embed) (prems_of st')
- val Cth = implies_elim_list st' (map (elim o assume) emBs)
+ val Cth = implies_elim_list st' (map (elim o Thm.assume) emBs)
in (*restore the unknowns to the hypotheses*)
free_instantiate (map swap_ctpair insts @
map mk_subgoal_swap_ctpair subgoal_insts)
@@ -175,7 +175,7 @@
end
(*function to replace the current subgoal*)
fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state
- in Seq.maps next (tacf subprems (trivial (cterm concl))) end;
+ in Seq.maps next (tacf subprems (Thm.trivial (cterm concl))) end;
end;
@@ -308,7 +308,7 @@
(*discharges assumptions from state in the order they appear in goal;
checks (if requested) that resulting theorem is equivalent to goal. *)
fun mkresult (check,state) =
- let val state = Seq.hd (flexflex_rule state)
+ let val state = Seq.hd (Thm.flexflex_rule state)
handle THM _ => state (*smash flexflex pairs*)
val ngoals = nprems_of state
val ath = implies_intr_list cAs state
@@ -522,7 +522,7 @@
| i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
in
Drule.export_without_context (implies_intr_list As
- (check (Seq.pull (EVERY (f asms) (trivial B)))))
+ (check (Seq.pull (EVERY (f asms) (Thm.trivial B)))))
end;
--- a/src/Pure/proofterm.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/proofterm.ML Sun May 23 10:38:11 2010 +0100
@@ -118,11 +118,6 @@
arity_proof: theory -> string * sort list * class -> proof} -> unit
val axm_proof: string -> term -> proof
val oracle_proof: string -> term -> oracle * proof
- val promise_proof: theory -> serial -> term -> proof
- val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
- val thm_proof: theory -> string -> term list -> term ->
- (serial * proof_body future) list -> proof_body -> pthm * proof
- val get_name: term list -> term -> proof -> string
(** rewriting on proof terms **)
val add_prf_rrule: proof * proof -> theory -> theory
@@ -134,6 +129,17 @@
val rewrite_proof_notypes: (proof * proof) list *
(typ list -> proof -> (proof * proof) option) list -> proof -> proof
val rew_proof: theory -> proof -> proof
+
+ val promise_proof: theory -> serial -> term -> proof
+ val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
+ val unconstrain_thm_proofs: bool Unsynchronized.ref
+ val thm_proof: theory -> string -> sort list -> term list -> term ->
+ (serial * proof_body future) list -> proof_body -> pthm * proof
+ val unconstrain_thm_proof: theory -> sort list -> term ->
+ (serial * proof_body future) list -> proof_body -> pthm * proof
+ val get_name: term list -> term -> proof -> string
+ val get_name_unconstrained: sort list -> term list -> term -> proof -> string
+ val guess_name: proof -> string
end
structure Proofterm : PROOFTERM =
@@ -344,7 +350,7 @@
fun change_type opTs (PAxm (name, prop, _)) = PAxm (name, prop, opTs)
| change_type (SOME [T]) (OfClass (_, c)) = OfClass (T, c)
| change_type opTs (Oracle (name, prop, _)) = Oracle (name, prop, opTs)
- | change_type opTs (Promise _) = error "change_type: unexpected promise"
+ | change_type opTs (Promise _) = raise Fail "change_type: unexpected promise"
| change_type opTs (PThm (i, ((name, prop, _), body))) =
PThm (i, ((name, prop, opTs), body))
| change_type _ prf = prf;
@@ -1071,7 +1077,7 @@
| Oracle (_, prop, _) => prop
| Promise (_, prop, _) => prop
| PThm (_, ((_, prop, _), _)) => prop
- | _ => error "shrink: proof not in normal form");
+ | _ => raise Fail "shrink: proof not in normal form");
val vs = vars_of prop;
val (ts', ts'') = chop (length vs) ts;
val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts';
@@ -1348,9 +1354,9 @@
let
val _ = prop |> Term.exists_subterm (fn t =>
(Term.is_Free t orelse Term.is_Var t) andalso
- error ("promise_proof: illegal variable " ^ Syntax.string_of_term_global thy t));
+ raise Fail ("promise_proof: illegal variable " ^ Syntax.string_of_term_global thy t));
val _ = prop |> Term.exists_type (Term.exists_subtype
- (fn TFree (a, _) => error ("promise_proof: illegal type variable " ^ quote a)
+ (fn TFree (a, _) => raise Fail ("promise_proof: illegal type variable " ^ quote a)
| _ => false));
in Promise (i, prop, map TVar (Term.add_tvars prop [])) end;
@@ -1370,15 +1376,40 @@
val proof = rewrite_prf fst (rules, K fill :: procs) proof0;
in PBody {oracles = oracles, thms = thms, proof = proof} end;
-fun fulfill_proof_future _ [] body = Future.value body
- | fulfill_proof_future thy promises body =
+fun fulfill_proof_future _ [] postproc body = Future.value (postproc body)
+ | fulfill_proof_future thy promises postproc body =
Future.fork_deps (map snd promises) (fn () =>
- fulfill_norm_proof thy (map (apsnd Future.join) promises) body);
+ postproc (fulfill_norm_proof thy (map (apsnd Future.join) promises) body));
+
+
+(***** abstraction over sort constraints *****)
+
+fun unconstrainT_prf thy (atyp_map, constraints) =
+ let
+ fun hyp_map hyp =
+ (case AList.lookup (op =) constraints hyp of
+ SOME t => Hyp t
+ | NONE => raise Fail "unconstrainT_prf: missing constraint");
+
+ val typ = Term_Subst.map_atypsT_same (Type.strip_sorts o atyp_map);
+ fun ofclass (ty, c) =
+ let val ty' = Term.map_atyps atyp_map ty;
+ in the_single (of_sort_proof thy hyp_map (ty', [c])) end;
+ in
+ Same.commit (map_proof_same (Term_Subst.map_types_same typ) typ ofclass)
+ #> fold_rev (implies_intr_proof o snd) constraints
+ end;
+
+fun unconstrainT_body thy constrs (PBody {oracles, thms, proof}) =
+ PBody
+ {oracles = oracles, (* FIXME merge (!), unconstrain (!?!) *)
+ thms = thms, (* FIXME merge (!) *)
+ proof = unconstrainT_prf thy constrs proof};
(***** theorems *****)
-fun thm_proof thy name hyps concl promises body =
+fun prepare_thm_proof do_unconstrain thy name shyps hyps concl promises body =
let
val PBody {oracles = oracles0, thms = thms0, proof = prf} = body;
val prop = Logic.list_implies (hyps, concl);
@@ -1387,24 +1418,48 @@
if member (op =) nvs ixn then SOME v else NONE) (vars_of prop) @
map SOME (frees_of prop);
+ val (postproc, ofclasses, prop1, args1) =
+ if do_unconstrain then
+ let
+ val ((atyp_map, constraints), prop1) = Logic.unconstrainT shyps prop;
+ val postproc = unconstrainT_body thy (atyp_map, constraints);
+ val args1 =
+ (map o Option.map o Term.map_types o Term.map_atyps)
+ (Type.strip_sorts o atyp_map) args;
+ in (postproc, map (OfClass o fst) constraints, prop1, args1) end
+ else (I, [], prop, args);
+ val argsP = ofclasses @ map Hyp hyps;
+
val proof0 =
if ! proofs = 2 then
#4 (shrink_proof [] 0 (rew_proof thy (fold_rev implies_intr_proof hyps prf)))
else MinProof;
val body0 = PBody {oracles = oracles0, thms = thms0, proof = proof0};
- fun new_prf () = (serial (), name, prop, fulfill_proof_future thy promises body0);
- val (i, name, prop, body') =
+ fun new_prf () = (serial (), fulfill_proof_future thy promises postproc body0);
+ val (i, body') =
(case strip_combt (fst (strip_combP prf)) of
(PThm (i, ((old_name, prop', NONE), body')), args') =>
- if (old_name = "" orelse old_name = name) andalso prop = prop' andalso args = args'
- then (i, name, prop, body')
+ if (old_name = "" orelse old_name = name) andalso prop1 = prop' andalso args = args'
+ then (i, body')
else new_prf ()
| _ => new_prf ());
- val head = PThm (i, ((name, prop, NONE), body'));
- in
- ((i, (name, prop, body')), proof_combP (proof_combt' (head, args), map Hyp hyps))
- end;
+ val head = PThm (i, ((name, prop1, NONE), body'));
+ in ((i, (name, prop1, body')), head, args, argsP, args1) end;
+
+val unconstrain_thm_proofs = Unsynchronized.ref false;
+
+fun thm_proof thy name shyps hyps concl promises body =
+ let val (pthm, head, args, argsP, _) =
+ prepare_thm_proof (! unconstrain_thm_proofs) thy name shyps hyps concl promises body
+ in (pthm, proof_combP (proof_combt' (head, args), argsP)) end;
+
+fun unconstrain_thm_proof thy shyps concl promises body =
+ let
+ val (pthm, head, _, _, args) =
+ prepare_thm_proof true thy "" shyps [] concl promises body
+ in (pthm, proof_combt' (head, args)) end;
+
fun get_name hyps prop prf =
let val prop = Logic.list_implies (hyps, prop) in
@@ -1413,6 +1468,20 @@
| _ => "")
end;
+fun get_name_unconstrained shyps hyps prop prf =
+ let val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)) in
+ (case strip_combt (fst (strip_combP prf)) of
+ (PThm (_, ((name, prop', _), _)), _) => if prop = prop' then name else ""
+ | _ => "")
+ end;
+
+fun guess_name (PThm (_, ((name, _, _), _))) = name
+ | guess_name (prf %% Hyp _) = guess_name prf
+ | guess_name (prf %% OfClass _) = guess_name prf
+ | guess_name (prf % NONE) = guess_name prf
+ | guess_name (prf % SOME (Var _)) = guess_name prf
+ | guess_name _ = "";
+
end;
structure Basic_Proofterm : BASIC_PROOFTERM = Proofterm;
--- a/src/Pure/pure_setup.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/pure_setup.ML Sun May 23 10:38:11 2010 +0100
@@ -9,7 +9,7 @@
val theory = ThyInfo.get_theory;
Context.>> (Context.map_theory
- (OuterSyntax.process_file (Path.explode "Pure.thy") #>
+ (Outer_Syntax.process_file (Path.explode "Pure.thy") #>
Theory.end_theory));
structure Pure = struct val thy = ML_Context.the_global_context () end;
Context.set_thread_data NONE;
--- a/src/Pure/tactic.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/tactic.ML Sun May 23 10:38:11 2010 +0100
@@ -156,7 +156,7 @@
fun dmatch_tac rls = ematch_tac (map make_elim rls);
(*Smash all flex-flex disagreement pairs in the proof state.*)
-val flexflex_tac = PRIMSEQ flexflex_rule;
+val flexflex_tac = PRIMSEQ Thm.flexflex_rule;
(*Remove duplicate subgoals.*)
val perm_tac = PRIMITIVE oo Thm.permute_prems;
@@ -185,7 +185,7 @@
(*Determine print names of goal parameters (reversed)*)
fun innermost_params i st =
- let val (_, _, Bi, _) = dest_state (st, i)
+ let val (_, _, Bi, _) = Thm.dest_state (st, i)
in Term.rename_wrt_term Bi (Logic.strip_params Bi) end;
--- a/src/Pure/thm.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Pure/thm.ML Sun May 23 10:38:11 2010 +0100
@@ -139,7 +139,6 @@
val of_class: ctyp * class -> thm
val strip_shyps: thm -> thm
val unconstrainT: thm -> thm
- val legacy_unconstrainT: ctyp -> thm -> thm
val legacy_freezeT: thm -> thm
val assumption: int -> thm -> thm Seq.seq
val eq_assumption: int -> thm -> thm
@@ -586,17 +585,17 @@
(* closed derivations with official name *)
fun derivation_name thm =
- Pt.get_name (hyps_of thm) (prop_of thm) (Pt.proof_of (raw_body thm));
+ Pt.guess_name (Pt.proof_of (raw_body thm)); (* FIXME tmp *)
fun name_derivation name (thm as Thm (der, args)) =
let
val Deriv {promises, body} = der;
- val {thy_ref, hyps, prop, tpairs, ...} = args;
+ val {thy_ref, shyps, hyps, prop, tpairs, ...} = args;
val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]);
val ps = map (apsnd (Future.map proof_body_of)) promises;
val thy = Theory.deref thy_ref;
- val (pthm, proof) = Pt.thm_proof thy name hyps prop ps body;
+ val (pthm, proof) = Pt.thm_proof thy name shyps hyps prop ps body;
val der' = make_deriv [] [] [pthm] proof;
val _ = Theory.check_thy thy;
in Thm (der', args) end;
@@ -1222,48 +1221,33 @@
end;
(*Internalize sort constraints of type variables*)
-fun unconstrainT (thm as Thm (_, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
+fun unconstrainT (thm as Thm (der, args)) =
let
+ val Deriv {promises, body} = der;
+ val {thy_ref, shyps, hyps, tpairs, prop, ...} = args;
+
fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]);
-
val _ = null hyps orelse err "illegal hyps";
val _ = null tpairs orelse err "unsolved flex-flex constraints";
val tfrees = rev (Term.add_tfree_names prop []);
val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
- val (_, prop') = Logic.unconstrainT shyps prop;
+ val ps = map (apsnd (Future.map proof_body_of)) promises;
+ val thy = Theory.deref thy_ref;
+ val (pthm as (_, (_, prop', _)), proof) = Pt.unconstrain_thm_proof thy shyps prop ps body;
+ val der' = make_deriv [] [] [pthm] proof;
+ val _ = Theory.check_thy thy;
in
- Thm (deriv_rule0 (Pt.PAxm ("Pure.unconstrainT", prop, SOME [])),
+ Thm (der',
{thy_ref = thy_ref,
tags = [],
maxidx = maxidx_of_term prop',
shyps = [[]], (*potentially redundant*)
- hyps = hyps,
- tpairs = tpairs,
+ hyps = [],
+ tpairs = [],
prop = prop'})
end;
-fun legacy_unconstrainT
- (Ctyp {thy_ref = thy_ref1, T, ...})
- (th as Thm (_, {thy_ref = thy_ref2, maxidx, shyps, hyps, tpairs, prop, ...})) =
- let
- val ((x, i), S) = Term.dest_TVar T handle TYPE _ =>
- raise THM ("unconstrainT: not a type variable", 0, [th]);
- val T' = TVar ((x, i), []);
- val unconstrain = Term.map_types (Term.map_atyps (fn U => if U = T then T' else U));
- val constraints = Logic.mk_of_sort (T', S);
- in
- Thm (deriv_rule0 (Pt.PAxm ("Pure.unconstrainT", prop, SOME [])),
- {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2),
- tags = [],
- maxidx = Int.max (maxidx, i),
- shyps = Sorts.remove_sort S shyps,
- hyps = hyps,
- tpairs = map (pairself unconstrain) tpairs,
- prop = Logic.list_implies (constraints, unconstrain prop)})
- end;
-
-
(* Replace all TFrees not fixed or in the hyps by new TVars *)
fun varifyT_global' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
let
--- a/src/Tools/Code/code_eval.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_eval.ML Sun May 23 10:38:11 2010 +0100
@@ -204,26 +204,24 @@
local
-structure P = OuterParse
-and K = OuterKeyword
-
val datatypesK = "datatypes";
val functionsK = "functions";
val fileK = "file";
val andK = "and"
-val _ = List.app K.keyword [datatypesK, functionsK];
+val _ = List.app Keyword.keyword [datatypesK, functionsK];
-val parse_datatype = (P.name --| P.$$$ "=" -- (P.term ::: (Scan.repeat (P.$$$ "|" |-- P.term))));
+val parse_datatype =
+ Parse.name --| Parse.$$$ "=" -- (Parse.term ::: (Scan.repeat (Parse.$$$ "|" |-- Parse.term)));
in
val _ =
- OuterSyntax.command "code_reflect" "enrich runtime environment with generated code"
- K.thy_decl (P.name -- Scan.optional (P.$$$ datatypesK |-- (parse_datatype
- ::: Scan.repeat (P.$$$ andK |-- parse_datatype))) []
- -- Scan.optional (P.$$$ functionsK |-- Scan.repeat1 P.name) []
- -- Scan.option (P.$$$ fileK |-- P.name)
+ Outer_Syntax.command "code_reflect" "enrich runtime environment with generated code"
+ Keyword.thy_decl (Parse.name -- Scan.optional (Parse.$$$ datatypesK |-- (parse_datatype
+ ::: Scan.repeat (Parse.$$$ andK |-- parse_datatype))) []
+ -- Scan.optional (Parse.$$$ functionsK |-- Scan.repeat1 Parse.name) []
+ -- Scan.option (Parse.$$$ fileK |-- Parse.name)
>> (fn (((module_name, raw_datatypes), raw_functions), some_file) => Toplevel.theory
(code_reflect_cmd raw_datatypes raw_functions module_name some_file)));
--- a/src/Tools/Code/code_haskell.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_haskell.ML Sun May 23 10:38:11 2010 +0100
@@ -469,8 +469,8 @@
serialize_haskell module_prefix module_name string_classes));
val _ =
- OuterSyntax.command "code_monad" "define code syntax for monads" OuterKeyword.thy_decl (
- OuterParse.term_group -- OuterParse.name >> (fn (raw_bind, target) =>
+ Outer_Syntax.command "code_monad" "define code syntax for monads" Keyword.thy_decl (
+ Parse.term_group -- Parse.name >> (fn (raw_bind, target) =>
Toplevel.theory (add_monad target raw_bind))
);
--- a/src/Tools/Code/code_preproc.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_preproc.ML Sun May 23 10:38:11 2010 +0100
@@ -479,8 +479,8 @@
end;
val _ =
- OuterSyntax.improper_command "print_codeproc" "print code preprocessor setup"
- OuterKeyword.diag (Scan.succeed
+ Outer_Syntax.improper_command "print_codeproc" "print code preprocessor setup"
+ Keyword.diag (Scan.succeed
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
(print_codeproc o Toplevel.theory_of)));
--- a/src/Tools/Code/code_printer.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_printer.ML Sun May 23 10:38:11 2010 +0100
@@ -72,9 +72,9 @@
val parse_infix: ('a -> 'b) -> lrx * int -> string
-> int * ((fixity -> 'b -> Pretty.T)
-> fixity -> 'a list -> Pretty.T)
- val parse_syntax: ('a -> 'b) -> OuterParse.token list
+ val parse_syntax: ('a -> 'b) -> Token.T list
-> (int * ((fixity -> 'b -> Pretty.T)
- -> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list
+ -> fixity -> 'a list -> Pretty.T)) option * Token.T list
val simple_const_syntax: simple_const_syntax -> proto_const_syntax
val activate_const_syntax: theory -> literals
-> proto_const_syntax -> Code_Thingol.naming -> const_syntax * Code_Thingol.naming
@@ -329,15 +329,15 @@
fun parse_syntax prep_arg xs =
Scan.option ((
- ((OuterParse.$$$ infixK >> K X)
- || (OuterParse.$$$ infixlK >> K L)
- || (OuterParse.$$$ infixrK >> K R))
- -- OuterParse.nat >> parse_infix prep_arg
+ ((Parse.$$$ infixK >> K X)
+ || (Parse.$$$ infixlK >> K L)
+ || (Parse.$$$ infixrK >> K R))
+ -- Parse.nat >> parse_infix prep_arg
|| Scan.succeed (parse_mixfix prep_arg))
- -- OuterParse.string
+ -- Parse.string
>> (fn (parse, s) => parse s)) xs;
-val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK];
+val _ = List.app Keyword.keyword [infixK, infixlK, infixrK];
(** module name spaces **)
--- a/src/Tools/Code/code_target.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_target.ML Sun May 23 10:38:11 2010 +0100
@@ -19,14 +19,13 @@
type destination
type serialization
- val parse_args: (OuterLex.token list -> 'a * OuterLex.token list)
- -> OuterLex.token list -> 'a
+ val parse_args: 'a parser -> Token.T list -> 'a
val stmt_names_of_destination: destination -> string list
val mk_serialization: string -> ('a -> unit) option
-> (Path.T option -> 'a -> unit)
-> ('a -> string * string option list)
-> 'a -> serialization
- val serialize: theory -> string -> int option -> string option -> OuterLex.token list
+ val serialize: theory -> string -> int option -> string option -> Token.T list
-> Code_Thingol.naming -> Code_Thingol.program -> string list -> serialization
val serialize_custom: theory -> string * (serializer * literals)
-> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
@@ -105,7 +104,7 @@
type serializer =
string option (*module name*)
- -> OuterLex.token list (*arguments*)
+ -> Token.T list (*arguments*)
-> (string -> string) (*labelled_name*)
-> string list (*reserved symbols*)
-> (string * Pretty.T) list (*includes*)
@@ -464,9 +463,6 @@
local
-structure P = OuterParse
-and K = OuterKeyword
-
fun zip_list (x::xs) f g =
f
#-> (fn y =>
@@ -474,9 +470,9 @@
#-> (fn xys => pair ((x, y) :: xys)));
fun parse_multi_syntax parse_thing parse_syntax =
- P.and_list1 parse_thing
- #-> (fn things => Scan.repeat1 (P.$$$ "(" |-- P.name --
- (zip_list things parse_syntax (P.$$$ "and")) --| P.$$$ ")"));
+ Parse.and_list1 parse_thing
+ #-> (fn things => Scan.repeat1 (Parse.$$$ "(" |-- Parse.name --
+ (zip_list things parse_syntax (Parse.$$$ "and")) --| Parse.$$$ ")"));
in
@@ -498,7 +494,7 @@
val allow_abort_cmd = gen_allow_abort Code.read_const;
fun parse_args f args =
- case Scan.read OuterLex.stopper f args
+ case Scan.read Token.stopper f args
of SOME x => x
| NONE => error "Bad serializer arguments";
@@ -508,75 +504,79 @@
val (inK, module_nameK, fileK) = ("in", "module_name", "file");
val code_exprP =
- (Scan.repeat1 P.term_group
- -- Scan.repeat (P.$$$ inK |-- P.name
- -- Scan.option (P.$$$ module_nameK |-- P.name)
- -- Scan.option (P.$$$ fileK |-- P.name)
- -- Scan.optional (P.$$$ "(" |-- Args.parse --| P.$$$ ")") []
+ (Scan.repeat1 Parse.term_group
+ -- Scan.repeat (Parse.$$$ inK |-- Parse.name
+ -- Scan.option (Parse.$$$ module_nameK |-- Parse.name)
+ -- Scan.option (Parse.$$$ fileK |-- Parse.name)
+ -- Scan.optional (Parse.$$$ "(" |-- Args.parse --| Parse.$$$ ")") []
) >> (fn (raw_cs, seris) => export_code_cmd raw_cs seris));
-val _ = List.app OuterKeyword.keyword [inK, module_nameK, fileK];
+val _ = List.app Keyword.keyword [inK, module_nameK, fileK];
val _ =
- OuterSyntax.command "code_class" "define code syntax for class" K.thy_decl (
- parse_multi_syntax P.xname (Scan.option P.string)
+ Outer_Syntax.command "code_class" "define code syntax for class" Keyword.thy_decl (
+ parse_multi_syntax Parse.xname (Scan.option Parse.string)
>> (Toplevel.theory oo fold) (fn (target, syns) =>
fold (fn (raw_class, syn) => add_syntax_class_cmd target raw_class syn) syns)
);
val _ =
- OuterSyntax.command "code_instance" "define code syntax for instance" K.thy_decl (
- parse_multi_syntax (P.xname --| P.$$$ "::" -- P.xname) (Scan.option (P.minus >> K ()))
+ Outer_Syntax.command "code_instance" "define code syntax for instance" Keyword.thy_decl (
+ parse_multi_syntax (Parse.xname --| Parse.$$$ "::" -- Parse.xname)
+ (Scan.option (Parse.minus >> K ()))
>> (Toplevel.theory oo fold) (fn (target, syns) =>
fold (fn (raw_inst, add_del) => add_syntax_inst_cmd target raw_inst add_del) syns)
);
val _ =
- OuterSyntax.command "code_type" "define code syntax for type constructor" K.thy_decl (
- parse_multi_syntax P.xname (Code_Printer.parse_syntax I)
+ Outer_Syntax.command "code_type" "define code syntax for type constructor" Keyword.thy_decl (
+ parse_multi_syntax Parse.xname (Code_Printer.parse_syntax I)
>> (Toplevel.theory oo fold) (fn (target, syns) =>
fold (fn (raw_tyco, syn) => add_syntax_tyco_cmd target raw_tyco syn) syns)
);
val _ =
- OuterSyntax.command "code_const" "define code syntax for constant" K.thy_decl (
- parse_multi_syntax P.term_group (Code_Printer.parse_syntax fst)
+ Outer_Syntax.command "code_const" "define code syntax for constant" Keyword.thy_decl (
+ parse_multi_syntax Parse.term_group (Code_Printer.parse_syntax fst)
>> (Toplevel.theory oo fold) (fn (target, syns) =>
fold (fn (raw_const, syn) => add_syntax_const_simple_cmd target raw_const syn) syns)
);
val _ =
- OuterSyntax.command "code_reserved" "declare words as reserved for target language" K.thy_decl (
- P.name -- Scan.repeat1 P.name
+ Outer_Syntax.command "code_reserved" "declare words as reserved for target language"
+ Keyword.thy_decl (
+ Parse.name -- Scan.repeat1 Parse.name
>> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)
);
val _ =
- OuterSyntax.command "code_include" "declare piece of code to be included in generated code" K.thy_decl (
- P.name -- P.name -- (P.text :|-- (fn "-" => Scan.succeed NONE
- | s => Scan.optional (P.$$$ "attach" |-- Scan.repeat1 P.term) [] >> pair s >> SOME))
+ Outer_Syntax.command "code_include" "declare piece of code to be included in generated code"
+ Keyword.thy_decl (
+ Parse.name -- Parse.name -- (Parse.text :|-- (fn "-" => Scan.succeed NONE
+ | s => Scan.optional (Parse.$$$ "attach" |-- Scan.repeat1 Parse.term) [] >> pair s >> SOME))
>> (fn ((target, name), content_consts) =>
(Toplevel.theory o add_include_cmd target) (name, content_consts))
);
val _ =
- OuterSyntax.command "code_modulename" "alias module to other name" K.thy_decl (
- P.name -- Scan.repeat1 (P.name -- P.name)
+ Outer_Syntax.command "code_modulename" "alias module to other name" Keyword.thy_decl (
+ Parse.name -- Scan.repeat1 (Parse.name -- Parse.name)
>> (fn (target, modlnames) => (Toplevel.theory o fold (add_module_alias target)) modlnames)
);
val _ =
- OuterSyntax.command "code_abort" "permit constant to be implemented as program abort" K.thy_decl (
- Scan.repeat1 P.term_group >> (Toplevel.theory o fold allow_abort_cmd)
+ Outer_Syntax.command "code_abort" "permit constant to be implemented as program abort"
+ Keyword.thy_decl (
+ Scan.repeat1 Parse.term_group >> (Toplevel.theory o fold allow_abort_cmd)
);
val _ =
- OuterSyntax.command "export_code" "generate executable code for constants"
- K.diag (P.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
+ Outer_Syntax.command "export_code" "generate executable code for constants"
+ Keyword.diag (Parse.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.theory_of)));
fun shell_command thyname cmd = Toplevel.program (fn _ =>
- (use_thy thyname; case Scan.read OuterLex.stopper (P.!!! code_exprP)
- ((filter OuterLex.is_proper o OuterSyntax.scan Position.none) cmd)
+ (use_thy thyname; case Scan.read Token.stopper (Parse.!!! code_exprP)
+ ((filter Token.is_proper o Outer_Syntax.scan Position.none) cmd)
of SOME f => (writeln "Now generating code..."; f (theory thyname))
| NONE => error ("Bad directive " ^ quote cmd)))
handle Runtime.TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure;
--- a/src/Tools/Code/code_thingol.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Code/code_thingol.ML Sun May 23 10:38:11 2010 +0100
@@ -915,23 +915,21 @@
local
-structure P = OuterParse
-and K = OuterKeyword
-
fun code_thms_cmd thy = code_thms thy o op @ o read_const_exprs thy;
fun code_deps_cmd thy = code_deps thy o op @ o read_const_exprs thy;
in
val _ =
- OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag
- (Scan.repeat1 P.term_group
+ Outer_Syntax.improper_command "code_thms" "print system of code equations for code" Keyword.diag
+ (Scan.repeat1 Parse.term_group
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
val _ =
- OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag
- (Scan.repeat1 P.term_group
+ Outer_Syntax.improper_command "code_deps" "visualize dependencies of code equations for code"
+ Keyword.diag
+ (Scan.repeat1 Parse.term_group
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
--- a/src/Tools/Compute_Oracle/compute.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/Compute_Oracle/compute.ML Sun May 23 10:38:11 2010 +0100
@@ -391,9 +391,9 @@
fun export_thm thy hyps shyps prop =
let
val th = export_oracle (thy, hyps, shyps, prop)
- val hyps = map (fn h => assume (cterm_of thy h)) hyps
+ val hyps = map (fn h => Thm.assume (cterm_of thy h)) hyps
in
- fold (fn h => fn p => implies_elim p h) hyps th
+ fold (fn h => fn p => Thm.implies_elim p h) hyps th
end
(* --------- Rewrite ----------- *)
--- a/src/Tools/IsaPlanner/rw_inst.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/IsaPlanner/rw_inst.ML Sun May 23 10:38:11 2010 +0100
@@ -54,13 +54,13 @@
(* beta contract the theorem *)
fun beta_contract thm =
- equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm;
+ Thm.equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm;
(* beta-eta contract the theorem *)
fun beta_eta_contract thm =
let
- val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
- val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
+ val thm2 = Thm.equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
+ val thm3 = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
in thm3 end;
--- a/src/Tools/WWW_Find/find_theorems.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/WWW_Find/find_theorems.ML Sun May 23 10:38:11 2010 +0100
@@ -158,21 +158,17 @@
end;
-structure P = OuterParse
- and K = OuterKeyword
- and FT = Find_Theorems;
-
val criterion =
- P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> Find_Theorems.Name ||
- P.reserved "intro" >> K Find_Theorems.Intro ||
- P.reserved "elim" >> K Find_Theorems.Elim ||
- P.reserved "dest" >> K Find_Theorems.Dest ||
- P.reserved "solves" >> K Find_Theorems.Solves ||
- P.reserved "simp" |-- P.!!! (P.$$$ ":" |-- P.term) >> Find_Theorems.Simp ||
- P.term >> Find_Theorems.Pattern;
+ Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.xname) >> Find_Theorems.Name ||
+ Parse.reserved "intro" >> K Find_Theorems.Intro ||
+ Parse.reserved "elim" >> K Find_Theorems.Elim ||
+ Parse.reserved "dest" >> K Find_Theorems.Dest ||
+ Parse.reserved "solves" >> K Find_Theorems.Solves ||
+ Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Find_Theorems.Simp ||
+ Parse.term >> Find_Theorems.Pattern;
val parse_query =
- Scan.repeat (((Scan.option P.minus >> is_none) -- criterion));
+ Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion));
fun app_index f xs = fold_index (fn x => K (f x)) xs ();
@@ -194,8 +190,8 @@
fun get_query () =
query
|> (fn s => s ^ ";")
- |> OuterSyntax.scan Position.start
- |> filter OuterLex.is_proper
+ |> Outer_Syntax.scan Position.start
+ |> filter Token.is_proper
|> Scan.error parse_query
|> fst;
--- a/src/Tools/WWW_Find/unicode_symbols.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/WWW_Find/unicode_symbols.ML Sun May 23 10:38:11 2010 +0100
@@ -114,8 +114,6 @@
local (* Parser *)
-structure P = OuterParse;
-
fun $$$ kw = Scan.one (is_keyword kw) >> str_of_token;
val hex_code = Scan.one is_hex_code >> int_of_code;
val ascii_sym = Scan.one is_ascii_sym >> str_of_token;
@@ -129,7 +127,7 @@
in
-val line = (symbol -- unicode --| font -- abbr) >> P.triple1;
+val line = (symbol -- unicode --| font -- abbr) >> Parse.triple1;
val symbols_path =
[getenv "ISABELLE_HOME", "etc", "symbols"]
--- a/src/Tools/coherent.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/coherent.ML Sun May 23 10:38:11 2010 +0100
@@ -41,17 +41,13 @@
val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1);
-local open Conv in
-
fun rulify_elim_conv ct =
- if is_atomic (Logic.strip_imp_concl (term_of ct)) then all_conv ct
- else concl_conv (length (Logic.strip_imp_prems (term_of ct)))
- (rewr_conv (symmetric Data.atomize_elimL) then_conv
- MetaSimplifier.rewrite true (map symmetric
+ if is_atomic (Logic.strip_imp_concl (term_of ct)) then Conv.all_conv ct
+ else Conv.concl_conv (length (Logic.strip_imp_prems (term_of ct)))
+ (Conv.rewr_conv (Thm.symmetric Data.atomize_elimL) then_conv
+ MetaSimplifier.rewrite true (map Thm.symmetric
[Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct
-end;
-
fun rulify_elim th = Simplifier.norm_hhf (Conv.fconv_rule rulify_elim_conv th);
(* Decompose elimination rule of the form
--- a/src/Tools/eqsubst.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/eqsubst.ML Sun May 23 10:38:11 2010 +0100
@@ -556,7 +556,7 @@
(Scan.succeed false);
val ith_syntax =
- Scan.optional (Args.parens (Scan.repeat OuterParse.nat)) [0];
+ Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0];
(* combination method that takes a flag (true indicates that subst
should be done to an assumption, false = apply to the conclusion of
--- a/src/Tools/induct.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/induct.ML Sun May 23 10:38:11 2010 +0100
@@ -254,8 +254,8 @@
end;
val _ =
- OuterSyntax.improper_command "print_induct_rules" "print induction and cases rules"
- OuterKeyword.diag (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
+ Outer_Syntax.improper_command "print_induct_rules" "print induction and cases rules"
+ Keyword.diag (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
Toplevel.keep (print_rules o Toplevel.context_of)));
@@ -845,8 +845,6 @@
(** concrete syntax **)
-structure P = OuterParse;
-
val arbitraryN = "arbitrary";
val takingN = "taking";
val ruleN = "rule";
@@ -892,7 +890,7 @@
Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan;
val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |--
- P.and_list1' (Scan.repeat (unless_more_args free))) [];
+ Parse.and_list1' (Scan.repeat (unless_more_args free))) [];
val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |--
Scan.repeat1 (unless_more_args inst)) [];
@@ -902,7 +900,7 @@
val cases_setup =
Method.setup @{binding cases}
(Args.mode no_simpN --
- (P.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >>
+ (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >>
(fn (no_simp, (insts, opt_rule)) => fn ctxt =>
METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL
(cases_tac ctxt (not no_simp) insts opt_rule facts)))))
@@ -910,11 +908,12 @@
val induct_setup =
Method.setup @{binding induct}
- (Args.mode no_simpN -- (P.and_list' (Scan.repeat (unless_more_args def_inst)) --
+ (Args.mode no_simpN -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
(arbitrary -- taking -- Scan.option induct_rule)) >>
(fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt =>
RAW_METHOD_CASES (fn facts =>
- Seq.DETERM (HEADGOAL (induct_tac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))))
+ Seq.DETERM
+ (HEADGOAL (induct_tac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))))
"induction on types or predicates/sets";
val coinduct_setup =
--- a/src/Tools/induct_tacs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/induct_tacs.ML Sun May 23 10:38:11 2010 +0100
@@ -116,8 +116,7 @@
val opt_rules = Scan.option (rule_spec |-- Attrib.thms);
val varss =
- OuterParse.and_list'
- (Scan.repeat (Scan.unless rule_spec (Scan.lift (Args.maybe Args.name_source))));
+ Parse.and_list' (Scan.repeat (Scan.unless rule_spec (Scan.lift (Args.maybe Args.name_source))));
in
--- a/src/Tools/intuitionistic.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/intuitionistic.ML Sun May 23 10:38:11 2010 +0100
@@ -71,7 +71,7 @@
val destN = "dest";
fun modifier name kind kind' att =
- Args.$$$ name |-- (kind >> K NONE || kind' |-- OuterParse.nat --| Args.colon >> SOME)
+ Args.$$$ name |-- (kind >> K NONE || kind' |-- Parse.nat --| Args.colon >> SOME)
>> (pair (I: Proof.context -> Proof.context) o att);
val modifiers =
@@ -87,7 +87,7 @@
fun method_setup name =
Method.setup name
- (Scan.lift (Scan.option OuterParse.nat) --| Method.sections modifiers >>
+ (Scan.lift (Scan.option Parse.nat) --| Method.sections modifiers >>
(fn opt_lim => fn ctxt =>
SIMPLE_METHOD' (Object_Logic.atomize_prems_tac THEN' prover_tac ctxt opt_lim)))
"intuitionistic proof search";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/dist-template/etc/isabelle-jedit.css Sun May 23 10:38:11 2010 +0100
@@ -0,0 +1,13 @@
+/* additional style file for Isabelle/jEdit output */
+
+pre.message { margin-top: 0.3ex; background-color: #F0F0F0; }
+
+.writeln { }
+.priority { }
+.tracing { background-color: #EAF8FF; }
+.warning { background-color: #EEE8AA; }
+.error { background-color: #FFC1C1; }
+.debug { background-color: #FFE4E1; }
+
+.hilite { background-color: #FFFACD; }
+
--- a/src/Tools/jEdit/dist-template/modes/isabelle.xml Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/dist-template/modes/isabelle.xml Sun May 23 10:38:11 2010 +0100
@@ -7,10 +7,8 @@
<PROPERTY NAME="commentStart" VALUE="(*"/>
<PROPERTY NAME="commentEnd" VALUE="*)"/>
<PROPERTY NAME="noWordSep" VALUE="_'.?"/>
- <PROPERTY NAME="indentOpenBrackets" VALUE="{"/>
- <PROPERTY NAME="indentCloseBrackets" VALUE="}"/>
- <PROPERTY NAME="unalignedOpenBrackets" VALUE="(" />
- <PROPERTY NAME="unalignedCloseBrackets" VALUE=")" />
+ <PROPERTY NAME="unalignedOpenBrackets" VALUE="{[(" />
+ <PROPERTY NAME="unalignedCloseBrackets" VALUE=")]}" />
<PROPERTY NAME="tabSize" VALUE="2" />
<PROPERTY NAME="indentSize" VALUE="2" />
</PROPS>
--- a/src/Tools/jEdit/src/jedit/document_view.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/document_view.scala Sun May 23 10:38:11 2010 +0100
@@ -25,11 +25,11 @@
{
def choose_color(command: Command, doc: Document): Color =
{
- doc.current_state(command).map(_.status) match {
- case Some(Command.Status.UNPROCESSED) => new Color(255, 228, 225)
- case Some(Command.Status.FINISHED) => new Color(234, 248, 255)
- case Some(Command.Status.FAILED) => new Color(255, 193, 193)
- case _ => Color.red
+ doc.current_state(command).status match {
+ case Command.Status.UNPROCESSED => new Color(255, 228, 225)
+ case Command.Status.FINISHED => new Color(234, 248, 255)
+ case Command.Status.FAILED => new Color(255, 193, 193)
+ case Command.Status.UNDEFINED => Color.red
}
}
@@ -146,7 +146,7 @@
val offset = model.from_current(document, text_area.xyToOffset(x, y))
document.command_at(offset) match {
case Some((command, command_start)) =>
- document.current_state(command).get.type_at(offset - command_start).getOrElse(null)
+ document.current_state(command).type_at(offset - command_start) getOrElse null
case None => null
}
}
--- a/src/Tools/jEdit/src/jedit/html_panel.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/html_panel.scala Sun May 23 10:38:11 2010 +0100
@@ -23,7 +23,6 @@
import org.lobobrowser.html.domimpl.{HTMLDocumentImpl, HTMLStyleElementImpl, NodeImpl}
import org.lobobrowser.html.test.{SimpleHtmlRendererContext, SimpleUserAgentContext}
-import scala.io.Source
import scala.actors.Actor._
@@ -38,12 +37,12 @@
}
-class HTML_Panel(
- sys: Isabelle_System,
- initial_font_size: Int,
- handler: PartialFunction[HTML_Panel.Event, Unit]) extends HtmlPanel
+class HTML_Panel(system: Isabelle_System, initial_font_size: Int) extends HtmlPanel
{
- // global logging
+ /** Lobo setup **/
+
+ /* global logging */
+
Logger.getLogger("org.lobobrowser").setLevel(Level.WARNING)
@@ -57,51 +56,15 @@
def raw_px(lobo_px: Int): Int = (lobo_px * screen_resolution + 95) / 96
- /* document template */
-
- private def try_file(name: String): String =
- {
- val file = sys.platform_file(name)
- if (file.isFile) Source.fromFile(file).mkString else ""
- }
+ /* contexts and event handling */
- private def template(font_size: Int): String =
- {
- """<?xml version="1.0" encoding="utf-8"?>
-<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN"
- "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
-<html xmlns="http://www.w3.org/1999/xhtml">
-<head>
-<style media="all" type="text/css">
-""" +
- try_file("$ISABELLE_HOME/lib/html/isabelle.css") + "\n" +
- try_file("$ISABELLE_HOME_USER/etc/isabelle.css") + "\n" +
- "body { font-family: " + sys.font_family + "; font-size: " + raw_px(font_size) + "px }" +
-"""
-</style>
-</head>
-<body/>
-</html>
-"""
- }
-
- private def font_metrics(font_size: Int): FontMetrics =
- Swing_Thread.now { getFontMetrics(sys.get_font(font_size)) }
-
- private def panel_width(font_size: Int): Int =
- Swing_Thread.now {
- (getWidth() / (font_metrics(font_size).charWidth(Symbol.spc) max 1) - 4) max 20
- }
-
-
- /* actor with local state */
+ protected val handler: PartialFunction[HTML_Panel.Event, Unit] = Library.undefined
private val ucontext = new SimpleUserAgentContext
-
private val rcontext = new SimpleHtmlRendererContext(this, ucontext)
{
private def handle(event: HTML_Panel.Event): Boolean =
- if (handler != null && handler.isDefinedAt(event)) { handler(event); true }
+ if (handler.isDefinedAt(event)) { handler(event); true }
else false
override def onContextMenu(elem: HTMLElement, event: MouseEvent): Boolean =
@@ -118,54 +81,107 @@
private val builder = new DocumentBuilderImpl(ucontext, rcontext)
- private case class Init(font_size: Int)
- private case class Render(divs: List[XML.Tree])
+
+ /* document template with style sheets */
+
+ private val template_head =
+ """<?xml version="1.0" encoding="utf-8"?>
+<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN"
+ "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
+<html xmlns="http://www.w3.org/1999/xhtml">
+<head>
+<style media="all" type="text/css">
+""" +
+ system.try_read("$ISABELLE_HOME/lib/html/isabelle.css") + "\n" +
+ system.try_read("$JEDIT_HOME/etc/isabelle-jedit.css") + "\n" +
+ system.try_read("$ISABELLE_HOME_USER/etc/isabelle.css") + "\n" +
+ system.try_read("$ISABELLE_HOME_USER/etc/isabelle-jedit.css") + "\n"
+
+ private val template_tail =
+"""
+</style>
+</head>
+<body/>
+</html>
+"""
+
+ private def template(font_size: Int): String =
+ template_head +
+ "body { font-family: " + system.font_family + "; font-size: " + raw_px(font_size) + "px }" +
+ template_tail
+
+
+ /** main actor **/
+
+ /* internal messages */
+
+ private case class Resize(font_size: Int)
+ private case class Render(body: List[XML.Tree])
+ private case object Refresh
private val main_actor = actor {
- // crude double buffering
- var doc1: org.w3c.dom.Document = null
- var doc2: org.w3c.dom.Document = null
+
+ /* internal state */
+
+ var current_font_metrics: FontMetrics = null
+ var current_font_size: Int = 0
+ var current_margin: Int = 0
+ var current_body: List[XML.Tree] = Nil
- var current_font_size = 16
- var current_font_metrics: FontMetrics = null
+ def resize(font_size: Int)
+ {
+ val (font_metrics, margin) =
+ Swing_Thread.now {
+ val metrics = getFontMetrics(system.get_font(lobo_px(raw_px(font_size))))
+ (metrics, (getWidth() / (metrics.charWidth(Symbol.spc) max 1) - 4) max 20)
+ }
+ if (current_font_metrics == null ||
+ current_font_size != font_size ||
+ current_margin != margin)
+ {
+ current_font_metrics = font_metrics
+ current_font_size = font_size
+ current_margin = margin
+ refresh()
+ }
+ }
+
+ def refresh() { render(current_body) }
+
+ def render(body: List[XML.Tree])
+ {
+ current_body = body
+ val html_body =
+ current_body.flatMap(div =>
+ Pretty.formatted(List(div), current_margin, Pretty.font_metric(current_font_metrics))
+ .map(t => XML.Elem(HTML.PRE, List((Markup.CLASS, Markup.MESSAGE)), HTML.spans(t))))
+ val doc =
+ builder.parse(
+ new InputSourceImpl(new StringReader(template(current_font_size)), "http://localhost"))
+ doc.removeChild(doc.getLastChild())
+ doc.appendChild(XML.document_node(doc, XML.elem(HTML.BODY, html_body)))
+ Swing_Thread.later { setDocument(doc, rcontext) }
+ }
+
+
+ /* main loop */
+
+ resize(initial_font_size)
loop {
react {
- case Init(font_size) =>
- current_font_size = font_size
- current_font_metrics = font_metrics(lobo_px(raw_px(font_size)))
-
- val src = template(font_size)
- def parse() =
- builder.parse(new InputSourceImpl(new StringReader(src), "http://localhost"))
- doc1 = parse()
- doc2 = parse()
- Swing_Thread.now { setDocument(doc1, rcontext) }
-
- case Render(divs) =>
- val doc = doc2
- val html_body =
- divs.flatMap(div =>
- Pretty.formatted(List(div), panel_width(current_font_size),
- Pretty.font_metric(current_font_metrics))
- .map(t => XML.elem(HTML.PRE, HTML.spans(t))))
- val node = XML.document_node(doc, XML.elem(HTML.BODY, html_body))
- doc.removeChild(doc.getLastChild())
- doc.appendChild(node)
- doc2 = doc1
- doc1 = doc
- Swing_Thread.now { setDocument(doc1, rcontext) }
-
+ case Resize(font_size) => resize(font_size)
+ case Refresh => refresh()
+ case Render(body) => render(body)
case bad => System.err.println("main_actor: ignoring bad message " + bad)
}
}
}
- /* main method wrappers */
+ /* external methods */
- def init(font_size: Int) { main_actor ! Init(font_size) }
- def render(divs: List[XML.Tree]) { main_actor ! Render(divs) }
-
- init(initial_font_size)
+ def resize(font_size: Int) { main_actor ! Resize(font_size) }
+ def refresh() { main_actor ! Refresh }
+ def render(body: List[XML.Tree]) { main_actor ! Render(body) }
}
--- a/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Sun May 23 10:38:11 2010 +0100
@@ -46,7 +46,7 @@
val offset = model.from_current(document, original_offset)
document.command_at(offset) match {
case Some((command, command_start)) =>
- document.current_state(command).get.ref_at(offset - command_start) match {
+ document.current_state(command).ref_at(offset - command_start) match {
case Some(ref) =>
val begin = model.to_current(document, command_start + ref.start)
val line = buffer.getLineOfOffset(begin)
--- a/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Sun May 23 10:38:11 2010 +0100
@@ -130,7 +130,7 @@
val root = data.root
val document = model.recent_document()
for ((command, command_start) <- document.command_range(0) if !stopped) {
- root.add(document.current_state(command).get.markup_root.swing_tree((node: Markup_Node) =>
+ root.add(document.current_state(command).markup_root.swing_tree((node: Markup_Node) =>
{
val content = command.source(node.start, node.stop).replace('\n', ' ')
val id = command.id
--- a/src/Tools/jEdit/src/jedit/isabelle_token_marker.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/isabelle_token_marker.scala Sun May 23 10:38:11 2010 +0100
@@ -116,7 +116,7 @@
var next_x = start
for {
(command, command_start) <- document.command_range(from(start), from(stop))
- markup <- document.current_state(command).get.highlight.flatten
+ markup <- document.current_state(command).highlight.flatten
val abs_start = to(command_start + markup.start)
val abs_stop = to(command_start + markup.stop)
if (abs_stop > start)
--- a/src/Tools/jEdit/src/jedit/output_dockable.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/output_dockable.scala Sun May 23 10:38:11 2010 +0100
@@ -11,8 +11,12 @@
import scala.actors.Actor._
+import scala.swing.{FlowPanel, Button, CheckBox}
+import scala.swing.event.ButtonClicked
+
import javax.swing.JPanel
import java.awt.{BorderLayout, Dimension}
+import java.awt.event.{ComponentEvent, ComponentAdapter}
import org.gjt.sp.jedit.View
import org.gjt.sp.jedit.gui.DockableWindowManager
@@ -24,25 +28,92 @@
if (position == DockableWindowManager.FLOATING)
setPreferredSize(new Dimension(500, 250))
- val html_panel = new HTML_Panel(Isabelle.system, Isabelle.font_size(), null)
+ val html_panel = new HTML_Panel(Isabelle.system, scala.math.round(Isabelle.font_size()))
add(html_panel, BorderLayout.CENTER)
+ /* controls */
+
+ private val zoom = new Library.Zoom_Box(factor => { zoom_factor = factor; handle_resize() })
+ zoom.tooltip = "Zoom factor for basic font size"
+
+ private val update = new Button("Update") {
+ reactions += { case ButtonClicked(_) => handle_update() }
+ }
+ update.tooltip =
+ "<html>Update display according to the<br>state of command at caret position</html>"
+
+ private val tracing = new CheckBox("Tracing") {
+ reactions += { case ButtonClicked(_) => handle_update() }
+ }
+ tracing.tooltip =
+ "<html>Indicate output of tracing messages<br>(also needs to be enabled on the prover side)</html>"
+
+ private val debug = new CheckBox("Debug") {
+ reactions += { case ButtonClicked(_) => handle_update() }
+ }
+ debug.tooltip =
+ "<html>Indicate output of debug messages<br>(also needs to be enabled on the prover side)</html>"
+
+ private val follow = new CheckBox("Follow")
+ follow.selected = true
+ follow.tooltip =
+ "<html>Indicate automatic update following<br>caret movement or state changes</html>"
+
+ private def filtered_results(document: Document, cmd: Command): List[XML.Tree] =
+ {
+ val show_tracing = tracing.selected
+ val show_debug = debug.selected
+ document.current_state(cmd).results filter {
+ case XML.Elem(Markup.TRACING, _, _) => show_tracing
+ case XML.Elem(Markup.DEBUG, _, _) => show_debug
+ case _ => true
+ }
+ }
+
+ private case class Render(body: List[XML.Tree])
+
+ private def handle_update()
+ {
+ Swing_Thread.now {
+ Document_Model(view.getBuffer) match {
+ case Some(model) =>
+ val document = model.recent_document
+ document.command_at(view.getTextArea.getCaretPosition) match {
+ case Some((cmd, _)) =>
+ output_actor ! Render(filtered_results(document, cmd))
+ case None =>
+ }
+ case None =>
+ }
+ }
+ }
+
+ private var zoom_factor = 100
+
+ private def handle_resize() =
+ Swing_Thread.now {
+ html_panel.resize(scala.math.round(Isabelle.font_size() * zoom_factor / 100))
+ }
+
+
/* actor wiring */
private val output_actor = actor {
loop {
react {
+ case Session.Global_Settings => handle_resize()
+ case Render(body) => html_panel.render(body)
+
case cmd: Command =>
- Swing_Thread.now { Document_Model(view.getBuffer) } match {
+ Swing_Thread.now {
+ if (follow.selected) Document_Model(view.getBuffer) else None
+ } match {
case None =>
case Some(model) =>
- val body = model.recent_document.current_state(cmd).map(_.results) getOrElse Nil
- html_panel.render(body)
+ html_panel.render(filtered_results(model.recent_document, cmd))
}
- case Session.Global_Settings => html_panel.init(Isabelle.font_size())
-
case bad => System.err.println("output_actor: ignoring bad message " + bad)
}
}
@@ -61,4 +132,20 @@
Isabelle.session.global_settings -= output_actor
super.removeNotify()
}
+
+
+ /* resize */
+
+ addComponentListener(new ComponentAdapter {
+ val delay = Swing_Thread.delay_last(500) { handle_resize() }
+ override def componentResized(e: ComponentEvent) { delay() }
+ })
+
+
+ /* init controls */
+
+ val controls = new FlowPanel(FlowPanel.Alignment.Right)(zoom, update, debug, tracing, follow)
+ add(controls.peer, BorderLayout.NORTH)
+
+ handle_update()
}
--- a/src/Tools/jEdit/src/jedit/plugin.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/plugin.scala Sun May 23 10:38:11 2010 +0100
@@ -70,8 +70,9 @@
jEdit.setIntegerProperty(OPTION_PREFIX + name, value)
}
- def font_size(): Int =
- (jEdit.getIntegerProperty("view.fontsize", 16) * Int_Property("relative-font-size", 100)) / 100
+ def font_size(): Float =
+ (jEdit.getIntegerProperty("view.fontsize", 16) *
+ Int_Property("relative-font-size", 100)).toFloat / 100
/* settings */
--- a/src/Tools/jEdit/src/jedit/protocol_dockable.scala Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/protocol_dockable.scala Sun May 23 10:38:11 2010 +0100
@@ -33,7 +33,7 @@
loop {
react {
case result: Isabelle_Process.Result =>
- Swing_Thread.now { text_area.append(result.toString + "\n") }
+ Swing_Thread.now { text_area.append(result.message.toString + "\n") }
case bad => System.err.println("raw_actor: ignoring bad message " + bad)
}
--- a/src/Tools/more_conv.ML Sun May 23 10:37:43 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-(* Title: Tools/more_conv.ML
- Author: Sascha Boehme, TU Muenchen
-
-Further conversions and conversionals.
-*)
-
-signature MORE_CONV =
-sig
- val rewrs_conv: thm list -> conv
-
- val sub_conv: (Proof.context -> conv) -> Proof.context -> conv
- val bottom_conv: (Proof.context -> conv) -> Proof.context -> conv
- val top_conv: (Proof.context -> conv) -> Proof.context -> conv
- val top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv
-
- val binder_conv: (Proof.context -> conv) -> Proof.context -> conv
-end
-
-structure More_Conv : MORE_CONV =
-struct
-
-fun rewrs_conv eqs = Conv.first_conv (map Conv.rewr_conv eqs)
-
-
-fun sub_conv conv ctxt =
- Conv.comb_conv (conv ctxt) else_conv
- Conv.abs_conv (fn (_, cx) => conv cx) ctxt else_conv
- Conv.all_conv
-
-fun bottom_conv conv ctxt ct =
- (sub_conv (bottom_conv conv) ctxt then_conv conv ctxt) ct
-
-fun top_conv conv ctxt ct =
- (conv ctxt then_conv sub_conv (top_conv conv) ctxt) ct
-
-fun top_sweep_conv conv ctxt ct =
- (conv ctxt else_conv sub_conv (top_sweep_conv conv) ctxt) ct
-
-
-fun binder_conv cv ctxt =
- Conv.arg_conv (Conv.abs_conv (fn (_, cx) => cv cx) ctxt)
-
-end
--- a/src/Tools/nbe.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/nbe.ML Sun May 23 10:38:11 2010 +0100
@@ -76,7 +76,7 @@
val get_triv_classes = map fst o Triv_Class_Data.get;
val (_, triv_of_class) = Context.>>> (Context.map_theory_result
- (Thm.add_oracle (Binding.name "triv_of_class", fn (thy, (v, T), class) =>
+ (Thm.add_oracle (Binding.name "triv_of_class", fn (thy, T, class) =>
Thm.cterm_of thy (Logic.mk_of_class (T, class)))));
in
@@ -84,37 +84,46 @@
fun lift_triv_classes_conv thy conv ct =
let
val algebra = Sign.classes_of thy;
+ val certT = Thm.ctyp_of thy;
val triv_classes = get_triv_classes thy;
- val certT = Thm.ctyp_of thy;
- fun critical_classes sort = filter_out (fn class => Sign.subsort thy (sort, [class])) triv_classes;
- val vs = Term.add_tfrees (Thm.term_of ct) []
- |> map_filter (fn (v, sort) => case critical_classes sort
- of [] => NONE
- | classes => SOME (v, ((sort, classes), Sorts.inter_sort algebra (triv_classes, sort))));
- val of_classes = maps (fn (v, ((sort, classes), _)) => map (fn class =>
- ((v, class), triv_of_class (thy, (v, TVar ((v, 0), sort)), class))) classes
- @ map (fn class => ((v, class), Thm.of_class (certT (TVar ((v, 0), sort)), class)))
- sort) vs;
+ fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes;
+ fun mk_entry (v, sort) =
+ let
+ val T = TFree (v, sort);
+ val cT = certT T;
+ val triv_sort = additional_classes sort;
+ in
+ (v, (Sorts.inter_sort algebra (sort, triv_sort),
+ (cT, AList.make (fn class => Thm.of_class (cT, class)) sort
+ @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort)))
+ end;
+ val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []);
+ fun instantiate thm =
+ let
+ val cert_tvars = map (certT o TVar) (Term.add_tvars
+ ((fst o Logic.dest_equals o Logic.strip_imp_concl o Thm.prop_of) thm) []);
+ val instantiation =
+ map2 (fn cert_tvar => fn (_, (_, (cT, _))) => (cert_tvar, cT)) cert_tvars vs_tab;
+ in Thm.instantiate (instantiation, []) thm end;
+ fun of_class (TFree (v, _), class) =
+ the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class)
+ | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ_global thy T);
fun strip_of_class thm =
let
- val prem_props = (Logic.strip_imp_prems o Thm.prop_of) thm;
- val prem_thms = map (the o AList.lookup (op =) of_classes
- o apfst (fst o fst o dest_TVar) o Logic.dest_of_class) prem_props;
- in Drule.implies_elim_list thm prem_thms end;
+ val prems_of_class = Thm.prop_of thm
+ |> Logic.strip_imp_prems
+ |> map (Logic.dest_of_class #> of_class);
+ in fold Thm.elim_implies prems_of_class thm end;
in
- (* FIXME avoid legacy operations *)
ct
- |> Drule.cterm_rule Thm.varifyT_global
- |> Thm.instantiate_cterm (Thm.certify_inst thy (map (fn (v, ((sort, _), sort')) =>
- (((v, 0), sort), TFree (v, sort'))) vs, []))
- |> Drule.cterm_rule Thm.legacy_freezeT
+ |> (Drule.cterm_fun o map_types o map_type_tfree)
+ (fn (v, sort) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v))
|> conv
+ |> Thm.strip_shyps
|> Thm.varifyT_global
- |> fold (fn (v, (_, sort')) => Thm.legacy_unconstrainT (certT (TVar ((v, 0), sort')))) vs
- |> Thm.certify_instantiate (map (fn (v, ((sort, _), _)) =>
- (((v, 0), []), TVar ((v, 0), sort))) vs, [])
+ |> Thm.unconstrainT
+ |> instantiate
|> strip_of_class
- |> Thm.legacy_freezeT
end;
fun lift_triv_classes_rew thy rew t =
@@ -365,7 +374,7 @@
(* code compilation *)
-fun compile_eqnss _ gr raw_deps [] = []
+fun compile_eqnss ctxt gr raw_deps [] = []
| compile_eqnss ctxt gr raw_deps eqnss =
let
val (deps, deps_vals) = split_list (map_filter
@@ -552,7 +561,7 @@
|> type_infer
|> traced (fn t => "Types inferred:\n" ^ string_of_term t)
|> check_tvars
- |> traced (fn t => "---\n")
+ |> traced (fn _ => "---\n")
end;
@@ -625,15 +634,12 @@
val setup = Value.add_evaluator ("nbe", norm o ProofContext.theory_of);
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
+val opt_modes =
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
val _ =
- OuterSyntax.improper_command "normal_form" "normalize term by evaluation" K.diag
- (opt_modes -- P.term >> (Toplevel.keep o norm_print_term_cmd));
-
-end;
+ Outer_Syntax.improper_command "normal_form" "normalize term by evaluation" Keyword.diag
+ (opt_modes -- Parse.term >> (Toplevel.keep o norm_print_term_cmd));
end;
\ No newline at end of file
--- a/src/Tools/quickcheck.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/quickcheck.ML Sun May 23 10:38:11 2010 +0100
@@ -315,27 +315,25 @@
test_goal quiet report generator_name size iterations default_type no_assms insts i assms state
end;
-fun quickcheck args i state = fst (gen_quickcheck args i state)
+fun quickcheck args i state = fst (gen_quickcheck args i state);
fun quickcheck_cmd args i state =
gen_quickcheck args i (Toplevel.proof_of state)
|> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state);
-local structure P = OuterParse and K = OuterKeyword in
+val parse_arg = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- Parse.name) "true");
-val parse_arg = P.name -- (Scan.optional (P.$$$ "=" |-- P.name) "true")
-
-val parse_args = P.$$$ "[" |-- P.list1 parse_arg --| P.$$$ "]"
+val parse_args = Parse.$$$ "[" |-- Parse.list1 parse_arg --| Parse.$$$ "]"
|| Scan.succeed [];
-val _ = OuterSyntax.command "quickcheck_params" "set parameters for random testing" K.thy_decl
- (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args)));
+val _ =
+ Outer_Syntax.command "quickcheck_params" "set parameters for random testing" Keyword.thy_decl
+ (parse_args >> (fn args => Toplevel.theory (quickcheck_params_cmd args)));
-val _ = OuterSyntax.improper_command "quickcheck" "try to find counterexample for subgoal" K.diag
- (parse_args -- Scan.optional P.nat 1
- >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i)));
-
-end; (*local*)
+val _ =
+ Outer_Syntax.improper_command "quickcheck" "try to find counterexample for subgoal" Keyword.diag
+ (parse_args -- Scan.optional Parse.nat 1
+ >> (fn (args, i) => Toplevel.no_timing o Toplevel.keep (quickcheck_cmd args i)));
end;
--- a/src/Tools/value.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/Tools/value.ML Sun May 23 10:38:11 2010 +0100
@@ -52,15 +52,13 @@
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
in Pretty.writeln p end;
-local structure P = OuterParse and K = OuterKeyword in
-
-val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
+val opt_modes =
+ Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
-val _ = OuterSyntax.improper_command "value" "evaluate and print term" K.diag
- (Scan.option (P.$$$ "[" |-- P.xname --| P.$$$ "]") -- opt_modes -- P.term
- >> (fn ((some_name, modes), t) => Toplevel.no_timing o Toplevel.keep
- (value_cmd some_name modes t)));
-
-end; (*local*)
+val _ =
+ Outer_Syntax.improper_command "value" "evaluate and print term" Keyword.diag
+ (Scan.option (Parse.$$$ "[" |-- Parse.xname --| Parse.$$$ "]") -- opt_modes -- Parse.term
+ >> (fn ((some_name, modes), t) => Toplevel.no_timing o Toplevel.keep
+ (value_cmd some_name modes t)));
end;
--- a/src/ZF/Tools/datatype_package.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/datatype_package.ML Sun May 23 10:38:11 2010 +0100
@@ -419,29 +419,26 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
fun mk_datatype ((((dom, dts), monos), type_intrs), type_elims) =
#1 o add_datatype (dom, map fst dts) (map snd dts) (monos, type_intrs, type_elims);
val con_decl =
- P.name -- Scan.optional (P.$$$ "(" |-- P.list1 P.term --| P.$$$ ")") [] -- P.opt_mixfix
- >> P.triple1;
+ Parse.name -- Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.term --| Parse.$$$ ")") [] --
+ Parse.opt_mixfix >> Parse.triple1;
val datatype_decl =
- (Scan.optional ((P.$$$ "\<subseteq>" || P.$$$ "<=") |-- P.!!! P.term) "") --
- P.and_list1 (P.term -- (P.$$$ "=" |-- P.enum1 "|" con_decl)) --
- Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) [] --
- Scan.optional (P.$$$ "type_intros" |-- P.!!! SpecParse.xthms1) [] --
- Scan.optional (P.$$$ "type_elims" |-- P.!!! SpecParse.xthms1) []
+ (Scan.optional ((Parse.$$$ "\<subseteq>" || Parse.$$$ "<=") |-- Parse.!!! Parse.term) "") --
+ Parse.and_list1 (Parse.term -- (Parse.$$$ "=" |-- Parse.enum1 "|" con_decl)) --
+ Scan.optional (Parse.$$$ "monos" |-- Parse.!!! Parse_Spec.xthms1) [] --
+ Scan.optional (Parse.$$$ "type_intros" |-- Parse.!!! Parse_Spec.xthms1) [] --
+ Scan.optional (Parse.$$$ "type_elims" |-- Parse.!!! Parse_Spec.xthms1) []
>> (Toplevel.theory o mk_datatype);
val coind_prefix = if coind then "co" else "";
-val _ = OuterSyntax.command (coind_prefix ^ "datatype")
- ("define " ^ coind_prefix ^ "datatype") K.thy_decl datatype_decl;
+val _ =
+ Outer_Syntax.command (coind_prefix ^ "datatype")
+ ("define " ^ coind_prefix ^ "datatype") Keyword.thy_decl datatype_decl;
end;
-end;
-
--- a/src/ZF/Tools/ind_cases.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/ind_cases.ML Sun May 23 10:38:11 2010 +0100
@@ -64,15 +64,11 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
val _ =
- OuterSyntax.command "inductive_cases"
- "create simplified instances of elimination rules (improper)" K.thy_script
- (P.and_list1 (SpecParse.opt_thm_name ":" -- Scan.repeat1 P.prop)
+ Outer_Syntax.command "inductive_cases"
+ "create simplified instances of elimination rules (improper)" Keyword.thy_script
+ (Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Scan.repeat1 Parse.prop)
>> (Toplevel.theory o inductive_cases));
end;
-end;
-
--- a/src/ZF/Tools/induct_tacs.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/induct_tacs.ML Sun May 23 10:38:11 2010 +0100
@@ -89,7 +89,7 @@
fun exhaust_induct_tac exh ctxt var i state =
let
val thy = ProofContext.theory_of ctxt
- val (_, _, Bi, _) = dest_state (state, i)
+ val (_, _, Bi, _) = Thm.dest_state (state, i)
val tn = find_tname var Bi
val rule =
if exh then #exhaustion (datatype_info thy tn)
@@ -186,25 +186,20 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword ["elimination", "induction", "case_eqns", "recursor_eqns"];
+val _ = List.app Keyword.keyword ["elimination", "induction", "case_eqns", "recursor_eqns"];
val rep_datatype_decl =
- (P.$$$ "elimination" |-- P.!!! SpecParse.xthm) --
- (P.$$$ "induction" |-- P.!!! SpecParse.xthm) --
- (P.$$$ "case_eqns" |-- P.!!! SpecParse.xthms1) --
- Scan.optional (P.$$$ "recursor_eqns" |-- P.!!! SpecParse.xthms1) []
+ (Parse.$$$ "elimination" |-- Parse.!!! Parse_Spec.xthm) --
+ (Parse.$$$ "induction" |-- Parse.!!! Parse_Spec.xthm) --
+ (Parse.$$$ "case_eqns" |-- Parse.!!! Parse_Spec.xthms1) --
+ Scan.optional (Parse.$$$ "recursor_eqns" |-- Parse.!!! Parse_Spec.xthms1) []
>> (fn (((x, y), z), w) => rep_datatype x y z w);
val _ =
- OuterSyntax.command "rep_datatype" "represent existing set inductively" K.thy_decl
+ Outer_Syntax.command "rep_datatype" "represent existing set inductively" Keyword.thy_decl
(rep_datatype_decl >> Toplevel.theory);
end;
-end;
-
-
val exhaust_tac = DatatypeTactics.exhaust_tac;
val induct_tac = DatatypeTactics.induct_tac;
--- a/src/ZF/Tools/inductive_package.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/inductive_package.ML Sun May 23 10:38:11 2010 +0100
@@ -576,29 +576,26 @@
(* outer syntax *)
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ = List.app OuterKeyword.keyword
+val _ = List.app Keyword.keyword
["domains", "intros", "monos", "con_defs", "type_intros", "type_elims"];
fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) =
- #1 o add_inductive doms (map P.triple_swap intrs) (monos, con_defs, type_intrs, type_elims);
+ #1 o add_inductive doms (map Parse.triple_swap intrs) (monos, con_defs, type_intrs, type_elims);
val ind_decl =
- (P.$$$ "domains" |-- P.!!! (P.enum1 "+" P.term --
- ((P.$$$ "\<subseteq>" || P.$$$ "<=") |-- P.term))) --
- (P.$$$ "intros" |--
- P.!!! (Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop))) --
- Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) [] --
- Scan.optional (P.$$$ "con_defs" |-- P.!!! SpecParse.xthms1) [] --
- Scan.optional (P.$$$ "type_intros" |-- P.!!! SpecParse.xthms1) [] --
- Scan.optional (P.$$$ "type_elims" |-- P.!!! SpecParse.xthms1) []
+ (Parse.$$$ "domains" |-- Parse.!!! (Parse.enum1 "+" Parse.term --
+ ((Parse.$$$ "\<subseteq>" || Parse.$$$ "<=") |-- Parse.term))) --
+ (Parse.$$$ "intros" |--
+ Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) --
+ Scan.optional (Parse.$$$ "monos" |-- Parse.!!! Parse_Spec.xthms1) [] --
+ Scan.optional (Parse.$$$ "con_defs" |-- Parse.!!! Parse_Spec.xthms1) [] --
+ Scan.optional (Parse.$$$ "type_intros" |-- Parse.!!! Parse_Spec.xthms1) [] --
+ Scan.optional (Parse.$$$ "type_elims" |-- Parse.!!! Parse_Spec.xthms1) []
>> (Toplevel.theory o mk_ind);
-val _ = OuterSyntax.command (co_prefix ^ "inductive")
- ("define " ^ co_prefix ^ "inductive sets") K.thy_decl ind_decl;
+val _ =
+ Outer_Syntax.command (co_prefix ^ "inductive")
+ ("define " ^ co_prefix ^ "inductive sets") Keyword.thy_decl ind_decl;
end;
-end;
-
--- a/src/ZF/Tools/primrec_package.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/primrec_package.ML Sun May 23 10:38:11 2010 +0100
@@ -194,12 +194,11 @@
(* outer syntax *)
-structure P = OuterParse and K = OuterKeyword;
-
val _ =
- OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
- (Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
- >> (Toplevel.theory o (#1 oo (add_primrec o map P.triple_swap))));
+ Outer_Syntax.command "primrec" "define primitive recursive functions on datatypes"
+ Keyword.thy_decl
+ (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)
+ >> (Toplevel.theory o (#1 oo (add_primrec o map Parse.triple_swap))));
end;
--- a/src/ZF/Tools/typechk.ML Sun May 23 10:37:43 2010 +0100
+++ b/src/ZF/Tools/typechk.ML Sun May 23 10:38:11 2010 +0100
@@ -118,7 +118,7 @@
"ZF type-checking";
val _ =
- OuterSyntax.improper_command "print_tcset" "print context of ZF typecheck" OuterKeyword.diag
+ Outer_Syntax.improper_command "print_tcset" "print context of ZF typecheck" Keyword.diag
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
Toplevel.keep (print_tcset o Toplevel.context_of)));