# HG changeset patch # User blanchet # Date 1256565469 -3600 # Node ID 79bd3fbf5d618bcc434aadc46e7907310a6629df # Parent 4705b7323a7d8a721d515675a5b6eb0a9225ec16# Parent 322d928d9f8fa219c24e03277af91d6ce9b669e4 merged diff -r 4705b7323a7d -r 79bd3fbf5d61 CONTRIBUTORS --- a/CONTRIBUTORS Mon Oct 26 14:54:43 2009 +0100 +++ b/CONTRIBUTORS Mon Oct 26 14:57:49 2009 +0100 @@ -7,6 +7,9 @@ Contributions to this Isabelle version -------------------------------------- +* October 2009: Jasmin Blanchette, TUM + Nitpick: yet another counterexample generator for Isabelle/HOL + * October 2009: Sascha Boehme, TUM Extension of SMT method: proof-reconstruction for the SMT solver Z3. diff -r 4705b7323a7d -r 79bd3fbf5d61 NEWS --- a/NEWS Mon Oct 26 14:54:43 2009 +0100 +++ b/NEWS Mon Oct 26 14:57:49 2009 +0100 @@ -50,6 +50,9 @@ this method is proof-producing. Certificates are provided to avoid calling the external solvers solely for re-checking proofs. +* New counterexample generator tool "nitpick" based on the Kodkod +relational model finder. + * Reorganization of number theory: * former session NumberTheory now named Old_Number_Theory * new session Number_Theory by Jeremy Avigad; if possible, prefer this. diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/Dirs --- a/doc-src/Dirs Mon Oct 26 14:54:43 2009 +0100 +++ b/doc-src/Dirs Mon Oct 26 14:57:49 2009 +0100 @@ -1,1 +1,1 @@ -Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions Main +Intro Ref System Logics HOL ZF Inductive TutorialI IsarOverview IsarRef IsarImplementation Locales LaTeXsugar Classes Codegen Functions Nitpick Main diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/Makefile.in --- a/doc-src/Makefile.in Mon Oct 26 14:54:43 2009 +0100 +++ b/doc-src/Makefile.in Mon Oct 26 14:57:49 2009 +0100 @@ -45,6 +45,9 @@ isabelle_zf.eps: test -r isabelle_zf.eps || ln -s ../gfx/isabelle_zf.eps . +isabelle_nitpick.eps: + test -r isabelle_nitpick.eps || ln -s ../gfx/isabelle_nitpick.eps . + isabelle.pdf: test -r isabelle.pdf || ln -s ../gfx/isabelle.pdf . @@ -58,6 +61,9 @@ isabelle_zf.pdf: test -r isabelle_zf.pdf || ln -s ../gfx/isabelle_zf.pdf . +isabelle_nitpick.pdf: + test -r isabelle_nitpick.pdf || ln -s ../gfx/isabelle_nitpick.pdf . + typedef.ps: test -r typedef.ps || ln -s ../gfx/typedef.ps . diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/Nitpick/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Nitpick/Makefile Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,36 @@ +# +# $Id$ +# + +## targets + +default: dvi + + +## dependencies + +include ../Makefile.in + +NAME = nitpick +FILES = nitpick.tex ../iman.sty ../manual.bib + +dvi: $(NAME).dvi + +$(NAME).dvi: $(FILES) isabelle_nitpick.eps + $(LATEX) $(NAME) + $(BIBTEX) $(NAME) + $(LATEX) $(NAME) + $(LATEX) $(NAME) + $(SEDINDEX) $(NAME) + $(LATEX) $(NAME) + +pdf: $(NAME).pdf + +$(NAME).pdf: $(FILES) isabelle_nitpick.pdf + $(PDFLATEX) $(NAME) + $(BIBTEX) $(NAME) + $(PDFLATEX) $(NAME) + $(PDFLATEX) $(NAME) + $(SEDINDEX) $(NAME) + $(FIXBOOKMARKS) $(NAME).out + $(PDFLATEX) $(NAME) diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/Nitpick/nitpick.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/Nitpick/nitpick.tex Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,2486 @@ +\documentclass[a4paper,12pt]{article} +\usepackage[T1]{fontenc} +\usepackage{amsmath} +\usepackage{amssymb} +\usepackage[french,english]{babel} +\usepackage{color} +\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\undef{\textit{undefined}} +\def\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?$}}}} + +\hyphenation{Mini-Sat size-change First-Steps grand-parent nit-pick +counter-example counter-examples data-type data-types co-data-type +co-data-types in-duc-tive co-in-duc-tive} + +\urlstyle{tt} + +\begin{document} + +\title{\includegraphics[scale=0.5]{isabelle_nitpick} \\[4ex] +Picking Nits \\[\smallskipamount] +\Large A User's Guide to Nitpick for Isabelle/HOL 2010} +\author{\hbox{} \\ +Jasmin Christian Blanchette \\ +{\normalsize Fakult\"at 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} + +Nitpick \cite{blanchette-nipkow-2009} 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 +first-order relational model finder developed by the Software Design Group at +MIT. It is conceptually similar to Refute \cite{weber-2008}, from which it +borrows many ideas and code fragments, but it benefits from Kodkod's +optimizations and a new encoding scheme. The name Nitpick is shamelessly +appropriated from a now retired Alloy precursor. + +Nitpick is easy to use---you simply enter \textbf{nitpick} after a putative +theorem and wait a few seconds. Nonetheless, there are situations where knowing +how it works under the hood and how it reacts to various options helps +increase the test coverage. This manual also explains how to install the tool on +your workstation. Should the motivation fail you, think of the many hours of +hard work Nitpick will save you. Proving non-theorems is \textsl{hard work}. + +Another common use of Nitpick is to find out whether the axioms of a locale are +satisfiable, while the locale is being developed. To check this, it suffices to +write + +\prew +\textbf{lemma}~``$\textit{False}$'' \\ +\textbf{nitpick}~[\textit{show\_all}] +\postw + +after the locale's \textbf{begin} keyword. To falsify \textit{False}, Nitpick +must find a model for the axioms. If it finds no model, we have an indication +that the axioms might be unsatisfiable. + +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. + +\newbox\boxA +\setbox\boxA=\hbox{\texttt{nospam}} + +The known bugs and limitations at the time of writing are listed in +\S\ref{known-bugs-and-limitations}. Comments and bug reports concerning Nitpick +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. +% and Perry James for reporting a typo. + +\section{First Steps} +\label{first-steps} + +This section introduces Nitpick by presenting small examples. If possible, you +should try out the examples on your workstation. Your theory file should start +the standard way: + +\prew +\textbf{theory}~\textit{Scratch} \\ +\textbf{imports}~\textit{Main} \\ +\textbf{begin} +\postw + +The results presented here were obtained using the JNI version of MiniSat and +with multithreading disabled to reduce nondeterminism. This was done by adding +the line + +\prew +\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSatJNI}, \,\textit{max\_threads}~= 1] +\postw + +after the \textbf{begin} keyword. The JNI version of MiniSat is bundled with +Kodkodi and is precompiled for the major platforms. Other SAT solvers can also +be installed, as explained in \S\ref{optimizations}. If you have already +configured SAT solvers in Isabelle (e.g., for Refute), these will also be +available to Nitpick. + +Throughout this manual, we will explicitly invoke the \textbf{nitpick} command. +Nitpick also provides an automatic mode that can be enabled by specifying + +\prew +\textbf{nitpick\_params} [\textit{auto}] +\postw + +at the beginning of the theory file. In this mode, Nitpick is run for up to 5 +seconds (by default) on every newly entered theorem, much like Auto Quickcheck. + +\subsection{Propositional Logic} +\label{propositional-logic} + +Let's start with a trivial example from propositional logic: + +\prew +\textbf{lemma}~``$P \longleftrightarrow Q$'' \\ +\textbf{nitpick} +\postw + +You should get the following output: + +\prew +\slshape +Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \textit{True}$ \\ +\hbox{}\qquad\qquad $Q = \textit{False}$ +\postw + +Nitpick can also be invoked on individual subgoals, as in the example below: + +\prew +\textbf{apply}~\textit{auto} \\[2\smallskipamount] +{\slshape goal (2 subgoals): \\ +\ 1. $P\,\Longrightarrow\, Q$ \\ +\ 2. $Q\,\Longrightarrow\, P$} \\[2\smallskipamount] +\textbf{nitpick}~1 \\[2\smallskipamount] +{\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \textit{True}$ \\ +\hbox{}\qquad\qquad $Q = \textit{False}$} \\[2\smallskipamount] +\textbf{nitpick}~2 \\[2\smallskipamount] +{\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \textit{False}$ \\ +\hbox{}\qquad\qquad $Q = \textit{True}$} \\[2\smallskipamount] +\textbf{oops} +\postw + +\subsection{Type Variables} +\label{type-variables} + +If you are left unimpressed by the previous example, don't worry. The next +one is more mind- and computer-boggling: + +\prew +\textbf{lemma} ``$P~x\,\Longrightarrow\, P~(\textrm{THE}~y.\;P~y)$'' +\postw +\pagebreak[2] %% TYPESETTING + +The putative lemma involves the definite description operator, {THE}, presented +in section 5.10.1 of the Isabelle tutorial \cite{isa-tutorial}. The +operator is defined by the axiom $(\textrm{THE}~x.\; x = a) = a$. The putative +lemma is merely asserting the indefinite description operator axiom with {THE} +substituted for {SOME}. + +The free variable $x$ and the bound variable $y$ have type $'a$. For formulas +containing type variables, Nitpick enumerates the possible domains for each type +variable, up to a given cardinality (8 by default), looking for a finite +countermodel: + +\prew +\textbf{nitpick} [\textit{verbose}] \\[2\smallskipamount] +\slshape +Trying 8 scopes: \nopagebreak \\ +\hbox{}\qquad \textit{card}~$'a$~= 1; \\ +\hbox{}\qquad \textit{card}~$'a$~= 2; \\ +\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount] +\hbox{}\qquad \textit{card}~$'a$~= 8. \\[2\smallskipamount] +Nitpick found a counterexample for \textit{card} $'a$~= 3: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \{a_2,\, a_3\}$ \\ +\hbox{}\qquad\qquad $x = a_3$ \\[2\smallskipamount] +Total time: 580 ms. +\postw + +Nitpick found a counterexample in which $'a$ has cardinality 3. (For +cardinalities 1 and 2, the formula holds.) In the counterexample, the three +values of type $'a$ are written $a_1$, $a_2$, and $a_3$. + +The message ``Trying $n$ scopes: {\ldots}''\ is shown only if the option +\textit{verbose} is enabled. You can specify \textit{verbose} each time you +invoke \textbf{nitpick}, or you can set it globally using the command + +\prew +\textbf{nitpick\_params} [\textit{verbose}] +\postw + +This command also displays the current default values for all of the options +supported by Nitpick. The options are listed in \S\ref{option-reference}. + +\subsection{Constants} +\label{constants} + +By just looking at Nitpick's output, it might not be clear why the +counterexample in \S\ref{type-variables} is genuine. Let's invoke Nitpick again, +this time telling it to show the values of the constants that occur in the +formula: + +\prew +\textbf{lemma}~``$P~x\,\Longrightarrow\, P~(\textrm{THE}~y.\;P~y)$'' \\ +\textbf{nitpick}~[\textit{show\_consts}] \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 3: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \{a_2,\, a_3\}$ \\ +\hbox{}\qquad\qquad $x = a_3$ \\ +\hbox{}\qquad Constant: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{The}~\textsl{fallback} = a_1$ +\postw + +We can see more clearly now. Since the predicate $P$ isn't true for a unique +value, $\textrm{THE}~y.\;P~y$ can denote any value of type $'a$, even +$a_1$. Since $P~a_1$ is false, the entire formula is falsified. + +As an optimization, Nitpick's preprocessor introduced the special constant +``\textit{The} fallback'' corresponding to $\textrm{THE}~y.\;P~y$ (i.e., +$\mathit{The}~(\lambda y.\;P~y)$) when there doesn't exist a unique $y$ +satisfying $P~y$. We disable this optimization by passing the +\textit{full\_descrs} option: + +\prew +\textbf{nitpick}~[\textit{full\_descrs},\, \textit{show\_consts}] \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 3: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \{a_2,\, a_3\}$ \\ +\hbox{}\qquad\qquad $x = a_3$ \\ +\hbox{}\qquad Constant: \nopagebreak \\ +\hbox{}\qquad\qquad $\hbox{\slshape THE}~y.\;P~y = a_1$ +\postw + +As the result of another optimization, Nitpick directly assigned a value to the +subterm $\textrm{THE}~y.\;P~y$, rather than to the \textit{The} constant. If we +disable this second optimization by using the command + +\prew +\textbf{nitpick}~[\textit{dont\_specialize},\, \textit{full\_descrs},\, +\textit{show\_consts}] +\postw + +we finally get \textit{The}: + +\prew +\slshape Constant: \nopagebreak \\ +\hbox{}\qquad $\mathit{The} = \undef{} + (\!\begin{aligned}[t]% + & \{\} := a_3,\> \{a_3\} := a_3,\> \{a_2\} := a_2, \\[-2pt] %% TYPESETTING + & \{a_2, a_3\} := a_1,\> \{a_1\} := a_1,\> \{a_1, a_3\} := a_3, \\[-2pt] + & \{a_1, a_2\} := a_3,\> \{a_1, a_2, a_3\} := a_3)\end{aligned}$ +\postw + +Notice that $\textit{The}~(\lambda y.\;P~y) = \textit{The}~\{a_2, a_3\} = a_1$, +just like before.\footnote{The \undef{} symbol's presence is explained as +follows: In higher-order logic, any function can be built from the undefined +function using repeated applications of the function update operator $f(x := +y)$, just like any list can be built from the empty list using $x \mathbin{\#} +xs$.} + +Our misadventures with THE suggest adding `$\exists!x{.}$' (``there exists a +unique $x$ such that'') at the front of our putative lemma's assumption: + +\prew +\textbf{lemma}~``$\exists {!}x.\; P~x\,\Longrightarrow\, P~(\textrm{THE}~y.\;P~y)$'' +\postw + +The fix appears to work: + +\prew +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found no counterexample. +\postw + +We can further increase our confidence in the formula by exhausting all +cardinalities up to 50: + +\prew +\textbf{nitpick} [\textit{card} $'a$~= 1--50]\footnote{The symbol `--' +can be entered as \texttt{-} (hyphen) or +\texttt{\char`\\\char`\}.} \\[2\smallskipamount] +\slshape Nitpick found no counterexample. +\postw + +Let's see if Sledgehammer \cite{sledgehammer-2009} can find a proof: + +\prew +\textbf{sledgehammer} \\[2\smallskipamount] +{\slshape Sledgehammer: external prover ``$e$'' for subgoal 1: \\ +$\exists{!}x.\; P~x\,\Longrightarrow\, P~(\hbox{\slshape THE}~y.\; P~y)$ \\ +Try this command: \textrm{apply}~(\textit{metis~the\_equality})} \\[2\smallskipamount] +\textbf{apply}~(\textit{metis~the\_equality\/}) \nopagebreak \\[2\smallskipamount] +{\slshape No subgoals!}% \\[2\smallskipamount] +%\textbf{done} +\postw + +This must be our lucky day. + +\subsection{Skolemization} +\label{skolemization} + +Are all invertible functions onto? Let's find out: + +\prew +\textbf{lemma} ``$\exists g.\; \forall x.~g~(f~x) = x + \,\Longrightarrow\, \forall y.\; \exists x.~y = f~x$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 2 and \textit{card} $'b$~=~1: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $f = \undef{}(b_1 := a_1)$ \\ +\hbox{}\qquad Skolem constants: \nopagebreak \\ +\hbox{}\qquad\qquad $g = \undef{}(a_1 := b_1,\> a_2 := b_1)$ \\ +\hbox{}\qquad\qquad $y = a_2$ +\postw + +Although $f$ is the only free variable occurring in the formula, Nitpick also +displays values for the bound variables $g$ and $y$. These values are available +to Nitpick because it performs skolemization as a preprocessing step. + +In the previous example, skolemization only affected the outermost quantifiers. +This is not always the case, as illustrated below: + +\prew +\textbf{lemma} ``$\exists x.\; \forall f.\; f~x = x$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 2: \\[2\smallskipamount] +\hbox{}\qquad Skolem constant: \nopagebreak \\ +\hbox{}\qquad\qquad $\lambda x.\; f = + \undef{}(\!\begin{aligned}[t] + & a_1 := \undef{}(a_1 := a_2,\> a_2 := a_1), \\[-2pt] + & a_2 := \undef{}(a_1 := a_1,\> a_2 := a_1))\end{aligned}$ +\postw + +The variable $f$ is bound within the scope of $x$; therefore, $f$ depends on +$x$, as suggested by the notation $\lambda x.\,f$. If $x = a_1$, then $f$ is the +function that maps $a_1$ to $a_2$ and vice versa; otherwise, $x = a_2$ and $f$ +maps both $a_1$ and $a_2$ to $a_1$. In both cases, $f~x \not= x$. + +The source of the Skolem constants is sometimes more obscure: + +\prew +\textbf{lemma} ``$\mathit{refl}~r\,\Longrightarrow\, \mathit{sym}~r$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 2: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $r = \{(a_1, a_1),\, (a_2, a_1),\, (a_2, a_2)\}$ \\ +\hbox{}\qquad Skolem constants: \nopagebreak \\ +\hbox{}\qquad\qquad $\mathit{sym}.x = a_2$ \\ +\hbox{}\qquad\qquad $\mathit{sym}.y = a_1$ +\postw + +What happened here is that Nitpick expanded the \textit{sym} constant to its +definition: + +\prew +$\mathit{sym}~r \,\equiv\, + \forall x\> y.\,\> (x, y) \in r \longrightarrow (y, x) \in r.$ +\postw + +As their names suggest, the Skolem constants $\mathit{sym}.x$ and +$\mathit{sym}.y$ are simply the bound variables $x$ and $y$ +from \textit{sym}'s definition. + +Although skolemization is a useful optimization, you can disable it by invoking +Nitpick with \textit{dont\_skolemize}. See \S\ref{optimizations} for details. + +\subsection{Natural Numbers and Integers} +\label{natural-numbers-and-integers} + +Because of the axiom of infinity, the type \textit{nat} does not admit any +finite models. To deal with this, Nitpick considers prefixes $\{0,\, 1,\, +\ldots,\, K - 1\}$ of \textit{nat} (where $K = \textit{card}~\textit{nat}$) and +maps all other numbers to the undefined value ($\unk$). The type \textit{int} is +handled in a similar way: If $K = \textit{card}~\textit{int}$, the subset of +\textit{int} known to Nitpick is $\{-\lceil K/2 \rceil + 1,\, \ldots,\, +\lfloor +K/2 \rfloor\}$. Undefined values lead to a three-valued logic. + +Here is an example involving \textit{int}: + +\prew +\textbf{lemma} ``$\lbrakk i \le j;\> n \le (m{\Colon}\mathit{int})\rbrakk \,\Longrightarrow\, i * n + j * m \le i * m + j * n$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $i = 0$ \\ +\hbox{}\qquad\qquad $j = 1$ \\ +\hbox{}\qquad\qquad $m = 1$ \\ +\hbox{}\qquad\qquad $n = 0$ +\postw + +With infinite types, we don't always have the luxury of a genuine counterexample +and must often content ourselves with a potential one. The tedious task of +finding out whether the potential counterexample is in fact genuine can be +outsourced to \textit{auto} by passing the option \textit{check\_potential}. For +example: + +\prew +\textbf{lemma} ``$\forall n.\; \textit{Suc}~n \mathbin{\not=} n \,\Longrightarrow\, P$'' \\ +\textbf{nitpick} [\textit{card~nat}~= 100,\, \textit{check\_potential}] \\[2\smallskipamount] +\slshape Nitpick found a potential counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \textit{False}$ \\[2\smallskipamount] +Confirmation by ``\textit{auto}'': The above counterexample is genuine. +\postw + +You might wonder why the counterexample is first reported as potential. The root +of the problem is that the bound variable in $\forall n.\; \textit{Suc}~n +\mathbin{\not=} n$ ranges over an infinite type. If Nitpick finds an $n$ such +that $\textit{Suc}~n \mathbin{=} n$, it evaluates the assumption to +\textit{False}; but otherwise, it does not know anything about values of $n \ge +\textit{card~nat}$ and must therefore evaluate the assumption to $\unk$, not +\textit{True}. Since the assumption can never be satisfied, the putative lemma +can never be falsified. + +Incidentally, if you distrust the so-called genuine counterexamples, you can +enable \textit{check\_\allowbreak genuine} to verify them as well. However, be +aware that \textit{auto} will often fail to prove that the counterexample is +genuine or spurious. + +Some conjectures involving elementary number theory make Nitpick look like a +giant with feet of clay: + +\prew +\textbf{lemma} ``$P~\textit{Suc}$'' \\ +\textbf{nitpick} [\textit{card} = 1--6] \\[2\smallskipamount] +\slshape +Nitpick found no counterexample. +\postw + +For any cardinality $k$, \textit{Suc} is the partial function $\{0 \mapsto 1,\, +1 \mapsto 2,\, \ldots,\, k - 1 \mapsto \unk\}$, which evaluates to $\unk$ when +it is passed as argument to $P$. As a result, $P~\textit{Suc}$ is always $\unk$. +The next example is similar: + +\prew +\textbf{lemma} ``$P~(\textit{op}~{+}\Colon +\textit{nat}\mathbin{\Rightarrow}\textit{nat}\mathbin{\Rightarrow}\textit{nat})$'' \\ +\textbf{nitpick} [\textit{card nat} = 1] \\[2\smallskipamount] +{\slshape Nitpick found a counterexample:} \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \{\}$ \\[2\smallskipamount] +\textbf{nitpick} [\textit{card nat} = 2] \\[2\smallskipamount] +{\slshape Nitpick found no counterexample.} +\postw + +The problem here is that \textit{op}~+ is total when \textit{nat} is taken to be +$\{0\}$ but becomes partial as soon as we add $1$, because $1 + 1 \notin \{0, +1\}$. + +Because numbers are infinite and are approximated using a three-valued logic, +there is usually no need to systematically enumerate domain sizes. If Nitpick +cannot find a genuine counterexample for \textit{card~nat}~= $k$, it is very +unlikely that one could be found for smaller domains. (The $P~(\textit{op}~{+})$ +example above is an exception to this principle.) Nitpick nonetheless enumerates +all cardinalities from 1 to 8 for \textit{nat}, mainly because smaller +cardinalities are fast to handle and give rise to simpler counterexamples. This +is explained in more detail in \S\ref{scope-monotonicity}. + +\subsection{Inductive Datatypes} +\label{inductive-datatypes} + +Like natural numbers and integers, inductive datatypes with recursive +constructors admit no finite models and must be approximated by a subterm-closed +subset. For example, using a cardinality of 10 for ${'}a~\textit{list}$, +Nitpick looks for all counterexamples that can be built using at most 10 +different lists. + +Let's see with an example involving \textit{hd} (which returns the first element +of a list) and $@$ (which concatenates two lists): + +\prew +\textbf{lemma} ``$\textit{hd}~(\textit{xs} \mathbin{@} [y, y]) = \textit{hd}~\textit{xs}$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample for \textit{card} $'a$~= 3: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{xs} = []$ \\ +\hbox{}\qquad\qquad $\textit{y} = a_3$ +\postw + +To see why the counterexample is genuine, we enable \textit{show\_consts} +and \textit{show\_\allowbreak datatypes}: + +\prew +{\slshape Datatype:} \\ +\hbox{}\qquad $'a$~\textit{list}~= $\{[],\, [a_3, a_3],\, [a_3],\, \unr\}$ \\ +{\slshape Constants:} \\ +\hbox{}\qquad $\lambda x_1.\; x_1 \mathbin{@} [y, y] = \undef([] := [a_3, a_3],\> [a_3, a_3] := \unk,\> [a_3] := \unk)$ \\ +\hbox{}\qquad $\textit{hd} = \undef([] := a_2,\> [a_3, a_3] := a_3,\> [a_3] := a_3)$ +\postw + +Since $\mathit{hd}~[]$ is undefined in the logic, it may be given any value, +including $a_2$. + +The second constant, $\lambda x_1.\; x_1 \mathbin{@} [y, y]$, is simply the +append operator whose second argument is fixed to be $[y, y]$. Appending $[a_3, +a_3]$ to $[a_3]$ would normally give $[a_3, a_3, a_3]$, but this value is not +representable in the subset of $'a$~\textit{list} considered by Nitpick, which +is shown under the ``Datatype'' heading; hence the result is $\unk$. Similarly, +appending $[a_3, a_3]$ to itself gives $\unk$. + +Given \textit{card}~$'a = 3$ and \textit{card}~$'a~\textit{list} = 3$, Nitpick +considers the following subsets: + +\kern-.5\smallskipamount %% TYPESETTING + +\prew +\begin{multicols}{3} +$\{[],\, [a_1],\, [a_2]\}$; \\ +$\{[],\, [a_1],\, [a_3]\}$; \\ +$\{[],\, [a_2],\, [a_3]\}$; \\ +$\{[],\, [a_1],\, [a_1, a_1]\}$; \\ +$\{[],\, [a_1],\, [a_2, a_1]\}$; \\ +$\{[],\, [a_1],\, [a_3, a_1]\}$; \\ +$\{[],\, [a_2],\, [a_1, a_2]\}$; \\ +$\{[],\, [a_2],\, [a_2, a_2]\}$; \\ +$\{[],\, [a_2],\, [a_3, a_2]\}$; \\ +$\{[],\, [a_3],\, [a_1, a_3]\}$; \\ +$\{[],\, [a_3],\, [a_2, a_3]\}$; \\ +$\{[],\, [a_3],\, [a_3, a_3]\}$. +\end{multicols} +\postw + +\kern-2\smallskipamount %% TYPESETTING + +All subterm-closed subsets of $'a~\textit{list}$ consisting of three values +are listed and only those. As an example of a non-subterm-closed subset, +consider $\mathcal{S} = \{[],\, [a_1],\,\allowbreak [a_1, a_3]\}$, and observe +that $[a_1, a_3]$ (i.e., $a_1 \mathbin{\#} [a_3]$) has $[a_3] \notin +\mathcal{S}$ as a subterm. + +Here's another m\"ochtegern-lemma that Nitpick can refute without a blink: + +\prew +\textbf{lemma} ``$\lbrakk \textit{length}~\textit{xs} = 1;\> \textit{length}~\textit{ys} = 1 +\rbrakk \,\Longrightarrow\, \textit{xs} = \textit{ys}$'' +\\ +\textbf{nitpick} [\textit{show\_datatypes}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample for \textit{card} $'a$~= 3: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{xs} = [a_2]$ \\ +\hbox{}\qquad\qquad $\textit{ys} = [a_3]$ \\ +\hbox{}\qquad Datatypes: \\ +\hbox{}\qquad\qquad $\textit{nat} = \{0,\, 1,\, 2,\, \unr\}$ \\ +\hbox{}\qquad\qquad $'a$~\textit{list} = $\{[],\, [a_3],\, [a_2],\, \unr\}$ +\postw + +Because datatypes are approximated using a three-valued logic, there is usually +no need to systematically enumerate cardinalities: If Nitpick cannot find a +genuine counterexample for \textit{card}~$'a~\textit{list}$~= 10, it is very +unlikely that one could be found for smaller cardinalities. + +\subsection{Typedefs, Records, Rationals, and Reals} +\label{typedefs-records-rationals-and-reals} + +Nitpick generally treats types declared using \textbf{typedef} as datatypes +whose single constructor is the corresponding \textit{Abs\_\kern.1ex} function. +For example: + +\prew +\textbf{typedef}~\textit{three} = ``$\{0\Colon\textit{nat},\, 1,\, 2\}$'' \\ +\textbf{by}~\textit{blast} \\[2\smallskipamount] +\textbf{definition}~$A \mathbin{\Colon} \textit{three}$ \textbf{where} ``\kern-.1em$A \,\equiv\, \textit{Abs\_\allowbreak three}~0$'' \\ +\textbf{definition}~$B \mathbin{\Colon} \textit{three}$ \textbf{where} ``$B \,\equiv\, \textit{Abs\_three}~1$'' \\ +\textbf{definition}~$C \mathbin{\Colon} \textit{three}$ \textbf{where} ``$C \,\equiv\, \textit{Abs\_three}~2$'' \\[2\smallskipamount] +\textbf{lemma} ``$\lbrakk P~A;\> P~B\rbrakk \,\Longrightarrow\, P~x$'' \\ +\textbf{nitpick} [\textit{show\_datatypes}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $P = \{\Abs{1},\, \Abs{0}\}$ \\ +\hbox{}\qquad\qquad $x = \Abs{2}$ \\ +\hbox{}\qquad Datatypes: \\ +\hbox{}\qquad\qquad $\textit{nat} = \{0,\, 1,\, 2,\, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{three} = \{\Abs{2},\, \Abs{1},\, \Abs{0},\, \unr\}$ +\postw + +%% MARK +In the output above, $\Abs{n}$ abbreviates $\textit{Abs\_three}~n$. + +%% MARK +Records, which are implemented as \textbf{typedef}s behind the scenes, are +handled in much the same way: + +\prew +\textbf{record} \textit{point} = \\ +\hbox{}\quad $\textit{Xcoord} \mathbin{\Colon} \textit{int}$ \\ +\hbox{}\quad $\textit{Ycoord} \mathbin{\Colon} \textit{int}$ \\[2\smallskipamount] +\textbf{lemma} ``$\textit{Xcoord}~(p\Colon\textit{point}) = \textit{Xcoord}~(q\Colon\textit{point})$'' \\ +\textbf{nitpick} [\textit{show\_datatypes}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $p = \lparr\textit{Xcoord} = 0,\> \textit{Ycoord} = 0\rparr$ \\ +\hbox{}\qquad\qquad $q = \lparr\textit{Xcoord} = 1,\> \textit{Ycoord} = 1\rparr$ \\ +\hbox{}\qquad Datatypes: \\ +\hbox{}\qquad\qquad $\textit{int} = \{0,\, 1,\, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{point} = \{\lparr\textit{Xcoord} = 1,\> +\textit{Ycoord} = 1\rparr,\> \lparr\textit{Xcoord} = 0,\> \textit{Ycoord} = 0\rparr,\, \unr\}$\kern-1pt %% QUIET +\postw + +Finally, Nitpick provides rudimentary support for rationals and reals using a +similar approach: + +\prew +\textbf{lemma} ``$4 * x + 3 * (y\Colon\textit{real}) \not= 1/2$'' \\ +\textbf{nitpick} [\textit{show\_datatypes}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $x = 1/2$ \\ +\hbox{}\qquad\qquad $y = -1/2$ \\ +\hbox{}\qquad Datatypes: \\ +\hbox{}\qquad\qquad $\textit{nat} = \{0,\, 1,\, 2,\, 3,\, 4,\, 5,\, 6,\, 7,\, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{int} = \{0,\, 1,\, 2,\, 3,\, 4,\, -3,\, -2,\, -1,\, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{real} = \{1,\, 0,\, 4,\, -3/2,\, 3,\, 2,\, 1/2,\, -1/2,\, \unr\}$ +\postw + +\subsection{Inductive and Coinductive Predicates} +\label{inductive-and-coinductive-predicates} + +Inductively defined predicates (and sets) are particularly problematic for +counterexample generators. They can make Quickcheck~\cite{berghofer-nipkow-2004} +loop forever and Refute~\cite{weber-2008} run out of resources. The crux of +the problem is that they are defined using a least fixed point construction. + +Nitpick's philosophy is that not all inductive predicates are equal. Consider +the \textit{even} predicate below: + +\prew +\textbf{inductive}~\textit{even}~\textbf{where} \\ +``\textit{even}~0'' $\,\mid$ \\ +``\textit{even}~$n\,\Longrightarrow\, \textit{even}~(\textit{Suc}~(\textit{Suc}~n))$'' +\postw + +This predicate enjoys the desirable property of being well-founded, which means +that the introduction rules don't give rise to infinite chains of the form + +\prew +$\cdots\,\Longrightarrow\, \textit{even}~k'' + \,\Longrightarrow\, \textit{even}~k' + \,\Longrightarrow\, \textit{even}~k.$ +\postw + +For \textit{even}, this is obvious: Any chain ending at $k$ will be of length +$k/2 + 1$: + +\prew +$\textit{even}~0\,\Longrightarrow\, \textit{even}~2\,\Longrightarrow\, \cdots + \,\Longrightarrow\, \textit{even}~(k - 2) + \,\Longrightarrow\, \textit{even}~k.$ +\postw + +Wellfoundedness is desirable because it enables Nitpick to use a very efficient +fixed point computation.% +\footnote{If an inductive predicate is +well-founded, then it has exactly one fixed point, which is simultaneously the +least and the greatest fixed point. In these circumstances, the computation of +the least fixed point amounts to the computation of an arbitrary fixed point, +which can be performed using a straightforward recursive equation.} +Moreover, Nitpick can prove wellfoundedness of most well-founded predicates, +just as Isabelle's \textbf{function} package usually discharges termination +proof obligations automatically. + +Let's try an example: + +\prew +\textbf{lemma} ``$\exists n.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\ +\textbf{nitpick}~[\textit{card nat}~= 100,\, \textit{verbose}] \\[2\smallskipamount] +\slshape The inductive predicate ``\textit{even}'' was proved well-founded. +Nitpick can compute it efficiently. \\[2\smallskipamount] +Trying 1 scope: \\ +\hbox{}\qquad \textit{card nat}~= 100. \\[2\smallskipamount] +Nitpick found a potential counterexample for \textit{card nat}~= 100: \\[2\smallskipamount] +\hbox{}\qquad Empty assignment \\[2\smallskipamount] +Nitpick could not find a better counterexample. \\[2\smallskipamount] +Total time: 2274 ms. +\postw + +No genuine counterexample is possible because Nitpick cannot rule out the +existence of a natural number $n \ge 100$ such that both $\textit{even}~n$ and +$\textit{even}~(\textit{Suc}~n)$ are true. To help Nitpick, we can bound the +existential quantifier: + +\prew +\textbf{lemma} ``$\exists n \mathbin{\le} 99.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\ +\textbf{nitpick}~[\textit{card nat}~= 100] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Empty assignment +\postw + +So far we were blessed by the wellfoundedness of \textit{even}. What happens if +we use the following definition instead? + +\prew +\textbf{inductive} $\textit{even}'$ \textbf{where} \\ +``$\textit{even}'~(0{\Colon}\textit{nat})$'' $\,\mid$ \\ +``$\textit{even}'~2$'' $\,\mid$ \\ +``$\lbrakk\textit{even}'~m;\> \textit{even}'~n\rbrakk \,\Longrightarrow\, \textit{even}'~(m + n)$'' +\postw + +This definition is not well-founded: From $\textit{even}'~0$ and +$\textit{even}'~0$, we can derive that $\textit{even}'~0$. Nonetheless, the +predicates $\textit{even}$ and $\textit{even}'$ are equivalent. + +Let's check a property involving $\textit{even}'$. To make up for the +foreseeable computational hurdles entailed by non-wellfoundedness, we decrease +\textit{nat}'s cardinality to a mere 10: + +\prew +\textbf{lemma}~``$\exists n \in \{0, 2, 4, 6, 8\}.\; +\lnot\;\textit{even}'~n$'' \\ +\textbf{nitpick}~[\textit{card nat}~= 10,\, \textit{verbose},\, \textit{show\_consts}] \\[2\smallskipamount] +\slshape +The inductive predicate ``$\textit{even}'\!$'' could not be proved well-founded. +Nitpick might need to unroll it. \\[2\smallskipamount] +Trying 6 scopes: \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 0; \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 1; \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 2; \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 4; \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 8; \\ +\hbox{}\qquad \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 9. \\[2\smallskipamount] +Nitpick found a counterexample for \textit{card nat}~= 10 and \textit{iter} $\textit{even}'$~= 2: \\[2\smallskipamount] +\hbox{}\qquad Constant: \nopagebreak \\ +\hbox{}\qquad\qquad $\lambda i.\; \textit{even}'$ = $\undef(\!\begin{aligned}[t] +& 2 := \{0, 2, 4, 6, 8, 1^\Q, 3^\Q, 5^\Q, 7^\Q, 9^\Q\}, \\[-2pt] +& 1 := \{0, 2, 4, 1^\Q, 3^\Q, 5^\Q, 6^\Q, 7^\Q, 8^\Q, 9^\Q\}, \\[-2pt] +& 0 := \{0, 2, 1^\Q, 3^\Q, 4^\Q, 5^\Q, 6^\Q, 7^\Q, 8^\Q, 9^\Q\})\end{aligned}$ \\[2\smallskipamount] +Total time: 1140 ms. +\postw + +Nitpick's output is very instructive. First, it tells us that the predicate is +unrolled, meaning that it is computed iteratively from the empty set. Then it +lists six scopes specifying different bounds on the numbers of iterations:\ 0, +1, 2, 4, 8, and~9. + +The output also shows how each iteration contributes to $\textit{even}'$. The +notation $\lambda i.\; \textit{even}'$ indicates that the value of the +predicate depends on an iteration counter. Iteration 0 provides the basis +elements, $0$ and $2$. Iteration 1 contributes $4$ ($= 2 + 2$). Iteration 2 +throws $6$ ($= 2 + 4 = 4 + 2$) and $8$ ($= 4 + 4$) into the mix. Further +iterations would not contribute any new elements. + +Some values are marked with superscripted question +marks~(`\lower.2ex\hbox{$^\Q$}'). These are the elements for which the +predicate evaluates to $\unk$. Thus, $\textit{even}'$ evaluates to either +\textit{True} or $\unk$, never \textit{False}. + +When unrolling a predicate, Nitpick tries 0, 1, 2, 4, 8, 12, 16, and 24 +iterations. However, these numbers are bounded by the cardinality of the +predicate's domain. With \textit{card~nat}~= 10, no more than 9 iterations are +ever needed to compute the value of a \textit{nat} predicate. You can specify +the number of iterations using the \textit{iter} option, as explained in +\S\ref{scope-of-search}. + +In the next formula, $\textit{even}'$ occurs both positively and negatively: + +\prew +\textbf{lemma} ``$\textit{even}'~(n - 2) \,\Longrightarrow\, \textit{even}'~n$'' \\ +\textbf{nitpick} [\textit{card nat} = 10,\, \textit{show\_consts}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $n = 1$ \\ +\hbox{}\qquad Constants: \nopagebreak \\ +\hbox{}\qquad\qquad $\lambda i.\; \textit{even}'$ = $\undef(\!\begin{aligned}[t] +& 0 := \{0, 2, 1^\Q, 3^\Q, 4^\Q, 5^\Q, 6^\Q, 7^\Q, 8^\Q, 9^\Q\})\end{aligned}$ \\ +\hbox{}\qquad\qquad $\textit{even}' \subseteq \{0, 2, 4, 6, 8, \unr\}$ +\postw + +Notice the special constraint $\textit{even}' \subseteq \{0,\, 2,\, 4,\, 6,\, +8,\, \unr\}$ in the output, whose right-hand side represents an arbitrary +fixed point (not necessarily the least one). It is used to falsify +$\textit{even}'~n$. In contrast, the unrolled predicate is used to satisfy +$\textit{even}'~(n - 2)$. + +Coinductive predicates are handled dually. For example: + +\prew +\textbf{coinductive} \textit{nats} \textbf{where} \\ +``$\textit{nats}~(x\Colon\textit{nat}) \,\Longrightarrow\, \textit{nats}~x$'' \\[2\smallskipamount] +\textbf{lemma} ``$\textit{nats} = \{0, 1, 2, 3, 4\}$'' \\ +\textbf{nitpick}~[\textit{card nat} = 10,\, \textit{show\_consts}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: +\\[2\smallskipamount] +\hbox{}\qquad Constants: \nopagebreak \\ +\hbox{}\qquad\qquad $\lambda i.\; \textit{nats} = \undef(0 := \{\!\begin{aligned}[t] +& 0^\Q, 1^\Q, 2^\Q, 3^\Q, 4^\Q, 5^\Q, 6^\Q, 7^\Q, 8^\Q, 9^\Q, \\[-2pt] +& \unr\})\end{aligned}$ \\ +\hbox{}\qquad\qquad $nats \supseteq \{9, 5^\Q, 6^\Q, 7^\Q, 8^\Q, \unr\}$ +\postw + +As a special case, Nitpick uses Kodkod's transitive closure operator to encode +negative occurrences of non-well-founded ``linear inductive predicates,'' i.e., +inductive predicates for which each the predicate occurs in at most one +assumption of each introduction rule. For example: + +\prew +\textbf{inductive} \textit{odd} \textbf{where} \\ +``$\textit{odd}~1$'' $\,\mid$ \\ +``$\lbrakk \textit{odd}~m;\>\, \textit{even}~n\rbrakk \,\Longrightarrow\, \textit{odd}~(m + n)$'' \\[2\smallskipamount] +\textbf{lemma}~``$\textit{odd}~n \,\Longrightarrow\, \textit{odd}~(n - 2)$'' \\ +\textbf{nitpick}~[\textit{card nat} = 10,\, \textit{show\_consts}] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: +\\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $n = 1$ \\ +\hbox{}\qquad Constants: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{even} = \{0, 2, 4, 6, 8, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{odd}_{\textsl{base}} = \{1, \unr\}$ \\ +\hbox{}\qquad\qquad $\textit{odd}_{\textsl{step}} = \! +\!\begin{aligned}[t] + & \{(0, 0), (0, 2), (0, 4), (0, 6), (0, 8), (1, 1), (1, 3), (1, 5), \\[-2pt] + & \phantom{\{} (1, 7), (1, 9), (2, 2), (2, 4), (2, 6), (2, 8), (3, 3), + (3, 5), \\[-2pt] + & \phantom{\{} (3, 7), (3, 9), (4, 4), (4, 6), (4, 8), (5, 5), (5, 7), (5, 9), \\[-2pt] + & \phantom{\{} (6, 6), (6, 8), (7, 7), (7, 9), (8, 8), (9, 9), \unr\}\end{aligned}$ \\ +\hbox{}\qquad\qquad $\textit{odd} \subseteq \{1, 3, 5, 7, 9, 8^\Q, \unr\}$ +\postw + +\noindent +In the output, $\textit{odd}_{\textrm{base}}$ represents the base elements and +$\textit{odd}_{\textrm{step}}$ is a transition relation that computes new +elements from known ones. The set $\textit{odd}$ consists of all the values +reachable through the reflexive transitive closure of +$\textit{odd}_{\textrm{step}}$ starting with any element from +$\textit{odd}_{\textrm{base}}$, namely 1, 3, 5, 7, and 9. Using Kodkod's +transitive closure to encode linear predicates is normally either more thorough +or more efficient than unrolling (depending on the value of \textit{iter}), but +for those cases where it isn't you can disable it by passing the +\textit{dont\_star\_linear\_preds} option. + +\subsection{Coinductive Datatypes} +\label{coinductive-datatypes} + +While Isabelle regrettably lacks a high-level mechanism for defining coinductive +datatypes, the \textit{Coinductive\_List} theory provides a coinductive ``lazy +list'' datatype, $'a~\textit{llist}$, defined the hard way. Nitpick supports +these lazy lists seamlessly and provides a hook, described in +\S\ref{registration-of-coinductive-datatypes}, to register custom coinductive +datatypes. + +(Co)intuitively, a coinductive datatype is similar to an inductive datatype but +allows infinite objects. Thus, the infinite lists $\textit{ps}$ $=$ $[a, a, a, +\ldots]$, $\textit{qs}$ $=$ $[a, b, a, b, \ldots]$, and $\textit{rs}$ $=$ $[0, +1, 2, 3, \ldots]$ can be defined as lazy lists using the +$\textit{LNil}\mathbin{\Colon}{'}a~\textit{llist}$ and +$\textit{LCons}\mathbin{\Colon}{'}a \mathbin{\Rightarrow} {'}a~\textit{llist} +\mathbin{\Rightarrow} {'}a~\textit{llist}$ constructors. + +Although it is otherwise no friend of infinity, Nitpick can find counterexamples +involving cyclic lists such as \textit{ps} and \textit{qs} above as well as +finite lists: + +\prew +\textbf{lemma} ``$\textit{xs} \not= \textit{LCons}~a~\textit{xs}$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample for {\itshape card}~$'a$ = 1: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{a} = a_1$ \\ +\hbox{}\qquad\qquad $\textit{xs} = \textsl{THE}~\omega.\; \omega = \textit{LCons}~a_1~\omega$ +\postw + +The notation $\textrm{THE}~\omega.\; \omega = t(\omega)$ stands +for the infinite term $t(t(t(\ldots)))$. Hence, \textit{xs} is simply the +infinite list $[a_1, a_1, a_1, \ldots]$. + +The next example is more interesting: + +\prew +\textbf{lemma}~``$\lbrakk\textit{xs} = \textit{LCons}~a~\textit{xs};\>\, +\textit{ys} = \textit{iterates}~(\lambda b.\> a)~b\rbrakk \,\Longrightarrow\, \textit{xs} = \textit{ys}$'' \\ +\textbf{nitpick} [\textit{verbose}] \\[2\smallskipamount] +\slshape The type ``\kern1pt$'a$'' passed the monotonicity test. Nitpick might be able to skip +some scopes. \\[2\smallskipamount] +Trying 8 scopes: \\ +\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} ``\kern1pt$'a~\textit{list}$''~= 1, +and \textit{bisim\_depth}~= 0. \\ +\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount] +\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} ``\kern1pt$'a~\textit{list}$''~= 8, +and \textit{bisim\_depth}~= 7. \\[2\smallskipamount] +Nitpick found a counterexample for {\itshape card}~$'a$ = 2, +\textit{card}~``\kern1pt$'a~\textit{list}$''~= 2, and \textit{bisim\_\allowbreak +depth}~= 1: +\\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{a} = a_2$ \\ +\hbox{}\qquad\qquad $\textit{b} = a_1$ \\ +\hbox{}\qquad\qquad $\textit{xs} = \textsl{THE}~\omega.\; \omega = \textit{LCons}~a_2~\omega$ \\ +\hbox{}\qquad\qquad $\textit{ys} = \textit{LCons}~a_1~(\textsl{THE}~\omega.\; \omega = \textit{LCons}~a_2~\omega)$ \\[2\smallskipamount] +Total time: 726 ms. +\postw + +The lazy list $\textit{xs}$ is simply $[a_2, a_2, a_2, \ldots]$, whereas +$\textit{ys}$ is $[a_1, a_2, a_2, a_2, \ldots]$, i.e., a lasso-shaped list with +$[a_1]$ as its stem and $[a_2]$ as its cycle. In general, the list segment +within the scope of the {THE} binder corresponds to the lasso's cycle, whereas +the segment leading to the binder is the stem. + +A salient property of coinductive datatypes is that two objects are considered +equal if and only if they lead to the same observations. For example, the lazy +lists $\textrm{THE}~\omega.\; \omega = +\textit{LCons}~a~(\textit{LCons}~b~\omega)$ and +$\textit{LCons}~a~(\textrm{THE}~\omega.\; \omega = +\textit{LCons}~b~(\textit{LCons}~a~\omega))$ are identical, because both lead +to the sequence of observations $a$, $b$, $a$, $b$, \hbox{\ldots} (or, +equivalently, both encode the infinite list $[a, b, a, b, \ldots]$). This +concept of equality for coinductive datatypes is called bisimulation and is +defined coinductively. + +Internally, Nitpick encodes the coinductive bisimilarity predicate as part of +the Kodkod problem to ensure that distinct objects lead to different +observations. This precaution is somewhat expensive and often unnecessary, so it +can be disabled by setting the \textit{bisim\_depth} option to $-1$. The +bisimilarity check is then performed \textsl{after} the counterexample has been +found to ensure correctness. If this after-the-fact check fails, the +counterexample is tagged as ``likely genuine'' and Nitpick recommends to try +again with \textit{bisim\_depth} set to a nonnegative integer. Disabling the +check for the previous example saves approximately 150~milli\-seconds; the speed +gains can be more significant for larger scopes. + +The next formula illustrates the need for bisimilarity (either as a Kodkod +predicate or as an after-the-fact check) to prevent spurious counterexamples: + +\prew +\textbf{lemma} ``$\lbrakk xs = \textit{LCons}~a~\textit{xs};\>\, \textit{ys} = \textit{LCons}~a~\textit{ys}\rbrakk +\,\Longrightarrow\, \textit{xs} = \textit{ys}$'' \\ +\textbf{nitpick} [\textit{bisim\_depth} = $-1$,\, \textit{show\_datatypes}] \\[2\smallskipamount] +\slshape Nitpick found a likely genuine counterexample for $\textit{card}~'a$ = 2: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $a = a_2$ \\ +\hbox{}\qquad\qquad $\textit{xs} = \textsl{THE}~\omega.\; \omega = +\textit{LCons}~a_2~\omega$ \\ +\hbox{}\qquad\qquad $\textit{ys} = \textsl{THE}~\omega.\; \omega = \textit{LCons}~a_2~\omega$ \\ +\hbox{}\qquad Codatatype:\strut \nopagebreak \\ +\hbox{}\qquad\qquad $'a~\textit{llist} = +\{\!\begin{aligned}[t] + & \textsl{THE}~\omega.\; \omega = \textit{LCons}~a_2~\omega, \\[-2pt] + & \textsl{THE}~\omega.\; \omega = \textit{LCons}~a_2~\omega,\> \unr\}\end{aligned}$ +\\[2\smallskipamount] +Try again with ``\textit{bisim\_depth}'' set to a nonnegative value to confirm +that the counterexample is genuine. \\[2\smallskipamount] +{\upshape\textbf{nitpick}} \\[2\smallskipamount] +\slshape Nitpick found no counterexample. +\postw + +In the first \textbf{nitpick} invocation, the after-the-fact check discovered +that the two known elements of type $'a~\textit{llist}$ are bisimilar. + +A compromise between leaving out the bisimilarity predicate from the Kodkod +problem and performing the after-the-fact check is to specify a lower +nonnegative \textit{bisim\_depth} value than the default one provided by +Nitpick. In general, a value of $K$ means that Nitpick will require all lists to +be distinguished from each other by their prefixes of length $K$. Be aware that +setting $K$ to a too low value can overconstrain Nitpick, preventing it from +finding any counterexamples. + +\subsection{Boxing} +\label{boxing} + +Nitpick normally maps function and product types directly to the corresponding +Kodkod concepts. As a consequence, if $'a$ has cardinality 3 and $'b$ has +cardinality 4, then $'a \times {'}b$ has cardinality 12 ($= 4 \times 3$) and $'a +\Rightarrow {'}b$ has cardinality 64 ($= 4^3$). In some circumstances, it pays +off to treat these types in the same way as plain datatypes, by approximating +them by a subset of a given cardinality. This technique is called ``boxing'' and +is particularly useful for functions passed as arguments to other functions, for +high-arity functions, and for large tuples. Under the hood, boxing involves +wrapping occurrences of the types $'a \times {'}b$ and $'a \Rightarrow {'}b$ in +isomorphic datatypes, as can be seen by enabling the \textit{debug} option. + +To illustrate boxing, we consider a formalization of $\lambda$-terms represented +using de Bruijn's notation: + +\prew +\textbf{datatype} \textit{tm} = \textit{Var}~\textit{nat}~$\mid$~\textit{Lam}~\textit{tm} $\mid$ \textit{App~tm~tm} +\postw + +The $\textit{lift}~t~k$ function increments all variables with indices greater +than or equal to $k$ by one: + +\prew +\textbf{primrec} \textit{lift} \textbf{where} \\ +``$\textit{lift}~(\textit{Var}~j)~k = \textit{Var}~(\textrm{if}~j < k~\textrm{then}~j~\textrm{else}~j + 1)$'' $\mid$ \\ +``$\textit{lift}~(\textit{Lam}~t)~k = \textit{Lam}~(\textit{lift}~t~(k + 1))$'' $\mid$ \\ +``$\textit{lift}~(\textit{App}~t~u)~k = \textit{App}~(\textit{lift}~t~k)~(\textit{lift}~u~k)$'' +\postw + +The $\textit{loose}~t~k$ predicate returns \textit{True} if and only if +term $t$ has a loose variable with index $k$ or more: + +\prew +\textbf{primrec}~\textit{loose} \textbf{where} \\ +``$\textit{loose}~(\textit{Var}~j)~k = (j \ge k)$'' $\mid$ \\ +``$\textit{loose}~(\textit{Lam}~t)~k = \textit{loose}~t~(\textit{Suc}~k)$'' $\mid$ \\ +``$\textit{loose}~(\textit{App}~t~u)~k = (\textit{loose}~t~k \mathrel{\lor} \textit{loose}~u~k)$'' +\postw + +Next, the $\textit{subst}~\sigma~t$ function applies the substitution $\sigma$ +on $t$: + +\prew +\textbf{primrec}~\textit{subst} \textbf{where} \\ +``$\textit{subst}~\sigma~(\textit{Var}~j) = \sigma~j$'' $\mid$ \\ +``$\textit{subst}~\sigma~(\textit{Lam}~t) = {}$\phantom{''} \\ +\phantom{``}$\textit{Lam}~(\textit{subst}~(\lambda n.\> \textrm{case}~n~\textrm{of}~0 \Rightarrow \textit{Var}~0 \mid \textit{Suc}~m \Rightarrow \textit{lift}~(\sigma~m)~1)~t)$'' $\mid$ \\ +``$\textit{subst}~\sigma~(\textit{App}~t~u) = \textit{App}~(\textit{subst}~\sigma~t)~(\textit{subst}~\sigma~u)$'' +\postw + +A substitution is a function that maps variable indices to terms. Observe that +$\sigma$ is a function passed as argument and that Nitpick can't optimize it +away, because the recursive call for the \textit{Lam} case involves an altered +version. Also notice the \textit{lift} call, which increments the variable +indices when moving under a \textit{Lam}. + +A reasonable property to expect of substitution is that it should leave closed +terms unchanged. Alas, even this simple property does not hold: + +\pre +\textbf{lemma}~``$\lnot\,\textit{loose}~t~0 \,\Longrightarrow\, \textit{subst}~\sigma~t = t$'' \\ +\textbf{nitpick} [\textit{verbose}] \\[2\smallskipamount] +\slshape +Trying 8 scopes: \nopagebreak \\ +\hbox{}\qquad \textit{card~nat}~= 1, \textit{card tm}~= 1, and \textit{card} ``$\textit{nat} \Rightarrow \textit{tm}$'' = 1; \\ +\hbox{}\qquad \textit{card~nat}~= 2, \textit{card tm}~= 2, and \textit{card} ``$\textit{nat} \Rightarrow \textit{tm}$'' = 2; \\ +\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount] +\hbox{}\qquad \textit{card~nat}~= 8, \textit{card tm}~= 8, and \textit{card} ``$\textit{nat} \Rightarrow \textit{tm}$'' = 8. \\[2\smallskipamount] +Nitpick found a counterexample for \textit{card~nat}~= 6, \textit{card~tm}~= 6, +and \textit{card}~``$\textit{nat} \Rightarrow \textit{tm}$''~= 6: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\sigma = \undef(\!\begin{aligned}[t] +& 0 := \textit{Var}~0,\> + 1 := \textit{Var}~0,\> + 2 := \textit{Var}~0, \\[-2pt] +& 3 := \textit{Var}~0,\> + 4 := \textit{Var}~0,\> + 5 := \textit{Var}~0)\end{aligned}$ \\ +\hbox{}\qquad\qquad $t = \textit{Lam}~(\textit{Lam}~(\textit{Var}~1))$ \\[2\smallskipamount] +Total time: $4679$ ms. +\postw + +Using \textit{eval}, we find out that $\textit{subst}~\sigma~t = +\textit{Lam}~(\textit{Lam}~(\textit{Var}~0))$. Using the traditional +$\lambda$-term notation, $t$~is +$\lambda x\, y.\> x$ whereas $\textit{subst}~\sigma~t$ is $\lambda x\, y.\> y$. +The bug is in \textit{subst}: The $\textit{lift}~(\sigma~m)~1$ call should be +replaced with $\textit{lift}~(\sigma~m)~0$. + +An interesting aspect of Nitpick's verbose output is that it assigned inceasing +cardinalities from 1 to 8 to the type $\textit{nat} \Rightarrow \textit{tm}$. +For the formula of interest, knowing 6 values of that type was enough to find +the counterexample. Without boxing, $46\,656$ ($= 6^6$) values must be +considered, a hopeless undertaking: + +\prew +\textbf{nitpick} [\textit{dont\_box}] \\[2\smallskipamount] +{\slshape Nitpick ran out of time after checking 4 of 8 scopes.} +\postw + +{\looseness=-1 +Boxing can be enabled or disabled globally or on a per-type basis using the +\textit{box} option. Moreover, setting the cardinality of a function or +product type implicitly enables boxing for that type. Nitpick usually performs +reasonable choices about which types should be boxed, but option tweaking +sometimes helps. + +} + +\subsection{Scope Monotonicity} +\label{scope-monotonicity} + +The \textit{card} option (together with \textit{iter}, \textit{bisim\_depth}, +and \textit{max}) controls which scopes are actually tested. In general, to +exhaust all models below a certain cardinality bound, the number of scopes that +Nitpick must consider increases exponentially with the number of type variables +(and \textbf{typedecl}'d types) occurring in the formula. Given the default +cardinality specification of 1--8, no fewer than $8^4 = 4096$ scopes must be +considered for a formula involving $'a$, $'b$, $'c$, and $'d$. + +Fortunately, many formulas exhibit a property called \textsl{scope +monotonicity}, meaning that if the formula is falsifiable for a given scope, +it is also falsifiable for all larger scopes \cite[p.~165]{jackson-2006}. + +Consider the formula + +\prew +\textbf{lemma}~``$\textit{length~xs} = \textit{length~ys} \,\Longrightarrow\, \textit{rev}~(\textit{zip~xs~ys}) = \textit{zip~xs}~(\textit{rev~ys})$'' +\postw + +where \textit{xs} is of type $'a~\textit{list}$ and \textit{ys} is of type +$'b~\textit{list}$. A priori, Nitpick would need to consider 512 scopes to +exhaust the specification \textit{card}~= 1--8. However, our intuition tells us +that any counterexample found with a small scope would still be a counterexample +in a larger scope---by simply ignoring the fresh $'a$ and $'b$ values provided +by the larger scope. Nitpick comes to the same conclusion after a careful +inspection of the formula and the relevant definitions: + +\prew +\textbf{nitpick}~[\textit{verbose}] \\[2\smallskipamount] +\slshape +The types ``\kern1pt$'a$'' and ``\kern1pt$'b$'' passed the monotonicity test. +Nitpick might be able to skip some scopes. + \\[2\smallskipamount] +Trying 8 scopes: \\ +\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} $'b$~= 1, +\textit{card} \textit{nat}~= 1, \textit{card} ``$('a \times {'}b)$ +\textit{list}''~= 1, \\ +\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 1, and +\textit{card} ``\kern1pt$'b$ \textit{list}''~= 1. \\ +\hbox{}\qquad \textit{card} $'a$~= 2, \textit{card} $'b$~= 2, +\textit{card} \textit{nat}~= 2, \textit{card} ``$('a \times {'}b)$ +\textit{list}''~= 2, \\ +\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 2, and +\textit{card} ``\kern1pt$'b$ \textit{list}''~= 2. \\ +\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount] +\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} $'b$~= 8, +\textit{card} \textit{nat}~= 8, \textit{card} ``$('a \times {'}b)$ +\textit{list}''~= 8, \\ +\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 8, and +\textit{card} ``\kern1pt$'b$ \textit{list}''~= 8. +\\[2\smallskipamount] +Nitpick found a counterexample for +\textit{card} $'a$~= 5, \textit{card} $'b$~= 5, +\textit{card} \textit{nat}~= 5, \textit{card} ``$('a \times {'}b)$ +\textit{list}''~= 5, \textit{card} ``\kern1pt$'a$ \textit{list}''~= 5, and +\textit{card} ``\kern1pt$'b$ \textit{list}''~= 5: +\\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $\textit{xs} = [a_4, a_5]$ \\ +\hbox{}\qquad\qquad $\textit{ys} = [b_3, b_3]$ \\[2\smallskipamount] +Total time: 1636 ms. +\postw + +In theory, it should be sufficient to test a single scope: + +\prew +\textbf{nitpick}~[\textit{card}~= 8] +\postw + +However, this is often less efficient in practice and may lead to overly complex +counterexamples. + +If the monotonicity check fails but we believe that the formula is monotonic (or +we don't mind missing some counterexamples), we can pass the +\textit{mono} option. To convince yourself that this option is risky, +simply consider this example from \S\ref{skolemization}: + +\prew +\textbf{lemma} ``$\exists g.\; \forall x\Colon 'b.~g~(f~x) = x + \,\Longrightarrow\, \forall y\Colon {'}a.\; \exists x.~y = f~x$'' \\ +\textbf{nitpick} [\textit{mono}] \\[2\smallskipamount] +{\slshape Nitpick found no counterexample.} \\[2\smallskipamount] +\textbf{nitpick} \\[2\smallskipamount] +\slshape +Nitpick found a counterexample for \textit{card} $'a$~= 2 and \textit{card} $'b$~=~1: \\ +\hbox{}\qquad $\vdots$ +\postw + +(It turns out the formula holds if and only if $\textit{card}~'a \le +\textit{card}~'b$.) Although this is rarely advisable, the automatic +monotonicity checks can be disabled by passing \textit{non\_mono} +(\S\ref{optimizations}). + +As insinuated in \S\ref{natural-numbers-and-integers} and +\S\ref{inductive-datatypes}, \textit{nat}, \textit{int}, and inductive datatypes +are normally monotonic and treated as such. The same is true for record types, +\textit{rat}, \textit{real}, and some \textbf{typedef}'d types. Thus, given the +cardinality specification 1--8, a formula involving \textit{nat}, \textit{int}, +\textit{int~list}, \textit{rat}, and \textit{rat~list} will lead Nitpick to +consider only 8~scopes instead of $32\,768$. + +\section{Case Studies} +\label{case-studies} + +As a didactic device, the previous section focused mostly on toy formulas whose +validity can easily be assessed just by looking at the formula. We will now +review two somewhat more realistic case studies that are within Nitpick's +reach:\ a context-free grammar modeled by mutually inductive sets and a +functional implementation of AA trees. The results presented in this +section were produced with the following settings: + +\prew +\textbf{nitpick\_params} [\textit{max\_potential}~= 0,\, \textit{max\_threads} = 2] +\postw + +\subsection{A Context-Free Grammar} +\label{a-context-free-grammar} + +Our first case study is taken from section 7.4 in the Isabelle tutorial +\cite{isa-tutorial}. The following grammar, originally due to Hopcroft and +Ullman, produces all strings with an equal number of $a$'s and $b$'s: + +\prew +\begin{tabular}{@{}r@{$\;\,$}c@{$\;\,$}l@{}} +$S$ & $::=$ & $\epsilon \mid bA \mid aB$ \\ +$A$ & $::=$ & $aS \mid bAA$ \\ +$B$ & $::=$ & $bS \mid aBB$ +\end{tabular} +\postw + +The intuition behind the grammar is that $A$ generates all string with one more +$a$ than $b$'s and $B$ generates all strings with one more $b$ than $a$'s. + +The alphabet consists exclusively of $a$'s and $b$'s: + +\prew +\textbf{datatype} \textit{alphabet}~= $a$ $\mid$ $b$ +\postw + +Strings over the alphabet are represented by \textit{alphabet list}s. +Nonterminals in the grammar become sets of strings. The production rules +presented above can be expressed as a mutually inductive definition: + +\prew +\textbf{inductive\_set} $S$ \textbf{and} $A$ \textbf{and} $B$ \textbf{where} \\ +\textit{R1}:\kern.4em ``$[] \in S$'' $\,\mid$ \\ +\textit{R2}:\kern.4em ``$w \in A\,\Longrightarrow\, b \mathbin{\#} w \in S$'' $\,\mid$ \\ +\textit{R3}:\kern.4em ``$w \in B\,\Longrightarrow\, a \mathbin{\#} w \in S$'' $\,\mid$ \\ +\textit{R4}:\kern.4em ``$w \in S\,\Longrightarrow\, a \mathbin{\#} w \in A$'' $\,\mid$ \\ +\textit{R5}:\kern.4em ``$w \in S\,\Longrightarrow\, b \mathbin{\#} w \in S$'' $\,\mid$ \\ +\textit{R6}:\kern.4em ``$\lbrakk v \in B;\> v \in B\rbrakk \,\Longrightarrow\, a \mathbin{\#} v \mathbin{@} w \in B$'' +\postw + +The conversion of the grammar into the inductive definition was done manually by +Joe Blow, an underpaid undergraduate student. As a result, some errors might +have sneaked in. + +Debugging faulty specifications is at the heart of Nitpick's \textsl{raison +d'\^etre}. A good approach is to state desirable properties of the specification +(here, that $S$ is exactly the set of strings over $\{a, b\}$ with as many $a$'s +as $b$'s) and check them with Nitpick. If the properties are correctly stated, +counterexamples will point to bugs in the specification. For our grammar +example, we will proceed in two steps, separating the soundness and the +completeness of the set $S$. First, soundness: + +\prew +\textbf{theorem}~\textit{S\_sound}: \\ +``$w \in S \longrightarrow \textit{length}~[x\mathbin{\leftarrow} w.\; x = a] = + \textit{length}~[x\mathbin{\leftarrow} w.\; x = b]$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $w = [b]$ +\postw + +It would seem that $[b] \in S$. How could this be? An inspection of the +introduction rules reveals that the only rule with a right-hand side of the form +$b \mathbin{\#} {\ldots} \in S$ that could have introduced $[b]$ into $S$ is +\textit{R5}: + +\prew +``$w \in S\,\Longrightarrow\, b \mathbin{\#} w \in S$'' +\postw + +On closer inspection, we can see that this rule is wrong. To match the +production $B ::= bS$, the second $S$ should be a $B$. We fix the typo and try +again: + +\prew +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $w = [a, a, b]$ +\postw + +Some detective work is necessary to find out what went wrong here. To get $[a, +a, b] \in S$, we need $[a, b] \in B$ by \textit{R3}, which in turn can only come +from \textit{R6}: + +\prew +``$\lbrakk v \in B;\> v \in B\rbrakk \,\Longrightarrow\, a \mathbin{\#} v \mathbin{@} w \in B$'' +\postw + +Now, this formula must be wrong: The same assumption occurs twice, and the +variable $w$ is unconstrained. Clearly, one of the two occurrences of $v$ in +the assumptions should have been a $w$. + +With the correction made, we don't get any counterexample from Nitpick. Let's +move on and check completeness: + +\prew +\textbf{theorem}~\textit{S\_complete}: \\ +``$\textit{length}~[x\mathbin{\leftarrow} w.\; x = a] = + \textit{length}~[x\mathbin{\leftarrow} w.\; x = b] + \longrightarrow w \in S$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variable: \nopagebreak \\ +\hbox{}\qquad\qquad $w = [b, b, a, a]$ +\postw + +Apparently, $[b, b, a, a] \notin S$, even though it has the same numbers of +$a$'s and $b$'s. But since our inductive definition passed the soundness check, +the introduction rules we have are probably correct. Perhaps we simply lack an +introduction rule. Comparing the grammar with the inductive definition, our +suspicion is confirmed: Joe Blow simply forgot the production $A ::= bAA$, +without which the grammar cannot generate two or more $b$'s in a row. So we add +the rule + +\prew +``$\lbrakk v \in A;\> w \in A\rbrakk \,\Longrightarrow\, b \mathbin{\#} v \mathbin{@} w \in A$'' +\postw + +With this last change, we don't get any counterexamples from Nitpick for either +soundness or completeness. We can even generalize our result to cover $A$ and +$B$ as well: + +\prew +\textbf{theorem} \textit{S\_A\_B\_sound\_and\_complete}: \\ +``$w \in S \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = b]$'' \\ +``$w \in A \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] + 1$'' \\ +``$w \in B \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] + 1$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found no counterexample. +\postw + +\subsection{AA Trees} +\label{aa-trees} + +AA trees are a kind of balanced trees discovered by Arne Andersson that provide +similar performance to red-black trees, but with a simpler implementation +\cite{andersson-1993}. They can be used to store sets of elements equipped with +a total order $<$. We start by defining the datatype and some basic extractor +functions: + +\prew +\textbf{datatype} $'a$~\textit{tree} = $\Lambda$ $\mid$ $N$ ``\kern1pt$'a\Colon \textit{linorder}$'' \textit{nat} ``\kern1pt$'a$ \textit{tree}'' ``\kern1pt$'a$ \textit{tree}'' \\[2\smallskipamount] +\textbf{primrec} \textit{data} \textbf{where} \\ +``$\textit{data}~\Lambda = \undef$'' $\,\mid$ \\ +``$\textit{data}~(N~x~\_~\_~\_) = x$'' \\[2\smallskipamount] +\textbf{primrec} \textit{dataset} \textbf{where} \\ +``$\textit{dataset}~\Lambda = \{\}$'' $\,\mid$ \\ +``$\textit{dataset}~(N~x~\_~t~u) = \{x\} \cup \textit{dataset}~t \mathrel{\cup} \textit{dataset}~u$'' \\[2\smallskipamount] +\textbf{primrec} \textit{level} \textbf{where} \\ +``$\textit{level}~\Lambda = 0$'' $\,\mid$ \\ +``$\textit{level}~(N~\_~k~\_~\_) = k$'' \\[2\smallskipamount] +\textbf{primrec} \textit{left} \textbf{where} \\ +``$\textit{left}~\Lambda = \Lambda$'' $\,\mid$ \\ +``$\textit{left}~(N~\_~\_~t~\_) = t$'' \\[2\smallskipamount] +\textbf{primrec} \textit{right} \textbf{where} \\ +``$\textit{right}~\Lambda = \Lambda$'' $\,\mid$ \\ +``$\textit{right}~(N~\_~\_~\_~u) = u$'' +\postw + +The wellformedness criterion for AA trees is fairly complex. Wikipedia states it +as follows \cite{wikipedia-2009-aa-trees}: + +\kern.2\parskip %% TYPESETTING + +\pre +Each node has a level field, and the following invariants must remain true for +the tree to be valid: + +\raggedright + +\kern-.4\parskip %% TYPESETTING + +\begin{enum} +\item[] +\begin{enum} +\item[1.] The level of a leaf node is one. +\item[2.] The level of a left child is strictly less than that of its parent. +\item[3.] The level of a right child is less than or equal to that of its parent. +\item[4.] The level of a right grandchild is strictly less than that of its grandparent. +\item[5.] Every node of level greater than one must have two children. +\end{enum} +\end{enum} +\post + +\kern.4\parskip %% TYPESETTING + +The \textit{wf} predicate formalizes this description: + +\prew +\textbf{primrec} \textit{wf} \textbf{where} \\ +``$\textit{wf}~\Lambda = \textit{True}$'' $\,\mid$ \\ +``$\textit{wf}~(N~\_~k~t~u) =$ \\ +\phantom{``}$(\textrm{if}~t = \Lambda~\textrm{then}$ \\ +\phantom{``$(\quad$}$k = 1 \mathrel{\land} (u = \Lambda \mathrel{\lor} (\textit{level}~u = 1 \mathrel{\land} \textit{left}~u = \Lambda \mathrel{\land} \textit{right}~u = \Lambda))$ \\ +\phantom{``$($}$\textrm{else}$ \\ +\hbox{}\phantom{``$(\quad$}$\textit{wf}~t \mathrel{\land} \textit{wf}~u +\mathrel{\land} u \not= \Lambda \mathrel{\land} \textit{level}~t < k +\mathrel{\land} \textit{level}~u \le k$ \\ +\hbox{}\phantom{``$(\quad$}${\land}\; \textit{level}~(\textit{right}~u) < k)$'' +\postw + +Rebalancing the tree upon insertion and removal of elements is performed by two +auxiliary functions called \textit{skew} and \textit{split}, defined below: + +\prew +\textbf{primrec} \textit{skew} \textbf{where} \\ +``$\textit{skew}~\Lambda = \Lambda$'' $\,\mid$ \\ +``$\textit{skew}~(N~x~k~t~u) = {}$ \\ +\phantom{``}$(\textrm{if}~t \not= \Lambda \mathrel{\land} k = +\textit{level}~t~\textrm{then}$ \\ +\phantom{``(\quad}$N~(\textit{data}~t)~k~(\textit{left}~t)~(N~x~k~ +(\textit{right}~t)~u)$ \\ +\phantom{``(}$\textrm{else}$ \\ +\phantom{``(\quad}$N~x~k~t~u)$'' +\postw + +\prew +\textbf{primrec} \textit{split} \textbf{where} \\ +``$\textit{split}~\Lambda = \Lambda$'' $\,\mid$ \\ +``$\textit{split}~(N~x~k~t~u) = {}$ \\ +\phantom{``}$(\textrm{if}~u \not= \Lambda \mathrel{\land} k = +\textit{level}~(\textit{right}~u)~\textrm{then}$ \\ +\phantom{``(\quad}$N~(\textit{data}~u)~(\textit{Suc}~k)~ +(N~x~k~t~(\textit{left}~u))~(\textit{right}~u)$ \\ +\phantom{``(}$\textrm{else}$ \\ +\phantom{``(\quad}$N~x~k~t~u)$'' +\postw + +Performing a \textit{skew} or a \textit{split} should have no impact on the set +of elements stored in the tree: + +\prew +\textbf{theorem}~\textit{dataset\_skew\_split}:\\ +``$\textit{dataset}~(\textit{skew}~t) = \textit{dataset}~t$'' \\ +``$\textit{dataset}~(\textit{split}~t) = \textit{dataset}~t$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +{\slshape Nitpick ran out of time after checking 7 of 8 scopes.} +\postw + +Furthermore, applying \textit{skew} or \textit{split} to a well-formed tree +should not alter the tree: + +\prew +\textbf{theorem}~\textit{wf\_skew\_split}:\\ +``$\textit{wf}~t\,\Longrightarrow\, \textit{skew}~t = t$'' \\ +``$\textit{wf}~t\,\Longrightarrow\, \textit{split}~t = t$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +{\slshape Nitpick found no counterexample.} +\postw + +Insertion is implemented recursively. It preserves the sort order: + +\prew +\textbf{primrec}~\textit{insort} \textbf{where} \\ +``$\textit{insort}~\Lambda~x = N~x~1~\Lambda~\Lambda$'' $\,\mid$ \\ +``$\textit{insort}~(N~y~k~t~u)~x =$ \\ +\phantom{``}$({*}~(\textit{split} \circ \textit{skew})~{*})~(N~y~k~(\textrm{if}~x < y~\textrm{then}~\textit{insort}~t~x~\textrm{else}~t)$ \\ +\phantom{``$({*}~(\textit{split} \circ \textit{skew})~{*})~(N~y~k~$}$(\textrm{if}~x > y~\textrm{then}~\textit{insort}~u~x~\textrm{else}~u))$'' +\postw + +Notice that we deliberately commented out the application of \textit{skew} and +\textit{split}. Let's see if this causes any problems: + +\prew +\textbf{theorem}~\textit{wf\_insort}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\slshape Nitpick found a counterexample for \textit{card} $'a$ = 4: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $t = N~a_3~1~\Lambda~\Lambda$ \\ +\hbox{}\qquad\qquad $x = a_4$ \\[2\smallskipamount] +Hint: Maybe you forgot a type constraint? +\postw + +It's hard to see why this is a counterexample. The hint is of no help here. To +improve readability, we will restrict the theorem to \textit{nat}, so that we +don't need to look up the value of the $\textit{op}~{<}$ constant to find out +which element is smaller than the other. In addition, we will tell Nitpick to +display the value of $\textit{insort}~t~x$ using the \textit{eval} option. This +gives + +\prew +\textbf{theorem} \textit{wf\_insort\_nat}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~(x\Colon\textit{nat}))$'' \\ +\textbf{nitpick} [\textit{eval} = ``$\textit{insort}~t~x$''] \\[2\smallskipamount] +\slshape Nitpick found a counterexample: \\[2\smallskipamount] +\hbox{}\qquad Free variables: \nopagebreak \\ +\hbox{}\qquad\qquad $t = N~1~1~\Lambda~\Lambda$ \\ +\hbox{}\qquad\qquad $x = 0$ \\ +\hbox{}\qquad Evaluated term: \\ +\hbox{}\qquad\qquad $\textit{insort}~t~x = N~1~1~(N~0~1~\Lambda~\Lambda)~\Lambda$ +\postw + +Nitpick's output reveals that the element $0$ was added as a left child of $1$, +where both have a level of 1. This violates the second AA tree invariant, which +states that a left child's level must be less than its parent's. This shouldn't +come as a surprise, considering that we commented out the tree rebalancing code. +Reintroducing the code seems to solve the problem: + +\prew +\textbf{theorem}~\textit{wf\_insort}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +{\slshape Nitpick ran out of time after checking 6 of 8 scopes.} +\postw + +Insertion should transform the set of elements represented by the tree in the +obvious way: + +\prew +\textbf{theorem} \textit{dataset\_insort}:\kern.4em +``$\textit{dataset}~(\textit{insort}~t~x) = \{x\} \cup \textit{dataset}~t$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +{\slshape Nitpick ran out of time after checking 5 of 8 scopes.} +\postw + +We could continue like this and sketch a complete theory of AA trees without +performing a single proof. Once the definitions and main theorems are in place +and have been thoroughly tested using Nitpick, we could start working on the +proofs. Developing theories this way usually saves time, because faulty theorems +and definitions are discovered much earlier in the process. + +\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\ops#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]} +\def\opt#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\quad [\textit{#3}]} \nopagebreak\\[\parskip]} +\def\opu#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]} +\def\opusmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} + +Nitpick's behavior can be influenced by various options, which can be specified +in brackets after the \textbf{nitpick} command. Default values can be set +using \textbf{nitpick\_\allowbreak params}. For example: + +\prew +\textbf{nitpick\_params} [\textit{verbose}, \,\textit{timeout} = 60$\,s$] +\postw + +The options are categorized as follows:\ mode of operation +(\S\ref{mode-of-operation}), scope of search (\S\ref{scope-of-search}), output +format (\S\ref{output-format}), automatic counterexample checks +(\S\ref{authentication}), optimizations +(\S\ref{optimizations}), and timeouts (\S\ref{timeouts}). + +The number of options can be overwhelming at first glance. Do not let that worry +you: Nitpick's defaults have been chosen so that it almost always does the right +thing, and the most important options have been covered in context in +\S\ref{first-steps}. + +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. Negative integers are prefixed with a hyphen. +\item[$\bullet$] \qtybf{int\_or\_smart}: An integer or \textit{smart}. +\item[$\bullet$] \qtybf{int\_range}: An integer (e.g., 3) or a range +of nonnegative integers (e.g., $1$--$4$). The range symbol `--' can be entered as \texttt{-} (hyphen) or \texttt{\char`\\\char`\}. + +\item[$\bullet$] \qtybf{int\_seq}: A comma-separated sequence of ranges of integers (e.g.,~1{,}3{,}\allowbreak6--8). +\item[$\bullet$] \qtybf{time}: An integer followed by $\textit{min}$ (minutes), $s$ (seconds), or \textit{ms} +(milliseconds), or the keyword \textit{none} ($\infty$ years). +\item[$\bullet$] \qtybf{const}: The name of a HOL constant. +\item[$\bullet$] \qtybf{term}: A HOL term (e.g., ``$f~x$''). +\item[$\bullet$] \qtybf{term\_list}: A space-separated list of HOL terms (e.g., +``$f~x$''~``$g~y$''). +\item[$\bullet$] \qtybf{type}: A HOL type. +\end{enum} + +Default values are indicated in square brackets. Boolean options have a negated +counterpart (e.g., \textit{auto} vs.\ \textit{no\_auto}). When setting Boolean +options, ``= \textit{true}'' may be omitted. + +\subsection{Mode of Operation} +\label{mode-of-operation} + +\begin{enum} +\opfalse{auto}{no\_auto} +Specifies whether Nitpick should be run automatically on newly entered theorems. +For automatic runs, \textit{user\_axioms} (\S\ref{mode-of-operation}) and +\textit{assms} (\S\ref{mode-of-operation}) are implicitly enabled, +\textit{blocking} (\S\ref{mode-of-operation}), \textit{verbose} +(\S\ref{output-format}), and \textit{debug} (\S\ref{output-format}) are +disabled, \textit{max\_potential} (\S\ref{output-format}) is taken to be 0, and +\textit{auto\_timeout} (\S\ref{timeouts}) is used as the time limit instead of +\textit{timeout} (\S\ref{timeouts}). The output is also more concise. + +\nopagebreak +{\small See also \textit{auto\_timeout} (\S\ref{timeouts}).} + +\optrue{blocking}{non\_blocking} +Specifies whether the \textbf{nitpick} command should operate synchronously. +The asynchronous (non-blocking) mode lets the user start proving the putative +theorem while Nitpick looks for a counterexample, but it can also be more +confusing. For technical reasons, automatic runs currently always block. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}).} + +\optrue{falsify}{satisfy} +Specifies whether Nitpick should look for falsifying examples (countermodels) or +satisfying examples (models). This manual assumes throughout that +\textit{falsify} is enabled. + +\opsmart{user\_axioms}{no\_user\_axioms} +Specifies whether the user-defined axioms (specified using +\textbf{axiomatization} and \textbf{axioms}) should be considered. If the option +is set to \textit{smart}, Nitpick performs an ad hoc axiom selection based on +the constants that occur in the formula to falsify. The option is implicitly set +to \textit{true} for automatic runs. + +\textbf{Warning:} If the option is set to \textit{true}, Nitpick might +nonetheless ignore some polymorphic axioms. Counterexamples generated under +these conditions are tagged as ``likely genuine.'' The \textit{debug} +(\S\ref{output-format}) option can be used to find out which axioms were +considered. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}), \textit{assms} +(\S\ref{mode-of-operation}), and \textit{debug} (\S\ref{output-format}).} + +\optrue{assms}{no\_assms} +Specifies whether the relevant assumptions in structured proof should be +considered. The option is implicitly enabled for automatic runs. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}) +and \textit{user\_axioms} (\S\ref{mode-of-operation}).} + +\opfalse{overlord}{no\_overlord} +Specifies whether Nitpick should put its temporary files in +\texttt{\$ISABELLE\_\allowbreak HOME\_\allowbreak USER}, which is useful for +debugging Nitpick but also unsafe if several instances of the tool are run +simultaneously. + +\nopagebreak +{\small See also \textit{debug} (\S\ref{output-format}).} +\end{enum} + +\subsection{Scope of Search} +\label{scope-of-search} + +\begin{enum} +\opu{card}{type}{int\_seq} +Specifies the sequence of cardinalities to use for a given type. For +\textit{nat} and \textit{int}, the cardinality fully specifies the subset used +to approximate the type. For example: +% +$$\hbox{\begin{tabular}{@{}rll@{}}% +\textit{card nat} = 4 & induces & $\{0,\, 1,\, 2,\, 3\}$ \\ +\textit{card int} = 4 & induces & $\{-1,\, 0,\, +1,\, +2\}$ \\ +\textit{card int} = 5 & induces & $\{-2,\, -1,\, 0,\, +1,\, +2\}.$% +\end{tabular}}$$ +% +In general: +% +$$\hbox{\begin{tabular}{@{}rll@{}}% +\textit{card nat} = $K$ & induces & $\{0,\, \ldots,\, K - 1\}$ \\ +\textit{card int} = $K$ & induces & $\{-\lceil K/2 \rceil + 1,\, \ldots,\, +\lfloor K/2 \rfloor\}.$% +\end{tabular}}$$ +% +For free types, and often also for \textbf{typedecl}'d types, it usually makes +sense to specify cardinalities as a range of the form \textit{$1$--$n$}. +Although function and product types are normally mapped directly to the +corresponding Kodkod concepts, setting +the cardinality of such types is also allowed and implicitly enables ``boxing'' +for them, as explained in the description of the \textit{box}~\qty{type} +and \textit{box} (\S\ref{scope-of-search}) options. + +\nopagebreak +{\small See also \textit{mono} (\S\ref{scope-of-search}).} + +\opt{card}{int\_seq}{$\mathbf{1}$--$\mathbf{8}$} +Specifies the default sequence of cardinalities to use. This can be overridden +on a per-type basis using the \textit{card}~\qty{type} option described above. + +\opu{max}{const}{int\_seq} +Specifies the sequence of maximum multiplicities to use for a given +(co)in\-duc\-tive datatype constructor. A constructor's multiplicity is the +number of distinct values that it can construct. Nonsensical values (e.g., +\textit{max}~[]~$=$~2) are silently repaired. This option is only available for +datatypes equipped with several constructors. + +\ops{max}{int\_seq} +Specifies the default sequence of maximum multiplicities to use for +(co)in\-duc\-tive datatype constructors. This can be overridden on a per-constructor +basis using the \textit{max}~\qty{const} option described above. + +\opusmart{wf}{const}{non\_wf} +Specifies whether the specified (co)in\-duc\-tively defined predicate is +well-founded. The option can take the following values: + +\begin{enum} +\item[$\bullet$] \textbf{\textit{true}}: Tentatively treat the (co)in\-duc\-tive +predicate as if it were well-founded. Since this is generally not sound when the +predicate is not well-founded, the counterexamples are tagged as ``likely +genuine.'' + +\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 +predicate is well-founded using Isabelle's \textit{lexicographic\_order} and +\textit{sizechange} tactics. If this succeeds (or the predicate occurs with an +appropriate polarity in the formula to falsify), use an efficient fixed point +equation as specification of the predicate; otherwise, unroll the predicates +according to the \textit{iter}~\qty{const} and \textit{iter} options. +\end{enum} + +\nopagebreak +{\small See also \textit{iter} (\S\ref{scope-of-search}), +\textit{star\_linear\_preds} (\S\ref{optimizations}), and \textit{tac\_timeout} +(\S\ref{timeouts}).} + +\opsmart{wf}{non\_wf} +Specifies the default wellfoundedness setting to use. This can be overridden on +a per-predicate basis using the \textit{wf}~\qty{const} option above. + +\opu{iter}{const}{int\_seq} +Specifies the sequence of iteration counts to use when unrolling a given +(co)in\-duc\-tive predicate. By default, unrolling is applied for inductive +predicates that occur negatively and coinductive predicates that occur +positively in the formula to falsify and that cannot be proved to be +well-founded, but this behavior is influenced by the \textit{wf} option. The +iteration counts are automatically bounded by the cardinality of the predicate's +domain. + +{\small See also \textit{wf} (\S\ref{scope-of-search}) and +\textit{star\_linear\_preds} (\S\ref{optimizations}).} + +\opt{iter}{int\_seq}{$\mathbf{1{,}2{,}4{,}8{,}12{,}16{,}24{,}32}$} +Specifies the sequence of iteration counts to use when unrolling (co)in\-duc\-tive +predicates. This can be overridden on a per-predicate basis using the +\textit{iter} \qty{const} option above. + +\opt{bisim\_depth}{int\_seq}{$\mathbf{7}$} +Specifies the sequence of iteration counts to use when unrolling the +bisimilarity predicate generated by Nitpick for coinductive datatypes. A value +of $-1$ means that no predicate is generated, in which case Nitpick performs an +after-the-fact check to see if the known coinductive datatype values are +bidissimilar. If two values are found to be bisimilar, the counterexample is +tagged as ``likely genuine.'' The iteration counts are automatically bounded by +the sum of the cardinalities of the coinductive datatypes occurring in the +formula to falsify. + +\opusmart{box}{type}{dont\_box} +Specifies whether Nitpick should attempt to wrap (``box'') a given function or +product type in an isomorphic datatype internally. Boxing is an effective mean +to reduce the search space and speed up Nitpick, because the isomorphic datatype +is approximated by a subset of the possible function or pair values; +like other drastic optimizations, it can also prevent the discovery of +counterexamples. The option can take the following values: + +\begin{enum} +\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 +is likely to help. For example, $n$-tuples where $n > 2$ and arguments to +higher-order functions are good candidates for boxing. +\end{enum} + +Setting the \textit{card}~\qty{type} option for a function or product type +implicitly enables boxing for that type. + +\nopagebreak +{\small See also \textit{verbose} (\S\ref{output-format}) +and \textit{debug} (\S\ref{output-format}).} + +\opsmart{box}{dont\_box} +Specifies the default boxing setting to use. This can be overridden on a +per-type basis using the \textit{box}~\qty{type} option described above. + +\opusmart{mono}{type}{non\_mono} +Specifies whether the specified type should be considered monotonic when +enumerating scopes. If the option is set to \textit{smart}, Nitpick performs a +monotonicity check on the type. Setting this option to \textit{true} can reduce +the number of scopes tried, but it also diminishes the theoretical chance of +finding a counterexample, as demonstrated in \S\ref{scope-monotonicity}. + +\nopagebreak +{\small See also \textit{card} (\S\ref{scope-of-search}), +\textit{coalesce\_type\_vars} (\S\ref{scope-of-search}), and \textit{verbose} +(\S\ref{output-format}).} + +\opsmart{mono}{non\_box} +Specifies the default monotonicity setting to use. This can be overridden on a +per-type basis using the \textit{mono}~\qty{type} option described above. + +\opfalse{coalesce\_type\_vars}{dont\_coalesce\_type\_vars} +Specifies whether type variables with the same sort constraints should be +merged. Setting this option to \textit{true} can reduce the number of scopes +tried and the size of the generated Kodkod formulas, but it also diminishes the +theoretical chance of finding a counterexample. + +{\small See also \textit{mono} (\S\ref{scope-of-search}).} +\end{enum} + +\subsection{Output Format} +\label{output-format} + +\begin{enum} +\opfalse{verbose}{quiet} +Specifies whether the \textbf{nitpick} command should explain what it does. This +option is useful to determine which scopes are tried or which SAT solver is +used. This option is implicitly disabled for automatic runs. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}).} + +\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} and \textit{show\_all} behind the scenes. The \textit{debug} +option is implicitly disabled for automatic runs. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}), \textit{overlord} +(\S\ref{mode-of-operation}), and \textit{batch\_size} (\S\ref{optimizations}).} + +\optrue{show\_skolems}{hide\_skolem} +Specifies whether the values of Skolem constants should be displayed as part of +counterexamples. Skolem constants correspond to bound variables in the original +formula and usually help us to understand why the counterexample falsifies the +formula. + +\nopagebreak +{\small See also \textit{skolemize} (\S\ref{optimizations}).} + +\opfalse{show\_datatypes}{hide\_datatypes} +Specifies whether the subsets used to approximate (co)in\-duc\-tive datatypes should +be displayed as part of counterexamples. Such subsets are sometimes helpful when +investigating whether a potential counterexample is genuine or spurious, but +their potential for clutter is real. + +\opfalse{show\_consts}{hide\_consts} +Specifies whether the values of constants occurring in the formula (including +its axioms) should be displayed along with any counterexample. These values are +sometimes helpful when investigating why a counterexample is +genuine, but they can clutter the output. + +\opfalse{show\_all}{dont\_show\_all} +Enabling this option effectively enables \textit{show\_skolems}, +\textit{show\_datatypes}, and \textit{show\_consts}. + +\opt{max\_potential}{int}{$\mathbf{1}$} +Specifies the maximum number of potential counterexamples to display. Setting +this option to 0 speeds up the search for a genuine counterexample. This option +is implicitly set to 0 for automatic runs. If you set this option to a value +greater than 1, you will need an incremental SAT solver: For efficiency, it is +recommended to install the JNI version of MiniSat and set \textit{sat\_solver} = +\textit{MiniSatJNI}. Also be aware that many of the counterexamples may look +identical, unless the \textit{show\_all} (\S\ref{output-format}) option is +enabled. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}), +\textit{check\_potential} (\S\ref{authentication}), and +\textit{sat\_solver} (\S\ref{optimizations}).} + +\opt{max\_genuine}{int}{$\mathbf{1}$} +Specifies the maximum number of genuine counterexamples to display. If you set +this option to a value greater than 1, you will need an incremental SAT solver: +For efficiency, it is recommended to install the JNI version of MiniSat and set +\textit{sat\_solver} = \textit{MiniSatJNI}. Also be aware that many of the +counterexamples may look identical, unless the \textit{show\_all} +(\S\ref{output-format}) option is enabled. + +\nopagebreak +{\small See also \textit{check\_genuine} (\S\ref{authentication}) and +\textit{sat\_solver} (\S\ref{optimizations}).} + +\ops{eval}{term\_list} +Specifies the list of terms whose values should be displayed along with +counterexamples. This option suffers from an ``observer effect'': Nitpick might +find different counterexamples for different values of this option. + +\opu{format}{term}{int\_seq} +Specifies how to uncurry the value displayed for a variable or constant. +Uncurrying sometimes increases the readability of the output for high-arity +functions. For example, given the variable $y \mathbin{\Colon} {'a}\Rightarrow +{'b}\Rightarrow {'c}\Rightarrow {'d}\Rightarrow {'e}\Rightarrow {'f}\Rightarrow +{'g}$, setting \textit{format}~$y$ = 3 tells Nitpick to group the last three +arguments, as if the type had been ${'a}\Rightarrow {'b}\Rightarrow +{'c}\Rightarrow {'d}\times {'e}\times {'f}\Rightarrow {'g}$. In general, a list +of values $n_1,\ldots,n_k$ tells Nitpick to show the last $n_k$ arguments as an +$n_k$-tuple, the previous $n_{k-1}$ arguments as an $n_{k-1}$-tuple, and so on; +arguments that are not accounted for are left alone, as if the specification had +been $1,\ldots,1,n_1,\ldots,n_k$. + +\nopagebreak +{\small See also \textit{uncurry} (\S\ref{optimizations}).} + +\opt{format}{int\_seq}{$\mathbf{1}$} +Specifies the default format to use. Irrespective of the default format, the +extra arguments to a Skolem constant corresponding to the outer bound variables +are kept separated from the remaining arguments, the \textbf{for} arguments of +an inductive definitions are kept separated from the remaining arguments, and +the iteration counter of an unrolled inductive definition is shown alone. The +default format can be overridden on a per-variable or per-constant basis using +the \textit{format}~\qty{term} option described above. +\end{enum} + +%% MARK: Authentication +\subsection{Authentication} +\label{authentication} + +\begin{enum} +\opfalse{check\_potential}{trust\_potential} +Specifies whether potential counterexamples should be given to Isabelle's +\textit{auto} tactic to assess their validity. If a potential counterexample is +shown to be genuine, Nitpick displays a message to this effect and terminates. + +\nopagebreak +{\small See also \textit{max\_potential} (\S\ref{output-format}) and +\textit{auto\_timeout} (\S\ref{timeouts}).} + +\opfalse{check\_genuine}{trust\_genuine} +Specifies whether genuine and likely genuine counterexamples should be given to +Isabelle's \textit{auto} tactic to assess their validity. If a ``genuine'' +counterexample is shown to be spurious, the user is kindly asked to send a bug +report to the author at +\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@in.tum.de}. + +\nopagebreak +{\small See also \textit{max\_genuine} (\S\ref{output-format}) and +\textit{auto\_timeout} (\S\ref{timeouts}).} + +\ops{expect}{string} +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{likely\_genuine}}: Nitpick found a ``likely +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., +Kodkod ran out of memory). +\end{enum} + +Nitpick emits an error if the actual outcome differs from the expected outcome. +This option is useful for regression testing. +\end{enum} + +\subsection{Optimizations} +\label{optimizations} + +\def\cpp{C\nobreak\raisebox{.1ex}{+}\nobreak\raisebox{.1ex}{+}} + +\sloppy + +\begin{enum} +\opt{sat\_solver}{string}{smart} +Specifies which SAT solver to use. SAT solvers implemented in C or \cpp{} tend +to be faster than their Java counterparts, but they can be more difficult to +install. Also, if you set the \textit{max\_potential} (\S\ref{output-format}) or +\textit{max\_genuine} (\S\ref{output-format}) option to a value greater than 1, +you will need an incremental SAT solver, such as \textit{MiniSatJNI} +(recommended) or \textit{SAT4J}. + +The supported solvers are listed below: + +\begin{enum} + +\item[$\bullet$] \textbf{\textit{MiniSat}}: MiniSat is an efficient solver +written in \cpp{}. To use MiniSat, set the environment variable +\texttt{MINISAT\_HOME} to the directory that contains the \texttt{minisat} +executable. The \cpp{} sources and executables for MiniSat are available at +\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{MiniSatJNI}}: The JNI (Java Native Interface) +version of MiniSat is bundled in \texttt{nativesolver.\allowbreak tgz}, which +you will find on Kodkod's web site \cite{kodkod-2009}. Unlike the standard +version of MiniSat, the JNI version can be used incrementally. + +\item[$\bullet$] \textbf{\textit{PicoSAT}}: PicoSAT is an efficient solver +written in C. It is bundled with Kodkodi and requires no further installation or +configuration steps. Alternatively, you can install a standard version of +PicoSAT and set the environment variable \texttt{PICOSAT\_HOME} to the directory +that contains the \texttt{picosat} executable. The C sources for PicoSAT are +available at \url{http://fmv.jku.at/picosat/} and are also bundled with Kodkodi. +Nitpick has been tested with version 913. + +\item[$\bullet$] \textbf{\textit{zChaff}}: zChaff is an efficient solver written +in \cpp{}. To use zChaff, set the environment variable \texttt{ZCHAFF\_HOME} to +the directory that contains the \texttt{zchaff} executable. The \cpp{} sources +and executables for zChaff are available at +\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{zChaffJNI}}: The JNI version of zChaff is +bundled in \texttt{native\-solver.\allowbreak tgz}, which you will find on +Kodkod's web site \cite{kodkod-2009}. + +\item[$\bullet$] \textbf{\textit{RSat}}: RSat is an efficient solver written in +\cpp{}. To use RSat, set the environment variable \texttt{RSAT\_HOME} to the +directory that contains the \texttt{rsat} executable. The \cpp{} sources for +RSat are available at \url{http://reasoning.cs.ucla.edu/rsat/}. Nitpick has been +tested with version 2.01. + +\item[$\bullet$] \textbf{\textit{BerkMin}}: BerkMin561 is an efficient solver +written in C. To use BerkMin, set the environment variable +\texttt{BERKMIN\_HOME} to the directory that contains the \texttt{BerkMin561} +executable. The BerkMin executables are available at +\url{http://eigold.tripod.com/BerkMin.html}. + +\item[$\bullet$] \textbf{\textit{BerkMinAlloy}}: 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. + +\item[$\bullet$] \textbf{\textit{Jerusat}}: Jerusat 1.3 is an efficient solver +written in C. To use Jerusat, set the environment variable +\texttt{JERUSAT\_HOME} to the directory that contains the \texttt{Jerusat1.3} +executable. The C sources for Jerusat are available at +\url{http://www.cs.tau.ac.il/~ale1/Jerusat1.3.tgz}. + +\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{SAT4JLight}}: 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. 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 +\textit{smart}, Nitpick selects the first solver among MiniSat, PicoSAT, zChaff, +RSat, BerkMin, BerkMinAlloy, and Jerusat that is recognized by Isabelle. If none +is found, it falls back on SAT4J, which should always be available. If +\textit{verbose} is enabled, Nitpick displays which SAT solver was chosen. + +\end{enum} +\fussy + +\opt{batch\_size}{int\_or\_smart}{smart} +Specifies the maximum number of Kodkod problems that should be lumped together +when invoking Kodkodi. Each problem corresponds to one scope. Lumping problems +together ensures that Kodkodi is launched less often, but it makes the verbose +output less readable and is sometimes detrimental to performance. If +\textit{batch\_size} is set to \textit{smart}, the actual value used is 1 if +\textit{debug} (\S\ref{output-format}) is set and 64 otherwise. + +\optrue{destroy\_constrs}{dont\_destroy\_constrs} +Specifies whether formulas involving (co)in\-duc\-tive datatype constructors should +be rewritten to use (automatically generated) discriminators and destructors. +This optimization can drastically reduce the size of the Boolean formulas given +to the SAT solver. + +\nopagebreak +{\small See also \textit{debug} (\S\ref{output-format}).} + +\optrue{specialize}{dont\_specialize} +Specifies whether functions invoked with static arguments should be specialized. +This optimization can drastically reduce the search space, especially for +higher-order functions. + +\nopagebreak +{\small See also \textit{debug} (\S\ref{output-format}) and +\textit{show\_consts} (\S\ref{output-format}).} + +\optrue{skolemize}{dont\_skolemize} +Specifies whether the formula should be skolemized. For performance reasons, +(positive) $\forall$-quanti\-fiers that occur in the scope of a higher-order +(positive) $\exists$-quanti\-fier are left unchanged. + +\nopagebreak +{\small See also \textit{debug} (\S\ref{output-format}) and +\textit{show\_skolems} (\S\ref{output-format}).} + +\optrue{star\_linear\_preds}{dont\_star\_linear\_preds} +Specifies whether Nitpick should use Kodkod's transitive closure operator to +encode non-well-founded ``linear inductive predicates,'' i.e., inductive +predicates for which each the predicate occurs in at most one assumption of each +introduction rule. Using the reflexive transitive closure is in principle +equivalent to setting \textit{iter} to the cardinality of the predicate's +domain, but it is usually more efficient. + +{\small See also \textit{wf} (\S\ref{scope-of-search}), \textit{debug} +(\S\ref{output-format}), and \textit{iter} (\S\ref{scope-of-search}).} + +\optrue{uncurry}{dont\_uncurry} +Specifies whether Nitpick should uncurry functions. Uncurrying has on its own no +tangible effect on efficiency, but it creates opportunities for the boxing +optimization. + +\nopagebreak +{\small See also \textit{box} (\S\ref{scope-of-search}), \textit{debug} +(\S\ref{output-format}), and \textit{format} (\S\ref{output-format}).} + +\optrue{fast\_descrs}{full\_descrs} +Specifies whether Nitpick should optimize the definite and indefinite +description operators (THE and SOME). The optimized versions usually help +Nitpick generate more counterexamples or at least find them faster, but only the +unoptimized versions are complete when all types occurring in the formula are +finite. + +{\small See also \textit{debug} (\S\ref{output-format}).} + +\optrue{peephole\_optim}{no\_peephole\_optim} +Specifies whether Nitpick should simplify the generated Kodkod formulas using a +peephole optimizer. These optimizations can make a significant difference. +Unless you are tracking down a bug in Nitpick or distrust the peephole +optimizer, you should leave this option enabled. + +\opt{sym\_break}{int}{20} +Specifies an upper bound on the number of relations for which Kodkod generates +symmetry breaking predicates. According to the Kodkod documentation +\cite{kodkod-2009-options}, ``in general, the higher this value, the more +symmetries will be broken, and the faster the formula will be solved. But, +setting the value too high may have the opposite effect and slow down the +solving.'' + +\opt{sharing\_depth}{int}{3} +Specifies the depth to which Kodkod should check circuits for equivalence during +the translation to SAT. The default of 3 is the same as in Alloy. The minimum +allowed depth is 1. Increasing the sharing may result in a smaller SAT problem, +but can also slow down Kodkod. + +\opfalse{flatten\_props}{dont\_flatten\_props} +Specifies whether Kodkod should try to eliminate intermediate Boolean variables. +Although this might sound like a good idea, in practice it can drastically slow +down Kodkod. + +\opt{max\_threads}{int}{0} +Specifies the maximum number of threads to use in Kodkod. If this option is set +to 0, Kodkod will compute an appropriate value based on the number of processor +cores available. + +\nopagebreak +{\small See also \textit{batch\_size} (\S\ref{optimizations}) and +\textit{timeout} (\S\ref{timeouts}).} +\end{enum} + +\subsection{Timeouts} +\label{timeouts} + +\begin{enum} +\opt{timeout}{time}{$\mathbf{30}$ s} +Specifies the maximum amount of time that the \textbf{nitpick} command should +spend looking for a counterexample. Nitpick tries to honor this constraint as +well as it can but offers no guarantees. For automatic runs, +\textit{auto\_timeout} is used instead. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}) +and \textit{max\_threads} (\S\ref{optimizations}).} + +\opt{auto\_timeout}{time}{$\mathbf{5}$ s} +Specifies the maximum amount of time that Nitpick should use to find a +counterexample when running automatically. Nitpick tries to honor this +constraint as well as it can but offers no guarantees. + +\nopagebreak +{\small See also \textit{auto} (\S\ref{mode-of-operation}).} + +\opt{tac\_timeout}{time}{$\mathbf{500}$ ms} +Specifies the maximum amount of time that the \textit{auto} tactic should use +when checking a counterexample, and similarly that \textit{lexicographic\_order} +and \textit{sizechange} should use when checking whether a (co)in\-duc\-tive +predicate is well-founded. Nitpick tries to honor this constraint as well as it +can but offers no guarantees. + +\nopagebreak +{\small See also \textit{wf} (\S\ref{scope-of-search}), +\textit{check\_potential} (\S\ref{authentication}), +and \textit{check\_genuine} (\S\ref{authentication}).} +\end{enum} + +\section{Attribute Reference} +\label{attribute-reference} + +Nitpick needs to consider the definitions of all constants occurring in a +formula in order to falsify it. For constants introduced using the +\textbf{definition} command, the definition is simply the associated +\textit{\_def} axiom. In contrast, instead of using the internal representation +of functions synthesized by Isabelle's \textbf{primrec}, \textbf{function}, and +\textbf{nominal\_primrec} packages, Nitpick relies on the more natural +equational specification entered by the user. + +Behind the scenes, Isabelle's built-in packages and theories rely on the +following attributes to affect Nitpick's behavior: + +\begin{itemize} +\flushitem{\textit{nitpick\_def}} + +\nopagebreak +This attribute specifies an alternative definition of a constant. The +alternative definition should be logically equivalent to the constant's actual +axiomatic definition and should be of the form + +\qquad $c~{?}x_1~\ldots~{?}x_n \,\equiv\, t$, + +where ${?}x_1, \ldots, {?}x_n$ are distinct variables and $c$ does not occur in +$t$. + +\flushitem{\textit{nitpick\_simp}} + +\nopagebreak +This attribute specifies the equations that constitute the specification of a +constant. For functions defined using the \textbf{primrec}, \textbf{function}, +and \textbf{nominal\_\allowbreak primrec} packages, this corresponds to the +\textit{simps} rules. The equations must be of the form + +\qquad $c~t_1~\ldots\ t_n \,=\, u.$ + +\flushitem{\textit{nitpick\_psimp}} + +\nopagebreak +This attribute specifies the equations that constitute the partial specification +of a constant. For functions defined using the \textbf{function} package, this +corresponds to the \textit{psimps} rules. The conditional equations must be of +the form + +\qquad $\lbrakk P_1;\> \ldots;\> P_m\rbrakk \,\Longrightarrow\, c\ t_1\ \ldots\ t_n \,=\, u$. + +\flushitem{\textit{nitpick\_intro}} + +\nopagebreak +This attribute specifies the introduction rules of a (co)in\-duc\-tive predicate. +For predicates defined using the \textbf{inductive} or \textbf{coinductive} +command, this corresponds to the \textit{intros} rules. The introduction rules +must be of the form + +\qquad $\lbrakk P_1;\> \ldots;\> P_m;\> M~(c\ t_{11}\ \ldots\ t_{1n});\> +\ldots;\> M~(c\ t_{k1}\ \ldots\ t_{kn})\rbrakk \,\Longrightarrow\, c\ u_1\ +\ldots\ u_n$, + +where the $P_i$'s are side conditions that do not involve $c$ and $M$ is an +optional monotonic operator. The order of the assumptions is irrelevant. + +\end{itemize} + +When faced with a constant, Nitpick proceeds as follows: + +\begin{enum} +\item[1.] If the \textit{nitpick\_simp} set associated with the constant +is not empty, Nitpick uses these rules as the specification of the constant. + +\item[2.] Otherwise, if the \textit{nitpick\_psimp} set associated with +the constant is not empty, it uses these rules as the specification of the +constant. + +\item[3.] Otherwise, it looks up the definition of the constant: + +\begin{enum} +\item[1.] If the \textit{nitpick\_def} set associated with the constant +is not empty, it uses the latest rule added to the set as the definition of the +constant; otherwise it uses the actual definition axiom. +\item[2.] If the definition is of the form + +\qquad $c~{?}x_1~\ldots~{?}x_m \,\equiv\, \lambda y_1~\ldots~y_n.\; \textit{lfp}~(\lambda f.\; t)$, + +then Nitpick assumes that the definition was made using an inductive package and +based on the introduction rules marked with \textit{nitpick\_\allowbreak +ind\_\allowbreak intros} tries to determine whether the definition is +well-founded. +\end{enum} +\end{enum} + +As an illustration, consider the inductive definition + +\prew +\textbf{inductive}~\textit{odd}~\textbf{where} \\ +``\textit{odd}~1'' $\,\mid$ \\ +``\textit{odd}~$n\,\Longrightarrow\, \textit{odd}~(\textit{Suc}~(\textit{Suc}~n))$'' +\postw + +Isabelle automatically attaches the \textit{nitpick\_intro} attribute to +the above rules. Nitpick then uses the \textit{lfp}-based definition in +conjunction with these rules. To override this, we can specify an alternative +definition as follows: + +\prew +\textbf{lemma} $\mathit{odd\_def}'$ [\textit{nitpick\_def}]: ``$\textit{odd}~n \,\equiv\, n~\textrm{mod}~2 = 1$'' +\postw + +Nitpick then expands all occurrences of $\mathit{odd}~n$ to $n~\textrm{mod}~2 += 1$. Alternatively, we can specify an equational specification of the constant: + +\prew +\textbf{lemma} $\mathit{odd\_simp}'$ [\textit{nitpick\_simp}]: ``$\textit{odd}~n = (n~\textrm{mod}~2 = 1)$'' +\postw + +Such tweaks should be done with great care, because Nitpick will assume that the +constant is completely defined by its equational specification. For example, if +you make ``$\textit{odd}~(2 * k + 1)$'' a \textit{nitpick\_simp} rule and neglect to provide rules to handle the $2 * k$ case, Nitpick will define +$\textit{odd}~n$ arbitrarily for even values of $n$. The \textit{debug} +(\S\ref{output-format}) option is extremely useful to understand what is going +on when experimenting with \textit{nitpick\_} attributes. + +\section{Standard ML Interface} +\label{standard-ml-interface} + +Nitpick provides a rich Standard ML interface used mainly for internal purposes +and debugging. Among the most interesting functions exported by Nitpick are +those that let you invoke the tool programmatically and those that let you +register and unregister custom coinductive datatypes. + +\subsection{Invocation of Nitpick} +\label{invocation-of-nitpick} + +The \textit{Nitpick} structure offers the following functions for invoking your +favorite counterexample generator: + +\prew +$\textbf{val}\,~\textit{pick\_nits\_in\_term} : \\ +\hbox{}\quad\textit{Proof.state} \rightarrow \textit{params} \rightarrow \textit{bool} \rightarrow \textit{term~list} \rightarrow \textit{term} \\ +\hbox{}\quad{\rightarrow}\; \textit{string} * \textit{Proof.state}$ \\ +$\textbf{val}\,~\textit{pick\_nits\_in\_subgoal} : \\ +\hbox{}\quad\textit{Proof.state} \rightarrow \textit{params} \rightarrow \textit{bool} \rightarrow \textit{int} \rightarrow \textit{string} * \textit{Proof.state}$ +\postw + +The return value is a new proof state paired with an outcome string +(``genuine'', ``likely\_genuine'', ``potential'', ``none'', or ``unknown''). The +\textit{params} type is a large record that lets you set Nitpick's options. The +current default options can be retrieved by calling the following function +defined in the \textit{NitpickIsar} structure: + +\prew +$\textbf{val}\,~\textit{default\_params} :\, +\textit{theory} \rightarrow (\textit{string} * \textit{string})~\textit{list} \rightarrow \textit{params}$ +\postw + +The second argument lets you override option values before they are parsed and +put into a \textit{params} record. Here is an example: + +\prew +$\textbf{val}\,~\textit{params} = \textit{NitpickIsar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\ +$\textbf{val}\,~(\textit{outcome},\, \textit{state}') = \textit{Nitpick.pick\_nits\_in\_subgoal}~\begin{aligned}[t] +& \textit{state}~\textit{params}~\textit{false} \\[-2pt] +& \textit{subgoal}\end{aligned}$ +\postw + +\subsection{Registration of Coinductive Datatypes} +\label{registration-of-coinductive-datatypes} + +\let\antiq=\textrm + +If you have defined a custom coinductive datatype, you can tell Nitpick about +it, so that it can use an efficient Kodkod axiomatization similar to the one it +uses for lazy lists. The interface for registering and unregistering coinductive +datatypes consists of the following pair of functions defined in the +\textit{Nitpick} structure: + +\prew +$\textbf{val}\,~\textit{register\_codatatype} :\, +\textit{typ} \rightarrow \textit{string} \rightarrow \textit{styp~list} \rightarrow \textit{theory} \rightarrow \textit{theory}$ \\ +$\textbf{val}\,~\textit{unregister\_codatatype} :\, +\textit{typ} \rightarrow \textit{theory} \rightarrow \textit{theory}$ +\postw + +The type $'a~\textit{llist}$ of lazy lists is already registered; had it +not been, you could have told Nitpick about it by adding the following line +to your theory file: + +\prew +$\textbf{setup}~\,\{{*}\,~\!\begin{aligned}[t] +& \textit{Nitpick.register\_codatatype} \\[-2pt] +& \qquad @\{\antiq{typ}~``\kern1pt'a~\textit{llist}\textrm{''}\}~@\{\antiq{const\_name}~ \textit{llist\_case}\} \\[-2pt] %% TYPESETTING +& \qquad (\textit{map}~\textit{dest\_Const}~[@\{\antiq{term}~\textit{LNil}\},\, @\{\antiq{term}~\textit{LCons}\}])\,\ {*}\}\end{aligned}$ +\postw + +The \textit{register\_codatatype} function takes a coinductive type, its case +function, and the list of its constructors. The case function must take its +arguments in the order that the constructors are listed. If no case function +with the correct signature is available, simply pass the empty string. + +On the other hand, if your goal is to cripple Nitpick, add the following line to +your theory file and try to check a few conjectures about lazy lists: + +\prew +$\textbf{setup}~\,\{{*}\,~\textit{Nitpick.unregister\_codatatype}~@\{\antiq{typ}~`` +\kern1pt'a~\textit{list}\textrm{''}\}\ \,{*}\}$ +\postw + +\section{Known Bugs and Limitations} +\label{known-bugs-and-limitations} + +Here are the known bugs and limitations in Nitpick at the time of writing: + +\begin{enum} +\item[$\bullet$] Underspecified functions defined using the \textbf{primrec}, +\textbf{function}, or \textbf{nominal\_\allowbreak primrec} packages can lead +Nitpick to generate spurious counterexamples for theorems that refer to values +for which the function is not defined. For example: + +\prew +\textbf{primrec} \textit{prec} \textbf{where} \\ +``$\textit{prec}~(\textit{Suc}~n) = n$'' \\[2\smallskipamount] +\textbf{lemma} ``$\textit{prec}~0 = \undef$'' \\ +\textbf{nitpick} \\[2\smallskipamount] +\quad{\slshape Nitpick found a counterexample for \textit{card nat}~= 2: +\nopagebreak +\\[2\smallskipamount] +\hbox{}\qquad Empty assignment} \nopagebreak\\[2\smallskipamount] +\textbf{by}~(\textit{auto simp}: \textit{prec\_def}) +\postw + +Such theorems are considered bad style because they rely on the internal +representation of functions synthesized by Isabelle, which is an implementation +detail. + +\item[$\bullet$] Nitpick produces spurious counterexamples when invoked after a +\textbf{guess} command in a structured proof. + +\item[$\bullet$] The \textit{nitpick\_} attributes and the +\textit{Nitpick.register\_} functions can cause havoc if used improperly. + +\item[$\bullet$] Local definitions are not supported and result in an error. + +\item[$\bullet$] All constants and types whose names start with +\textit{Nitpick}{.} are reserved for internal use. +\end{enum} + +\let\em=\sl +\bibliography{../manual}{} +\bibliographystyle{abbrv} + +\end{document} diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/gfx/isabelle_nitpick.eps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/gfx/isabelle_nitpick.eps Mon Oct 26 14:57:49 2009 +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 (Nitpick) 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 diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/gfx/isabelle_nitpick.pdf Binary file doc-src/gfx/isabelle_nitpick.pdf has changed diff -r 4705b7323a7d -r 79bd3fbf5d61 doc-src/manual.bib --- a/doc-src/manual.bib Mon Oct 26 14:54:43 2009 +0100 +++ b/doc-src/manual.bib Mon Oct 26 14:57:49 2009 +0100 @@ -49,7 +49,7 @@ @Unpublished{abrial93, author = {J. R. Abrial and G. Laffitte}, - title = {Towards the Mechanization of the Proofs of some Classical + title = {Towards the Mechanization of the Proofs of Some Classical Theorems of Set Theory}, note = {preprint}, year = 1993, @@ -73,6 +73,17 @@ crossref = {types93}, pages = {213-237}} +@inproceedings{andersson-1993, + author = "Arne Andersson", + title = "Balanced Search Trees Made Simple", + editor = "F. K. H. A. Dehne and N. Santoro and S. Whitesides", + booktitle = "WADS 1993", + series = LNCS, + volume = {709}, + pages = "61--70", + year = 1993, + publisher = Springer} + @book{andrews86, author = "Peter Andrews", title = "An Introduction to Mathematical Logic and Type Theory: to Truth @@ -167,6 +178,15 @@ author = "Stefan Berghofer and Tobias Nipkow", pages = "38--52"} +@inproceedings{berghofer-nipkow-2004, + author = {Stefan Berghofer and Tobias Nipkow}, + title = {Random Testing in {I}sabelle/{HOL}}, + pages = {230--239}, + editor = "J. Cuellar and Z. Liu", + booktitle = {{SEFM} 2004}, + publisher = IEEE, + year = 2004} + @InProceedings{Berghofer-Nipkow:2002, author = {Stefan Berghofer and Tobias Nipkow}, title = {Executing Higher Order Logic}, @@ -200,6 +220,14 @@ 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)", + author = "Jasmin Christian Blanchette and Tobias Nipkow", + booktitle = "{TAP} 2009: Short Papers", + editor = "Catherine Dubois", + publisher = "ETH Technical Report 630", + year = 2009} + @Article{boyer86, author = {Robert Boyer and Ewing Lusk and William McCune and Ross Overbeek and Mark Stickel and Lawrence Wos}, @@ -241,7 +269,7 @@ } @InProceedings{bulwahn-et-al:2008:imperative, - author = {Lukas Bulwahn and Alexander Krauss and Florian Haftmann and Levent Erkök and John Matthews}, + author = {Lukas Bulwahn and Alexander Krauss and Florian Haftmann and Levent Erkök and John Matthews}, title = {Imperative Functional Programming with {Isabelle/HOL}}, crossref = {tphols2008}, } @@ -597,6 +625,12 @@ year = 2003, note = {\url{http://www.haskell.org/definition/}}} +@book{jackson-2006, + author = "Daniel Jackson", + title = "Software Abstractions: Logic, Language, and Analysis", + publisher = MIT, + year = 2006} + %K @InProceedings{kammueller-locales, @@ -878,10 +912,11 @@ @Book{isa-tutorial, author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, - title = {Isabelle/HOL: A Proof Assistant for Higher-Order Logic}, - publisher = {Springer}, + title = {Isabelle/{HOL}: A Proof Assistant for Higher-Order Logic}, + publisher = Springer, year = 2002, - note = {LNCS Tutorial 2283}} + series = LNCS, + volume = 2283} @Article{noel, author = {Philippe No{\"e}l}, @@ -1021,7 +1056,7 @@ Essays in Honor of {Robin Milner}}, booktitle = {Proof, Language, and Interaction: Essays in Honor of {Robin Milner}}, - publisher = {MIT Press}, + publisher = MIT, year = 2000, editor = {Gordon Plotkin and Colin Stirling and Mads Tofte}} @@ -1236,6 +1271,12 @@ number = 4 } +@misc{sledgehammer-2009, + key = "Sledgehammer", + title = "The {S}ledgehammer: Let Automatic Theorem Provers +Write Your {I}s\-a\-belle Scripts", + note = "\url{http://www.cl.cam.ac.uk/research/hvg/Isabelle/sledgehammer.html}"} + @inproceedings{slind-tfl, author = {Konrad Slind}, title = {Function Definition in Higher Order Logic}, @@ -1295,6 +1336,27 @@ title={Haskell: The Craft of Functional Programming}, publisher={Addison-Wesley},year=1999} +@misc{kodkod-2009, + author = "Emina Torlak", + title = {Kodkod: Constraint Solver for Relational Logic}, + note = "\url{http://alloy.mit.edu/kodkod/}"} + +@misc{kodkod-2009-options, + author = "Emina Torlak", + title = "Kodkod {API}: Class {Options}", + note = "\url{http://alloy.mit.edu/kodkod/docs/kodkod/engine/config/Options.html}"} + +@inproceedings{torlak-jackson-2007, + title = "Kodkod: A Relational Model Finder", + author = "Emina Torlak and Daniel Jackson", + editor = "Orna Grumberg and Michael Huth", + booktitle = "TACAS 2007", + series = LNCS, + volume = {4424}, + pages = "632--647", + year = 2007, + publisher = Springer} + @Unpublished{Trybulec:1993:MizarFeatures, author = {A. Trybulec}, title = {Some Features of the {Mizar} Language}, @@ -1320,6 +1382,13 @@ year = 1989 } +@phdthesis{weber-2008, + author = "Tjark Weber", + title = "SAT-Based Finite Model Generation for Higher-Order Logic", + school = {Dept.\ of Informatics, T.U. M\"unchen}, + type = "{Ph.D.}\ thesis", + year = 2008} + @Misc{x-symbol, author = {Christoph Wedler}, title = {Emacs package ``{X-Symbol}''}, @@ -1570,7 +1639,7 @@ Essays in Honor of {Larry Wos}}, booktitle = {Automated Reasoning and its Applications: Essays in Honor of {Larry Wos}}, - publisher = {MIT Press}, + publisher = MIT, year = 1997, editor = {Robert Veroff}} @@ -1669,3 +1738,8 @@ title = {{ML} Modules and {Haskell} Type Classes: A Constructive Comparison}, author = {Stefan Wehr et. al.} } + +@misc{wikipedia-2009-aa-trees, + key = "Wikipedia", + title = "Wikipedia: {AA} Tree", + note = "\url{http://en.wikipedia.org/wiki/AA_tree}"} diff -r 4705b7323a7d -r 79bd3fbf5d61 etc/isar-keywords-ZF.el --- a/etc/isar-keywords-ZF.el Mon Oct 26 14:54:43 2009 +0100 +++ b/etc/isar-keywords-ZF.el Mon Oct 26 14:57:49 2009 +0100 @@ -9,12 +9,9 @@ "\\.\\." "Isabelle\\.command" "Isar\\.begin_document" - "Isar\\.command" "Isar\\.define_command" "Isar\\.edit_document" "Isar\\.end_document" - "Isar\\.insert" - "Isar\\.remove" "ML" "ML_command" "ML_prf" @@ -252,12 +249,9 @@ (defconst isar-keywords-control '("Isabelle\\.command" "Isar\\.begin_document" - "Isar\\.command" "Isar\\.define_command" "Isar\\.edit_document" "Isar\\.end_document" - "Isar\\.insert" - "Isar\\.remove" "ProofGeneral\\.inform_file_processed" "ProofGeneral\\.inform_file_retracted" "ProofGeneral\\.kill_proof" diff -r 4705b7323a7d -r 79bd3fbf5d61 etc/isar-keywords.el --- a/etc/isar-keywords.el Mon Oct 26 14:54:43 2009 +0100 +++ b/etc/isar-keywords.el Mon Oct 26 14:57:49 2009 +0100 @@ -9,12 +9,9 @@ "\\.\\." "Isabelle\\.command" "Isar\\.begin_document" - "Isar\\.command" "Isar\\.define_command" "Isar\\.edit_document" "Isar\\.end_document" - "Isar\\.insert" - "Isar\\.remove" "ML" "ML_command" "ML_prf" @@ -135,6 +132,8 @@ "method_setup" "moreover" "next" + "nitpick" + "nitpick_params" "no_notation" "no_syntax" "no_translations" @@ -317,12 +316,9 @@ (defconst isar-keywords-control '("Isabelle\\.command" "Isar\\.begin_document" - "Isar\\.command" "Isar\\.define_command" "Isar\\.edit_document" "Isar\\.end_document" - "Isar\\.insert" - "Isar\\.remove" "ProofGeneral\\.inform_file_processed" "ProofGeneral\\.inform_file_retracted" "ProofGeneral\\.kill_proof" @@ -360,6 +356,7 @@ "header" "help" "kill_thy" + "nitpick" "normal_form" "pr" "pretty_setmargin" @@ -482,6 +479,7 @@ "local_setup" "locale" "method_setup" + "nitpick_params" "no_notation" "no_syntax" "no_translations" diff -r 4705b7323a7d -r 79bd3fbf5d61 lib/jedit/isabelle.xml --- a/lib/jedit/isabelle.xml Mon Oct 26 14:54:43 2009 +0100 +++ b/lib/jedit/isabelle.xml Mon Oct 26 14:57:49 2009 +0100 @@ -36,12 +36,9 @@ .. Isabelle.command Isar.begin_document - Isar.command Isar.define_command Isar.edit_document Isar.end_document - Isar.insert - Isar.remove ML ML_prf @@ -196,6 +193,8 @@ moreover morphisms next + + nitpick_params no_notation no_syntax no_translations @@ -246,6 +245,7 @@ + diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/GCD.thy --- a/src/HOL/GCD.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/GCD.thy Mon Oct 26 14:57:49 2009 +0100 @@ -1702,4 +1702,12 @@ show ?thesis by(simp add: Gcd_def fold_set gcd_commute_int) qed +lemma gcd_eq_nitpick_gcd [nitpick_def]: "gcd x y \ Nitpick.nat_gcd x y" +apply (rule eq_reflection) +apply (induct x y rule: nat_gcd.induct) +by (simp add: gcd_nat.simps Nitpick.nat_gcd.simps) + +lemma lcm_eq_nitpick_lcm [nitpick_def]: "lcm x y \ Nitpick.nat_lcm x y" +by (simp only: lcm_nat_def Nitpick.nat_lcm_def gcd_eq_nitpick_gcd) + end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Induct/LList.thy --- a/src/HOL/Induct/LList.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/Induct/LList.thy Mon Oct 26 14:57:49 2009 +0100 @@ -905,4 +905,9 @@ lemma lappend_assoc': "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)" by (rule_tac l = l1 in llist_fun_equalityI, auto) +setup {* +Nitpick.register_codatatype @{typ "'a llist"} @{const_name llist_case} + (map dest_Const [@{term LNil}, @{term LCons}]) +*} + end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/IsaMakefile Mon Oct 26 14:57:49 2009 +0100 @@ -35,6 +35,7 @@ HOL-Modelcheck \ HOL-Multivariate_Analysis \ HOL-NanoJava \ + HOL-Nitpick_Examples \ HOL-Nominal-Examples \ HOL-Number_Theory \ HOL-Old_Number_Theory \ @@ -44,9 +45,9 @@ HOL-SMT-Examples \ HOL-Statespace \ HOL-Subst \ - TLA-Buffer \ - TLA-Inc \ - TLA-Memory \ + TLA-Buffer \ + TLA-Inc \ + TLA-Memory \ HOL-UNITY \ HOL-Unix \ HOL-W0 \ @@ -132,6 +133,7 @@ Inductive.thy \ Lattices.thy \ Nat.thy \ + Nitpick.thy \ Option.thy \ OrderedGroup.thy \ Orderings.thy \ @@ -178,6 +180,21 @@ 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_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 \ @@ -566,6 +583,21 @@ @$(ISABELLE_TOOL) usedir $(OUT)/HOL Metis_Examples +## HOL-Nitpick_Examples + +HOL-Nitpick_Examples: HOL $(LOG)/HOL-Nitpick_Examples.gz + +$(LOG)/HOL-Nitpick_Examples.gz: $(OUT)/HOL Nitpick_Examples/ROOT.ML \ + Nitpick_Examples/Core_Nits.thy Nitpick_Examples/Datatype_Nits.thy \ + Nitpick_Examples/Induct_Nits.thy Nitpick_Examples/Manual_Nits.thy \ + Nitpick_Examples/Mini_Nits.thy Nitpick_Examples/Mono_Nits.thy \ + Nitpick_Examples/Nitpick_Examples.thy Nitpick_Examples/Pattern_Nits.thy \ + Nitpick_Examples/Record_Nits.thy Nitpick_Examples/Refute_Nits.thy \ + Nitpick_Examples/Special_Nits.thy Nitpick_Examples/Tests_Nits.thy \ + Nitpick_Examples/Typedef_Nits.thy + @$(ISABELLE_TOOL) usedir $(OUT)/HOL Nitpick_Examples + + ## HOL-Algebra HOL-Algebra: HOL $(LOG)/HOL-Algebra.gz diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Library/Coinductive_List.thy --- a/src/HOL/Library/Coinductive_List.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/Library/Coinductive_List.thy Mon Oct 26 14:57:49 2009 +0100 @@ -200,6 +200,7 @@ [code del]: "llist_case c d l = List_case c (\x y. d (inv Datatype.Leaf x) (Abs_llist y)) (Rep_llist l)" + syntax (* FIXME? *) LNil :: logic LCons :: logic @@ -848,4 +849,9 @@ qed qed +setup {* +Nitpick.register_codatatype @{typ "'a llist"} @{const_name llist_case} + (map dest_Const [@{term LNil}, @{term LCons}]) +*} + end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Main.thy --- a/src/HOL/Main.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/Main.thy Mon Oct 26 14:57:49 2009 +0100 @@ -1,7 +1,7 @@ header {* Main HOL *} theory Main -imports Plain Quickcheck Map Recdef SAT +imports Plain Nitpick Quickcheck Recdef begin text {* diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,240 @@ +(* Title: HOL/Nitpick.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Nitpick: Yet another counterexample generator for Isabelle/HOL. +*) + +header {* Nitpick: Yet Another Counterexample Generator for Isabelle/HOL *} + +theory Nitpick +imports Map SAT +uses ("Tools/Nitpick/kodkod.ML") + ("Tools/Nitpick/kodkod_sat.ML") + ("Tools/Nitpick/nitpick_util.ML") + ("Tools/Nitpick/nitpick_hol.ML") + ("Tools/Nitpick/nitpick_mono.ML") + ("Tools/Nitpick/nitpick_scope.ML") + ("Tools/Nitpick/nitpick_peephole.ML") + ("Tools/Nitpick/nitpick_rep.ML") + ("Tools/Nitpick/nitpick_nut.ML") + ("Tools/Nitpick/nitpick_kodkod.ML") + ("Tools/Nitpick/nitpick_model.ML") + ("Tools/Nitpick/nitpick.ML") + ("Tools/Nitpick/nitpick_isar.ML") + ("Tools/Nitpick/nitpick_tests.ML") + ("Tools/Nitpick/minipick.ML") +begin + +typedecl bisim_iterator + +(* FIXME: use axiomatization (here and elsewhere) *) +axiomatization unknown :: 'a + and undefined_fast_The :: 'a + and undefined_fast_Eps :: 'a + and bisim :: "bisim_iterator \ 'a \ 'a \ bool" + and bisim_iterator_max :: bisim_iterator + and Tha :: "('a \ bool) \ 'a" + +datatype ('a, 'b) pair_box = PairBox 'a 'b +datatype ('a, 'b) fun_box = FunBox "'a \ 'b" + +text {* +Alternative definitions. +*} + +lemma If_def [nitpick_def]: +"(if P then Q else R) \ (P \ Q) \ (\ P \ R)" +by (rule eq_reflection) (rule if_bool_eq_conj) + +lemma Ex1_def [nitpick_def]: +"Ex1 P \ \x. P = {x}" +apply (rule eq_reflection) +apply (simp add: Ex1_def expand_set_eq) +apply (rule iffI) + apply (erule exE) + apply (erule conjE) + apply (rule_tac x = x in exI) + apply (rule allI) + apply (rename_tac y) + apply (erule_tac x = y in allE) +by (auto simp: mem_def) + +lemma rtrancl_def [nitpick_def]: "r\<^sup>* \ (r\<^sup>+)\<^sup>=" +by simp + +lemma rtranclp_def [nitpick_def]: +"rtranclp r a b \ (a = b \ tranclp r a b)" +by (rule eq_reflection) (auto dest: rtranclpD) + +lemma tranclp_def [nitpick_def]: +"tranclp r a b \ trancl (split r) (a, b)" +by (simp add: trancl_def Collect_def mem_def) + +definition refl' :: "('a \ 'a \ bool) \ bool" where +"refl' r \ \x. (x, x) \ r" + +definition wf' :: "('a \ 'a \ bool) \ bool" where +"wf' r \ acyclic r \ (finite r \ unknown)" + +axiomatization wf_wfrec :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" + +definition wf_wfrec' :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where +[nitpick_simp]: "wf_wfrec' R F x = F (Recdef.cut (wf_wfrec R F) R x) x" + +definition wfrec' :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where +"wfrec' R F x \ if wf R then wf_wfrec' R F x + else THE y. wfrec_rel R (%f x. F (Recdef.cut f R x) x) x y" + +definition card' :: "('a \ bool) \ nat" where +"card' X \ length (SOME xs. set xs = X \ distinct xs)" + +definition setsum' :: "('a \ 'b\comm_monoid_add) \ ('a \ bool) \ 'b" where +"setsum' f A \ if finite A then listsum (map f (SOME xs. set xs = A \ distinct xs)) else 0" + +inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ ('a \ bool) \ 'b \ bool" where +"fold_graph' f z {} z" | +"\x \ A; fold_graph' f z (A - {x}) y\ \ fold_graph' f z A (f x y)" + +text {* +The following lemmas are not strictly necessary but they help the +\textit{special\_level} optimization. +*} + +lemma The_psimp [nitpick_psimp]: +"P = {x} \ The P = x" +by (subgoal_tac "{x} = (\y. y = x)") (auto simp: mem_def) + +lemma Eps_psimp [nitpick_psimp]: +"\P x; \ P y; Eps P = y\ \ Eps P = x" +apply (case_tac "P (Eps P)") + apply auto +apply (erule contrapos_np) +by (rule someI) + +lemma unit_case_def [nitpick_def]: +"unit_case x u \ x" +apply (subgoal_tac "u = ()") + apply (simp only: unit.cases) +by simp + +lemma nat_case_def [nitpick_def]: +"nat_case x f n \ if n = 0 then x else f (n - 1)" +apply (rule eq_reflection) +by (case_tac n) auto + +lemmas dvd_def = dvd_eq_mod_eq_0 [THEN eq_reflection, nitpick_def] + +lemma list_size_simp [nitpick_simp]: +"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)))" +by (case_tac xs) auto + +text {* +Auxiliary definitions used to provide an alternative representation for +@{text rat} and @{text real}. +*} + +function nat_gcd :: "nat \ nat \ nat" where +[simp del]: "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))" +by auto +termination +apply (relation "measure (\(x, y). x + y + (if y > x then 1 else 0))") + apply auto + apply (metis mod_less_divisor xt1(9)) +by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10)) + +definition nat_lcm :: "nat \ nat \ nat" where +"nat_lcm x y = x * y div (nat_gcd x y)" + +definition int_gcd :: "int \ int \ int" where +"int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))" + +definition int_lcm :: "int \ int \ int" where +"int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))" + +definition Frac :: "int \ int \ bool" where +"Frac \ \(a, b). b > 0 \ int_gcd a b = 1" + +axiomatization Abs_Frac :: "int \ int \ 'a" + and Rep_Frac :: "'a \ int \ int" + +definition zero_frac :: 'a where +"zero_frac \ Abs_Frac (0, 1)" + +definition one_frac :: 'a where +"one_frac \ Abs_Frac (1, 1)" + +definition num :: "'a \ int" where +"num \ fst o Rep_Frac" + +definition denom :: "'a \ int" where +"denom \ snd o Rep_Frac" + +function norm_frac :: "int \ int \ int \ int" where +[simp del]: "norm_frac a b = (if b < 0 then norm_frac (- a) (- b) + else if a = 0 \ b = 0 then (0, 1) + else let c = int_gcd a b in (a div c, b div c))" +by pat_completeness auto +termination by (relation "measure (\(_, b). if b < 0 then 1 else 0)") auto + +definition frac :: "int \ int \ 'a" where +"frac a b \ Abs_Frac (norm_frac a b)" + +definition plus_frac :: "'a \ 'a \ 'a" where +[nitpick_simp]: +"plus_frac q r = (let d = int_lcm (denom q) (denom r) in + frac (num q * (d div denom q) + num r * (d div denom r)) d)" + +definition times_frac :: "'a \ 'a \ 'a" where +[nitpick_simp]: +"times_frac q r = frac (num q * num r) (denom q * denom r)" + +definition uminus_frac :: "'a \ 'a" where +"uminus_frac q \ Abs_Frac (- num q, denom q)" + +definition number_of_frac :: "int \ 'a" where +"number_of_frac n \ Abs_Frac (n, 1)" + +definition inverse_frac :: "'a \ 'a" where +"inverse_frac q \ frac (denom q) (num q)" + +definition less_eq_frac :: "'a \ 'a \ bool" where +[nitpick_simp]: +"less_eq_frac q r \ num (plus_frac q (uminus_frac r)) \ 0" + +definition of_frac :: "'a \ 'b\{inverse,ring_1}" where +"of_frac q \ of_int (num q) / of_int (denom q)" + +use "Tools/Nitpick/kodkod.ML" +use "Tools/Nitpick/kodkod_sat.ML" +use "Tools/Nitpick/nitpick_util.ML" +use "Tools/Nitpick/nitpick_hol.ML" +use "Tools/Nitpick/nitpick_mono.ML" +use "Tools/Nitpick/nitpick_scope.ML" +use "Tools/Nitpick/nitpick_peephole.ML" +use "Tools/Nitpick/nitpick_rep.ML" +use "Tools/Nitpick/nitpick_nut.ML" +use "Tools/Nitpick/nitpick_kodkod.ML" +use "Tools/Nitpick/nitpick_model.ML" +use "Tools/Nitpick/nitpick.ML" +use "Tools/Nitpick/nitpick_isar.ML" +use "Tools/Nitpick/nitpick_tests.ML" +use "Tools/Nitpick/minipick.ML" + +hide (open) const unknown undefined_fast_The undefined_fast_Eps bisim + bisim_iterator_max Tha refl' wf' wf_wfrec wf_wfrec' wfrec' card' setsum' + fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac + one_frac num denom norm_frac frac plus_frac times_frac uminus_frac + number_of_frac inverse_frac less_eq_frac of_frac +hide (open) type bisim_iterator pair_box fun_box +hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def + wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def + The_psimp Eps_psimp unit_case_def nat_case_def dvd_def list_size_simp + nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def + one_frac_def num_def denom_def norm_frac_def frac_def plus_frac_def + times_frac_def uminus_frac_def number_of_frac_def inverse_frac_def + less_eq_frac_def of_frac_def + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Core_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Core_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,1123 @@ +(* Title: HOL/Nitpick_Examples/Core_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick's functional core. +*) + +header {* Examples Featuring Nitpick's Functional Core *} + +theory Core_Nits +imports Main +begin + +subsection {* Curry in a Hurry *} + +lemma "(\f x y. (curry o split) f x y) = (\f x y. (\x. x) f x y)" +nitpick [card = 1\4, expect = none] +nitpick [card = 100, expect = none, timeout = none] +by auto + +lemma "(\f p. (split o curry) f p) = (\f p. (\x. x) f p)" +nitpick [card = 2] +nitpick [card = 1\4, expect = none] +nitpick [card = 10, expect = none] +by auto + +lemma "split (curry f) = f" +nitpick [card = 1\4, expect = none] +nitpick [card = 10, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "curry (split f) = f" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "(split o curry) f = f" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "(curry o split) f = f" +nitpick [card = 1\4, expect = none] +nitpick [card = 1000, expect = none] +by auto + +lemma "(split o curry) f = (\x. x) f" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "(curry o split) f = (\x. x) f" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "((split o curry) f) p = ((\x. x) f) p" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +lemma "((curry o split) f) x = ((\x. x) f) x" +nitpick [card = 1\4, expect = none] +nitpick [card = 1000, expect = none] +by auto + +lemma "((curry o split) f) x y = ((\x. x) f) x y" +nitpick [card = 1\4, expect = none] +nitpick [card = 1000, expect = none] +by auto + +lemma "split o curry = (\x. x)" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +apply (rule ext)+ +by auto + +lemma "curry o split = (\x. x)" +nitpick [card = 1\4, expect = none] +nitpick [card = 100, expect = none] +apply (rule ext)+ +by auto + +lemma "split (\x y. f (x, y)) = f" +nitpick [card = 1\4, expect = none] +nitpick [card = 40, expect = none] +by auto + +subsection {* Representations *} + +lemma "\f. f = (\x. x) \ f y = y" +nitpick [expect = none] +by auto + +lemma "(\g. \x. g (f x) = x) \ (\y. \x. y = f x)" +nitpick [card 'a = 35, card 'b = 34, expect = genuine] +nitpick [card = 1\15, mono, expect = none] +oops + +lemma "\f. f = (\x. x) \ f y \ y" +nitpick [card = 1, expect = genuine] +nitpick [card = 2, expect = genuine] +nitpick [card = 5, expect = genuine] +oops + +lemma "P (\x. x)" +nitpick [card = 1, expect = genuine] +nitpick [card = 5, expect = genuine] +oops + +lemma "{(a\'a\'a, b\'b)}^-1 = {(b, a)}" +nitpick [card = 1\6, expect = none] +nitpick [card = 20, expect = none] +by auto + +lemma "fst (a, b) = a" +nitpick [card = 1\20, expect = none] +by auto + +lemma "\P. P = Id" +nitpick [card = 1\4, expect = none] +by auto + +lemma "(a\'a\'b, a) \ Id\<^sup>*" +nitpick [card = 1\3, expect = none] +by auto + +lemma "(a\'a\'a, a) \ Id\<^sup>* \ {(a, b)}\<^sup>*" +nitpick [card = 1\6, expect = none] +by auto + +lemma "Id (a, a)" +nitpick [card = 1\100, expect = none] +by (auto simp: Id_def Collect_def) + +lemma "Id ((a\'a, b\'a), (a, b))" +nitpick [card = 1\20, expect = none] +by (auto simp: Id_def Collect_def) + +lemma "UNIV (x\'a\'a)" +nitpick [card = 1\50, expect = none] +sorry + +lemma "{} = A - A" +nitpick [card = 1\100, expect = none] +by auto + +lemma "g = Let (A \ B)" +nitpick [card = 1, expect = none] +nitpick [card = 2, expect = genuine] +nitpick [card = 20, expect = genuine] +oops + +lemma "(let a_or_b = A \ B in a_or_b \ \ a_or_b)" +nitpick [expect = none] +by auto + +lemma "A \ B" +nitpick [card = 100, expect = genuine] +oops + +lemma "A = {b}" +nitpick [card = 100, expect = genuine] +oops + +lemma "{a, b} = {b}" +nitpick [card = 100, expect = genuine] +oops + +lemma "(a\'a\'a, a\'a\'a) \ R" +nitpick [card = 1, expect = genuine] +nitpick [card = 2, expect = genuine] +nitpick [card = 4, expect = genuine] +nitpick [card = 20, expect = genuine] +nitpick [card = 10, dont_box, expect = genuine] +oops + +lemma "f (g\'a\'a) = x" +nitpick [card = 3, expect = genuine] +nitpick [card = 3, dont_box, expect = genuine] +nitpick [card = 5, expect = genuine] +nitpick [card = 10, expect = genuine] +oops + +lemma "f (a, b) = x" +nitpick [card = 3, expect = genuine] +nitpick [card = 10, expect = genuine] +nitpick [card = 16, expect = genuine] +nitpick [card = 30, expect = genuine] +oops + +lemma "f (a, a) = f (c, d)" +nitpick [card = 4, expect = genuine] +nitpick [card = 20, expect = genuine] +oops + +lemma "(x\'a) = (\a. \b. \c. if c then a else b) x x True" +nitpick [card = 2, expect = none] +by auto + +lemma "\F. F a b = G a b" +nitpick [card = 3, expect = none] +by auto + +lemma "f = split" +nitpick [card = 1, expect = none] +nitpick [card = 2, expect = genuine] +oops + +lemma "(A\'a\'a, B\'a\'a) \ R \ (A, B) \ R" +nitpick [card = 20, expect = none] +by auto + +lemma "(A, B) \ R \ (\C. (A, C) \ R \ (C, B) \ R) \ + A = B \ (A, B) \ R \ (\C. (A, C) \ R \ (C, B) \ R)" +nitpick [card = 1\50, expect = none] +by auto + +lemma "f = (\x\'a\'b. x)" +nitpick [card = 3, expect = genuine] +nitpick [card = 4, expect = genuine] +nitpick [card = 8, expect = genuine] +oops + +subsection {* Quantifiers *} + +lemma "x = y" +nitpick [card 'a = 1, expect = none] +nitpick [card 'a = 2, expect = genuine] +nitpick [card 'a = 100, expect = genuine] +nitpick [card 'a = 1000, expect = genuine] +oops + +lemma "\x. x = y" +nitpick [card 'a = 1, expect = none] +nitpick [card 'a = 2, expect = genuine] +nitpick [card 'a = 100, expect = genuine] +nitpick [card 'a = 1000, expect = genuine] +oops + +lemma "\x\'a \ bool. x = y" +nitpick [card 'a = 1, expect = genuine] +nitpick [card 'a = 2, expect = genuine] +nitpick [card 'a = 100, expect = genuine] +nitpick [card 'a = 1000, expect = genuine] +oops + +lemma "\x\'a \ bool. x = y" +nitpick [card 'a = 1\10, expect = none] +by auto + +lemma "\x y\'a \ bool. x = y" +nitpick [card = 1\40, expect = none] +by auto + +lemma "\x. \y. f x y = f x (g x)" +nitpick [card = 1\5, expect = none] +by auto + +lemma "\u. \v. \w. \x. f u v w x = f u (g u) w (h u w)" +nitpick [card = 1\5, expect = none] +by auto + +lemma "\u. \v. \w. \x. f u v w x = f u (g u w) w (h u)" +nitpick [card = 1\2, expect = genuine] +nitpick [card = 3, expect = genuine] +oops + +lemma "\u. \v. \w. \x. \y. \z. + f u v w x y z = f u (g u) w (h u w) y (k u w y)" +nitpick [card = 1\2, expect = none] +nitpick [card = 3, expect = none] +nitpick [card = 4, expect = none] +sorry + +lemma "\u. \v. \w. \x. \y. \z. + f u v w x y z = f u (g u) w (h u w y) y (k u w y)" +nitpick [card = 1\2, expect = genuine] +oops + +lemma "\u. \v. \w. \x. \y. \z. + f u v w x y z = f u (g u w) w (h u w) y (k u w y)" +nitpick [card = 1\2, expect = genuine] +oops + +lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. + f u v w x = f u (g u) w (h u w)" +nitpick [card = 1\2, expect = none] +sorry + +lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. + f u v w x = f u (g u w) w (h u)" +nitpick [card = 1\2, dont_box, expect = genuine] +oops + +lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. + f u v w x = f u (g u) w (h u w)" +nitpick [card = 1\2, dont_box, expect = none] +sorry + +lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. + f u v w x = f u (g u w) w (h u)" +nitpick [card = 1\2, dont_box, expect = genuine] +oops + +lemma "\x. if (\y. x = y) then False else True" +nitpick [card = 1, expect = genuine] +nitpick [card = 2\5, expect = none] +oops + +lemma "\x\'a\'b. if (\y. x = y) then False else True" +nitpick [card = 1, expect = genuine] +nitpick [card = 2, expect = none] +oops + +lemma "\x. if (\y. x = y) then True else False" +nitpick [expect = none] +sorry + +lemma "\x\'a\'b. if (\y. x = y) then True else False" +nitpick [expect = none] +sorry + +lemma "(\ (\x. P x)) \ (\x. \ P x)" +nitpick [expect = none] +by auto + +lemma "(\ \ (\x. P x)) \ (\ (\x. \ P x))" +nitpick [expect = none] +by auto + +lemma "(\x\'a. \y. P x y) \ (\x\'a \ 'a. \y. P y x)" +nitpick [card 'a = 1, expect = genuine] +nitpick [card 'a = 2, expect = genuine] +nitpick [card 'a = 3, expect = genuine] +nitpick [card 'a = 4, expect = genuine] +nitpick [card 'a = 5, expect = genuine] +oops + +lemma "\x. if x = y then (\y. y = x \ y \ x) + else (\y. y = (x, x) \ y \ (x, x))" +nitpick [expect = none] +by auto + +lemma "\x. if x = y then (\y. y = x \ y \ x) + else (\y. y = (x, x) \ y \ (x, x))" +nitpick [expect = none] +by auto + +lemma "let x = (\x. P x) in if x then x else \ x" +nitpick [expect = none] +by auto + +lemma "let x = (\x\'a \ 'b. P x) in if x then x else \ x" +nitpick [expect = none] +by auto + +subsection {* Schematic Variables *} + +lemma "x = ?x" +nitpick [expect = none] +by auto + +lemma "\x. x = ?x" +nitpick [expect = genuine] +oops + +lemma "\x. x = ?x" +nitpick [expect = none] +by auto + +lemma "\x\'a \ 'b. x = ?x" +nitpick [expect = none] +by auto + +lemma "\x. ?x = ?y" +nitpick [expect = none] +by auto + +lemma "\x. ?x = ?y" +nitpick [expect = none] +by auto + +subsection {* Known Constants *} + +lemma "x \ all \ False" +nitpick [card = 1, expect = genuine] +nitpick [card = 1, box "('a \ prop) \ prop", expect = genuine] +nitpick [card = 2, expect = genuine] +nitpick [card = 8, expect = genuine] +nitpick [card = 10, expect = unknown] +oops + +lemma "\x. f x y = f x y" +nitpick [expect = none] +oops + +lemma "\x. f x y = f y x" +nitpick [expect = genuine] +oops + +lemma "all (\x. Trueprop (f x y = f x y)) \ Trueprop True" +nitpick [expect = none] +by auto + +lemma "all (\x. Trueprop (f x y = f x y)) \ Trueprop False" +nitpick [expect = genuine] +oops + +lemma "I = (\x. x) \ all P \ all (\x. P (I x))" +nitpick [expect = none] +by auto + +lemma "x \ (op \) \ False" +nitpick [card = 1, expect = genuine] +nitpick [card = 2, expect = genuine] +nitpick [card = 3, expect = genuine] +nitpick [card = 4, expect = genuine] +nitpick [card = 5, expect = genuine] +nitpick [card = 100, expect = genuine] +oops + +lemma "I = (\x. x) \ (op \ x) \ (\y. (x \ I y))" +nitpick [expect = none] +by auto + +lemma "P x \ P x" +nitpick [card = 1\10, expect = none] +by auto + +lemma "P x \ Q x \ P x = Q x" +nitpick [card = 1\10, expect = none] +by auto + +lemma "P x = Q x \ P x \ Q x" +nitpick [card = 1\10, expect = none] +by auto + +lemma "x \ (op \) \ False" +nitpick [expect = genuine] +oops + +lemma "I \ (\x. x) \ (op \ x) \ (\y. (op \ x (I y)))" +nitpick [expect = none] +by auto + +lemma "P x \ P x" +nitpick [card = 1\10, expect = none] +by auto + +lemma "True \ True" "False \ True" "False \ False" +nitpick [expect = none] +by auto + +lemma "True \ False" +nitpick [expect = genuine] +oops + +lemma "x = Not" +nitpick [expect = genuine] +oops + +lemma "I = (\x. x) \ Not = (\x. Not (I x))" +nitpick [expect = none] +by auto + +lemma "x = True" +nitpick [expect = genuine] +oops + +lemma "x = False" +nitpick [expect = genuine] +oops + +lemma "x = undefined" +nitpick [expect = genuine] +oops + +lemma "(False, ()) = undefined \ ((), False) = undefined" +nitpick [expect = genuine] +oops + +lemma "undefined = undefined" +nitpick [expect = none] +by auto + +lemma "f undefined = f undefined" +nitpick [expect = none] +by auto + +lemma "f undefined = g undefined" +nitpick [card = 33, expect = genuine] +oops + +lemma "\!x. x = undefined" +nitpick [card = 30, expect = none] +by auto + +lemma "x = All \ False" +nitpick [card = 1, dont_box, expect = genuine] +nitpick [card = 2, dont_box, expect = genuine] +nitpick [card = 8, dont_box, expect = genuine] +nitpick [card = 10, dont_box, expect = unknown] +oops + +lemma "\x. f x y = f x y" +nitpick [expect = none] +oops + +lemma "\x. f x y = f y x" +nitpick [expect = genuine] +oops + +lemma "All (\x. f x y = f x y) = True" +nitpick [expect = none] +by auto + +lemma "All (\x. f x y = f x y) = False" +nitpick [expect = genuine] +oops + +lemma "I = (\x. x) \ All P = All (\x. P (I x))" +nitpick [expect = none] +by auto + +lemma "x = Ex \ False" +nitpick [card = 1, dont_box, expect = genuine] +nitpick [card = 2, dont_box, expect = genuine] +nitpick [card = 8, dont_box, expect = genuine] +nitpick [card = 10, dont_box, expect = unknown] +oops + +lemma "\x. f x y = f x y" +nitpick [expect = none] +oops + +lemma "\x. f x y = f y x" +nitpick [expect = none] +oops + +lemma "Ex (\x. f x y = f x y) = True" +nitpick [expect = none] +by auto + +lemma "Ex (\x. f x y = f y x) = True" +nitpick [expect = none] +by auto + +lemma "Ex (\x. f x y = f x y) = False" +nitpick [expect = genuine] +oops + +lemma "Ex (\x. f x y = f y x) = False" +nitpick [expect = genuine] +oops + +lemma "Ex (\x. f x y \ f x y) = False" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ Ex P = Ex (\x. P (I x))" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ (op =) = (\x. (op= (I x)))" + "I = (\x. x) \ (op =) = (\x y. x = (I y))" +nitpick [expect = none] +by auto + +lemma "x = y \ y = x" +nitpick [expect = none] +by auto + +lemma "x = y \ f x = f y" +nitpick [expect = none] +by auto + +lemma "x = y \ y = z \ x = z" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ (op &) = (\x. op & (I x))" + "I = (\x. x) \ (op &) = (\x y. x & (I y))" +nitpick [expect = none] +by auto + +lemma "(a \ b) = (\ (\ a \ \ b))" +nitpick [expect = none] +by auto + +lemma "a \ b \ a" "a \ b \ b" +nitpick [expect = none] +by auto + +lemma "\ a \ \ (a \ b)" "\ b \ \ (a \ b)" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ (op \) = (\x. op \ (I x))" + "I = (\x. x) \ (op \) = (\x y. x \ (I y))" +nitpick [expect = none] +by auto + +lemma "a \ a \ b" "b \ a \ b" +nitpick [expect = none] +by auto + +lemma "\ (a \ b) \ \ a" "\ (a \ b) \ \ b" +nitpick [expect = none] +by auto + +lemma "(op \) = (\x. op\ x)" "(op\ ) = (\x y. x \ y)" +nitpick [expect = none] +by auto + +lemma "\a \ a \ b" "b \ a \ b" +nitpick [expect = none] +by auto + +lemma "\a; \ b\ \ \ (a \ b)" +nitpick [expect = none] +by auto + +lemma "((if a then b else c) = d) = ((a \ (b = d)) \ (\ a \ (c = d)))" +nitpick [expect = none] +by auto + +lemma "(if a then b else c) = (THE d. (a \ (d = b)) \ (\ a \ (d = c)))" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ If = (\x. If (I x))" + "J = (\x. x) \ If = (\x y. If x (J y))" + "K = (\x. x) \ If = (\x y z. If x y (K z))" +nitpick [expect = none] +by auto + +lemma "fst (x, y) = x" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd (x, y) = y" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "fst (x\'a\'b, y) = x" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd (x\'a\'b, y) = y" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "fst (x, y\'a\'b) = x" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd (x, y\'a\'b) = y" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "fst (x\'a\'b, y) = x" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd (x\'a\'b, y) = y" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "fst (x, y\'a\'b) = x" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd (x, y\'a\'b) = y" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "fst p = (THE a. \b. p = Pair a b)" +nitpick [expect = none] +by (simp add: fst_def) + +lemma "snd p = (THE b. \a. p = Pair a b)" +nitpick [expect = none] +by (simp add: snd_def) + +lemma "I = (\x. x) \ fst = (\x. fst (I x))" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ snd = (\x. snd (I x))" +nitpick [expect = none] +by auto + +lemma "fst (x, y) = snd (y, x)" +nitpick [expect = none] +by auto + +lemma "(x, x) \ Id" +nitpick [expect = none] +by auto + +lemma "(x, y) \ Id \ x = y" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ Id = (\x. Id (I x))" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ curry Id = (\x y. Id (x, I y))" +nitpick [expect = none] +by (simp add: curry_def) + +lemma "{} = (\x. False)" +nitpick [expect = none] +by (metis Collect_def empty_def) + +lemma "x \ {}" +nitpick [expect = genuine] +oops + +lemma "{a, b} = {b}" +nitpick [expect = genuine] +oops + +lemma "{a, b} \ {b}" +nitpick [expect = genuine] +oops + +lemma "{a} = {b}" +nitpick [expect = genuine] +oops + +lemma "{a} \ {b}" +nitpick [expect = genuine] +oops + +lemma "{a, b, c} = {c, b, a}" +nitpick [expect = none] +by auto + +lemma "UNIV = (\x. True)" +nitpick [expect = none] +by (simp only: UNIV_def Collect_def) + +lemma "UNIV x = True" +nitpick [expect = none] +by (simp only: UNIV_def Collect_def) + +lemma "x \ UNIV" +nitpick [expect = genuine] +oops + +lemma "op \ = (\x P. P x)" +nitpick [expect = none] +apply (rule ext) +apply (rule ext) +by (simp add: mem_def) + +lemma "I = (\x. x) \ op \ = (\x. (op \ (I x)))" +nitpick [expect = none] +apply (rule ext) +apply (rule ext) +by (simp add: mem_def) + +lemma "P x = (x \ P)" +nitpick [expect = none] +by (simp add: mem_def) + +lemma "I = (\x. x) \ insert = (\x. insert (I x))" +nitpick [expect = none] +by simp + +lemma "insert = (\x y. insert x (y \ y))" +nitpick [expect = none] +by simp + +lemma "I = (\x. x) \ trancl = (\x. trancl (I x))" +nitpick [card = 1\2, expect = none] +by auto + +lemma "rtrancl = (\x. rtrancl x \ {(y, y)})" +nitpick [card = 1\3, expect = none] +apply (rule ext) +by auto + +lemma "(x, x) \ rtrancl {(y, y)}" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ rtrancl = (\x. rtrancl (I x))" +nitpick [card = 1\2, expect = none] +by auto + +lemma "((x, x), (x, x)) \ rtrancl {}" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x. op \ (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x y. op \ x (I y))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "a \ A \ a \ (A \ B)" "b \ B \ b \ (A \ B)" +nitpick [expect = none] +by auto + +lemma "a \ (A \ B) \ a \ A \ a \ B" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x. op \ (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x y. op \ x (I y))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "a \ A \ a \ (A \ B)" "b \ B \ b \ (A \ B)" +nitpick [card = 1\5, expect = none] +by auto + +lemma "a \ (A \ B) \ a \ A \ a \ B" +nitpick [expect = none] +by auto + +lemma "I = (\x. x) \ op - = (\x\'a set. op - (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op - = (\x y\'a set. op - x (I y))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "x \ ((A\'a set) - B) \ x \ A \ x \ B" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x. op \ (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x y. op \ x (I y))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "A \ B \ (\a \ A. a \ B) \ (\b \ B. b \ A)" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x. op \ (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "I = (\x. x) \ op \ = (\x y. op \ x (I y))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "A \ B \ \a \ A. a \ B" +nitpick [card = 1\5, expect = none] +by auto + +lemma "A \ B \ A \ B" +nitpick [card = 5, expect = genuine] +oops + +lemma "A \ B \ A \ B" +nitpick [expect = none] +by auto + +lemma "I = (\x\'a set. x) \ uminus = (\x. uminus (I x))" +nitpick [expect = none] +by auto + +lemma "A \ - A = UNIV" +nitpick [expect = none] +by auto + +lemma "A \ - A = {}" +nitpick [expect = none] +by auto + +lemma "A = -(A\'a set)" +nitpick [card 'a = 10, expect = genuine] +oops + +lemma "I = (\x. x) \ finite = (\x. finite (I x))" +nitpick [expect = none] +oops + +lemma "finite A" +nitpick [expect = none] +oops + +lemma "finite A \ finite B" +nitpick [expect = none] +oops + +lemma "All finite" +nitpick [expect = none] +oops + +subsection {* The and Eps *} + +lemma "x = The" +nitpick [card = 5, expect = genuine] +oops + +lemma "\x. x = The" +nitpick [card = 1\3] +by auto + +lemma "P x \ (\y. P y \ y = x) \ The P = x" +nitpick [expect = none] +by auto + +lemma "P x \ P y \ x \ y \ The P = z" +nitpick [expect = genuine] +oops + +lemma "P x \ P y \ x \ y \ The P = x \ The P = y" +nitpick [card = 2, expect = none] +nitpick [card = 3\5, expect = genuine] +oops + +lemma "P x \ P (The P)" +nitpick [card = 1, expect = none] +nitpick [card = 1\2, expect = none] +nitpick [card = 3\5, expect = genuine] +nitpick [card = 8, expect = genuine] +oops + +lemma "(\x. \ P x) \ The P = y" +nitpick [expect = genuine] +oops + +lemma "I = (\x. x) \ The = (\x. The (I x))" +nitpick [card = 1\5, expect = none] +by auto + +lemma "x = Eps" +nitpick [card = 5, expect = genuine] +oops + +lemma "\x. x = Eps" +nitpick [card = 1\3, expect = none] +by auto + +lemma "P x \ (\y. P y \ y = x) \ Eps P = x" +nitpick [expect = none] +by auto + +lemma "P x \ P y \ x \ y \ Eps P = z" +nitpick [expect = genuine] +apply auto +oops + +lemma "P x \ P (Eps P)" +nitpick [card = 1\8, expect = none] +by (metis exE_some) + +lemma "\x. \ P x \ Eps P = y" +nitpick [expect = genuine] +oops + +lemma "P (Eps P)" +nitpick [expect = genuine] +oops + +lemma "(P\nat set) (Eps P)" +nitpick [expect = genuine] +oops + +lemma "\ P (Eps P)" +nitpick [expect = genuine] +oops + +lemma "\ (P\nat set) (Eps P)" +nitpick [expect = genuine] +oops + +lemma "P \ {} \ P (Eps P)" +nitpick [expect = none] +sorry + +lemma "(P\nat set) \ {} \ P (Eps P)" +nitpick [expect = none] +sorry + +lemma "P (The P)" +nitpick [expect = genuine] +oops + +lemma "(P\nat set) (The P)" +nitpick [expect = genuine] +oops + +lemma "\ P (The P)" +nitpick [expect = genuine] +oops + +lemma "\ (P\nat set) (The P)" +nitpick [expect = genuine] +oops + +lemma "The P \ x" +nitpick [expect = genuine] +oops + +lemma "The P \ (x\nat)" +nitpick [expect = genuine] +oops + +lemma "P x \ P (The P)" +nitpick [expect = genuine] +oops + +lemma "P (x\nat) \ P (The P)" +nitpick [expect = genuine] +oops + +lemma "P = {x} \ P (The P)" +nitpick [expect = none] +oops + +lemma "P = {x\nat} \ P (The P)" +nitpick [expect = none] +oops + +consts Q :: 'a + +lemma "Q (Eps Q)" +nitpick [expect = genuine] +oops + +lemma "(Q\nat set) (Eps Q)" +nitpick [expect = none] +oops + +lemma "\ Q (Eps Q)" +nitpick [expect = genuine] +oops + +lemma "\ (Q\nat set) (Eps Q)" +nitpick [expect = genuine] +oops + +lemma "(Q\'a set) \ {} \ (Q\'a set) (Eps Q)" +nitpick [expect = none] +sorry + +lemma "(Q\nat set) \ {} \ (Q\nat set) (Eps Q)" +nitpick [expect = none] +sorry + +lemma "Q (The Q)" +nitpick [expect = genuine] +oops + +lemma "(Q\nat set) (The Q)" +nitpick [expect = genuine] +oops + +lemma "\ Q (The Q)" +nitpick [expect = genuine] +oops + +lemma "\ (Q\nat set) (The Q)" +nitpick [expect = genuine] +oops + +lemma "The Q \ x" +nitpick [expect = genuine] +oops + +lemma "The Q \ (x\nat)" +nitpick [expect = genuine] +oops + +lemma "Q x \ Q (The Q)" +nitpick [expect = genuine] +oops + +lemma "Q (x\nat) \ Q (The Q)" +nitpick [expect = genuine] +oops + +lemma "Q = {x\'a} \ (Q\'a set) (The Q)" +nitpick [expect = none] +oops + +lemma "Q = {x\nat} \ (Q\nat set) (The Q)" +nitpick [expect = none] +oops + +subsection {* Destructors and Recursors *} + +lemma "(x\'a) = (case True of True \ x | False \ x)" +nitpick [card = 2, expect = none] +by auto + +lemma "bool_rec x y True = x" +nitpick [card = 2, expect = none] +by auto + +lemma "bool_rec x y False = y" +nitpick [card = 2, expect = none] +by auto + +lemma "(x\bool) = bool_rec x x True" +nitpick [card = 2, expect = none] +by auto + +lemma "x = (case (x, y) of (x', y') \ x')" +nitpick [expect = none] +sorry + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Datatype_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,104 @@ +(* Title: HOL/Nitpick_Examples/Datatype_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick applied to datatypes. +*) + +header {* Examples Featuring Nitpick Applied to Datatypes *} + +theory Datatype_Nits +imports Main +begin + +primrec rot where +"rot Nibble0 = Nibble1" | +"rot Nibble1 = Nibble2" | +"rot Nibble2 = Nibble3" | +"rot Nibble3 = Nibble4" | +"rot Nibble4 = Nibble5" | +"rot Nibble5 = Nibble6" | +"rot Nibble6 = Nibble7" | +"rot Nibble7 = Nibble8" | +"rot Nibble8 = Nibble9" | +"rot Nibble9 = NibbleA" | +"rot NibbleA = NibbleB" | +"rot NibbleB = NibbleC" | +"rot NibbleC = NibbleD" | +"rot NibbleD = NibbleE" | +"rot NibbleE = NibbleF" | +"rot NibbleF = Nibble0" + +lemma "rot n \ n" +nitpick [card = 1\16, expect = none] +sorry + +lemma "rot Nibble2 \ Nibble3" +nitpick [card = 1, expect = none] +nitpick [card = 2, expect = genuine] +nitpick [card = 2, max Nibble2 = 0, expect = none] +nitpick [card = 2, max Nibble3 = 0, expect = none] +oops + +lemma "(rot ^^ 15) n \ n" +nitpick [card = 17, expect = none] +sorry + +lemma "(rot ^^ 15) n = n" +nitpick [card = 17, expect = genuine] +oops + +lemma "(rot ^^ 16) n = n" +nitpick [card = 17, expect = none] +oops + +datatype ('a, 'b) pd = Pd "'a \ 'b" + +fun fs where +"fs (Pd (a, _)) = a" + +fun sn where +"sn (Pd (_, b)) = b" + +lemma "fs (Pd p) = fst p" +nitpick [card = 20, expect = none] +sorry + +lemma "fs (Pd p) = snd p" +nitpick [expect = genuine] +oops + +lemma "sn (Pd p) = snd p" +nitpick [card = 20, expect = none] +sorry + +lemma "sn (Pd p) = fst p" +nitpick [expect = genuine] +oops + +lemma "fs (Pd ((a, b), (c, d))) = (a, b)" +nitpick [card = 1\12, expect = none] +sorry + +lemma "fs (Pd ((a, b), (c, d))) = (c, d)" +nitpick [expect = genuine] +oops + +datatype ('a, 'b) fn = Fn "'a \ 'b" + +fun app where +"app (Fn f) x = f x" + +lemma "app (Fn g) y = g y" +nitpick [card = 1\12, expect = none] +sorry + +lemma "app (Fn g) y = g' y" +nitpick [expect = genuine] +oops + +lemma "app (Fn g) y = g y'" +nitpick [expect = genuine] +oops + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Induct_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Induct_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,171 @@ +(* Title: HOL/Nitpick_Examples/Induct_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick applied to (co)inductive definitions. +*) + +header {* Examples Featuring Nitpick Applied to (Co)inductive Definitions *} + +theory Induct_Nits +imports Main +begin + +nitpick_params [show_all] + +inductive p1 :: "nat \ bool" where +"p1 0" | +"p1 n \ p1 (n + 2)" + +coinductive q1 :: "nat \ bool" where +"q1 0" | +"q1 n \ q1 (n + 2)" + +lemma "p1 = q1" +nitpick [expect = none] +nitpick [wf, expect = none] +nitpick [non_wf, expect = none] +nitpick [non_wf, dont_star_linear_preds, expect = none] +oops + +lemma "p1 \ q1" +nitpick [expect = potential] +nitpick [wf, expect = potential] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_star_linear_preds, expect = potential] +oops + +lemma "p1 (n - 2) \ p1 n" +nitpick [expect = genuine] +nitpick [non_wf, expect = genuine] +nitpick [non_wf, dont_star_linear_preds, expect = genuine] +oops + +lemma "q1 (n - 2) \ q1 n" +nitpick [expect = genuine] +nitpick [non_wf, expect = genuine] +nitpick [non_wf, dont_star_linear_preds, expect = genuine] +oops + +inductive p2 :: "nat \ bool" where +"p2 n \ p2 n" + +coinductive q2 :: "nat \ bool" where +"q2 n \ q2 n" + +lemma "p2 = {}" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +sorry + +lemma "q2 = {}" +nitpick [expect = genuine] +nitpick [dont_star_linear_preds, expect = genuine] +nitpick [wf, expect = likely_genuine] +oops + +lemma "p2 = UNIV" +nitpick [expect = genuine] +nitpick [dont_star_linear_preds, expect = genuine] +oops + +lemma "q2 = UNIV" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +nitpick [wf, expect = likely_genuine] +sorry + +lemma "p2 = q2" +nitpick [expect = genuine] +nitpick [dont_star_linear_preds, expect = genuine] +oops + +lemma "p2 n" +nitpick [expect = genuine] +nitpick [dont_star_linear_preds, expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +lemma "q2 n" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +sorry + +lemma "\ p2 n" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +sorry + +lemma "\ q2 n" +nitpick [expect = genuine] +nitpick [dont_star_linear_preds, expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +inductive p3 and p4 where +"p3 0" | +"p3 n \ p4 (Suc n)" | +"p4 n \ p3 (Suc n)" + +coinductive q3 and q4 where +"q3 0" | +"q3 n \ q4 (Suc n)" | +"q4 n \ q3 (Suc n)" + +lemma "p3 = q3" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_box, expect = none] +nitpick [non_wf, dont_star_linear_preds, expect = none] +sorry + +lemma "p4 = q4" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_box, expect = none] +nitpick [non_wf, dont_star_linear_preds, expect = none] +sorry + +lemma "p3 = UNIV - p4" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_box, expect = none] +nitpick [non_wf, dont_star_linear_preds, expect = none] +sorry + +lemma "q3 = UNIV - q4" +nitpick [expect = none] +nitpick [dont_star_linear_preds, expect = none] +nitpick [non_wf, expect = none] +nitpick [non_wf, dont_box, expect = none] +nitpick [non_wf, dont_star_linear_preds, expect = none] +sorry + +lemma "p3 \ q4 \ {}" +nitpick [expect = potential] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_star_linear_preds, expect = potential] +sorry + +lemma "q3 \ p4 \ {}" +nitpick [expect = potential] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_star_linear_preds, expect = potential] +sorry + +lemma "p3 \ q4 \ UNIV" +nitpick [expect = potential] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_star_linear_preds, expect = potential] +sorry + +lemma "q3 \ p4 \ UNIV" +nitpick [expect = potential] +nitpick [non_wf, expect = potential] +nitpick [non_wf, dont_star_linear_preds, expect = potential] +sorry + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Manual_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,389 @@ +(* Title: HOL/Nitpick_Examples/Manual_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples from the Nitpick manual. +*) + +header {* Examples from the Nitpick Manual *} + +theory Manual_Nits +imports Main Coinductive_List RealDef +begin + +chapter {* 3. First Steps *} + +nitpick_params [sat_solver = MiniSatJNI, max_threads = 1] + +subsection {* 3.1. Propositional Logic *} + +lemma "P \ Q" +nitpick +apply auto +nitpick 1 +nitpick 2 +oops + +subsection {* 3.2. Type Variables *} + +lemma "P x \ P (THE y. P y)" +nitpick [verbose] +oops + +subsection {* 3.3. Constants *} + +lemma "P x \ P (THE y. P y)" +nitpick [show_consts] +nitpick [full_descrs, show_consts] +nitpick [dont_specialize, full_descrs, show_consts] +oops + +lemma "\!x. P x \ P (THE y. P y)" +nitpick +nitpick [card 'a = 1-50] +(* sledgehammer *) +apply (metis the_equality) +done + +subsection {* 3.4. Skolemization *} + +lemma "\g. \x. g (f x) = x \ \y. \x. y = f x" +nitpick +oops + +lemma "\x. \f. f x = x" +nitpick +oops + +lemma "refl r \ sym r" +nitpick +oops + +subsection {* 3.5. Numbers *} + +lemma "\i \ j; n \ (m\int)\ \ i * n + j * m \ i * m + j * n" +nitpick +oops + +lemma "\n. Suc n \ n \ P" +nitpick [card nat = 100, check_potential] +oops + +lemma "P Suc" +nitpick [card = 1-6] +oops + +lemma "P (op +\nat\nat\nat)" +nitpick [card nat = 1] +nitpick [card nat = 2] +oops + +subsection {* 3.6. Inductive Datatypes *} + +lemma "hd (xs @ [y, y]) = hd xs" +nitpick +nitpick [show_consts, show_datatypes] +oops + +lemma "\length xs = 1; length ys = 1\ \ xs = ys" +nitpick [show_datatypes] +oops + +subsection {* 3.7. Typedefs, Records, Rationals, and Reals *} + +typedef three = "{0\nat, 1, 2}" +by blast + +definition A :: three where "A \ Abs_three 0" +definition B :: three where "B \ Abs_three 1" +definition C :: three where "C \ Abs_three 2" + +lemma "\P A; P B\ \ P x" +nitpick [show_datatypes] +oops + +record point = + Xcoord :: int + Ycoord :: int + +lemma "Xcoord (p\point) = Xcoord (q\point)" +nitpick [show_datatypes] +oops + +lemma "4 * x + 3 * (y\real) \ 1 / 2" +nitpick [show_datatypes] +oops + +subsection {* 3.8. Inductive and Coinductive Predicates *} + +inductive even where +"even 0" | +"even n \ even (Suc (Suc n))" + +lemma "\n. even n \ even (Suc n)" +nitpick [card nat = 100, verbose] +oops + +lemma "\n \ 99. even n \ even (Suc n)" +nitpick [card nat = 100] +oops + +inductive even' where +"even' (0\nat)" | +"even' 2" | +"\even' m; even' n\ \ even' (m + n)" + +lemma "\n \ {0, 2, 4, 6, 8}. \ even' n" +nitpick [card nat = 10, verbose, show_consts] +oops + +lemma "even' (n - 2) \ even' n" +nitpick [card nat = 10, show_consts] +oops + +coinductive nats where +"nats (x\nat) \ nats x" + +lemma "nats = {0, 1, 2, 3, 4}" +nitpick [card nat = 10, show_consts] +oops + +inductive odd where +"odd 1" | +"\odd m; even n\ \ odd (m + n)" + +lemma "odd n \ odd (n - 2)" +nitpick [card nat = 10, show_consts] +oops + +subsection {* 3.9. Coinductive Datatypes *} + +lemma "xs \ LCons a xs" +nitpick +oops + +lemma "\xs = LCons a xs; ys = iterates (\b. a) b\ \ xs = ys" +nitpick [verbose] +nitpick [bisim_depth = -1, verbose] +oops + +lemma "\xs = LCons a xs; ys = LCons a ys\ \ xs = ys" +nitpick [bisim_depth = -1, show_datatypes] +nitpick +sorry + +subsection {* 3.10. Boxing *} + +datatype tm = Var nat | Lam tm | App tm tm + +primrec lift where +"lift (Var j) k = Var (if j < k then j else j + 1)" | +"lift (Lam t) k = Lam (lift t (k + 1))" | +"lift (App t u) k = App (lift t k) (lift u k)" + +primrec loose where +"loose (Var j) k = (j \ k)" | +"loose (Lam t) k = loose t (Suc k)" | +"loose (App t u) k = (loose t k \ loose u k)" + +primrec subst\<^isub>1 where +"subst\<^isub>1 \ (Var j) = \ j" | +"subst\<^isub>1 \ (Lam t) = + Lam (subst\<^isub>1 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 1) t)" | +"subst\<^isub>1 \ (App t u) = App (subst\<^isub>1 \ t) (subst\<^isub>1 \ u)" + +lemma "\ loose t 0 \ subst\<^isub>1 \ t = t" +nitpick [verbose] +nitpick [eval = "subst\<^isub>1 \ t"] +nitpick [dont_box] +oops + +primrec subst\<^isub>2 where +"subst\<^isub>2 \ (Var j) = \ j" | +"subst\<^isub>2 \ (Lam t) = + Lam (subst\<^isub>2 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 0) t)" | +"subst\<^isub>2 \ (App t u) = App (subst\<^isub>2 \ t) (subst\<^isub>2 \ u)" + +lemma "\ loose t 0 \ subst\<^isub>2 \ t = t" +nitpick +sorry + +subsection {* 3.11. Scope Monotonicity *} + +lemma "length xs = length ys \ rev (zip xs ys) = zip xs (rev ys)" +nitpick [verbose] +nitpick [card = 8, verbose] +oops + +lemma "\g. \x\'b. g (f x) = x \ \y\'a. \x. y = f x" +nitpick [mono] +nitpick +oops + +section {* 4. Case Studies *} + +nitpick_params [max_potential = 0, max_threads = 2] + +subsection {* 4.1. A Context-Free Grammar *} + +datatype alphabet = a | b + +inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where + "[] \ S\<^isub>1" +| "w \ A\<^isub>1 \ b # w \ S\<^isub>1" +| "w \ B\<^isub>1 \ a # w \ S\<^isub>1" +| "w \ S\<^isub>1 \ a # w \ A\<^isub>1" +| "w \ S\<^isub>1 \ b # w \ S\<^isub>1" +| "\v \ B\<^isub>1; v \ B\<^isub>1\ \ a # v @ w \ B\<^isub>1" + +theorem S\<^isub>1_sound: +"w \ S\<^isub>1 \ length [x \ w. x = a] = length [x \ w. x = b]" +nitpick +oops + +inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where + "[] \ S\<^isub>2" +| "w \ A\<^isub>2 \ b # w \ S\<^isub>2" +| "w \ B\<^isub>2 \ a # w \ S\<^isub>2" +| "w \ S\<^isub>2 \ a # w \ A\<^isub>2" +| "w \ S\<^isub>2 \ b # w \ B\<^isub>2" +| "\v \ B\<^isub>2; v \ B\<^isub>2\ \ a # v @ w \ B\<^isub>2" + +theorem S\<^isub>2_sound: +"w \ S\<^isub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" +nitpick +oops + +inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where + "[] \ S\<^isub>3" +| "w \ A\<^isub>3 \ b # w \ S\<^isub>3" +| "w \ B\<^isub>3 \ a # w \ S\<^isub>3" +| "w \ S\<^isub>3 \ a # w \ A\<^isub>3" +| "w \ S\<^isub>3 \ b # w \ B\<^isub>3" +| "\v \ B\<^isub>3; w \ B\<^isub>3\ \ a # v @ w \ B\<^isub>3" + +theorem S\<^isub>3_sound: +"w \ S\<^isub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" +nitpick +sorry + +theorem S\<^isub>3_complete: +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>3" +nitpick +oops + +inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where + "[] \ S\<^isub>4" +| "w \ A\<^isub>4 \ b # w \ S\<^isub>4" +| "w \ B\<^isub>4 \ a # w \ S\<^isub>4" +| "w \ S\<^isub>4 \ a # w \ A\<^isub>4" +| "\v \ A\<^isub>4; w \ A\<^isub>4\ \ b # v @ w \ A\<^isub>4" +| "w \ S\<^isub>4 \ b # w \ B\<^isub>4" +| "\v \ B\<^isub>4; w \ B\<^isub>4\ \ a # v @ w \ B\<^isub>4" + +theorem S\<^isub>4_sound: +"w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +nitpick +sorry + +theorem S\<^isub>4_complete: +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>4" +nitpick +sorry + +theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete: +"w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +"w \ A\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b] + 1" +"w \ B\<^isub>4 \ length [x \ w. x = b] = length [x \ w. x = a] + 1" +nitpick +sorry + +subsection {* 4.2. AA Trees *} + +datatype 'a tree = \ | N "'a\linorder" nat "'a tree" "'a tree" + +primrec data where +"data \ = undefined" | +"data (N x _ _ _) = x" + +primrec dataset where +"dataset \ = {}" | +"dataset (N x _ t u) = {x} \ dataset t \ dataset u" + +primrec level where +"level \ = 0" | +"level (N _ k _ _) = k" + +primrec left where +"left \ = \" | +"left (N _ _ t\<^isub>1 _) = t\<^isub>1" + +primrec right where +"right \ = \" | +"right (N _ _ _ t\<^isub>2) = t\<^isub>2" + +fun wf where +"wf \ = True" | +"wf (N _ k t u) = + (if t = \ then + k = 1 \ (u = \ \ (level u = 1 \ left u = \ \ right u = \)) + else + wf t \ wf u \ u \ \ \ level t < k \ level u \ k \ level (right u) < k)" + +fun skew where +"skew \ = \" | +"skew (N x k t u) = + (if t \ \ \ k = level t then + N (data t) k (left t) (N x k (right t) u) + else + N x k t u)" + +fun split where +"split \ = \" | +"split (N x k t u) = + (if u \ \ \ k = level (right u) then + N (data u) (Suc k) (N x k t (left u)) (right u) + else + N x k t u)" + +theorem dataset_skew_split: +"dataset (skew t) = dataset t" +"dataset (split t) = dataset t" +nitpick +sorry + +theorem wf_skew_split: +"wf t \ skew t = t" +"wf t \ split t = t" +nitpick +sorry + +primrec insort\<^isub>1 where +"insort\<^isub>1 \ x = N x 1 \ \" | +"insort\<^isub>1 (N y k t u) x = + (* (split \ skew) *) (N y k (if x < y then insort\<^isub>1 t x else t) + (if x > y then insort\<^isub>1 u x else u))" + +theorem wf_insort\<^isub>1: "wf t \ wf (insort\<^isub>1 t x)" +nitpick +oops + +theorem wf_insort\<^isub>1_nat: "wf t \ wf (insort\<^isub>1 t (x\nat))" +nitpick [eval = "insort\<^isub>1 t x"] +oops + +primrec insort\<^isub>2 where +"insort\<^isub>2 \ x = N x 1 \ \" | +"insort\<^isub>2 (N y k t u) x = + (split \ skew) (N y k (if x < y then insort\<^isub>2 t x else t) + (if x > y then insort\<^isub>2 u x else u))" + +theorem wf_insort\<^isub>2: "wf t \ wf (insort\<^isub>2 t x)" +nitpick +sorry + +theorem dataset_insort\<^isub>2: "dataset (insort\<^isub>2 t x) = {x} \ dataset t" +nitpick +sorry + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Mini_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,107 @@ +(* Title: HOL/Nitpick_Examples/Mini_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Minipick, the minimalistic version of Nitpick. +*) + +header {* Examples Featuring Minipick, the Minimalistic Version of Nitpick *} + +theory Mini_Nits +imports Main +begin + +ML {* +exception FAIL + +(* int -> term -> string *) +fun minipick 0 _ = "none" + | minipick n t = + case minipick (n - 1) t of + "none" => Minipick.pick_nits_in_term @{context} (K n) t + | outcome_code => outcome_code +(* int -> term -> bool *) +fun none n t = (minipick n t = "none" orelse raise FAIL) +fun genuine n t = (minipick n t = "genuine" orelse raise FAIL) +fun unknown n t = (minipick n t = "unknown" orelse raise FAIL) +*} + +ML {* minipick 1 @{prop "\x\'a. \y\'b. f x = y"} *} + +ML {* genuine 1 @{prop "x = Not"} *} +ML {* none 1 @{prop "\x. x = Not"} *} +ML {* none 1 @{prop "\ False"} *} +ML {* genuine 1 @{prop "\ True"} *} +ML {* none 1 @{prop "\ \ b \ b"} *} +ML {* none 1 @{prop True} *} +ML {* genuine 1 @{prop False} *} +ML {* genuine 1 @{prop "True \ False"} *} +ML {* none 1 @{prop "True \ \ False"} *} +ML {* none 5 @{prop "\x. x = x"} *} +ML {* none 5 @{prop "\x. x = x"} *} +ML {* none 1 @{prop "\x. x = y"} *} +ML {* genuine 2 @{prop "\x. x = y"} *} +ML {* none 1 @{prop "\x. x = y"} *} +ML {* none 2 @{prop "\x. x = y"} *} +ML {* none 2 @{prop "\x\'a \ 'a. x = x"} *} +ML {* none 2 @{prop "\x\'a \ 'a. x = y"} *} +ML {* genuine 2 @{prop "\x\'a \ 'a. x = y"} *} +ML {* none 2 @{prop "\x\'a \ 'a. x = y"} *} +ML {* none 1 @{prop "All = Ex"} *} +ML {* genuine 2 @{prop "All = Ex"} *} +ML {* none 1 @{prop "All P = Ex P"} *} +ML {* genuine 2 @{prop "All P = Ex P"} *} +ML {* none 5 @{prop "x = y \ P x = P y"} *} +ML {* none 5 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} +ML {* none 2 @{prop "(x\'a \ 'a) = y \ P x y = P y x"} *} +ML {* none 5 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} +ML {* none 2 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} +ML {* none 2 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} +ML {* genuine 1 @{prop "(op =) X = Ex"} *} +ML {* none 2 @{prop "\x::'a \ 'a. x = x"} *} +ML {* none 1 @{prop "x = y"} *} +ML {* genuine 1 @{prop "x \ y"} *} +ML {* genuine 2 @{prop "x = y"} *} +ML {* genuine 1 @{prop "X \ Y"} *} +ML {* none 1 @{prop "P \ Q \ Q \ P"} *} +ML {* none 1 @{prop "P \ Q \ P"} *} +ML {* none 1 @{prop "P \ Q \ Q \ P"} *} +ML {* genuine 1 @{prop "P \ Q \ P"} *} +ML {* none 1 @{prop "(P \ Q) \ (\ P \ Q)"} *} +ML {* none 5 @{prop "{a} = {a, a}"} *} +ML {* genuine 2 @{prop "{a} = {a, b}"} *} +ML {* genuine 1 @{prop "{a} \ {a, b}"} *} +ML {* none 5 @{prop "{}\<^sup>+ = {}"} *} +ML {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML {* none 5 @{prop "a \ c \ {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML {* none 5 @{prop "A \ B = (\x. A x \ B x)"} *} +ML {* none 5 @{prop "A \ B = (\x. A x \ B x)"} *} +ML {* none 5 @{prop "A - B = (\x. A x \ \ B x)"} *} +ML {* none 5 @{prop "\a b. (a, b) = (b, a)"} *} +ML {* genuine 2 @{prop "(a, b) = (b, a)"} *} +ML {* genuine 2 @{prop "(a, b) \ (b, a)"} *} +ML {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} +ML {* genuine 2 @{prop "(a\'a \ 'a, b) = (b, a)"} *} +ML {* none 5 @{prop "\a b\'a \ 'a \ 'a. (a, b) = (b, a)"} *} +ML {* genuine 2 @{prop "(a\'a \ 'a \ 'a, b) \ (b, a)"} *} +ML {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} +ML {* genuine 1 @{prop "(a\'a \ 'a, b) \ (b, a)"} *} +ML {* none 8 @{prop "fst (a, b) = a"} *} +ML {* none 1 @{prop "fst (a, b) = b"} *} +ML {* genuine 2 @{prop "fst (a, b) = b"} *} +ML {* genuine 2 @{prop "fst (a, b) \ b"} *} +ML {* none 8 @{prop "snd (a, b) = b"} *} +ML {* none 1 @{prop "snd (a, b) = a"} *} +ML {* genuine 2 @{prop "snd (a, b) = a"} *} +ML {* genuine 2 @{prop "snd (a, b) \ a"} *} +ML {* genuine 1 @{prop P} *} +ML {* genuine 1 @{prop "(\x. P) a"} *} +ML {* genuine 1 @{prop "(\x y z. P y x z) a b c"} *} +ML {* none 5 @{prop "\f. f = (\x. x) \ f y = y"} *} +ML {* genuine 1 @{prop "\f. f p \ p \ (\a b. f (a, b) = (a, b))"} *} +ML {* none 2 @{prop "\f. \a b. f (a, b) = (a, b)"} *} +ML {* none 3 @{prop "f = (\a b. (b, a)) \ f x y = (y, x)"} *} +ML {* genuine 2 @{prop "f = (\a b. (b, a)) \ f x y = (x, y)"} *} + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Mono_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,98 @@ +(* Title: HOL/Nitpick_Examples/Mono_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick's monotonicity check. +*) + +header {* Examples Featuring Nitpick's Monotonicity Check *} + +theory Mono_Nits +imports Main +begin + +ML {* +exception FAIL + +val defs = NitpickHOL.all_axioms_of @{theory} |> #1 +val def_table = NitpickHOL.const_def_table @{context} defs +val ext_ctxt = + {thy = @{theory}, ctxt = @{context}, max_bisim_depth = ~1, boxes = [], + user_axioms = NONE, debug = false, wfs = [], destroy_constrs = false, + specialize = false, skolemize = false, star_linear_preds = false, + uncurry = false, fast_descrs = false, tac_timeout = NONE, evals = [], + case_names = [], def_table = def_table, nondef_table = Symtab.empty, + user_nondefs = [], simp_table = Unsynchronized.ref Symtab.empty, + psimp_table = Symtab.empty, intro_table = Symtab.empty, + ground_thm_table = Inttab.empty, ersatz_table = [], + skolems = Unsynchronized.ref [], special_funs = Unsynchronized.ref [], + unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref []} +(* term -> bool *) +val is_mono = NitpickMono.formulas_monotonic ext_ctxt @{typ 'a} [] [] +fun is_const t = + let val T = fastype_of t in + is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), + @{const False})) + end +fun mono t = is_mono t orelse raise FAIL +fun nonmono t = not (is_mono t) orelse raise FAIL +fun const t = is_const t orelse raise FAIL +fun nonconst t = not (is_const t) orelse raise FAIL +*} + +ML {* const @{term "A::('a\'b)"} *} +ML {* const @{term "(A::'a set) = A"} *} +ML {* const @{term "(A::'a set set) = A"} *} +ML {* const @{term "(\x::'a set. x a)"} *} +ML {* const @{term "{{a}} = C"} *} +ML {* const @{term "{f::'a\nat} = {g::'a\nat}"} *} +ML {* const @{term "A \ B"} *} +ML {* const @{term "P (a::'a)"} *} +ML {* const @{term "\a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *} +ML {* const @{term "\A::'a set. A a"} *} +ML {* const @{term "\A::'a set. P A"} *} +ML {* const @{term "P \ Q"} *} +ML {* const @{term "A \ B = C"} *} +ML {* const @{term "(if P then (A::'a set) else B) = C"} *} +ML {* const @{term "let A = C in A \ B"} *} +ML {* const @{term "THE x::'b. P x"} *} +ML {* const @{term "{}::'a set"} *} +ML {* const @{term "(\x::'a. True)"} *} +ML {* const @{term "Let a A"} *} +ML {* const @{term "A (a::'a)"} *} +ML {* const @{term "insert a A = B"} *} +ML {* const @{term "- (A::'a set)"} *} +ML {* const @{term "finite A"} *} +ML {* const @{term "\ finite A"} *} +ML {* const @{term "finite (A::'a set set)"} *} +ML {* const @{term "\a::'a. A a \ \ B a"} *} +ML {* const @{term "A < (B::'a set)"} *} +ML {* const @{term "A \ (B::'a set)"} *} +ML {* const @{term "[a::'a]"} *} +ML {* const @{term "[a::'a set]"} *} +ML {* const @{term "[A \ (B::'a set)]"} *} +ML {* const @{term "[A \ (B::'a set)] = [C]"} *} +ML {* const @{term "\P. P a"} *} + +ML {* nonconst @{term "{%x. True}"} *} +ML {* nonconst @{term "{(%x. x = a)} = C"} *} +ML {* nonconst @{term "\P (a::'a). P a"} *} +ML {* nonconst @{term "\a::'a. P a"} *} +ML {* nonconst @{term "(\a::'a. \ A a) = B"} *} +ML {* nonconst @{term "THE x. P x"} *} +ML {* nonconst @{term "SOME x. P x"} *} + +ML {* mono @{prop "Q (\x::'a set. P x)"} *} +ML {* mono @{prop "P (a::'a)"} *} +ML {* mono @{prop "{a} = {b}"} *} +ML {* mono @{prop "P (a::'a) \ P \ P = P"} *} +ML {* mono @{prop "\F::'a set set. P"} *} +ML {* mono @{prop "\ (\F f g (h::'a set). F f \ F g \ \ f a \ g a \ F h)"} *} +ML {* mono @{prop "\ Q (\x::'a set. P x)"} *} +ML {* mono @{prop "\ (\x. P x)"} *} + +ML {* nonmono @{prop "\x. P x"} *} +ML {* nonmono @{prop "\F f g (h::'a set). F f \ F g \ \ f a \ g a \ F h"} *} +ML {* nonmono @{prop "myall P = (P = (\x. True))"} *} + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Nitpick_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Nitpick_Examples.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,14 @@ +(* Title: HOL/Nitpick_Examples/Nitpick_Examples.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Nitpick examples. +*) + +theory Nitpick_Examples +imports Core_Nits Datatype_Nits Induct_Nits Manual_Nits Mini_Nits Mono_Nits + Pattern_Nits Record_Nits Refute_Nits Special_Nits Tests_Nits + Typedef_Nits +begin + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Pattern_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Pattern_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,138 @@ +(* Title: HOL/Nitpick_Examples/Pattern_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick's "destroy_constrs" optimization. +*) + +header {* Examples Featuring Nitpick's \textit{destroy\_constrs} Optimization *} + +theory Pattern_Nits +imports Main +begin + +nitpick_params [card = 14] + +lemma "x = (case u of () \ y)" +nitpick [expect = genuine] +oops + +lemma "x = (case b of True \ x | False \ y)" +nitpick [expect = genuine] +oops + +lemma "x = (case p of (x, y) \ y)" +nitpick [expect = genuine] +oops + +lemma "x = (case n of 0 \ x | Suc n \ n)" +nitpick [expect = genuine] +oops + +lemma "x = (case opt of None \ x | Some y \ y)" +nitpick [expect = genuine] +oops + +lemma "x = (case xs of [] \ x | y # ys \ y)" +nitpick [expect = genuine] +oops + +lemma "x = (case xs of + [] \ x + | y # ys \ + (case ys of + [] \ x + | z # zs \ + (case z of + None \ x + | Some p \ + (case p of + (a, b) \ b))))" +nitpick [expect = genuine] +oops + +fun f1 where +"f1 x () = x" + +lemma "x = f1 y u" +nitpick [expect = genuine] +oops + +fun f2 where +"f2 x _ True = x" | +"f2 _ y False = y" + +lemma "x = f2 x y b" +nitpick [expect = genuine] +oops + +fun f3 where +"f3 (_, y) = y" + +lemma "x = f3 p" +nitpick [expect = genuine] +oops + +fun f4 where +"f4 x 0 = x" | +"f4 _ (Suc n) = n" + +lemma "x = f4 x n" +nitpick [expect = genuine] +oops + +fun f5 where +"f5 x None = x" | +"f5 _ (Some y) = y" + +lemma "x = f5 x opt" +nitpick [expect = genuine] +oops + +fun f6 where +"f6 x [] = x" | +"f6 _ (y # ys) = y" + +lemma "x = f6 x xs" +nitpick [expect = genuine] +oops + +fun f7 where +"f7 _ (y # Some (a, b) # zs) = b" | +"f7 x (y # None # zs) = x" | +"f7 x [y] = x" | +"f7 x [] = x" + +lemma "x = f7 x xs" +nitpick [expect = genuine] +oops + +lemma "u = ()" +nitpick [expect = none] +sorry + +lemma "\y. (b::bool) = y" +nitpick [expect = none] +sorry + +lemma "\x y. p = (x, y)" +nitpick [expect = none] +sorry + +lemma "\x. n = Suc x" +nitpick [expect = genuine] +oops + +lemma "\y. x = Some y" +nitpick [expect = genuine] +oops + +lemma "\y ys. xs = y # ys" +nitpick [expect = genuine] +oops + +lemma "\y a b zs. x = (y # Some (a, b) # zs)" +nitpick [expect = genuine] +oops + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/ROOT.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,8 @@ +(* Title: HOL/Nitpick_Examples/ROOT.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Nitpick examples. +*) + +setmp_noncritical quick_and_dirty true (try use_thy) "Nitpick_Examples"; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Record_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Record_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,84 @@ +(* Title: HOL/Nitpick_Examples/Record_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick applied to records. +*) + +header {* Examples Featuring Nitpick Applied to Records *} + +theory Record_Nits +imports Main +begin + +record point2d = + xc :: int + yc :: int + +lemma "\xc = x, yc = y\ = p\xc := x, yc := y\" +nitpick [expect = none] +sorry + +lemma "\xc = x, yc = y\ = p\xc := x\" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y\ \ p" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y\ = p" +nitpick [expect = genuine] +oops + +record point3d = point2d + + zc :: int + +lemma "\xc = x, yc = y, zc = z\ = p\xc := x, yc := y, zc := z\" +nitpick [expect = none] +sorry + +lemma "\xc = x, yc = y, zc = z\ = p\xc := x\" +nitpick [expect = genuine] +oops + +lemma "\xc = x, yc = y, zc = z\ = p\zc := z\" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y, zc := z\ \ p" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y, zc := z\ = p" +nitpick [expect = genuine] +oops + +record point4d = point3d + + wc :: int + +lemma "\xc = x, yc = y, zc = z, wc = w\ = p\xc := x, yc := y, zc := z, wc := w\" +nitpick [expect = none] +sorry + +lemma "\xc = x, yc = y, zc = z, wc = w\ = p\xc := x\" +nitpick [expect = genuine] +oops + +lemma "\xc = x, yc = y, zc = z, wc = w\ = p\zc := z\" +nitpick [expect = genuine] +oops + +lemma "\xc = x, yc = y, zc = z, wc = w\ = p\wc := w\" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y, zc := z, wc := w\ \ p" +nitpick [expect = genuine] +oops + +lemma "p\xc := x, yc := y, zc := z, wc := w\ = p" +nitpick [expect = genuine] +oops + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Refute_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,1476 @@ +(* Title: HOL/Nitpick_Examples/Refute_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Refute examples adapted to Nitpick. +*) + +header {* Refute Examples Adapted to Nitpick *} + +theory Refute_Nits +imports Main +begin + +lemma "P \ Q" +apply (rule conjI) +nitpick [expect = genuine] 1 +nitpick [expect = genuine] 2 +nitpick [expect = genuine] +nitpick [card = 5, expect = genuine] +nitpick [sat_solver = MiniSat, expect = genuine] 2 +oops + +subsection {* Examples and Test Cases *} + +subsubsection {* Propositional logic *} + +lemma "True" +nitpick [expect = none] +apply auto +done + +lemma "False" +nitpick [expect = genuine] +oops + +lemma "P" +nitpick [expect = genuine] +oops + +lemma "\ P" +nitpick [expect = genuine] +oops + +lemma "P \ Q" +nitpick [expect = genuine] +oops + +lemma "P \ Q" +nitpick [expect = genuine] +oops + +lemma "P \ Q" +nitpick [expect = genuine] +oops + +lemma "(P\bool) = Q" +nitpick [expect = genuine] +oops + +lemma "(P \ Q) \ (P \ Q)" +nitpick [expect = genuine] +oops + +subsubsection {* Predicate logic *} + +lemma "P x y z" +nitpick [expect = genuine] +oops + +lemma "P x y \ P y x" +nitpick [expect = genuine] +oops + +lemma "P (f (f x)) \ P x \ P (f x)" +nitpick [expect = genuine] +oops + +subsubsection {* Equality *} + +lemma "P = True" +nitpick [expect = genuine] +oops + +lemma "P = False" +nitpick [expect = genuine] +oops + +lemma "x = y" +nitpick [expect = genuine] +oops + +lemma "f x = g x" +nitpick [expect = genuine] +oops + +lemma "(f\'a\'b) = g" +nitpick [expect = genuine] +oops + +lemma "(f\('d\'d)\('c\'d)) = g" +nitpick [expect = genuine] +oops + +lemma "distinct [a, b]" +nitpick [expect = genuine] +apply simp +nitpick [expect = genuine] +oops + +subsubsection {* First-Order Logic *} + +lemma "\x. P x" +nitpick [expect = genuine] +oops + +lemma "\x. P x" +nitpick [expect = genuine] +oops + +lemma "\!x. P x" +nitpick [expect = genuine] +oops + +lemma "Ex P" +nitpick [expect = genuine] +oops + +lemma "All P" +nitpick [expect = genuine] +oops + +lemma "Ex1 P" +nitpick [expect = genuine] +oops + +lemma "(\x. P x) \ (\x. P x)" +nitpick [expect = genuine] +oops + +lemma "(\x. \y. P x y) \ (\y. \x. P x y)" +nitpick [expect = genuine] +oops + +lemma "(\x. P x) \ (\!x. P x)" +nitpick [expect = genuine] +oops + +text {* A true statement (also testing names of free and bound variables being identical) *} + +lemma "(\x y. P x y \ P y x) \ (\x. P x y) \ P y x" +nitpick [expect = none] +apply fast +done + +text {* "A type has at most 4 elements." *} + +lemma "\ distinct [a, b, c, d, e]" +nitpick [expect = genuine] +apply simp +nitpick [expect = genuine] +oops + +lemma "distinct [a, b, c, d]" +nitpick [expect = genuine] +apply simp +nitpick [expect = genuine] +oops + +text {* "Every reflexive and symmetric relation is transitive." *} + +lemma "\\x. P x x; \x y. P x y \ P y x\ \ P x y \ P y z \ P x z" +nitpick [expect = genuine] +oops + +text {* The "Drinker's theorem" ... *} + +lemma "\x. f x = g x \ f = g" +nitpick [expect = none] +apply (auto simp add: ext) +done + +text {* ... and an incorrect version of it *} + +lemma "(\x. f x = g x) \ f = g" +nitpick [expect = genuine] +oops + +text {* "Every function has a fixed point." *} + +lemma "\x. f x = x" +nitpick [expect = genuine] +oops + +text {* "Function composition is commutative." *} + +lemma "f (g x) = g (f x)" +nitpick [expect = genuine] +oops + +text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *} + +lemma "((P\('a\'b)\bool) f = P g) \ (f x = g x)" +nitpick [expect = genuine] +oops + +subsubsection {* Higher-Order Logic *} + +lemma "\P. P" +nitpick [expect = none] +apply auto +done + +lemma "\P. P" +nitpick [expect = genuine] +oops + +lemma "\!P. P" +nitpick [expect = none] +apply auto +done + +lemma "\!P. P x" +nitpick [expect = genuine] +oops + +lemma "P Q \ Q x" +nitpick [expect = genuine] +oops + +lemma "x \ All" +nitpick [expect = genuine] +oops + +lemma "x \ Ex" +nitpick [expect = genuine] +oops + +lemma "x \ Ex1" +nitpick [expect = genuine] +oops + +text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *} + +constdefs +"trans" :: "('a \ 'a \ bool) \ bool" +"trans P \ (ALL x y z. P x y \ P y z \ P x z)" +"subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" +"subset P Q \ (ALL x y. P x y \ Q x y)" +"trans_closure" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" +"trans_closure P Q \ (subset Q P) \ (trans P) \ (ALL R. subset Q R \ trans R \ subset P R)" + +lemma "trans_closure T P \ (\x y. T x y)" +nitpick [expect = genuine] +oops + +text {* "The union of transitive closures is equal to the transitive closure of unions." *} + +lemma "(\x y. (P x y \ R x y) \ T x y) \ trans T \ (\Q. (\x y. (P x y \ R x y) \ Q x y) \ trans Q \ subset T Q) + \ trans_closure TP P + \ trans_closure TR R + \ (T x y = (TP x y \ TR x y))" +nitpick [expect = genuine] +oops + +text {* "Every surjective function is invertible." *} + +lemma "(\y. \x. y = f x) \ (\g. \x. g (f x) = x)" +nitpick [expect = genuine] +oops + +text {* "Every invertible function is surjective." *} + +lemma "(\g. \x. g (f x) = x) \ (\y. \x. y = f x)" +nitpick [expect = genuine] +oops + +text {* Every point is a fixed point of some function. *} + +lemma "\f. f x = x" +nitpick [card = 1\7, expect = none] +apply (rule_tac x = "\x. x" in exI) +apply simp +done + +text {* Axiom of Choice: first an incorrect version ... *} + +lemma "(\x. \y. P x y) \ (\!f. \x. P x (f x))" +nitpick [expect = genuine] +oops + +text {* ... and now two correct ones *} + +lemma "(\x. \y. P x y) \ (\f. \x. P x (f x))" +nitpick [card = 1-5, expect = none] +apply (simp add: choice) +done + +lemma "(\x. \!y. P x y) \ (\!f. \x. P x (f x))" +nitpick [card = 1-4, expect = none] +apply auto + apply (simp add: ex1_implies_ex choice) +apply (fast intro: ext) +done + +subsubsection {* Metalogic *} + +lemma "\x. P x" +nitpick [expect = genuine] +oops + +lemma "f x \ g x" +nitpick [expect = genuine] +oops + +lemma "P \ Q" +nitpick [expect = genuine] +oops + +lemma "\P; Q; R\ \ S" +nitpick [expect = genuine] +oops + +lemma "(x \ all) \ False" +nitpick [expect = genuine] +oops + +lemma "(x \ (op \)) \ False" +nitpick [expect = genuine] +oops + +lemma "(x \ (op \)) \ False" +nitpick [expect = genuine] +oops + +subsubsection {* Schematic Variables *} + +lemma "?P" +nitpick [expect = none] +apply auto +done + +lemma "x = ?y" +nitpick [expect = none] +apply auto +done + +subsubsection {* Abstractions *} + +lemma "(\x. x) = (\x. y)" +nitpick [expect = genuine] +oops + +lemma "(\f. f x) = (\f. True)" +nitpick [expect = genuine] +oops + +lemma "(\x. x) = (\y. y)" +nitpick [expect = none] +apply simp +done + +subsubsection {* Sets *} + +lemma "P (A\'a set)" +nitpick [expect = genuine] +oops + +lemma "P (A\'a set set)" +nitpick [expect = genuine] +oops + +lemma "{x. P x} = {y. P y}" +nitpick [expect = none] +apply simp +done + +lemma "x \ {x. P x}" +nitpick [expect = genuine] +oops + +lemma "P (op \)" +nitpick [expect = genuine] +oops + +lemma "P (op \ x)" +nitpick [expect = genuine] +oops + +lemma "P Collect" +nitpick [expect = genuine] +oops + +lemma "A Un B = A Int B" +nitpick [expect = genuine] +oops + +lemma "(A Int B) Un C = (A Un C) Int B" +nitpick [expect = genuine] +oops + +lemma "Ball A P \ Bex A P" +nitpick [expect = genuine] +oops + +subsubsection {* @{const undefined} *} + +lemma "undefined" +nitpick [expect = genuine] +oops + +lemma "P undefined" +nitpick [expect = genuine] +oops + +lemma "undefined x" +nitpick [expect = genuine] +oops + +lemma "undefined undefined" +nitpick [expect = genuine] +oops + +subsubsection {* @{const The} *} + +lemma "The P" +nitpick [expect = genuine] +oops + +lemma "P The" +nitpick [expect = genuine] +oops + +lemma "P (The P)" +nitpick [expect = genuine] +oops + +lemma "(THE x. x=y) = z" +nitpick [expect = genuine] +oops + +lemma "Ex P \ P (The P)" +nitpick [expect = genuine] +oops + +subsubsection {* @{const Eps} *} + +lemma "Eps P" +nitpick [expect = genuine] +oops + +lemma "P Eps" +nitpick [expect = genuine] +oops + +lemma "P (Eps P)" +nitpick [expect = genuine] +oops + +lemma "(SOME x. x=y) = z" +nitpick [expect = genuine] +oops + +lemma "Ex P \ P (Eps P)" +nitpick [expect = none] +apply (auto simp add: someI) +done + +subsubsection {* Operations on Natural Numbers *} + +lemma "(x\nat) + y = 0" +nitpick [expect = genuine] +oops + +lemma "(x\nat) = x + x" +nitpick [expect = genuine] +oops + +lemma "(x\nat) - y + y = x" +nitpick [expect = genuine] +oops + +lemma "(x\nat) = x * x" +nitpick [expect = genuine] +oops + +lemma "(x\nat) < x + y" +nitpick [card = 1, expect = genuine] +nitpick [card = 2-5, expect = genuine] +oops + +text {* \ *} + +lemma "P (x\'a\'b)" +nitpick [expect = genuine] +oops + +lemma "\x\'a\'b. P x" +nitpick [expect = genuine] +oops + +lemma "P (x, y)" +nitpick [expect = genuine] +oops + +lemma "P (fst x)" +nitpick [expect = genuine] +oops + +lemma "P (snd x)" +nitpick [expect = genuine] +oops + +lemma "P Pair" +nitpick [expect = genuine] +oops + +lemma "prod_rec f p = f (fst p) (snd p)" +nitpick [expect = none] +by (case_tac p) auto + +lemma "prod_rec f (a, b) = f a b" +nitpick [expect = none] +by auto + +lemma "P (prod_rec f x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Pair a b \ f a b)" +nitpick [expect = genuine] +oops + +subsubsection {* Subtypes (typedef), typedecl *} + +text {* A completely unspecified non-empty subset of @{typ "'a"}: *} + +typedef 'a myTdef = "insert (undefined\'a) (undefined\'a set)" +by auto + +lemma "(x\'a myTdef) = y" +nitpick [expect = genuine] +oops + +typedecl myTdecl + +typedef 'a T_bij = "{(f\'a\'a). \y. \!x. f x = y}" +by auto + +lemma "P (f\(myTdecl myTdef) T_bij)" +nitpick [expect = genuine] +oops + +subsubsection {* Inductive Datatypes *} + +text {* unit *} + +lemma "P (x\unit)" +nitpick [expect = genuine] +oops + +lemma "\x\unit. P x" +nitpick [expect = genuine] +oops + +lemma "P ()" +nitpick [expect = genuine] +oops + +lemma "unit_rec u x = u" +nitpick [expect = none] +apply simp +done + +lemma "P (unit_rec u x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of () \ u)" +nitpick [expect = genuine] +oops + +text {* option *} + +lemma "P (x\'a option)" +nitpick [expect = genuine] +oops + +lemma "\x\'a option. P x" +nitpick [expect = genuine] +oops + +lemma "P None" +nitpick [expect = genuine] +oops + +lemma "P (Some x)" +nitpick [expect = genuine] +oops + +lemma "option_rec n s None = n" +nitpick [expect = none] +apply simp +done + +lemma "option_rec n s (Some x) = s x" +nitpick [expect = none] +apply simp +done + +lemma "P (option_rec n s x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of None \ n | Some u \ s u)" +nitpick [expect = genuine] +oops + +text {* + *} + +lemma "P (x\'a+'b)" +nitpick [expect = genuine] +oops + +lemma "\x\'a+'b. P x" +nitpick [expect = genuine] +oops + +lemma "P (Inl x)" +nitpick [expect = genuine] +oops + +lemma "P (Inr x)" +nitpick [expect = genuine] +oops + +lemma "P Inl" +nitpick [expect = genuine] +oops + +lemma "sum_rec l r (Inl x) = l x" +nitpick [expect = none] +apply simp +done + +lemma "sum_rec l r (Inr x) = r x" +nitpick [expect = none] +apply simp +done + +lemma "P (sum_rec l r x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Inl a \ l a | Inr b \ r b)" +nitpick [expect = genuine] +oops + +text {* Non-recursive datatypes *} + +datatype T1 = A | B + +lemma "P (x\T1)" +nitpick [expect = genuine] +oops + +lemma "\x\T1. P x" +nitpick [expect = genuine] +oops + +lemma "P A" +nitpick [expect = genuine] +oops + +lemma "P B" +nitpick [expect = genuine] +oops + +lemma "T1_rec a b A = a" +nitpick [expect = none] +apply simp +done + +lemma "T1_rec a b B = b" +nitpick [expect = none] +apply simp +done + +lemma "P (T1_rec a b x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of A \ a | B \ b)" +nitpick [expect = genuine] +oops + +datatype 'a T2 = C T1 | D 'a + +lemma "P (x\'a T2)" +nitpick [expect = genuine] +oops + +lemma "\x\'a T2. P x" +nitpick [expect = genuine] +oops + +lemma "P D" +nitpick [expect = genuine] +oops + +lemma "T2_rec c d (C x) = c x" +nitpick [expect = none] +apply simp +done + +lemma "T2_rec c d (D x) = d x" +nitpick [expect = none] +apply simp +done + +lemma "P (T2_rec c d x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of C u \ c u | D v \ d v)" +nitpick [expect = genuine] +oops + +datatype ('a, 'b) T3 = E "'a \ 'b" + +lemma "P (x\('a, 'b) T3)" +nitpick [expect = genuine] +oops + +lemma "\x\('a, 'b) T3. P x" +nitpick [expect = genuine] +oops + +lemma "P E" +nitpick [expect = genuine] +oops + +lemma "T3_rec e (E x) = e x" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "P (T3_rec e x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of E f \ e f)" +nitpick [expect = genuine] +oops + +text {* Recursive datatypes *} + +text {* nat *} + +lemma "P (x\nat)" +nitpick [expect = genuine] +oops + +lemma "\x\nat. P x" +nitpick [expect = genuine] +oops + +lemma "P (Suc 0)" +nitpick [expect = genuine] +oops + +lemma "P Suc" +nitpick [card = 1\7, expect = none] +oops + +lemma "nat_rec zero suc 0 = zero" +nitpick [expect = none] +apply simp +done + +lemma "nat_rec zero suc (Suc x) = suc x (nat_rec zero suc x)" +nitpick [expect = none] +apply simp +done + +lemma "P (nat_rec zero suc x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of 0 \ zero | Suc n \ suc n)" +nitpick [expect = genuine] +oops + +text {* 'a list *} + +lemma "P (xs\'a list)" +nitpick [expect = genuine] +oops + +lemma "\xs\'a list. P xs" +nitpick [expect = genuine] +oops + +lemma "P [x, y]" +nitpick [expect = genuine] +oops + +lemma "list_rec nil cons [] = nil" +nitpick [expect = none] +apply simp +done + +lemma "list_rec nil cons (x#xs) = cons x xs (list_rec nil cons xs)" +nitpick [expect = none] +apply simp +done + +lemma "P (list_rec nil cons xs)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Nil \ nil | Cons a b \ cons a b)" +nitpick [expect = genuine] +oops + +lemma "(xs\'a list) = ys" +nitpick [expect = genuine] +oops + +lemma "a # xs = b # xs" +nitpick [expect = genuine] +oops + +datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList + +lemma "P (x\BitList)" +nitpick [expect = genuine] +oops + +lemma "\x\BitList. P x" +nitpick [expect = genuine] +oops + +lemma "P (Bit0 (Bit1 BitListNil))" +nitpick [expect = genuine] +oops + +lemma "BitList_rec nil bit0 bit1 BitListNil = nil" +nitpick [expect = none] +apply simp +done + +lemma "BitList_rec nil bit0 bit1 (Bit0 xs) = bit0 xs (BitList_rec nil bit0 bit1 xs)" +nitpick [expect = none] +apply simp +done + +lemma "BitList_rec nil bit0 bit1 (Bit1 xs) = bit1 xs (BitList_rec nil bit0 bit1 xs)" +nitpick [expect = none] +apply simp +done + +lemma "P (BitList_rec nil bit0 bit1 x)" +nitpick [expect = genuine] +oops + +datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree" + +lemma "P (x\'a BinTree)" +nitpick [expect = genuine] +oops + +lemma "\x\'a BinTree. P x" +nitpick [expect = genuine] +oops + +lemma "P (Node (Leaf x) (Leaf y))" +nitpick [expect = genuine] +oops + +lemma "BinTree_rec l n (Leaf x) = l x" +nitpick [expect = none] +apply simp +done + +lemma "BinTree_rec l n (Node x y) = n x y (BinTree_rec l n x) (BinTree_rec l n y)" +nitpick [card = 1\6, expect = none] +apply simp +done + +lemma "P (BinTree_rec l n x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Leaf a \ l a | Node a b \ n a b)" +nitpick [expect = genuine] +oops + +text {* Mutually recursive datatypes *} + +datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp" + and 'a bexp = Equal "'a aexp" "'a aexp" + +lemma "P (x\'a aexp)" +nitpick [expect = genuine] +oops + +lemma "\x\'a aexp. P x" +nitpick [expect = genuine] +oops + +lemma "P (ITE (Equal (Number x) (Number y)) (Number x) (Number y))" +nitpick [expect = genuine] +oops + +lemma "P (x\'a bexp)" +nitpick [expect = genuine] +oops + +lemma "\x\'a bexp. P x" +nitpick [expect = genuine] +oops + +lemma "aexp_bexp_rec_1 number ite equal (Number x) = number x" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "aexp_bexp_rec_1 number ite equal (ITE x y z) = ite x y z (aexp_bexp_rec_2 number ite equal x) (aexp_bexp_rec_1 number ite equal y) (aexp_bexp_rec_1 number ite equal z)" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "P (aexp_bexp_rec_1 number ite equal x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Number a \ number a | ITE b a1 a2 \ ite b a1 a2)" +nitpick [expect = genuine] +oops + +lemma "aexp_bexp_rec_2 number ite equal (Equal x y) = equal x y (aexp_bexp_rec_1 number ite equal x) (aexp_bexp_rec_1 number ite equal y)" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "P (aexp_bexp_rec_2 number ite equal x)" +nitpick [expect = genuine] +oops + +lemma "P (case x of Equal a1 a2 \ equal a1 a2)" +nitpick [expect = genuine] +oops + +datatype X = A | B X | C Y + and Y = D X | E Y | F + +lemma "P (x\X)" +nitpick [expect = genuine] +oops + +lemma "P (y\Y)" +nitpick [expect = genuine] +oops + +lemma "P (B (B A))" +nitpick [expect = genuine] +oops + +lemma "P (B (C F))" +nitpick [expect = genuine] +oops + +lemma "P (C (D A))" +nitpick [expect = genuine] +oops + +lemma "P (C (E F))" +nitpick [expect = genuine] +oops + +lemma "P (D (B A))" +nitpick [expect = genuine] +oops + +lemma "P (D (C F))" +nitpick [expect = genuine] +oops + +lemma "P (E (D A))" +nitpick [expect = genuine] +oops + +lemma "P (E (E F))" +nitpick [expect = genuine] +oops + +lemma "P (C (D (C F)))" +nitpick [expect = genuine] +oops + +lemma "X_Y_rec_1 a b c d e f A = a" +nitpick [expect = none] +apply simp +done + +lemma "X_Y_rec_1 a b c d e f (B x) = b x (X_Y_rec_1 a b c d e f x)" +nitpick [expect = none] +apply simp +done + +lemma "X_Y_rec_1 a b c d e f (C y) = c y (X_Y_rec_2 a b c d e f y)" +nitpick [expect = none] +apply simp +done + +lemma "X_Y_rec_2 a b c d e f (D x) = d x (X_Y_rec_1 a b c d e f x)" +nitpick [expect = none] +apply simp +done + +lemma "X_Y_rec_2 a b c d e f (E y) = e y (X_Y_rec_2 a b c d e f y)" +nitpick [expect = none] +apply simp +done + +lemma "X_Y_rec_2 a b c d e f F = f" +nitpick [expect = none] +apply simp +done + +lemma "P (X_Y_rec_1 a b c d e f x)" +nitpick [expect = genuine] +oops + +lemma "P (X_Y_rec_2 a b c d e f y)" +nitpick [expect = genuine] +oops + +text {* Other datatype examples *} + +text {* Indirect recursion is implemented via mutual recursion. *} + +datatype XOpt = CX "XOpt option" | DX "bool \ XOpt option" + +lemma "P (x\XOpt)" +nitpick [expect = genuine] +oops + +lemma "P (CX None)" +nitpick [expect = genuine] +oops + +lemma "P (CX (Some (CX None)))" +nitpick [expect = genuine] +oops + +lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (CX x) = cx x (XOpt_rec_2 cx dx n1 s1 n2 s2 x)" +nitpick [card = 1\6, expect = none] +apply simp +done + +lemma "XOpt_rec_1 cx dx n1 s1 n2 s2 (DX x) = dx x (\b. XOpt_rec_3 cx dx n1 s1 n2 s2 (x b))" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 None = n1" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "XOpt_rec_2 cx dx n1 s1 n2 s2 (Some x) = s1 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 None = n2" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "XOpt_rec_3 cx dx n1 s1 n2 s2 (Some x) = s2 x (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" +nitpick [card = 1\4, expect = none] +apply simp +done + +lemma "P (XOpt_rec_1 cx dx n1 s1 n2 s2 x)" +nitpick [expect = genuine] +oops + +lemma "P (XOpt_rec_2 cx dx n1 s1 n2 s2 x)" +nitpick [expect = genuine] +oops + +lemma "P (XOpt_rec_3 cx dx n1 s1 n2 s2 x)" +nitpick [expect = genuine] +oops + +datatype 'a YOpt = CY "('a \ 'a YOpt) option" + +lemma "P (x\'a YOpt)" +nitpick [expect = genuine] +oops + +lemma "P (CY None)" +nitpick [expect = genuine] +oops + +lemma "P (CY (Some (\a. CY None)))" +nitpick [expect = genuine] +oops + +lemma "YOpt_rec_1 cy n s (CY x) = cy x (YOpt_rec_2 cy n s x)" +nitpick [card = 1\2, expect = none] +apply simp +done + +lemma "YOpt_rec_2 cy n s None = n" +nitpick [card = 1\2, expect = none] +apply simp +done + +lemma "YOpt_rec_2 cy n s (Some x) = s x (\a. YOpt_rec_1 cy n s (x a))" +nitpick [card = 1\2, expect = none] +apply simp +done + +lemma "P (YOpt_rec_1 cy n s x)" +nitpick [expect = genuine] +oops + +lemma "P (YOpt_rec_2 cy n s x)" +nitpick [expect = genuine] +oops + +datatype Trie = TR "Trie list" + +lemma "P (x\Trie)" +nitpick [expect = genuine] +oops + +lemma "\x\Trie. P x" +nitpick [expect = genuine] +oops + +lemma "P (TR [TR []])" +nitpick [expect = genuine] +oops + +lemma "Trie_rec_1 tr nil cons (TR x) = tr x (Trie_rec_2 tr nil cons x)" +nitpick [card = 1\6, expect = none] +apply simp +done + +lemma "Trie_rec_2 tr nil cons [] = nil" +nitpick [card = 1\6, expect = none] +apply simp +done + +lemma "Trie_rec_2 tr nil cons (x#xs) = cons x xs (Trie_rec_1 tr nil cons x) (Trie_rec_2 tr nil cons xs)" +nitpick [card = 1\6, expect = none] +apply simp +done + +lemma "P (Trie_rec_1 tr nil cons x)" +nitpick [card = 1, expect = genuine] +oops + +lemma "P (Trie_rec_2 tr nil cons x)" +nitpick [card = 1, expect = genuine] +oops + +datatype InfTree = Leaf | Node "nat \ InfTree" + +lemma "P (x\InfTree)" +nitpick [expect = genuine] +oops + +lemma "\x\InfTree. P x" +nitpick [expect = genuine] +oops + +lemma "P (Node (\n. Leaf))" +nitpick [expect = genuine] +oops + +lemma "InfTree_rec leaf node Leaf = leaf" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "InfTree_rec leaf node (Node x) = node x (\n. InfTree_rec leaf node (x n))" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "P (InfTree_rec leaf node x)" +nitpick [expect = genuine] +oops + +datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \ 'a lambda" + +lemma "P (x\'a lambda)" +nitpick [expect = genuine] +oops + +lemma "\x\'a lambda. P x" +nitpick [expect = genuine] +oops + +lemma "P (Lam (\a. Var a))" +nitpick [card = 1\5, expect = none] +nitpick [card 'a = 4, card "'a lambda" = 5, expect = genuine] +oops + +lemma "lambda_rec var app lam (Var x) = var x" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "lambda_rec var app lam (App x y) = app x y (lambda_rec var app lam x) (lambda_rec var app lam y)" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "lambda_rec var app lam (Lam x) = lam x (\a. lambda_rec var app lam (x a))" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "P (lambda_rec v a l x)" +nitpick [expect = genuine] +oops + +text {* Taken from "Inductive datatypes in HOL", p. 8: *} + +datatype ('a, 'b) T = C "'a \ bool" | D "'b list" +datatype 'c U = E "('c, 'c U) T" + +lemma "P (x\'c U)" +nitpick [expect = genuine] +oops + +lemma "\x\'c U. P x" +nitpick [expect = genuine] +oops + +lemma "P (E (C (\a. True)))" +nitpick [expect = genuine] +oops + +lemma "U_rec_1 e c d nil cons (E x) = e x (U_rec_2 e c d nil cons x)" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "U_rec_2 e c d nil cons (C x) = c x" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "U_rec_2 e c d nil cons (D x) = d x (U_rec_3 e c d nil cons x)" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "U_rec_3 e c d nil cons [] = nil" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "U_rec_3 e c d nil cons (x#xs) = cons x xs (U_rec_1 e c d nil cons x) (U_rec_3 e c d nil cons xs)" +nitpick [card = 1\3, expect = none] +apply simp +done + +lemma "P (U_rec_1 e c d nil cons x)" +nitpick [expect = genuine] +oops + +lemma "P (U_rec_2 e c d nil cons x)" +nitpick [card = 1, expect = genuine] +oops + +lemma "P (U_rec_3 e c d nil cons x)" +nitpick [card = 1, expect = genuine] +oops + +subsubsection {* Records *} + +record ('a, 'b) point = + xpos :: 'a + ypos :: 'b + +lemma "(x\('a, 'b) point) = y" +nitpick [expect = genuine] +oops + +record ('a, 'b, 'c) extpoint = "('a, 'b) point" + + ext :: 'c + +lemma "(x\('a, 'b, 'c) extpoint) = y" +nitpick [expect = genuine] +oops + +subsubsection {* Inductively Defined Sets *} + +inductive_set undefinedSet :: "'a set" where +"undefined \ undefinedSet" + +lemma "x \ undefinedSet" +nitpick [expect = genuine] +oops + +inductive_set evenCard :: "'a set set" +where +"{} \ evenCard" | +"\S \ evenCard; x \ S; y \ S; x \ y\ \ S \ {x, y} \ evenCard" + +lemma "S \ evenCard" +nitpick [expect = genuine] +oops + +inductive_set +even :: "nat set" +and odd :: "nat set" +where +"0 \ even" | +"n \ even \ Suc n \ odd" | +"n \ odd \ Suc n \ even" + +lemma "n \ odd" +nitpick [expect = genuine] +oops + +consts f :: "'a \ 'a" + +inductive_set a_even :: "'a set" and a_odd :: "'a set" where +"undefined \ a_even" | +"x \ a_even \ f x \ a_odd" | +"x \ a_odd \ f x \ a_even" + +lemma "x \ a_odd" +nitpick [expect = genuine] +oops + +subsubsection {* Examples Involving Special Functions *} + +lemma "card x = 0" +nitpick [expect = genuine] +oops + +lemma "finite x" +nitpick [expect = none] +oops + +lemma "xs @ [] = ys @ []" +nitpick [expect = genuine] +oops + +lemma "xs @ ys = ys @ xs" +nitpick [expect = genuine] +oops + +lemma "f (lfp f) = lfp f" +nitpick [expect = genuine] +oops + +lemma "f (gfp f) = gfp f" +nitpick [expect = genuine] +oops + +lemma "lfp f = gfp f" +nitpick [expect = genuine] +oops + +subsubsection {* Axiomatic Type Classes and Overloading *} + +text {* A type class without axioms: *} + +axclass classA + +lemma "P (x\'a\classA)" +nitpick [expect = genuine] +oops + +text {* The axiom of this type class does not contain any type variables: *} + +axclass classB +classB_ax: "P \ \ P" + +lemma "P (x\'a\classB)" +nitpick [expect = genuine] +oops + +text {* An axiom with a type variable (denoting types which have at least two elements): *} + +axclass classC < type +classC_ax: "\x y. x \ y" + +lemma "P (x\'a\classC)" +nitpick [expect = genuine] +oops + +lemma "\x y. (x\'a\classC) \ y" +nitpick [expect = none] +sorry + +text {* A type class for which a constant is defined: *} + +consts +classD_const :: "'a \ 'a" + +axclass classD < type +classD_ax: "classD_const (classD_const x) = classD_const x" + +lemma "P (x\'a\classD)" +nitpick [expect = genuine] +oops + +text {* A type class with multiple superclasses: *} + +axclass classE < classC, classD + +lemma "P (x\'a\classE)" +nitpick [expect = genuine] +oops + +lemma "P (x\'a\{classB, classE})" +nitpick [expect = genuine] +oops + +text {* OFCLASS: *} + +lemma "OFCLASS('a\type, type_class)" +nitpick [expect = none] +apply intro_classes +done + +lemma "OFCLASS('a\classC, type_class)" +nitpick [expect = none] +apply intro_classes +done + +lemma "OFCLASS('a, classB_class)" +nitpick [expect = none] +apply intro_classes +apply simp +done + +lemma "OFCLASS('a\type, classC_class)" +nitpick [expect = genuine] +oops + +text {* Overloading: *} + +consts inverse :: "'a \ 'a" + +defs (overloaded) +inverse_bool: "inverse (b\bool) \ \ b" +inverse_set: "inverse (S\'a set) \ -S" +inverse_pair: "inverse p \ (inverse (fst p), inverse (snd p))" + +lemma "inverse b" +nitpick [expect = genuine] +oops + +lemma "P (inverse (S\'a set))" +nitpick [expect = genuine] +oops + +lemma "P (inverse (p\'a\'b))" +nitpick [expect = genuine] +oops + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Special_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Special_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,172 @@ +(* Title: HOL/Nitpick_Examples/Special_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick's "specialize" optimization. +*) + +header {* Examples Featuring Nitpick's \textit{specialize} Optimization *} + +theory Special_Nits +imports Main +begin + +nitpick_params [card = 4, debug, show_consts, timeout = 10 s] + +fun f1 :: "nat \ nat \ nat \ nat \ nat \ nat" where +"f1 a b c d e = a + b + c + d + e" + +lemma "f1 0 0 0 0 0 = f1 0 0 0 0 (1 - 1)" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "f1 u v w x y = f1 y x w v u" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +fun f2 :: "nat \ nat \ nat \ nat \ nat \ nat" where +"f2 a b c d (Suc e) = a + b + c + d + e" + +lemma "f2 0 0 0 0 0 = f2 (1 - 1) 0 0 0 0" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "f2 0 (v - v) 0 (x - x) 0 = f2 (u - u) 0 (w - w) 0 (y - y)" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "f2 1 0 0 0 0 = f2 0 1 0 0 0" +nitpick [expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +lemma "f2 0 0 0 0 0 = f2 0 0 0 0 0" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +fun f3 :: "nat \ nat \ nat \ nat \ nat \ nat" where +"f3 (Suc a) b 0 d (Suc e) = a + b + d + e" | +"f3 0 b 0 d 0 = b + d" + +lemma "f3 a b c d e = f3 e d c b a" +nitpick [expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +lemma "f3 a b c d a = f3 a d c d a" +nitpick [expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +lemma "\c < 1; a \ e; e \ a\ \ f3 a b c d a = f3 e d c b e" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "(\u. a = u \ f3 a a a a a = f3 u u u u u) + \ (\u. b = u \ f3 b b u b b = f3 u u b u u)" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +nitpick [dont_skolemize, expect = none] +nitpick [dont_specialize, dont_skolemize, expect = none] +sorry + +function f4 :: "nat \ nat \ nat" where +"f4 x x = 1" | +"f4 y z = (if y = z then 1 else 0)" +by auto +termination by lexicographic_order + +lemma "f4 a b = f4 b a" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "f4 a (Suc a) = f4 a a" +nitpick [expect = genuine] +nitpick [dont_specialize, expect = genuine] +oops + +fun f5 :: "(nat \ nat) \ nat \ nat" where +"f5 f (Suc a) = f a" + +lemma "\one \ {1}. \two \ {2}. + f5 (\a. if a = one then 1 else if a = two then 2 else a) (Suc x) = x" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "\two \ {2}. \one \ {1}. + f5 (\a. if a = one then 1 else if a = two then 2 else a) (Suc x) = x" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "\one \ {1}. \two \ {2}. + f5 (\a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x" +nitpick [expect = genuine] +sorry + +lemma "\two \ {2}. \one \ {1}. + f5 (\a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x" +nitpick [expect = genuine] +sorry + +lemma "\a. g a = a + \ \one \ {1}. \two \ {2}. f5 g x = + f5 (\a. if a = one then 1 else if a = two then 2 else a) x" +nitpick [expect = none] +nitpick [dont_specialize, expect = none] +sorry + +lemma "\a. g a = a + \ \one \ {2}. \two \ {1}. f5 g x = + f5 (\a. if a = one then 1 else if a = two then 2 else a) x" +nitpick [expect = potential] +nitpick [dont_specialize, expect = potential] +sorry + +lemma "\a. g a = a + \ \b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\nat). + b\<^isub>1 < b\<^isub>11 \ f5 g x = f5 (\a. if b\<^isub>1 < b\<^isub>11 then a else h b\<^isub>2) x" +nitpick [expect = potential] +nitpick [dont_specialize, expect = none] +nitpick [dont_box, expect = none] +nitpick [dont_box, dont_specialize, expect = none] +sorry + +lemma "\a. g a = a + \ \b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\nat). + b\<^isub>1 < b\<^isub>11 + \ f5 g x = f5 (\a. if b\<^isub>1 < b\<^isub>11 then + a + else + h b\<^isub>2 + h b\<^isub>3 + h b\<^isub>4 + h b\<^isub>5 + h b\<^isub>6 + h b\<^isub>7 + h b\<^isub>8 + + h b\<^isub>9 + h b\<^isub>10) x" +nitpick [card nat = 2, card 'a = 1, expect = none] +nitpick [card nat = 2, card 'a = 1, dont_box, expect = none] +nitpick [card nat = 2, card 'a = 1, dont_specialize, expect = none] +nitpick [card nat = 2, card 'a = 1, dont_box, dont_specialize, expect = none] +sorry + +lemma "\a. g a = a + \ \b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\nat). + b\<^isub>1 < b\<^isub>11 + \ f5 g x = f5 (\a. if b\<^isub>1 \ b\<^isub>11 then + a + else + h b\<^isub>2 + h b\<^isub>3 + h b\<^isub>4 + h b\<^isub>5 + h b\<^isub>6 + h b\<^isub>7 + h b\<^isub>8 + + h b\<^isub>9 + h b\<^isub>10) x" +nitpick [card nat = 2, card 'a = 1, expect = none] +nitpick [card nat = 2, card 'a = 1, dont_box, expect = potential] +nitpick [card nat = 2, card 'a = 1, dont_specialize, expect = potential] +nitpick [card nat = 2, card 'a = 1, dont_box, dont_specialize, + expect = potential] +oops + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Tests_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Tests_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,16 @@ +(* Title: HOL/Nitpick_Examples/Tests_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Nitpick tests. +*) + +header {* Nitpick Tests *} + +theory Tests_Nits +imports Main +begin + +ML {* NitpickTests.run_all_tests () *} + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Nitpick_Examples/Typedef_Nits.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,187 @@ +(* Title: HOL/Nitpick_Examples/Typedef_Nits.thy + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Examples featuring Nitpick applied to typedefs. +*) + +header {* Examples Featuring Nitpick Applied to Typedefs *} + +theory Typedef_Nits +imports Main Rational +begin + +nitpick_params [card = 1\4, timeout = 5 s] + +typedef three = "{0\nat, 1, 2}" +by blast + +definition A :: three where "A \ Abs_three 0" +definition B :: three where "B \ Abs_three 1" +definition C :: three where "C \ Abs_three 2" + +lemma "x = (y\three)" +nitpick [expect = genuine] +oops + +typedef 'a one_or_two = "{undefined False\'a, undefined True}" +by auto + +lemma "x = (y\unit one_or_two)" +nitpick [expect = none] +sorry + +lemma "x = (y\bool one_or_two)" +nitpick [expect = genuine] +oops + +lemma "undefined False \ undefined True \ x = (y\bool one_or_two)" +nitpick [expect = none] +sorry + +lemma "undefined False \ undefined True \ \x (y\bool one_or_two). x \ y" +nitpick [card = 1, expect = potential] (* unfortunate *) +oops + +lemma "\x (y\bool one_or_two). x \ y" +nitpick [card = 1, expect = potential] (* unfortunate *) +nitpick [card = 2, expect = none] +oops + +typedef 'a bounded = + "{n\nat. finite (UNIV\'a \ bool) \ n < card (UNIV\'a \ bool)}" +apply (rule_tac x = 0 in exI) +apply (case_tac "card UNIV = 0") +by auto + +lemma "x = (y\unit bounded)" +nitpick [expect = none] +sorry + +lemma "x = (y\bool bounded)" +nitpick [expect = genuine] +oops + +lemma "x \ (y\bool bounded) \ z = x \ z = y" +nitpick [expect = none] +sorry + +lemma "x \ (y\(bool \ bool) bounded) \ z = x \ z = y" +nitpick [card = 1\5, timeout = 10 s, expect = genuine] +oops + +lemma "True \ ((\x\bool. x) = (\x. x))" +nitpick [expect = none] +by (rule True_def) + +lemma "False \ \P. P" +nitpick [expect = none] +by (rule False_def) + +lemma "() = Abs_unit True" +nitpick [expect = none] +by (rule Unity_def) + +lemma "() = Abs_unit False" +nitpick [expect = none] +by simp + +lemma "Rep_unit () = True" +nitpick [expect = none] +by (insert Rep_unit) (simp add: unit_def) + +lemma "Rep_unit () = False" +nitpick [expect = genuine] +oops + +lemma "Pair a b \ Abs_Prod (Pair_Rep a b)" +nitpick [card = 1\2, expect = none] +by (rule Pair_def) + +lemma "Pair a b \ Abs_Prod (Pair_Rep b a)" +nitpick [card = 1\2, expect = none] +nitpick [dont_box, expect = genuine] +oops + +lemma "fst (Abs_Prod (Pair_Rep a b)) = a" +nitpick [card = 2, expect = none] +by (simp add: Pair_def [THEN symmetric]) + +lemma "fst (Abs_Prod (Pair_Rep a b)) = b" +nitpick [card = 1\2, expect = none] +nitpick [dont_box, expect = genuine] +oops + +lemma "a \ a' \ Pair_Rep a b \ Pair_Rep a' b" +nitpick [expect = none] +apply (rule ccontr) +apply simp +apply (drule subst [where P = "\r. Abs_Prod r = Abs_Prod (Pair_Rep a b)"]) + apply (rule refl) +by (simp add: Pair_def [THEN symmetric]) + +lemma "Abs_Prod (Rep_Prod a) = a" +nitpick [card = 2, expect = none] +by (rule Rep_Prod_inverse) + +lemma "Inl \ \a. Abs_Sum (Inl_Rep a)" +nitpick [card = 1, expect = none] +by (rule Inl_def) + +lemma "Inl \ \a. Abs_Sum (Inr_Rep a)" +nitpick [card = 1, card "'a + 'a" = 2, expect = genuine] +oops + +lemma "Inl_Rep a \ Inr_Rep a" +nitpick [expect = none] +by (rule Inl_Rep_not_Inr_Rep) + +lemma "Abs_Sum (Rep_Sum a) = a" +nitpick [card = 1\2, timeout = 30 s, expect = none] +by (rule Rep_Sum_inverse) + +lemma "0::nat \ Abs_Nat Zero_Rep" +nitpick [expect = none] +by (rule Zero_nat_def_raw) + +lemma "Suc \ \n. Abs_Nat (Suc_Rep (Rep_Nat n))" +nitpick [expect = none] +by (rule Suc_def) + +lemma "Suc \ \n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))" +nitpick [expect = genuine] +oops + +lemma "Abs_Nat (Rep_Nat a) = a" +nitpick [expect = none] +by (rule Rep_Nat_inverse) + +lemma "0 \ Abs_Integ (intrel `` {(0, 0)})" +nitpick [card = 1, timeout = 30 s, max_potential = 0, expect = none] +by (rule Zero_int_def_raw) + +lemma "Abs_Integ (Rep_Integ a) = a" +nitpick [card = 1, timeout = 30 s, max_potential = 0, expect = none] +by (rule Rep_Integ_inverse) + +lemma "Abs_list (Rep_list a) = a" +nitpick [card = 1\2, timeout = 30 s, expect = none] +by (rule Rep_list_inverse) + +record point = + Xcoord :: int + Ycoord :: int + +lemma "Abs_point_ext_type (Rep_point_ext_type a) = a" +nitpick [expect = unknown] +by (rule Rep_point_ext_type_inverse) + +lemma "Fract a b = of_int a / of_int b" +nitpick [card = 1\2, expect = unknown] +by (rule Fract_of_int_quotient) + +lemma "Abs_Rat (Rep_Rat a) = a" +nitpick [card = 1\2, expect = unknown] +by (rule Rep_Rat_inverse) + +end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Rational.thy --- a/src/HOL/Rational.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/Rational.thy Mon Oct 26 14:57:49 2009 +0100 @@ -1063,4 +1063,23 @@ fun rat_of_int i = (i, 1); *} +setup {* +Nitpick.register_frac_type @{type_name rat} + [(@{const_name zero_rat_inst.zero_rat}, @{const_name Nitpick.zero_frac}), + (@{const_name one_rat_inst.one_rat}, @{const_name Nitpick.one_frac}), + (@{const_name plus_rat_inst.plus_rat}, @{const_name Nitpick.plus_frac}), + (@{const_name times_rat_inst.times_rat}, @{const_name Nitpick.times_frac}), + (@{const_name uminus_rat_inst.uminus_rat}, @{const_name Nitpick.uminus_frac}), + (@{const_name number_rat_inst.number_of_rat}, @{const_name Nitpick.number_of_frac}), + (@{const_name inverse_rat_inst.inverse_rat}, @{const_name Nitpick.inverse_frac}), + (@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}), + (@{const_name field_char_0_class.of_rat}, @{const_name Nitpick.of_frac}), + (@{const_name field_char_0_class.Rats}, @{const_name UNIV})] +*} + +lemmas [nitpick_def] = inverse_rat_inst.inverse_rat + number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_eq_rat + plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat + zero_rat_inst.zero_rat + end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Mon Oct 26 14:54:43 2009 +0100 +++ b/src/HOL/RealDef.thy Mon Oct 26 14:57:49 2009 +0100 @@ -1185,4 +1185,22 @@ fun real_of_int i = (i, 1); *} +setup {* +Nitpick.register_frac_type @{type_name real} + [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}), + (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}), + (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}), + (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}), + (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}), + (@{const_name number_real_inst.number_of_real}, @{const_name Nitpick.number_of_frac}), + (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}), + (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})] +*} + +lemmas [nitpick_def] = inverse_real_inst.inverse_real + number_real_inst.number_of_real one_real_inst.one_real + ord_real_inst.less_eq_real plus_real_inst.plus_real + times_real_inst.times_real uminus_real_inst.uminus_real + zero_real_inst.zero_real + end diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/HISTORY --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/HISTORY Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,155 @@ +Version 2010 + + * Moved into Isabelle/HOL "Main" + * Renamed "nitpick_const_def" to "nitpick_def", "nitpick_const_simp" to + "nitpick_simp", "nitpick_const_psimp" to "nitpick_psimp", and + "nitpick_ind_intro" to "nitpick_intro" + * Replaced "special_depth" and "skolemize_depth" options by "specialize" + and "skolemize" + * Fixed monotonicity check + +Version 1.2.2 (16 Oct 2009) + + * Added and implemented "star_linear_preds" option + * Added and implemented "format" option + * Added and implemented "coalesce_type_vars" option + * Added and implemented "max_genuine" option + * Fixed soundness issues related to "set", "distinct", "image", "Sigma", + "-1::nat", subset, constructors, sort axioms, and partially applied + interpreted constants + * Fixed error in "show_consts" for boxed specialized constants + * Fixed errors in Kodkod encoding of "The", "Eps", and "ind" + * Fixed display of Skolem constants + * Fixed error in "check_potential" and "check_genuine" + * Added boxing support for higher-order constructor parameters + * Changed notation used for coinductive datatypes + * Optimized Kodkod encoding of "If", "card", and "setsum" + * Improved the monotonicity check + +Version 1.2.1 (25 Sep 2009) + + * Added explicit support for coinductive datatypes + * Added and implemented "box" option + * Added and implemented "fast_descrs" option + * Added and implemented "uncurry" option + * Renamed and generalized "sync_cards" and "inductive_mood" to "mono" and "wf" + * Fixed soundness issue related to nullary constructors + * Fixed soundness issue related to higher-order quantifiers + * Fixed soundness issue related to the "destroy_constrs" optimization + * Fixed soundness issues related to the "special_depth" optimization + * Added support for PicoSAT and incorporated it with the Nitpick package + * Added automatic detection of installed SAT solvers based on naming + convention + * Optimized handling of quantifiers by moving them inward whenever possible + * Optimized and improved precision of "wf" and "wfrec" + * Improved handling of definitions made in locales + * Fixed checked scope count in message shown upon interruption and timeout + * Added minimalistic Nitpick-like tool called Minipick + +Version 1.2.0 (27 Jul 2009) + + * Optimized Kodkod encoding of datatypes and records + * Optimized coinductive definitions + * Fixed soundness issues related to pairs of functions + * Fixed soundness issue in the peephole optimizer + * Improved precision of non-well-founded predicates occurring positively in + the formula to falsify + * Added and implemented "destroy_constrs" option + * Changed semantics of "inductive_mood" option to ensure soundness + * Fixed semantics of "lockstep" option (broken in 1.1.1) and renamed it + "sync_cards" + * Improved precision of "trancl" and "rtrancl" + * Optimized Kodkod encoding of "tranclp" and "rtranclp" + * Made detection of inconsistent scope specifications more robust + * Fixed a few Kodkod generation bugs that resulted in exceptions + +Version 1.1.1 (24 Jun 2009) + + * Added "show_all" option + * Fixed soundness issues related to type classes + * Improved precision of some set constructs + * Fiddled with the automatic monotonicity check + * Fixed performance issues related to scope enumeration + * Fixed a few Kodkod generation bugs that resulted in exceptions or stack + overflows + +Version 1.1.0 (16 Jun 2009) + + * Redesigned handling of datatypes to provide an interface baded on "card" and + "max", obsoleting "mult" + * Redesigned handling of typedefs, "rat", and "real" + * Made "lockstep" option a three-state option and added an automatic + monotonicity check + * Made "batch_size" a (n + 1)-state option whose default depends on whether + "debug" is enabled + * Made "debug" automatically enable "verbose" + * Added "destroy_equals" option + * Added "no_assms" option + * Fixed bug in computation of ground type + * Fixed performance issue related to datatype acyclicity constraint generation + * Fixed performance issue related to axiom selection + * Improved precision of some well-founded inductive predicates + * Added more checks to guard against very large cardinalities + * Improved hit rate of potential counterexamples + * Fixed several soundness issues + * Optimized the Kodkod encoding to benefit more from symmetry breaking + * Optimized relational composition, cartesian product, and converse + * Added support for HaifaSat + +Version 1.0.3 (17 Mar 2009) + + * Added "HOL-Nominal-Nitpick" as a target in addition to "HOL-Nitpick" + * Added "overlord" option to assist debugging + * Increased default value of "special_depth" option + * Fixed a bug that effectively disabled the "nitpick_const_def" attribute + * Ensured that no scopes are skipped when multithreading is enabled + * Fixed soundness issue in handling of "The", "Eps", and partial functions + defined using Isabelle's function package + * Fixed soundness issue in handling of non-definitional axioms + * Fixed soundness issue in handling of "Abs_" and "Rep_" functions for "unit", + "nat", "int", and "*" + * Fixed a few Kodkod generation bugs that resulted in exceptions + * Optimized "div", "mod", and "dvd" on "nat" and "int" + +Version 1.0.2 (12 Mar 2009) + + * Added support for non-definitional axioms + * Improved Isar integration + * Added support for multiplicities of 0 + * Added "max_threads" option and support for multithreaded Kodkodi + * Added "blocking" option to control whether Nitpick should be run + synchronously or asynchronously + * Merged "auto_timeout" and "wellfounded_timeout" into "tac_timeout" + * Added "auto" option to run Nitpick automatically (like Auto Quickcheck) + * Introduced "auto_timeout" to specify Auto Nitpick's time limit + * Renamed the possible values for the "expect" option + * Renamed the "peephole" option to "peephole_optim" + * Added negative versions of all Boolean options and made "= true" optional + * Altered order of automatic SAT solver selection + +Version 1.0.1 (6 Mar 2009) + + * Eliminated the need to import "Nitpick" to use "nitpick" + * Added "debug" option + * Replaced "specialize_funs" with the more general "special_depth" option + * Renamed "watch" option to "eval" + * Improved parsing of "card", "mult", and "iter" options + * Fixed a soundness bug in the "specialize_funs" optimization + * Increased the scope of the "specialize_funs" optimization + * Fixed a soundness bug in the treatment of "<" and "<=" on type "int" + * Fixed a soundness bug in the "subterm property" optimization for types of + cardinality 1 + * Improved the axiom selection for overloaded constants, which led to an + infinite loop for "Nominal.perm" + * Added support for multiple instantiations of non-well-founded inductive + predicates, which previously raised an exception + * Fixed a Kodkod generation bug that resulted in an exception + * Altered order of scope enumeration and automatic SAT solver selection + * Optimized "Eps", "nat_case", and "list_case" + * Improved output formatting + * Added checks to prevent infinite loops in axiom selector and constant + unfolding + +Version 1.0.0 (27 Feb 2009) + + * First release diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/kodkod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/kodkod.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,1087 @@ +(* Title: HOL/Nitpick/Tools/kodkod.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +ML interface on top of Kodkod. +*) + +signature KODKOD = +sig + type n_ary_index = int * int + type setting = string * string + + datatype tuple = + Tuple of int list | + TupleIndex of n_ary_index | + TupleReg of n_ary_index + + datatype tuple_set = + TupleUnion of tuple_set * tuple_set | + TupleDifference of tuple_set * tuple_set | + TupleIntersect of tuple_set * tuple_set | + TupleProduct of tuple_set * tuple_set | + TupleProject of tuple_set * int | + TupleSet of tuple list | + TupleRange of tuple * tuple | + TupleArea of tuple * tuple | + TupleAtomSeq of int * int | + TupleSetReg of n_ary_index + + datatype tuple_assign = + AssignTuple of n_ary_index * tuple | + AssignTupleSet of n_ary_index * tuple_set + + type bound = (n_ary_index * string) list * tuple_set list + type int_bound = int option * tuple_set list + + datatype formula = + All of decl list * formula | + Exist of decl list * formula | + FormulaLet of expr_assign list * formula | + FormulaIf of formula * formula * formula | + Or of formula * formula | + Iff of formula * formula | + Implies of formula * formula | + And of formula * formula | + Not of formula | + Acyclic of n_ary_index | + Function of n_ary_index * rel_expr * rel_expr | + Functional of n_ary_index * rel_expr * rel_expr | + TotalOrdering of n_ary_index * n_ary_index * n_ary_index * n_ary_index | + Subset of rel_expr * rel_expr | + RelEq of rel_expr * rel_expr | + IntEq of int_expr * int_expr | + LT of int_expr * int_expr | + LE of int_expr * int_expr | + No of rel_expr | + Lone of rel_expr | + One of rel_expr | + Some of rel_expr | + False | + True | + FormulaReg of int + and rel_expr = + RelLet of expr_assign list * rel_expr | + RelIf of formula * rel_expr * rel_expr | + Union of rel_expr * rel_expr | + Difference of rel_expr * rel_expr | + Override of rel_expr * rel_expr | + Intersect of rel_expr * rel_expr | + Product of rel_expr * rel_expr | + IfNo of rel_expr * rel_expr | + Project of rel_expr * int_expr list | + Join of rel_expr * rel_expr | + Closure of rel_expr | + ReflexiveClosure of rel_expr | + Transpose of rel_expr | + Comprehension of decl list * formula | + Bits of int_expr | + Int of int_expr | + Iden | + Ints | + None | + Univ | + Atom of int | + AtomSeq of int * int | + Rel of n_ary_index | + Var of n_ary_index | + RelReg of n_ary_index + and int_expr = + Sum of decl list * int_expr | + IntLet of expr_assign list * int_expr | + IntIf of formula * int_expr * int_expr | + SHL of int_expr * int_expr | + SHA of int_expr * int_expr | + SHR of int_expr * int_expr | + Add of int_expr * int_expr | + Sub of int_expr * int_expr | + Mult of int_expr * int_expr | + Div of int_expr * int_expr | + Mod of int_expr * int_expr | + Cardinality of rel_expr | + SetSum of rel_expr | + BitOr of int_expr * int_expr | + BitXor of int_expr * int_expr | + BitAnd of int_expr * int_expr | + BitNot of int_expr | + Neg of int_expr | + Absolute of int_expr | + Signum of int_expr | + Num of int | + IntReg of int + and decl = + DeclNo of n_ary_index * rel_expr | + DeclLone of n_ary_index * rel_expr | + DeclOne of n_ary_index * rel_expr | + DeclSome of n_ary_index * rel_expr | + DeclSet of n_ary_index * rel_expr + and expr_assign = + AssignFormulaReg of int * formula | + AssignRelReg of n_ary_index * rel_expr | + AssignIntReg of int * int_expr + + type 'a fold_expr_funcs = { + formula_func: formula -> 'a -> 'a, + rel_expr_func: rel_expr -> 'a -> 'a, + int_expr_func: int_expr -> 'a -> 'a + } + + val fold_formula : 'a fold_expr_funcs -> formula -> 'a -> 'a + val fold_rel_expr : 'a fold_expr_funcs -> rel_expr -> 'a -> 'a + val fold_int_expr : 'a fold_expr_funcs -> int_expr -> 'a -> 'a + val fold_decl : 'a fold_expr_funcs -> decl -> 'a -> 'a + val fold_expr_assign : 'a fold_expr_funcs -> expr_assign -> 'a -> 'a + + type 'a fold_tuple_funcs = { + tuple_func: tuple -> 'a -> 'a, + tuple_set_func: tuple_set -> 'a -> 'a + } + + val fold_tuple : 'a fold_tuple_funcs -> tuple -> 'a -> 'a + val fold_tuple_set : 'a fold_tuple_funcs -> tuple_set -> 'a -> 'a + val fold_tuple_assign : 'a fold_tuple_funcs -> tuple_assign -> 'a -> 'a + val fold_bound : + 'a fold_expr_funcs -> 'a fold_tuple_funcs -> bound -> 'a -> 'a + val fold_int_bound : 'a fold_tuple_funcs -> int_bound -> 'a -> 'a + + type problem = { + comment: string, + settings: setting list, + univ_card: int, + tuple_assigns: tuple_assign list, + bounds: bound list, + int_bounds: int_bound list, + expr_assigns: expr_assign list, + formula: formula} + + type raw_bound = n_ary_index * int list list + + datatype outcome = + Normal of (int * raw_bound list) list * int list | + TimedOut of int list | + Interrupted of int list option | + Error of string * int list + + exception SYNTAX of string * string + + val max_arity : int -> int + val arity_of_rel_expr : rel_expr -> int + val problems_equivalent : problem -> problem -> bool + val solve_any_problem : + bool -> Time.time option -> int -> int -> problem list -> outcome +end; + +structure Kodkod : KODKOD = +struct + +type n_ary_index = int * int + +type setting = string * string + +datatype tuple = + Tuple of int list | + TupleIndex of n_ary_index | + TupleReg of n_ary_index + +datatype tuple_set = + TupleUnion of tuple_set * tuple_set | + TupleDifference of tuple_set * tuple_set | + TupleIntersect of tuple_set * tuple_set | + TupleProduct of tuple_set * tuple_set | + TupleProject of tuple_set * int | + TupleSet of tuple list | + TupleRange of tuple * tuple | + TupleArea of tuple * tuple | + TupleAtomSeq of int * int | + TupleSetReg of n_ary_index + +datatype tuple_assign = + AssignTuple of n_ary_index * tuple | + AssignTupleSet of n_ary_index * tuple_set + +type bound = (n_ary_index * string) list * tuple_set list +type int_bound = int option * tuple_set list + +datatype formula = + All of decl list * formula | + Exist of decl list * formula | + FormulaLet of expr_assign list * formula | + FormulaIf of formula * formula * formula | + Or of formula * formula | + Iff of formula * formula | + Implies of formula * formula | + And of formula * formula | + Not of formula | + Acyclic of n_ary_index | + Function of n_ary_index * rel_expr * rel_expr | + Functional of n_ary_index * rel_expr * rel_expr | + TotalOrdering of n_ary_index * n_ary_index * n_ary_index * n_ary_index | + Subset of rel_expr * rel_expr | + RelEq of rel_expr * rel_expr | + IntEq of int_expr * int_expr | + LT of int_expr * int_expr | + LE of int_expr * int_expr | + No of rel_expr | + Lone of rel_expr | + One of rel_expr | + Some of rel_expr | + False | + True | + FormulaReg of int +and rel_expr = + RelLet of expr_assign list * rel_expr | + RelIf of formula * rel_expr * rel_expr | + Union of rel_expr * rel_expr | + Difference of rel_expr * rel_expr | + Override of rel_expr * rel_expr | + Intersect of rel_expr * rel_expr | + Product of rel_expr * rel_expr | + IfNo of rel_expr * rel_expr | + Project of rel_expr * int_expr list | + Join of rel_expr * rel_expr | + Closure of rel_expr | + ReflexiveClosure of rel_expr | + Transpose of rel_expr | + Comprehension of decl list * formula | + Bits of int_expr | + Int of int_expr | + Iden | + Ints | + None | + Univ | + Atom of int | + AtomSeq of int * int | + Rel of n_ary_index | + Var of n_ary_index | + RelReg of n_ary_index +and int_expr = + Sum of decl list * int_expr | + IntLet of expr_assign list * int_expr | + IntIf of formula * int_expr * int_expr | + SHL of int_expr * int_expr | + SHA of int_expr * int_expr | + SHR of int_expr * int_expr | + Add of int_expr * int_expr | + Sub of int_expr * int_expr | + Mult of int_expr * int_expr | + Div of int_expr * int_expr | + Mod of int_expr * int_expr | + Cardinality of rel_expr | + SetSum of rel_expr | + BitOr of int_expr * int_expr | + BitXor of int_expr * int_expr | + BitAnd of int_expr * int_expr | + BitNot of int_expr | + Neg of int_expr | + Absolute of int_expr | + Signum of int_expr | + Num of int | + IntReg of int +and decl = + DeclNo of n_ary_index * rel_expr | + DeclLone of n_ary_index * rel_expr | + DeclOne of n_ary_index * rel_expr | + DeclSome of n_ary_index * rel_expr | + DeclSet of n_ary_index * rel_expr +and expr_assign = + AssignFormulaReg of int * formula | + AssignRelReg of n_ary_index * rel_expr | + AssignIntReg of int * int_expr + +type problem = { + comment: string, + settings: setting list, + univ_card: int, + tuple_assigns: tuple_assign list, + bounds: bound list, + int_bounds: int_bound list, + expr_assigns: expr_assign list, + formula: formula} + +type raw_bound = n_ary_index * int list list + +datatype outcome = + Normal of (int * raw_bound list) list * int list | + TimedOut of int list | + Interrupted of int list option | + Error of string * int list + +exception SYNTAX of string * string + +type 'a fold_expr_funcs = { + formula_func: formula -> 'a -> 'a, + rel_expr_func: rel_expr -> 'a -> 'a, + int_expr_func: int_expr -> 'a -> 'a +} + +(* 'a fold_expr_funcs -> formula -> 'a -> 'a *) +fun fold_formula (F : 'a fold_expr_funcs) formula = + case formula of + All (ds, f) => fold (fold_decl F) ds #> fold_formula F f + | Exist (ds, f) => fold (fold_decl F) ds #> fold_formula F f + | FormulaLet (bs, f) => fold (fold_expr_assign F) bs #> fold_formula F f + | FormulaIf (f, f1, f2) => + fold_formula F f #> fold_formula F f1 #> fold_formula F f2 + | Or (f1, f2) => fold_formula F f1 #> fold_formula F f2 + | Iff (f1, f2) => fold_formula F f1 #> fold_formula F f2 + | Implies (f1, f2) => fold_formula F f1 #> fold_formula F f2 + | And (f1, f2) => fold_formula F f1 #> fold_formula F f2 + | Not f => fold_formula F f + | Acyclic x => fold_rel_expr F (Rel x) + | Function (x, r1, r2) => + fold_rel_expr F (Rel x) #> fold_rel_expr F r1 #> fold_rel_expr F r2 + | Functional (x, r1, r2) => + fold_rel_expr F (Rel x) #> fold_rel_expr F r1 #> fold_rel_expr F r2 + | TotalOrdering (x1, x2, x3, x4) => + fold_rel_expr F (Rel x1) #> fold_rel_expr F (Rel x2) + #> fold_rel_expr F (Rel x3) #> fold_rel_expr F (Rel x4) + | Subset (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | RelEq (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | IntEq (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | LT (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | LE (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | No r => fold_rel_expr F r + | Lone r => fold_rel_expr F r + | One r => fold_rel_expr F r + | Some r => fold_rel_expr F r + | False => #formula_func F formula + | True => #formula_func F formula + | FormulaReg _ => #formula_func F formula +(* 'a fold_expr_funcs -> rel_expr -> 'a -> 'a *) +and fold_rel_expr F rel_expr = + case rel_expr of + RelLet (bs, r) => fold (fold_expr_assign F) bs #> fold_rel_expr F r + | RelIf (f, r1, r2) => + fold_formula F f #> fold_rel_expr F r1 #> fold_rel_expr F r2 + | Union (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Difference (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Override (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Intersect (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Product (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | IfNo (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Project (r1, is) => fold_rel_expr F r1 #> fold (fold_int_expr F) is + | Join (r1, r2) => fold_rel_expr F r1 #> fold_rel_expr F r2 + | Closure r => fold_rel_expr F r + | ReflexiveClosure r => fold_rel_expr F r + | Transpose r => fold_rel_expr F r + | Comprehension (ds, f) => fold (fold_decl F) ds #> fold_formula F f + | Bits i => fold_int_expr F i + | Int i => fold_int_expr F i + | Iden => #rel_expr_func F rel_expr + | Ints => #rel_expr_func F rel_expr + | None => #rel_expr_func F rel_expr + | Univ => #rel_expr_func F rel_expr + | Atom _ => #rel_expr_func F rel_expr + | AtomSeq _ => #rel_expr_func F rel_expr + | Rel _ => #rel_expr_func F rel_expr + | Var _ => #rel_expr_func F rel_expr + | RelReg _ => #rel_expr_func F rel_expr +(* 'a fold_expr_funcs -> int_expr -> 'a -> 'a *) +and fold_int_expr F int_expr = + case int_expr of + Sum (ds, i) => fold (fold_decl F) ds #> fold_int_expr F i + | IntLet (bs, i) => fold (fold_expr_assign F) bs #> fold_int_expr F i + | IntIf (f, i1, i2) => + fold_formula F f #> fold_int_expr F i1 #> fold_int_expr F i2 + | SHL (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | SHA (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | SHR (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Add (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Sub (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Mult (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Div (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Mod (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | Cardinality r => fold_rel_expr F r + | SetSum r => fold_rel_expr F r + | BitOr (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | BitXor (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | BitAnd (i1, i2) => fold_int_expr F i1 #> fold_int_expr F i2 + | BitNot i => fold_int_expr F i + | Neg i => fold_int_expr F i + | Absolute i => fold_int_expr F i + | Signum i => fold_int_expr F i + | Num _ => #int_expr_func F int_expr + | IntReg _ => #int_expr_func F int_expr +(* 'a fold_expr_funcs -> decl -> 'a -> 'a *) +and fold_decl F decl = + case decl of + DeclNo (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r + | DeclLone (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r + | DeclOne (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r + | DeclSome (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r + | DeclSet (x, r) => fold_rel_expr F (Var x) #> fold_rel_expr F r +(* 'a fold_expr_funcs -> expr_assign -> 'a -> 'a *) +and fold_expr_assign F assign = + case assign of + AssignFormulaReg (x, f) => fold_formula F (FormulaReg x) #> fold_formula F f + | AssignRelReg (x, r) => fold_rel_expr F (RelReg x) #> fold_rel_expr F r + | AssignIntReg (x, i) => fold_int_expr F (IntReg x) #> fold_int_expr F i + +type 'a fold_tuple_funcs = { + tuple_func: tuple -> 'a -> 'a, + tuple_set_func: tuple_set -> 'a -> 'a +} + +(* 'a fold_tuple_funcs -> tuple -> 'a -> 'a *) +fun fold_tuple (F : 'a fold_tuple_funcs) = #tuple_func F +(* 'a fold_tuple_funcs -> tuple_set -> 'a -> 'a *) +fun fold_tuple_set F tuple_set = + case tuple_set of + TupleUnion (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 + | TupleDifference (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 + | TupleIntersect (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 + | TupleProduct (ts1, ts2) => fold_tuple_set F ts1 #> fold_tuple_set F ts2 + | TupleProject (ts, _) => fold_tuple_set F ts + | TupleSet ts => fold (fold_tuple F) ts + | TupleRange (t1, t2) => fold_tuple F t1 #> fold_tuple F t2 + | TupleArea (t1, t2) => fold_tuple F t1 #> fold_tuple F t2 + | TupleAtomSeq _ => #tuple_set_func F tuple_set + | TupleSetReg _ => #tuple_set_func F tuple_set +(* 'a fold_tuple_funcs -> tuple_assign -> 'a -> 'a *) +fun fold_tuple_assign F assign = + case assign of + AssignTuple (x, t) => fold_tuple F (TupleReg x) #> fold_tuple F t + | AssignTupleSet (x, ts) => + fold_tuple_set F (TupleSetReg x) #> fold_tuple_set F ts +(* 'a fold_expr_funcs -> 'a fold_tuple_funcs -> bound -> 'a -> 'a *) +fun fold_bound expr_F tuple_F (zs, tss) = + fold (fold_rel_expr expr_F) (map (Rel o fst) zs) + #> fold (fold_tuple_set tuple_F) tss +(* 'a fold_tuple_funcs -> int_bound -> 'a -> 'a *) +fun fold_int_bound F (_, tss) = fold (fold_tuple_set F) tss + +(* int -> int *) +fun max_arity univ_card = floor (Math.ln 2147483647.0 + / Math.ln (Real.fromInt univ_card)) +(* rel_expr -> int *) +fun arity_of_rel_expr (RelLet (_, r)) = arity_of_rel_expr r + | arity_of_rel_expr (RelIf (_, r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Union (r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Difference (r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Override (r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Intersect (r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Product (r1, r2)) = sum_arities_of_rel_exprs r1 r2 + | arity_of_rel_expr (IfNo (r1, _)) = arity_of_rel_expr r1 + | arity_of_rel_expr (Project (r, is)) = length is + | arity_of_rel_expr (Join (r1, r2)) = sum_arities_of_rel_exprs r1 r2 - 2 + | arity_of_rel_expr (Closure _) = 2 + | arity_of_rel_expr (ReflexiveClosure _) = 2 + | arity_of_rel_expr (Transpose _) = 2 + | arity_of_rel_expr (Comprehension (ds, _)) = + fold (curry op + o arity_of_decl) ds 0 + | arity_of_rel_expr (Bits _) = 1 + | arity_of_rel_expr (Int _) = 1 + | arity_of_rel_expr Iden = 2 + | arity_of_rel_expr Ints = 1 + | arity_of_rel_expr None = 1 + | arity_of_rel_expr Univ = 1 + | arity_of_rel_expr (Atom _) = 1 + | arity_of_rel_expr (AtomSeq _) = 1 + | arity_of_rel_expr (Rel (n, _)) = n + | arity_of_rel_expr (Var (n, _)) = n + | arity_of_rel_expr (RelReg (n, _)) = n +(* rel_expr -> rel_expr -> int *) +and sum_arities_of_rel_exprs r1 r2 = arity_of_rel_expr r1 + arity_of_rel_expr r2 +(* decl -> int *) +and arity_of_decl (DeclNo ((n, _), _)) = n + | arity_of_decl (DeclLone ((n, _), _)) = n + | arity_of_decl (DeclOne ((n, _), _)) = n + | arity_of_decl (DeclSome ((n, _), _)) = n + | arity_of_decl (DeclSet ((n, _), _)) = n + +(* string -> bool *) +val is_relevant_setting = not o member (op =) ["solver", "delay"] + +(* problem -> problem -> bool *) +fun problems_equivalent (p1 : problem) (p2 : problem) = + #univ_card p1 = #univ_card p2 + andalso #formula p1 = #formula p2 + andalso #bounds p1 = #bounds p2 + andalso #expr_assigns p1 = #expr_assigns p2 + andalso #tuple_assigns p1 = #tuple_assigns p2 + andalso #int_bounds p1 = #int_bounds p2 + andalso filter (is_relevant_setting o fst) (#settings p1) + = filter (is_relevant_setting o fst) (#settings p2) + +(* int -> string *) +fun base_name j = if j < 0 then Int.toString (~j - 1) ^ "'" else Int.toString j + +(* n_ary_index -> string -> string -> string -> string *) +fun n_ary_name (1, j) prefix _ _ = prefix ^ base_name j + | n_ary_name (2, j) _ prefix _ = prefix ^ base_name j + | n_ary_name (n, j) _ _ prefix = prefix ^ Int.toString n ^ "_" ^ base_name j + +(* int -> string *) +fun atom_name j = "A" ^ base_name j +fun atom_seq_name (k, 0) = "u" ^ base_name k + | atom_seq_name (k, j0) = "u" ^ base_name k ^ "@" ^ base_name j0 +fun formula_reg_name j = "$f" ^ base_name j +fun rel_reg_name j = "$e" ^ base_name j +fun int_reg_name j = "$i" ^ base_name j + +(* n_ary_index -> string *) +fun tuple_name x = n_ary_name x "A" "P" "T" +fun rel_name x = n_ary_name x "s" "r" "m" +fun var_name x = n_ary_name x "S" "R" "M" +fun tuple_reg_name x = n_ary_name x "$A" "$P" "$T" +fun tuple_set_reg_name x = n_ary_name x "$a" "$p" "$t" + +(* string -> string *) +fun inline_comment "" = "" + | inline_comment comment = + " /* " ^ translate_string (fn "\n" => " " | "*" => "* " | s => s) comment ^ + " */" +fun block_comment "" = "" + | block_comment comment = prefix_lines "// " comment ^ "\n" + +(* (n_ary_index * string) -> string *) +fun commented_rel_name (x, s) = rel_name x ^ inline_comment s + +(* tuple -> string *) +fun string_for_tuple (Tuple js) = "[" ^ commas (map atom_name js) ^ "]" + | string_for_tuple (TupleIndex x) = tuple_name x + | string_for_tuple (TupleReg x) = tuple_reg_name x + +val no_prec = 100 +val prec_TupleUnion = 1 +val prec_TupleIntersect = 2 +val prec_TupleProduct = 3 +val prec_TupleProject = 4 + +(* tuple_set -> int *) +fun precedence_ts (TupleUnion _) = prec_TupleUnion + | precedence_ts (TupleDifference _) = prec_TupleUnion + | precedence_ts (TupleIntersect _) = prec_TupleIntersect + | precedence_ts (TupleProduct _) = prec_TupleProduct + | precedence_ts (TupleProject _) = prec_TupleProject + | precedence_ts _ = no_prec + +(* tuple_set -> string *) +fun string_for_tuple_set tuple_set = + let + (* tuple_set -> int -> string *) + fun sub tuple_set outer_prec = + let + val prec = precedence_ts tuple_set + val need_parens = (prec < outer_prec) + in + (if need_parens then "(" else "") ^ + (case tuple_set of + TupleUnion (ts1, ts2) => sub ts1 prec ^ " + " ^ sub ts2 (prec + 1) + | TupleDifference (ts1, ts2) => + sub ts1 prec ^ " - " ^ sub ts1 (prec + 1) + | TupleIntersect (ts1, ts2) => sub ts1 prec ^ " & " ^ sub ts1 prec + | TupleProduct (ts1, ts2) => sub ts1 prec ^ "->" ^ sub ts2 prec + | TupleProject (ts, c) => sub ts prec ^ "[" ^ Int.toString c ^ "]" + | TupleSet ts => "{" ^ commas (map string_for_tuple ts) ^ "}" + | TupleRange (t1, t2) => + "{" ^ string_for_tuple t1 ^ + (if t1 = t2 then "" else " .. " ^ string_for_tuple t2) ^ "}" + | TupleArea (t1, t2) => + "{" ^ string_for_tuple t1 ^ " # " ^ string_for_tuple t2 ^ "}" + | TupleAtomSeq x => atom_seq_name x + | TupleSetReg x => tuple_set_reg_name x) ^ + (if need_parens then ")" else "") + end + in sub tuple_set 0 end + +(* tuple_assign -> string *) +fun string_for_tuple_assign (AssignTuple (x, t)) = + tuple_reg_name x ^ " := " ^ string_for_tuple t ^ "\n" + | string_for_tuple_assign (AssignTupleSet (x, ts)) = + tuple_set_reg_name x ^ " := " ^ string_for_tuple_set ts ^ "\n" + +(* bound -> string *) +fun string_for_bound (zs, tss) = + "bounds " ^ commas (map commented_rel_name zs) ^ ": " ^ + (if length tss = 1 then "" else "[") ^ commas (map string_for_tuple_set tss) ^ + (if length tss = 1 then "" else "]") ^ "\n" + +(* int_bound -> string *) +fun int_string_for_bound (opt_n, tss) = + (case opt_n of + SOME n => Int.toString n ^ ": " + | NONE => "") ^ "[" ^ commas (map string_for_tuple_set tss) ^ "]" + +val prec_All = 1 +val prec_Or = 2 +val prec_Iff = 3 +val prec_Implies = 4 +val prec_And = 5 +val prec_Not = 6 +val prec_Eq = 7 +val prec_Some = 8 +val prec_SHL = 9 +val prec_Add = 10 +val prec_Mult = 11 +val prec_Override = 12 +val prec_Intersect = 13 +val prec_Product = 14 +val prec_IfNo = 15 +val prec_Project = 17 +val prec_Join = 18 +val prec_BitNot = 19 + +(* formula -> int *) +fun precedence_f (All _) = prec_All + | precedence_f (Exist _) = prec_All + | precedence_f (FormulaLet _) = prec_All + | precedence_f (FormulaIf _) = prec_All + | precedence_f (Or _) = prec_Or + | precedence_f (Iff _) = prec_Iff + | precedence_f (Implies _) = prec_Implies + | precedence_f (And _) = prec_And + | precedence_f (Not _) = prec_Not + | precedence_f (Acyclic _) = no_prec + | precedence_f (Function _) = no_prec + | precedence_f (Functional _) = no_prec + | precedence_f (TotalOrdering _) = no_prec + | precedence_f (Subset _) = prec_Eq + | precedence_f (RelEq _) = prec_Eq + | precedence_f (IntEq _) = prec_Eq + | precedence_f (LT _) = prec_Eq + | precedence_f (LE _) = prec_Eq + | precedence_f (No _) = prec_Some + | precedence_f (Lone _) = prec_Some + | precedence_f (One _) = prec_Some + | precedence_f (Some _) = prec_Some + | precedence_f False = no_prec + | precedence_f True = no_prec + | precedence_f (FormulaReg _) = no_prec +(* rel_expr -> int *) +and precedence_r (RelLet _) = prec_All + | precedence_r (RelIf _) = prec_All + | precedence_r (Union _) = prec_Add + | precedence_r (Difference _) = prec_Add + | precedence_r (Override _) = prec_Override + | precedence_r (Intersect _) = prec_Intersect + | precedence_r (Product _) = prec_Product + | precedence_r (IfNo _) = prec_IfNo + | precedence_r (Project _) = prec_Project + | precedence_r (Join _) = prec_Join + | precedence_r (Closure _) = prec_BitNot + | precedence_r (ReflexiveClosure _) = prec_BitNot + | precedence_r (Transpose _) = prec_BitNot + | precedence_r (Comprehension _) = no_prec + | precedence_r (Bits _) = no_prec + | precedence_r (Int _) = no_prec + | precedence_r Iden = no_prec + | precedence_r Ints = no_prec + | precedence_r None = no_prec + | precedence_r Univ = no_prec + | precedence_r (Atom _) = no_prec + | precedence_r (AtomSeq _) = no_prec + | precedence_r (Rel _) = no_prec + | precedence_r (Var _) = no_prec + | precedence_r (RelReg _) = no_prec +(* int_expr -> int *) +and precedence_i (Sum _) = prec_All + | precedence_i (IntLet _) = prec_All + | precedence_i (IntIf _) = prec_All + | precedence_i (SHL _) = prec_SHL + | precedence_i (SHA _) = prec_SHL + | precedence_i (SHR _) = prec_SHL + | precedence_i (Add _) = prec_Add + | precedence_i (Sub _) = prec_Add + | precedence_i (Mult _) = prec_Mult + | precedence_i (Div _) = prec_Mult + | precedence_i (Mod _) = prec_Mult + | precedence_i (Cardinality _) = no_prec + | precedence_i (SetSum _) = no_prec + | precedence_i (BitOr _) = prec_Intersect + | precedence_i (BitXor _) = prec_Intersect + | precedence_i (BitAnd _) = prec_Intersect + | precedence_i (BitNot _) = prec_BitNot + | precedence_i (Neg _) = prec_BitNot + | precedence_i (Absolute _) = prec_BitNot + | precedence_i (Signum _) = prec_BitNot + | precedence_i (Num _) = no_prec + | precedence_i (IntReg _) = no_prec + +(* (string -> unit) -> problem list -> unit *) +fun write_problem_file out problems = + let + (* formula -> unit *) + fun out_outmost_f (And (f1, f2)) = + (out_outmost_f f1; out "\n && "; out_outmost_f f2) + | out_outmost_f f = out_f f prec_And + (* formula -> int -> unit *) + and out_f formula outer_prec = + let + val prec = precedence_f formula + val need_parens = (prec < outer_prec) + in + (if need_parens then out "(" else ()); + (case formula of + All (ds, f) => (out "all ["; out_decls ds; out "] | "; out_f f prec) + | Exist (ds, f) => + (out "some ["; out_decls ds; out "] | "; out_f f prec) + | FormulaLet (bs, f) => + (out "let ["; out_assigns bs; out "] | "; out_f f prec) + | FormulaIf (f, f1, f2) => + (out "if "; out_f f prec; out " then "; out_f f1 prec; out " else "; + out_f f2 prec) + | Or (f1, f2) => (out_f f1 prec; out " || "; out_f f2 prec) + | Iff (f1, f2) => (out_f f1 prec; out " <=> "; out_f f2 prec) + | Implies (f1, f2) => (out_f f1 (prec + 1); out " => "; out_f f2 prec) + | And (f1, f2) => (out_f f1 prec; out " && "; out_f f2 prec) + | Not f => (out "! "; out_f f prec) + | Acyclic x => out ("ACYCLIC(" ^ rel_name x ^ ")") + | Function (x, r1, r2) => + (out ("FUNCTION(" ^ rel_name x ^ ", "); out_r r1 0; out " -> one "; + out_r r2 0; out ")") + | Functional (x, r1, r2) => + (out ("FUNCTION(" ^ rel_name x ^ ", "); out_r r1 0; out " -> lone "; + out_r r2 0; out ")") + | TotalOrdering (x1, x2, x3, x4) => + out ("TOTAL_ORDERING(" ^ rel_name x1 ^ ", " ^ rel_name x2 ^ ", " + ^ rel_name x3 ^ ", " ^ rel_name x4 ^ ")") + | Subset (r1, r2) => (out_r r1 prec; out " in "; out_r r2 prec) + | RelEq (r1, r2) => (out_r r1 prec; out " = "; out_r r2 prec) + | IntEq (i1, i2) => (out_i i1 prec; out " = "; out_i i2 prec) + | LT (i1, i2) => (out_i i1 prec; out " < "; out_i i2 prec) + | LE (i1, i2) => (out_i i1 prec; out " <= "; out_i i2 prec) + | No r => (out "no "; out_r r prec) + | Lone r => (out "lone "; out_r r prec) + | One r => (out "one "; out_r r prec) + | Some r => (out "some "; out_r r prec) + | False => out "false" + | True => out "true" + | FormulaReg j => out (formula_reg_name j)); + (if need_parens then out ")" else ()) + end + (* rel_expr -> int -> unit *) + and out_r rel_expr outer_prec = + let + val prec = precedence_r rel_expr + val need_parens = (prec < outer_prec) + in + (if need_parens then out "(" else ()); + (case rel_expr of + RelLet (bs, r) => + (out "let ["; out_assigns bs; out "] | "; out_r r prec) + | RelIf (f, r1, r2) => + (out "if "; out_f f prec; out " then "; out_r r1 prec; + out " else "; out_r r2 prec) + | Union (r1, r2) => (out_r r1 prec; out " + "; out_r r2 (prec + 1)) + | Difference (r1, r2) => + (out_r r1 prec; out " - "; out_r r2 (prec + 1)) + | Override (r1, r2) => (out_r r1 prec; out " ++ "; out_r r2 prec) + | Intersect (r1, r2) => (out_r r1 prec; out " & "; out_r r2 prec) + | Product (r1, r2) => (out_r r1 prec; out "->"; out_r r2 prec) + | IfNo (r1, r2) => (out_r r1 prec; out "\\"; out_r r2 prec) + | Project (r1, is) => (out_r r1 prec; out "["; out_columns is; out "]") + | Join (r1, r2) => (out_r r1 prec; out "."; out_r r2 (prec + 1)) + | Closure r => (out "^"; out_r r prec) + | ReflexiveClosure r => (out "*"; out_r r prec) + | Transpose r => (out "~"; out_r r prec) + | Comprehension (ds, f) => + (out "{["; out_decls ds; out "] | "; out_f f 0; out "}") + | Bits i => (out "Bits["; out_i i 0; out "]") + | Int i => (out "Int["; out_i i 0; out "]") + | Iden => out "iden" + | Ints => out "ints" + | None => out "none" + | Univ => out "univ" + | Atom j => out (atom_name j) + | AtomSeq x => out (atom_seq_name x) + | Rel x => out (rel_name x) + | Var x => out (var_name x) + | RelReg (_, j) => out (rel_reg_name j)); + (if need_parens then out ")" else ()) + end + (* int_expr -> int -> unit *) + and out_i int_expr outer_prec = + let + val prec = precedence_i int_expr + val need_parens = (prec < outer_prec) + in + (if need_parens then out "(" else ()); + (case int_expr of + Sum (ds, i) => (out "sum ["; out_decls ds; out "] | "; out_i i prec) + | IntLet (bs, i) => + (out "let ["; out_assigns bs; out "] | "; out_i i prec) + | IntIf (f, i1, i2) => + (out "if "; out_f f prec; out " then "; out_i i1 prec; + out " else "; out_i i2 prec) + | SHL (i1, i2) => (out_i i1 prec; out " << "; out_i i2 (prec + 1)) + | SHA (i1, i2) => (out_i i1 prec; out " >> "; out_i i2 (prec + 1)) + | SHR (i1, i2) => (out_i i1 prec; out " >>> "; out_i i2 (prec + 1)) + | Add (i1, i2) => (out_i i1 prec; out " + "; out_i i2 (prec + 1)) + | Sub (i1, i2) => (out_i i1 prec; out " - "; out_i i2 (prec + 1)) + | Mult (i1, i2) => (out_i i1 prec; out " * "; out_i i2 (prec + 1)) + | Div (i1, i2) => (out_i i1 prec; out " / "; out_i i2 (prec + 1)) + | Mod (i1, i2) => (out_i i1 prec; out " % "; out_i i2 (prec + 1)) + | Cardinality r => (out "#("; out_r r 0; out ")") + | SetSum r => (out "sum("; out_r r 0; out ")") + | BitOr (i1, i2) => (out_i i1 prec; out " | "; out_i i2 prec) + | BitXor (i1, i2) => (out_i i1 prec; out " ^ "; out_i i2 prec) + | BitAnd (i1, i2) => (out_i i1 prec; out " & "; out_i i2 prec) + | BitNot i => (out "~"; out_i i prec) + | Neg i => (out "-"; out_i i prec) + | Absolute i => (out "abs "; out_i i prec) + | Signum i => (out "sgn "; out_i i prec) + | Num k => out (Int.toString k) + | IntReg j => out (int_reg_name j)); + (if need_parens then out ")" else ()) + end + (* decl list -> unit *) + and out_decls [] = () + | out_decls [d] = out_decl d + | out_decls (d :: ds) = (out_decl d; out ", "; out_decls ds) + (* decl -> unit *) + and out_decl (DeclNo (x, r)) = + (out (var_name x); out " : no "; out_r r 0) + | out_decl (DeclLone (x, r)) = + (out (var_name x); out " : lone "; out_r r 0) + | out_decl (DeclOne (x, r)) = + (out (var_name x); out " : one "; out_r r 0) + | out_decl (DeclSome (x, r)) = + (out (var_name x); out " : some "; out_r r 0) + | out_decl (DeclSet (x, r)) = + (out (var_name x); out " : set "; out_r r 0) + (* assign_expr list -> unit *) + and out_assigns [] = () + | out_assigns [b] = out_assign b + | out_assigns (b :: bs) = (out_assign b; out ", "; out_assigns bs) + (* assign_expr -> unit *) + and out_assign (AssignFormulaReg (j, f)) = + (out (formula_reg_name j); out " := "; out_f f 0) + | out_assign (AssignRelReg ((_, j), r)) = + (out (rel_reg_name j); out " := "; out_r r 0) + | out_assign (AssignIntReg (j, i)) = + (out (int_reg_name j); out " := "; out_i i 0) + (* int_expr list -> unit *) + and out_columns [] = () + | out_columns [i] = out_i i 0 + | out_columns (i :: is) = (out_i i 0; out ", "; out_columns is) + (* problem -> unit *) + and out_problem {comment, settings, univ_card, tuple_assigns, bounds, + int_bounds, expr_assigns, formula} = + (out ("\n" ^ block_comment comment ^ + implode (map (fn (key, value) => key ^ ": " ^ value ^ "\n") + settings) ^ + "univ: " ^ atom_seq_name (univ_card, 0) ^ "\n" ^ + implode (map string_for_tuple_assign tuple_assigns) ^ + implode (map string_for_bound bounds) ^ + (if int_bounds = [] then + "" + else + "int_bounds: " ^ + commas (map int_string_for_bound int_bounds) ^ "\n")); + map (fn b => (out_assign b; out ";")) expr_assigns; + out "solve "; out_outmost_f formula; out ";\n") + in + out ("// This file was generated by Isabelle (probably Nitpick)\n" ^ + "// " ^ Date.fmt "%Y-%m-%d %H:%M:%S" + (Date.fromTimeLocal (Time.now ())) ^ "\n"); + map out_problem problems + end + +(* string -> bool *) +fun is_ident_char s = + Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s + orelse s = "_" orelse s = "'" orelse s = "$" + +(* string list -> string list *) +fun strip_blanks [] = [] + | strip_blanks (" " :: ss) = strip_blanks ss + | strip_blanks [s1, " "] = [s1] + | strip_blanks (s1 :: " " :: s2 :: ss) = + if is_ident_char s1 andalso is_ident_char s2 then + s1 :: " " :: strip_blanks (s2 :: ss) + else + strip_blanks (s1 :: s2 :: ss) + | strip_blanks (s :: ss) = s :: strip_blanks ss + +(* (string list -> 'a * string list) -> string list -> 'a list * string list *) +fun scan_non_empty_list scan = scan ::: Scan.repeat ($$ "," |-- scan) +fun scan_list scan = scan_non_empty_list scan || Scan.succeed [] +(* string list -> int * string list *) +val scan_nat = Scan.repeat1 (Scan.one Symbol.is_ascii_digit) + >> (the o Int.fromString o space_implode "") +(* string list -> (int * int) * string list *) +val scan_rel_name = $$ "s" |-- scan_nat >> pair 1 + || $$ "r" |-- scan_nat >> pair 2 + || ($$ "m" |-- scan_nat --| $$ "_") -- scan_nat +(* string list -> int * string list *) +val scan_atom = $$ "A" |-- scan_nat +(* string list -> int list * string list *) +val scan_tuple = $$ "[" |-- scan_list scan_atom --| $$ "]" +(* string list -> int list list * string list *) +val scan_tuple_set = $$ "[" |-- scan_list scan_tuple --| $$ "]" +(* string list -> ((int * int) * int list list) * string list *) +val scan_assignment = (scan_rel_name --| $$ "=") -- scan_tuple_set +(* string list -> ((int * int) * int list list) list * string list *) +val scan_instance = Scan.this_string "relations:" |-- + $$ "{" |-- scan_list scan_assignment --| $$ "}" + +(* string -> raw_bound list *) +fun parse_instance inst = + Scan.finite Symbol.stopper + (Scan.error (!! (fn _ => raise SYNTAX ("Kodkod.parse_instance", + "ill-formed Kodkodi output")) + scan_instance)) + (strip_blanks (explode inst)) + |> fst + +val problem_marker = "*** PROBLEM " +val outcome_marker = "---OUTCOME---\n" +val instance_marker = "---INSTANCE---\n" + +(* string -> substring -> string *) +fun read_section_body marker = + Substring.string o fst o Substring.position "\n\n" + o Substring.triml (size marker) + +(* substring -> raw_bound list *) +fun read_next_instance s = + let val s = Substring.position instance_marker s |> snd in + if Substring.isEmpty s then + raise SYNTAX ("Kodkod.read_next_instance", "expected \"INSTANCE\" marker") + else + read_section_body instance_marker s |> parse_instance + end + +(* int -> substring * (int * raw_bound list) list * int list + -> substring * (int * raw_bound list) list * int list *) +fun read_next_outcomes j (s, ps, js) = + let val (s1, s2) = Substring.position outcome_marker s in + if Substring.isEmpty s2 + orelse not (Substring.isEmpty (Substring.position problem_marker s1 + |> snd)) then + (s, ps, js) + else + let + val outcome = read_section_body outcome_marker s2 + val s = Substring.triml (size outcome_marker) s2 + in + if String.isSuffix "UNSATISFIABLE" outcome then + read_next_outcomes j (s, ps, j :: js) + else if String.isSuffix "SATISFIABLE" outcome then + read_next_outcomes j (s, (j, read_next_instance s2) :: ps, js) + else + raise SYNTAX ("Kodkod.read_next_outcomes", + "unknown outcome " ^ quote outcome) + end + end + +(* substring * (int * raw_bound list) list * int list + -> (int * raw_bound list) list * int list *) +fun read_next_problems (s, ps, js) = + let val s = Substring.position problem_marker s |> snd in + if Substring.isEmpty s then + (ps, js) + else + let + val s = Substring.triml (size problem_marker) s + val j_plus_1 = s |> Substring.takel (not_equal #" ") |> Substring.string + |> Int.fromString |> the + val j = j_plus_1 - 1 + in read_next_problems (read_next_outcomes j (s, ps, js)) end + end + handle Option.Option => raise SYNTAX ("Kodkod.read_next_problems", + "expected number after \"PROBLEM\"") + +(* Path.T -> (int * raw_bound list) list * int list *) +fun read_output_file path = + read_next_problems (Substring.full (File.read path), [], []) |>> rev ||> rev + +(* The fudge term below is to account for Kodkodi's slow start-up time, which + is partly due to the JVM and partly due to the ML "system" function. *) +val fudge_ms = 250 + +(* bool -> Time.time option -> int -> int -> problem list -> outcome *) +fun solve_any_problem overlord deadline max_threads max_solutions problems = + let + val j = find_index (equal True o #formula) problems + val indexed_problems = if j >= 0 then + [(j, nth problems j)] + else + filter (not_equal False o #formula o snd) + (0 upto length problems - 1 ~~ problems) + val triv_js = filter_out (AList.defined (op =) indexed_problems) + (0 upto length problems - 1) + (* int -> int *) + val reindex = fst o nth indexed_problems + in + if null indexed_problems then + Normal ([], triv_js) + else + let + val (serial_str, tmp_path) = + if overlord then + ("", Path.append (Path.variable "ISABELLE_HOME_USER") o Path.base) + else + (serial_string (), File.tmp_path) + (* string -> string -> Path.T *) + fun path_for base suf = + tmp_path (Path.explode (base ^ serial_str ^ "." ^ suf)) + val in_path = path_for "isabelle" "kki" + val in_buf = Unsynchronized.ref Buffer.empty + (* string -> unit *) + fun out s = Unsynchronized.change in_buf (Buffer.add s) + val out_path = path_for "kodkodi" "out" + val err_path = path_for "kodkodi" "err" + val _ = write_problem_file out (map snd indexed_problems) + val _ = File.write_buffer in_path (!in_buf) + (* (int list -> outcome) -> outcome *) + fun stopped constr = + let val nontriv_js = map reindex (snd (read_output_file out_path)) in + constr (triv_js @ nontriv_js) + handle Exn.Interrupt => Interrupted NONE + end + in + let + val ms = + case deadline of + NONE => ~1 + | SOME time => + Int.max (0, Time.toMilliseconds (Time.- (time, Time.now ())) + - fudge_ms) + val outcome = + let + val code = + system ("env CLASSPATH=\"$KODKODI_CLASSPATH:$CLASSPATH\" \ + \\"$ISABELLE_TOOL\" java \ + \de.tum.in.isabelle.Kodkodi.Kodkodi" ^ + (if ms >= 0 then " -max-msecs " ^ Int.toString ms + else "") ^ + (if max_solutions > 1 then " -solve-all" else "") ^ + " -max-solutions " ^ Int.toString max_solutions ^ + (if max_threads > 0 then + " -max-threads " ^ Int.toString max_threads + else + "") ^ + " < " ^ Path.implode in_path ^ + " > " ^ Path.implode out_path ^ + " 2> " ^ Path.implode err_path) + val (ps, nontriv_js) = read_output_file out_path + |>> map (apfst reindex) ||> map reindex + val js = triv_js @ nontriv_js + val first_error = + File.fold_lines (fn line => fn "" => line | s => s) err_path "" + in + if null ps then + if code = 2 then + TimedOut js + else if first_error <> "" then + Error (first_error |> perhaps (try (unsuffix ".")) + |> perhaps (try (unprefix "Error: ")), js) + else if code <> 0 then + Error ("Unknown error", js) + else + Normal ([], js) + else + Normal (ps, js) + end + in + if overlord then () + else List.app File.rm [in_path, out_path, err_path]; + outcome + end + handle Exn.Interrupt => stopped (Interrupted o SOME) + end + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/kodkod_sat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/kodkod_sat.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,109 @@ +(* Title: HOL/Nitpick/Tools/kodkod_sat.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Kodkod SAT solver integration. +*) + +signature KODKOD_SAT = +sig + val configured_sat_solvers : bool -> string list + val smart_sat_solver_name : bool -> string + val sat_solver_spec : string -> string * string list +end; + +structure KodkodSAT : KODKOD_SAT = +struct + +datatype sink = ToStdout | ToFile + +datatype sat_solver_info = + Internal of bool * string list | + External of sink * string * string * string list | + ExternalV2 of sink * string * string * string list * string * string * string + +val berkmin_exec = getenv "BERKMIN_EXE" + +(* (string * sat_solver_info) list *) +val static_list = + [("MiniSat", ExternalV2 (ToFile, "MINISAT_HOME", "minisat", [], "SAT", "", + "UNSAT")), + ("PicoSAT", External (ToStdout, "PICOSAT_HOME", "picosat", [])), + ("zChaff", ExternalV2 (ToStdout, "ZCHAFF_HOME", "zchaff", [], + "Instance Satisfiable", "", + "Instance Unsatisfiable")), + ("RSat", ExternalV2 (ToStdout, "RSAT_HOME", "rsat", ["-s"], + "s SATISFIABLE", "v ", "s UNSATISFIABLE")), + ("BerkMin", ExternalV2 (ToStdout, "BERKMIN_HOME", + if berkmin_exec = "" then "BerkMin561" + else berkmin_exec, [], "Satisfiable !!", + "solution =", "UNSATISFIABLE !!")), + ("BerkMinAlloy", External (ToStdout, "BERKMINALLOY_HOME", "berkmin", [])), + ("Jerusat", External (ToStdout, "JERUSAT_HOME", "Jerusat1.3", [])), + ("SAT4J", Internal (true, ["DefaultSAT4J"])), + ("MiniSatJNI", Internal (true, ["MiniSat"])), + ("zChaffJNI", Internal (false, ["zChaff"])), + ("SAT4JLight", Internal (true, ["LightSAT4J"])), + ("HaifaSat", ExternalV2 (ToStdout, "HAIFASAT_HOME", "HaifaSat", ["-p", "1"], + "s SATISFIABLE", "v ", "s UNSATISFIABLE"))] + +val created_temp_dir = Unsynchronized.ref false + +(* string -> sink -> string -> string -> string list -> string list + -> (string * (unit -> string list)) option *) +fun dynamic_entry_for_external name dev home exec args markers = + case getenv home of + "" => NONE + | dir => SOME (name, fn () => + let + val temp_dir = getenv "ISABELLE_TMP" + val _ = if !created_temp_dir then + () + else + (created_temp_dir := true; + File.mkdir (Path.explode temp_dir)) + val temp = temp_dir ^ "/" ^ name ^ serial_string () + val out_file = temp ^ ".out" + in + [if null markers then "External" else "ExternalV2", + dir ^ "/" ^ exec, temp ^ ".cnf", + if dev = ToFile then out_file else ""] @ markers @ + (if dev = ToFile then [out_file] else []) @ args + end) +(* bool -> string * sat_solver_info + -> (string * (unit -> string list)) option *) +fun dynamic_entry_for_info false (name, Internal (_, ss)) = SOME (name, K ss) + | dynamic_entry_for_info false (name, External (dev, home, exec, args)) = + dynamic_entry_for_external name dev home exec args [] + | dynamic_entry_for_info false (name, ExternalV2 (dev, home, exec, args, + m1, m2, m3)) = + dynamic_entry_for_external name dev home exec args [m1, m2, m3] + | dynamic_entry_for_info true (name, Internal (true, ss)) = SOME (name, K ss) + | dynamic_entry_for_info true _ = NONE +(* bool -> (string * (unit -> string list)) list *) +fun dynamic_list incremental = + map_filter (dynamic_entry_for_info incremental) static_list + +(* bool -> string list *) +val configured_sat_solvers = map fst o dynamic_list + +(* bool -> string *) +val smart_sat_solver_name = dynamic_list #> hd #> fst + +(* (string * 'a) list -> string *) +fun enum_solvers xs = commas (map (quote o fst) xs |> distinct (op =)) +(* string -> string * string list *) +fun sat_solver_spec name = + let val dynamic_list = dynamic_list false in + (name, the (AList.lookup (op =) dynamic_list name) ()) + handle Option.Option => + error (if AList.defined (op =) static_list name then + "The SAT solver " ^ quote name ^ " is not configured. The \ + \following solvers are configured:\n" ^ + enum_solvers dynamic_list ^ "." + else + "Unknown SAT solver " ^ quote name ^ ". The following \ + \solvers are supported:\n" ^ enum_solvers static_list ^ ".") + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/minipick.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/minipick.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,322 @@ +(* Title: HOL/Nitpick/Tools/minipick.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Finite model generation for HOL formulas using Kodkod, minimalistic version. +*) + +signature MINIPICK = +sig + val pick_nits_in_term : Proof.context -> (typ -> int) -> term -> string +end; + +structure Minipick : MINIPICK = +struct + +open Kodkod +open NitpickUtil +open NitpickHOL +open NitpickPeephole +open NitpickKodkod + +(* theory -> typ -> unit *) +fun check_type ctxt (Type ("fun", Ts)) = List.app (check_type ctxt) Ts + | check_type ctxt (Type ("*", Ts)) = List.app (check_type ctxt) Ts + | check_type _ @{typ bool} = () + | check_type _ (TFree (_, @{sort "{}"})) = () + | check_type _ (TFree (_, @{sort HOL.type})) = () + | check_type ctxt T = + raise NOT_SUPPORTED ("type " ^ quote (Syntax.string_of_typ ctxt T)) + +(* (typ -> int) -> typ -> int *) +fun atom_schema_of_one scope (Type ("fun", [T1, T2])) = + replicate_list (scope T1) (atom_schema_of_one scope T2) + | atom_schema_of_one scope (Type ("*", [T1, T2])) = + atom_schema_of_one scope T1 @ atom_schema_of_one scope T2 + | atom_schema_of_one scope T = [scope T] +fun atom_schema_of_set scope (Type ("fun", [T1, @{typ bool}])) = + atom_schema_of_one scope T1 + | atom_schema_of_set scope (Type ("fun", [T1, T2])) = + atom_schema_of_one scope T1 @ atom_schema_of_set scope T2 + | atom_schema_of_set scope T = atom_schema_of_one scope T +val arity_of_one = length oo atom_schema_of_one +val arity_of_set = length oo atom_schema_of_set + +(* (typ -> int) -> typ list -> int -> int *) +fun index_for_bound_var _ [_] 0 = 0 + | index_for_bound_var scope (_ :: Ts) 0 = + index_for_bound_var scope Ts 0 + arity_of_one scope (hd Ts) + | index_for_bound_var scope Ts n = index_for_bound_var scope (tl Ts) (n - 1) +(* (typ -> int) -> typ list -> int -> rel_expr list *) +fun one_vars_for_bound_var scope Ts j = + map (curry Var 1) (index_seq (index_for_bound_var scope Ts j) + (arity_of_one scope (nth Ts j))) +fun set_vars_for_bound_var scope Ts j = + map (curry Var 1) (index_seq (index_for_bound_var scope Ts j) + (arity_of_set scope (nth Ts j))) +(* (typ -> int) -> typ list -> typ -> decl list *) +fun decls_for_one scope Ts T = + map2 (curry DeclOne o pair 1) + (index_seq (index_for_bound_var scope (T :: Ts) 0) + (arity_of_one scope (nth (T :: Ts) 0))) + (map (AtomSeq o rpair 0) (atom_schema_of_one scope T)) +fun decls_for_set scope Ts T = + map2 (curry DeclOne o pair 1) + (index_seq (index_for_bound_var scope (T :: Ts) 0) + (arity_of_set scope (nth (T :: Ts) 0))) + (map (AtomSeq o rpair 0) (atom_schema_of_set scope T)) + +(* int list -> rel_expr *) +val atom_product = foldl1 Product o map Atom + +val false_atom = Atom 0 +val true_atom = Atom 1 + +(* rel_expr -> formula *) +fun formula_from_atom r = RelEq (r, true_atom) +(* formula -> rel_expr *) +fun atom_from_formula f = RelIf (f, true_atom, false_atom) + +(* Proof.context -> (typ -> int) -> styp list -> term -> formula *) +fun kodkod_formula_for_term ctxt scope frees = + let + (* typ list -> int -> rel_expr *) + val one_from_bound_var = foldl1 Product oo one_vars_for_bound_var scope + val set_from_bound_var = foldl1 Product oo set_vars_for_bound_var scope + (* typ -> rel_expr -> rel_expr *) + fun set_from_one (T as Type ("fun", [T1, @{typ bool}])) r = + let + val jss = atom_schema_of_one scope T1 |> map (rpair 0) + |> all_combinations + in + map2 (fn i => fn js => + RelIf (RelEq (Project (r, [Num i]), true_atom), + atom_product js, empty_n_ary_rel (length js))) + (index_seq 0 (length jss)) jss + |> foldl1 Union + end + | set_from_one (Type ("fun", [T1, T2])) r = + let + val jss = atom_schema_of_one scope T1 |> map (rpair 0) + |> all_combinations + val arity2 = arity_of_one scope T2 + in + map2 (fn i => fn js => + Product (atom_product js, + Project (r, num_seq (i * arity2) arity2) + |> set_from_one T2)) + (index_seq 0 (length jss)) jss + |> foldl1 Union + end + | set_from_one _ r = r + (* typ list -> typ -> rel_expr -> rel_expr *) + fun one_from_set Ts (T as Type ("fun", _)) r = + Comprehension (decls_for_one scope Ts T, + RelEq (set_from_one T (one_from_bound_var (T :: Ts) 0), + r)) + | one_from_set _ _ r = r + (* typ list -> term -> formula *) + fun to_f Ts t = + (case t of + @{const Not} $ t1 => Not (to_f Ts t1) + | @{const False} => False + | @{const True} => True + | Const (@{const_name All}, _) $ Abs (s, T, t') => + All (decls_for_one scope Ts T, to_f (T :: Ts) t') + | (t0 as Const (@{const_name All}, _)) $ t1 => + to_f Ts (t0 $ eta_expand Ts t1 1) + | Const (@{const_name Ex}, _) $ Abs (s, T, t') => + Exist (decls_for_one scope Ts T, to_f (T :: Ts) t') + | (t0 as Const (@{const_name Ex}, _)) $ t1 => + to_f Ts (t0 $ eta_expand Ts t1 1) + | Const (@{const_name "op ="}, _) $ t1 $ t2 => + RelEq (to_set Ts t1, to_set Ts t2) + | Const (@{const_name ord_class.less_eq}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => + Subset (to_set Ts t1, to_set Ts t2) + | @{const "op &"} $ t1 $ t2 => And (to_f Ts t1, to_f Ts t2) + | @{const "op |"} $ t1 $ t2 => Or (to_f Ts t1, to_f Ts t2) + | @{const "op -->"} $ t1 $ t2 => Implies (to_f Ts t1, to_f Ts t2) + | t1 $ t2 => Subset (to_one Ts t2, to_set Ts t1) + | Free _ => raise SAME () + | Term.Var _ => raise SAME () + | Bound _ => raise SAME () + | Const (s, _) => raise NOT_SUPPORTED ("constant " ^ quote s) + | _ => raise TERM ("to_f", [t])) + handle SAME () => formula_from_atom (to_set Ts t) + (* typ list -> term -> rel_expr *) + and to_one Ts t = + case t of + Const (@{const_name Pair}, _) $ t1 $ t2 => + Product (to_one Ts t1, to_one Ts t2) + | Const (@{const_name Pair}, _) $ _ => to_one Ts (eta_expand Ts t 1) + | Const (@{const_name Pair}, _) => to_one Ts (eta_expand Ts t 2) + | Const (@{const_name fst}, _) $ t1 => + let val fst_arity = arity_of_one scope (fastype_of1 (Ts, t)) in + Project (to_one Ts t1, num_seq 0 fst_arity) + end + | Const (@{const_name fst}, _) => to_one Ts (eta_expand Ts t 1) + | Const (@{const_name snd}, _) $ t1 => + let + val pair_arity = arity_of_one scope (fastype_of1 (Ts, t1)) + val snd_arity = arity_of_one scope (fastype_of1 (Ts, t)) + val fst_arity = pair_arity - snd_arity + in Project (to_one Ts t1, num_seq fst_arity snd_arity) end + | Const (@{const_name snd}, _) => to_one Ts (eta_expand Ts t 1) + | Bound j => one_from_bound_var Ts j + | _ => one_from_set Ts (fastype_of1 (Ts, t)) (to_set Ts t) + (* term -> rel_expr *) + and to_set Ts t = + (case t of + @{const Not} => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name All}, _) => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name Ex}, _) => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name "op ="}, _) $ _ => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name "op ="}, _) => to_set Ts (eta_expand Ts t 2) + | Const (@{const_name ord_class.less_eq}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ _ => + to_set Ts (eta_expand Ts t 1) + | Const (@{const_name ord_class.less_eq}, _) => + to_set Ts (eta_expand Ts t 2) + | @{const "op &"} $ _ => to_set Ts (eta_expand Ts t 1) + | @{const "op &"} => to_set Ts (eta_expand Ts t 2) + | @{const "op |"} $ _ => to_set Ts (eta_expand Ts t 1) + | @{const "op |"} => to_set Ts (eta_expand Ts t 2) + | @{const "op -->"} $ _ => to_set Ts (eta_expand Ts t 1) + | @{const "op -->"} => to_set Ts (eta_expand Ts t 2) + | Const (@{const_name bot_class.bot}, + T as Type ("fun", [_, @{typ bool}])) => + empty_n_ary_rel (arity_of_set scope T) + | Const (@{const_name insert}, _) $ t1 $ t2 => + Union (to_one Ts t1, to_set Ts t2) + | Const (@{const_name insert}, _) $ _ => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name insert}, _) => to_set Ts (eta_expand Ts t 2) + | Const (@{const_name trancl}, _) $ t1 => + if arity_of_set scope (fastype_of1 (Ts, t1)) = 2 then + Closure (to_set Ts t1) + else + raise NOT_SUPPORTED "transitive closure for function or pair type" + | Const (@{const_name trancl}, _) => to_set Ts (eta_expand Ts t 1) + | Const (@{const_name lower_semilattice_class.inf}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => + Intersect (to_set Ts t1, to_set Ts t2) + | Const (@{const_name lower_semilattice_class.inf}, _) $ _ => + to_set Ts (eta_expand Ts t 1) + | Const (@{const_name lower_semilattice_class.inf}, _) => + to_set Ts (eta_expand Ts t 2) + | Const (@{const_name upper_semilattice_class.sup}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => + Union (to_set Ts t1, to_set Ts t2) + | Const (@{const_name upper_semilattice_class.sup}, _) $ _ => + to_set Ts (eta_expand Ts t 1) + | Const (@{const_name upper_semilattice_class.sup}, _) => + to_set Ts (eta_expand Ts t 2) + | Const (@{const_name minus_class.minus}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => + Difference (to_set Ts t1, to_set Ts t2) + | Const (@{const_name minus_class.minus}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ _ => + to_set Ts (eta_expand Ts t 1) + | Const (@{const_name minus_class.minus}, + Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) => + to_set Ts (eta_expand Ts t 2) + | Const (@{const_name Pair}, _) $ _ $ _ => raise SAME () + | Const (@{const_name Pair}, _) $ _ => raise SAME () + | Const (@{const_name Pair}, _) => raise SAME () + | Const (@{const_name fst}, _) $ _ => raise SAME () + | Const (@{const_name fst}, _) => raise SAME () + | Const (@{const_name snd}, _) $ _ => raise SAME () + | Const (@{const_name snd}, _) => raise SAME () + | Const (_, @{typ bool}) => atom_from_formula (to_f Ts t) + | Free (x as (_, T)) => + Rel (arity_of_set scope T, find_index (equal x) frees) + | Term.Var _ => raise NOT_SUPPORTED "schematic variables" + | Bound j => raise SAME () + | Abs (_, T, t') => + (case fastype_of1 (T :: Ts, t') of + @{typ bool} => Comprehension (decls_for_one scope Ts T, + to_f (T :: Ts) t') + | T' => Comprehension (decls_for_one scope Ts T @ + decls_for_set scope (T :: Ts) T', + Subset (set_from_bound_var (T' :: T :: Ts) 0, + to_set (T :: Ts) t'))) + | t1 $ t2 => + (case fastype_of1 (Ts, t) of + @{typ bool} => atom_from_formula (to_f Ts t) + | T => + let val T2 = fastype_of1 (Ts, t2) in + case arity_of_one scope T2 of + 1 => Join (to_one Ts t2, to_set Ts t1) + | n => + let + val arity2 = arity_of_one scope T2 + val res_arity = arity_of_set scope T + in + Project (Intersect + (Product (to_one Ts t2, + atom_schema_of_set scope T + |> map (AtomSeq o rpair 0) |> foldl1 Product), + to_set Ts t1), + num_seq arity2 res_arity) + end + end) + | _ => raise NOT_SUPPORTED ("term " ^ + quote (Syntax.string_of_term ctxt t))) + handle SAME () => set_from_one (fastype_of1 (Ts, t)) (to_one Ts t) + in to_f [] end + +(* (typ -> int) -> int -> styp -> bound *) +fun bound_for_free scope i (s, T) = + let val js = atom_schema_of_set scope T in + ([((length js, i), s)], + [TupleSet [], atom_schema_of_set scope T |> map (rpair 0) + |> tuple_set_from_atom_schema]) + end + +(* (typ -> int) -> typ list -> typ -> rel_expr -> formula *) +fun declarative_axiom_for_rel_expr scope Ts (Type ("fun", [T1, T2])) r = + if body_type T2 = bool_T then + True + else + All (decls_for_one scope Ts T1, + declarative_axiom_for_rel_expr scope (T1 :: Ts) T2 + (List.foldl Join r (one_vars_for_bound_var scope (T1 :: Ts) 0))) + | declarative_axiom_for_rel_expr _ _ _ r = One r + +(* (typ -> int) -> int -> styp -> formula *) +fun declarative_axiom_for_free scope i (_, T) = + declarative_axiom_for_rel_expr scope [] T (Rel (arity_of_set scope T, i)) + +(* Proof.context -> (typ -> int) -> term -> string *) +fun pick_nits_in_term ctxt raw_scope t = + let + val thy = ProofContext.theory_of ctxt + (* typ -> int *) + fun scope (Type ("fun", [T1, T2])) = reasonable_power (scope T2) (scope T1) + | scope (Type ("*", [T1, T2])) = scope T1 * scope T2 + | scope @{typ bool} = 2 + | scope T = Int.max (1, raw_scope T) + val neg_t = @{const Not} $ ObjectLogic.atomize_term thy t + val _ = fold_types (K o check_type ctxt) neg_t () + val frees = Term.add_frees neg_t [] + val bounds = map2 (bound_for_free scope) (index_seq 0 (length frees)) frees + val declarative_axioms = + map2 (declarative_axiom_for_free scope) (index_seq 0 (length frees)) + frees + val formula = kodkod_formula_for_term ctxt scope frees neg_t + |> fold_rev (curry And) declarative_axioms + val univ_card = univ_card 0 0 0 bounds formula + val problem = + {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [], + bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula} + in + case solve_any_problem true NONE 0 1 [problem] of + Normal ([], _) => "none" + | Normal _ => "genuine" + | TimedOut _ => "unknown" + | Interrupted _ => "unknown" + | Error (s, _) => error ("Kodkod error: " ^ s) + end + handle NOT_SUPPORTED details => + (warning ("Unsupported case: " ^ details ^ "."); "unknown") +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,857 @@ +(* Title: HOL/Nitpick/Tools/nitpick.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Finite model generation for HOL formulas using Kodkod. +*) + +signature NITPICK = +sig + type params = { + cards_assigns: (typ option * int list) list, + maxes_assigns: (styp option * int list) list, + iters_assigns: (styp option * int list) list, + bisim_depths: int list, + boxes: (typ option * bool option) list, + monos: (typ option * bool option) list, + wfs: (styp option * bool option) list, + sat_solver: string, + blocking: bool, + falsify: bool, + debug: bool, + verbose: bool, + overlord: bool, + user_axioms: bool option, + assms: bool, + coalesce_type_vars: bool, + destroy_constrs: bool, + specialize: bool, + skolemize: bool, + star_linear_preds: bool, + uncurry: bool, + fast_descrs: bool, + peephole_optim: bool, + timeout: Time.time option, + tac_timeout: Time.time option, + sym_break: int, + sharing_depth: int, + flatten_props: bool, + max_threads: int, + show_skolems: bool, + show_datatypes: bool, + show_consts: bool, + evals: term list, + formats: (term option * int list) list, + max_potential: int, + max_genuine: int, + check_potential: bool, + check_genuine: bool, + batch_size: int, + expect: string} + + val register_frac_type : string -> (string * string) list -> theory -> theory + val unregister_frac_type : string -> theory -> theory + val register_codatatype : typ -> string -> styp list -> theory -> theory + val unregister_codatatype : typ -> theory -> theory + val pick_nits_in_term : + Proof.state -> params -> bool -> term list -> term -> string * Proof.state + val pick_nits_in_subgoal : + Proof.state -> params -> bool -> int -> string * Proof.state +end; + +structure Nitpick : NITPICK = +struct + +open NitpickUtil +open NitpickHOL +open NitpickMono +open NitpickScope +open NitpickPeephole +open NitpickRep +open NitpickNut +open NitpickKodkod +open NitpickModel + +type params = { + cards_assigns: (typ option * int list) list, + maxes_assigns: (styp option * int list) list, + iters_assigns: (styp option * int list) list, + bisim_depths: int list, + boxes: (typ option * bool option) list, + monos: (typ option * bool option) list, + wfs: (styp option * bool option) list, + sat_solver: string, + blocking: bool, + falsify: bool, + debug: bool, + verbose: bool, + overlord: bool, + user_axioms: bool option, + assms: bool, + coalesce_type_vars: bool, + destroy_constrs: bool, + specialize: bool, + skolemize: bool, + star_linear_preds: bool, + uncurry: bool, + fast_descrs: bool, + peephole_optim: bool, + timeout: Time.time option, + tac_timeout: Time.time option, + sym_break: int, + sharing_depth: int, + flatten_props: bool, + max_threads: int, + show_skolems: bool, + show_datatypes: bool, + show_consts: bool, + evals: term list, + formats: (term option * int list) list, + max_potential: int, + max_genuine: int, + check_potential: bool, + check_genuine: bool, + batch_size: int, + expect: string} + +type problem_extension = { + free_names: nut list, + sel_names: nut list, + nonsel_names: nut list, + rel_table: nut NameTable.table, + liberal: bool, + scope: scope, + core: Kodkod.formula, + defs: Kodkod.formula list} + +type rich_problem = Kodkod.problem * problem_extension + +(* Proof.context -> string -> term list -> Pretty.T list *) +fun pretties_for_formulas _ _ [] = [] + | pretties_for_formulas ctxt s ts = + [Pretty.str (s ^ plural_s_for_list ts ^ ":"), + Pretty.indent indent_size (Pretty.chunks + (map2 (fn j => fn t => + Pretty.block [t |> shorten_const_names_in_term + |> Syntax.pretty_term ctxt, + Pretty.str (if j = 1 then "." else ";")]) + (length ts downto 1) ts))] + +val max_liberal_delay_ms = 200 +val max_liberal_delay_percent = 2 + +(* Time.time option -> int *) +fun liberal_delay_for_timeout NONE = max_liberal_delay_ms + | liberal_delay_for_timeout (SOME timeout) = + Int.max (0, Int.min (max_liberal_delay_ms, + Time.toMilliseconds timeout + * max_liberal_delay_percent div 100)) + +(* Time.time option -> bool *) +fun passed_deadline NONE = false + | passed_deadline (SOME time) = Time.compare (Time.now (), time) <> LESS + +(* ('a * bool option) list -> bool *) +fun none_true asgns = forall (not_equal (SOME true) o snd) asgns + +val weaselly_sorts = + [@{sort default}, @{sort zero}, @{sort one}, @{sort plus}, @{sort minus}, + @{sort uminus}, @{sort times}, @{sort inverse}, @{sort abs}, @{sort sgn}, + @{sort ord}, @{sort eq}, @{sort number}] +(* theory -> typ -> bool *) +fun is_tfree_with_weaselly_sort thy (TFree (_, S)) = + exists (curry (Sign.subsort thy) S) weaselly_sorts + | is_tfree_with_weaselly_sort _ _ = false +(* theory term -> bool *) +val has_weaselly_sorts = + exists_type o exists_subtype o is_tfree_with_weaselly_sort + +(* Time.time -> Proof.state -> params -> bool -> term -> string * Proof.state *) +fun pick_them_nits_in_term deadline state (params : params) auto orig_assm_ts + orig_t = + let + val timer = Timer.startRealTimer () + val thy = Proof.theory_of state + val ctxt = Proof.context_of state + val {cards_assigns, maxes_assigns, iters_assigns, bisim_depths, boxes, + monos, wfs, sat_solver, blocking, falsify, debug, verbose, overlord, + user_axioms, assms, coalesce_type_vars, destroy_constrs, specialize, + skolemize, star_linear_preds, uncurry, fast_descrs, peephole_optim, + tac_timeout, sym_break, sharing_depth, flatten_props, max_threads, + show_skolems, show_datatypes, show_consts, evals, formats, + max_potential, max_genuine, check_potential, check_genuine, batch_size, + ...} = + params + val state_ref = Unsynchronized.ref state + (* Pretty.T -> unit *) + val pprint = + if auto then + Unsynchronized.change state_ref o Proof.goal_message o K + o curry Pretty.blk 0 o cons (Pretty.str "") o single + o Pretty.mark Markup.hilite + else + priority o Pretty.string_of + (* (unit -> Pretty.T) -> unit *) + fun pprint_m f = () |> not auto ? pprint o f + fun pprint_v f = () |> verbose ? pprint o f + fun pprint_d f = () |> debug ? pprint o f + (* string -> unit *) + val print = pprint o curry Pretty.blk 0 o pstrs + (* (unit -> string) -> unit *) + fun print_m f = pprint_m (curry Pretty.blk 0 o pstrs o f) + fun print_v f = pprint_v (curry Pretty.blk 0 o pstrs o f) + fun print_d f = pprint_d (curry Pretty.blk 0 o pstrs o f) + + (* unit -> unit *) + fun check_deadline () = + if debug andalso passed_deadline deadline then raise TimeLimit.TimeOut + else () + (* unit -> 'a *) + fun do_interrupted () = + if passed_deadline deadline then raise TimeLimit.TimeOut + else raise Interrupt + + val _ = print_m (K "Nitpicking...") + val neg_t = if falsify then Logic.mk_implies (orig_t, @{prop False}) + else orig_t + val assms_t = if assms orelse auto then + Logic.mk_conjunction_list (neg_t :: orig_assm_ts) + else + neg_t + val (assms_t, evals) = + assms_t :: evals + |> coalesce_type_vars ? coalesce_type_vars_in_terms + |> hd pairf tl + val original_max_potential = max_potential + val original_max_genuine = max_genuine +(* + val _ = priority ("*** " ^ Syntax.string_of_term ctxt orig_t) + val _ = List.app (fn t => priority ("*** " ^ Syntax.string_of_term ctxt t)) + orig_assm_ts +*) + val max_bisim_depth = fold Integer.max bisim_depths ~1 + val case_names = case_const_names thy + val (defs, built_in_nondefs, user_nondefs) = all_axioms_of thy + val def_table = const_def_table ctxt defs + val nondef_table = const_nondef_table (built_in_nondefs @ user_nondefs) + val simp_table = Unsynchronized.ref (const_simp_table ctxt) + val psimp_table = const_psimp_table ctxt + val intro_table = inductive_intro_table ctxt def_table + val ground_thm_table = ground_theorem_table thy + val ersatz_table = ersatz_table thy + val (ext_ctxt as {skolems, special_funs, wf_cache, ...}) = + {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes, + user_axioms = user_axioms, debug = debug, wfs = wfs, + destroy_constrs = destroy_constrs, specialize = specialize, + skolemize = skolemize, star_linear_preds = star_linear_preds, + uncurry = uncurry, fast_descrs = fast_descrs, tac_timeout = tac_timeout, + evals = evals, case_names = case_names, def_table = def_table, + nondef_table = nondef_table, user_nondefs = user_nondefs, + simp_table = simp_table, psimp_table = psimp_table, + intro_table = intro_table, ground_thm_table = ground_thm_table, + ersatz_table = ersatz_table, skolems = Unsynchronized.ref [], + special_funs = Unsynchronized.ref [], + unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref []} + val frees = Term.add_frees assms_t [] + val _ = null (Term.add_tvars assms_t []) + orelse raise NOT_SUPPORTED "schematic type variables" + val (((def_ts, nondef_ts), (got_all_mono_user_axioms, no_poly_user_axioms)), + core_t) = preprocess_term ext_ctxt assms_t + val got_all_user_axioms = + got_all_mono_user_axioms andalso no_poly_user_axioms + + (* styp * (bool * bool) -> unit *) + fun print_wf (x, (gfp, wf)) = + pprint (Pretty.blk (0, + pstrs ("The " ^ (if gfp then "co" else "") ^ "inductive predicate \"") + @ Syntax.pretty_term ctxt (Const x) :: + pstrs (if wf then + "\" was proved well-founded. Nitpick can compute it \ + \efficiently." + else + "\" could not be proved well-founded. Nitpick might need to \ + \unroll it."))) + val _ = if verbose then List.app print_wf (!wf_cache) else () + val _ = + pprint_d (fn () => + Pretty.chunks + (pretties_for_formulas ctxt "Preprocessed formula" [core_t] @ + pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @ + pretties_for_formulas ctxt "Relevant nondefinitional axiom" + nondef_ts)) + val _ = List.app (ignore o Term.type_of) (core_t :: def_ts @ nondef_ts) + handle TYPE (_, Ts, ts) => + raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts) + + val unique_scope = forall (equal 1 o length o snd) cards_assigns + (* typ -> bool *) + fun is_free_type_monotonic T = + unique_scope orelse + case triple_lookup (type_match thy) monos T of + SOME (SOME b) => b + | _ => formulas_monotonic ext_ctxt T def_ts nondef_ts core_t + fun is_datatype_monotonic T = + unique_scope orelse + case triple_lookup (type_match thy) monos T of + SOME (SOME b) => b + | _ => + not (is_pure_typedef thy T) orelse is_univ_typedef thy T + orelse is_number_type thy T + orelse formulas_monotonic ext_ctxt T def_ts nondef_ts core_t + val Ts = ground_types_in_terms ext_ctxt (core_t :: def_ts @ nondef_ts) + |> sort TermOrd.typ_ord + val (all_dataTs, all_free_Ts) = + List.partition (is_integer_type orf is_datatype thy) Ts + val (mono_dataTs, nonmono_dataTs) = + List.partition is_datatype_monotonic all_dataTs + val (mono_free_Ts, nonmono_free_Ts) = + List.partition is_free_type_monotonic all_free_Ts + + val _ = + if not unique_scope andalso not (null mono_free_Ts) then + print_v (fn () => + let + val ss = map (quote o string_for_type ctxt) mono_free_Ts + in + "The type" ^ plural_s_for_list ss ^ " " ^ + space_implode " " (serial_commas "and" ss) ^ " " ^ + (if none_true monos then + "passed the monotonicity test" + else + (if length ss = 1 then "is" else "are") ^ + " considered monotonic") ^ + ". Nitpick might be able to skip some scopes." + end) + else + () + val mono_Ts = mono_dataTs @ mono_free_Ts + val nonmono_Ts = nonmono_dataTs @ nonmono_free_Ts + +(* + val _ = priority "Monotonic datatypes:" + val _ = List.app (priority o string_for_type ctxt) mono_dataTs + val _ = priority "Nonmonotonic datatypes:" + val _ = List.app (priority o string_for_type ctxt) nonmono_dataTs + val _ = priority "Monotonic free types:" + val _ = List.app (priority o string_for_type ctxt) mono_free_Ts + val _ = priority "Nonmonotonic free types:" + val _ = List.app (priority o string_for_type ctxt) nonmono_free_Ts +*) + + val core_u = nut_from_term thy fast_descrs (!special_funs) Eq core_t + val def_us = map (nut_from_term thy fast_descrs (!special_funs) DefEq) + def_ts + val nondef_us = map (nut_from_term thy fast_descrs (!special_funs) Eq) + nondef_ts + val (free_names, const_names) = + fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], []) + val nonsel_names = filter_out (is_sel o nickname_of) const_names + val would_be_genuine = got_all_user_axioms andalso none_true wfs +(* + val _ = List.app (priority o string_for_nut ctxt) + (core_u :: def_us @ nondef_us) +*) + val need_incremental = Int.max (max_potential, max_genuine) >= 2 + val effective_sat_solver = + if sat_solver <> "smart" then + if need_incremental andalso + not (sat_solver mem KodkodSAT.configured_sat_solvers true) then + (print_m (K ("An incremental SAT solver is required: \"SAT4J\" will \ + \be used instead of " ^ quote sat_solver ^ ".")); + "SAT4J") + else + sat_solver + else + KodkodSAT.smart_sat_solver_name need_incremental + val _ = + if sat_solver = "smart" then + print_v (fn () => "Using SAT solver " ^ quote effective_sat_solver ^ + ". The following" ^ + (if need_incremental then " incremental " else " ") ^ + "solvers are configured: " ^ + commas (map quote (KodkodSAT.configured_sat_solvers + need_incremental)) ^ ".") + else + () + + val too_big_scopes = Unsynchronized.ref [] + + (* bool -> scope -> rich_problem option *) + fun problem_for_scope liberal + (scope as {card_assigns, bisim_depth, datatypes, ofs, ...}) = + let + val _ = not (exists (fn other => scope_less_eq other scope) + (!too_big_scopes)) + orelse raise LIMIT ("Nitpick.pick_them_nits_in_term.\ + \problem_for_scope", "too big scope") +(* + val _ = priority "Offsets:" + val _ = List.app (fn (T, j0) => + priority (string_for_type ctxt T ^ " = " ^ + string_of_int j0)) + (Typtab.dest ofs) +*) + val all_precise = forall (is_precise_type datatypes) Ts + (* nut list -> rep NameTable.table -> nut list * rep NameTable.table *) + val repify_consts = choose_reps_for_consts scope all_precise + val main_j0 = offset_of_type ofs bool_T + val (nat_card, nat_j0) = spec_of_type scope nat_T + val (int_card, int_j0) = spec_of_type scope int_T + val _ = forall (equal main_j0) [nat_j0, int_j0] + orelse raise BAD ("Nitpick.pick_them_nits_in_term.\ + \problem_for_scope", "bad offsets") + val kk = kodkod_constrs peephole_optim nat_card int_card main_j0 + val (free_names, rep_table) = + choose_reps_for_free_vars scope free_names NameTable.empty + val (sel_names, rep_table) = choose_reps_for_all_sels scope rep_table + val (nonsel_names, rep_table) = repify_consts nonsel_names rep_table + val min_highest_arity = + NameTable.fold (curry Int.max o arity_of_rep o snd) rep_table 1 + val min_univ_card = + NameTable.fold (curry Int.max o min_univ_card_of_rep o snd) rep_table + (univ_card nat_card int_card main_j0 [] Kodkod.True) + val _ = check_arity min_univ_card min_highest_arity + + val core_u = choose_reps_in_nut scope liberal rep_table false core_u + val def_us = map (choose_reps_in_nut scope liberal rep_table true) + def_us + val nondef_us = map (choose_reps_in_nut scope liberal rep_table false) + nondef_us +(* + val _ = List.app (priority o string_for_nut ctxt) + (free_names @ sel_names @ nonsel_names @ + core_u :: def_us @ nondef_us) +*) + val (free_rels, pool, rel_table) = + rename_free_vars free_names initial_pool NameTable.empty + val (sel_rels, pool, rel_table) = + rename_free_vars sel_names pool rel_table + val (other_rels, pool, rel_table) = + rename_free_vars nonsel_names pool rel_table + val core_u = rename_vars_in_nut pool rel_table core_u + val def_us = map (rename_vars_in_nut pool rel_table) def_us + val nondef_us = map (rename_vars_in_nut pool rel_table) nondef_us + (* nut -> Kodkod.formula *) + val to_f = kodkod_formula_from_nut ofs liberal kk + val core_f = to_f core_u + val def_fs = map to_f def_us + val nondef_fs = map to_f nondef_us + val formula = fold (fold s_and) [def_fs, nondef_fs] core_f + val comment = (if liberal then "liberal" else "conservative") ^ "\n" ^ + PrintMode.setmp [] multiline_string_for_scope scope + val kodkod_sat_solver = KodkodSAT.sat_solver_spec effective_sat_solver + |> snd + val delay = if liberal then + Option.map (fn time => Time.- (time, Time.now ())) + deadline + |> liberal_delay_for_timeout + else + 0 + val settings = [("solver", commas (map quote kodkod_sat_solver)), + ("skolem_depth", "-1"), + ("bit_width", "16"), + ("symmetry_breaking", signed_string_of_int sym_break), + ("sharing", signed_string_of_int sharing_depth), + ("flatten", Bool.toString flatten_props), + ("delay", signed_string_of_int delay)] + val plain_rels = free_rels @ other_rels + val plain_bounds = map (bound_for_plain_rel ctxt debug) plain_rels + val plain_axioms = map (declarative_axiom_for_plain_rel kk) plain_rels + val sel_bounds = map (bound_for_sel_rel ctxt debug datatypes) sel_rels + val dtype_axioms = declarative_axioms_for_datatypes ext_ctxt ofs kk + rel_table datatypes + val declarative_axioms = plain_axioms @ dtype_axioms + val univ_card = univ_card nat_card int_card main_j0 + (plain_bounds @ sel_bounds) formula + val built_in_bounds = bounds_for_built_in_rels_in_formula debug + univ_card nat_card int_card main_j0 formula + val bounds = built_in_bounds @ plain_bounds @ sel_bounds + |> not debug ? merge_bounds + val highest_arity = + fold Integer.max (map (fst o fst) (maps fst bounds)) 0 + val formula = fold_rev s_and declarative_axioms formula + val _ = if formula = Kodkod.False then () + else check_arity univ_card highest_arity + in + SOME ({comment = comment, settings = settings, univ_card = univ_card, + tuple_assigns = [], bounds = bounds, + int_bounds = sequential_int_bounds univ_card, + expr_assigns = [], formula = formula}, + {free_names = free_names, sel_names = sel_names, + nonsel_names = nonsel_names, rel_table = rel_table, + liberal = liberal, scope = scope, core = core_f, + defs = nondef_fs @ def_fs @ declarative_axioms}) + end + handle LIMIT (loc, msg) => + if loc = "NitpickKodkod.check_arity" + andalso not (Typtab.is_empty ofs) then + problem_for_scope liberal + {ext_ctxt = ext_ctxt, card_assigns = card_assigns, + bisim_depth = bisim_depth, datatypes = datatypes, + ofs = Typtab.empty} + else if loc = "Nitpick.pick_them_nits_in_term.\ + \problem_for_scope" then + NONE + else + (Unsynchronized.change too_big_scopes (cons scope); + print_v (fn () => ("Limit reached: " ^ msg ^ + ". Dropping " ^ (if liberal then "potential" + else "genuine") ^ + " component of scope.")); + NONE) + + (* int -> (''a * int list list) list -> ''a -> Kodkod.tuple_set *) + fun tuple_set_for_rel univ_card = + Kodkod.TupleSet o map (kk_tuple debug univ_card) o the + oo AList.lookup (op =) + + val word_model = if falsify then "counterexample" else "model" + + val scopes = Unsynchronized.ref [] + val generated_scopes = Unsynchronized.ref [] + val generated_problems = Unsynchronized.ref [] + val checked_problems = Unsynchronized.ref (SOME []) + val met_potential = Unsynchronized.ref 0 + + (* rich_problem list -> int list -> unit *) + fun update_checked_problems problems = + List.app (Unsynchronized.change checked_problems o Option.map o cons + o nth problems) + + (* bool -> Kodkod.raw_bound list -> problem_extension -> bool option *) + fun print_and_check_model genuine bounds + ({free_names, sel_names, nonsel_names, rel_table, scope, ...} + : problem_extension) = + let + val (reconstructed_model, codatatypes_ok) = + reconstruct_hol_model {show_skolems = show_skolems, + show_datatypes = show_datatypes, + show_consts = show_consts} + scope formats frees free_names sel_names nonsel_names rel_table + bounds + val would_be_genuine = would_be_genuine andalso codatatypes_ok + in + pprint (Pretty.chunks + [Pretty.blk (0, + (pstrs ("Nitpick found a" ^ + (if not genuine then " potential " + else if would_be_genuine then " " + else " likely genuine ") ^ word_model) @ + (case pretties_for_scope scope verbose of + [] => [] + | pretties => pstrs " for " @ pretties) @ + [Pretty.str ":\n"])), + Pretty.indent indent_size reconstructed_model]); + if genuine then + (if check_genuine then + (case prove_hol_model scope tac_timeout free_names sel_names + rel_table bounds assms_t of + SOME true => print ("Confirmation by \"auto\": The above " ^ + word_model ^ " is really genuine.") + | SOME false => + if would_be_genuine then + error ("A supposedly genuine " ^ word_model ^ " was shown to\ + \be spurious by \"auto\".\nThis should never happen.\n\ + \Please send a bug report to blanchet\ + \te@in.tum.de.") + else + print ("Refutation by \"auto\": The above " ^ word_model ^ + " is spurious.") + | NONE => print "No confirmation by \"auto\".") + else + (); + if has_weaselly_sorts thy orig_t then + print "Hint: Maybe you forgot a type constraint?" + else + (); + if not would_be_genuine then + if no_poly_user_axioms then + let + val options = + [] |> not got_all_mono_user_axioms + ? cons ("user_axioms", "\"true\"") + |> not (none_true wfs) + ? cons ("wf", "\"smart\" or \"false\"") + |> not codatatypes_ok + ? cons ("bisim_depth", "a nonnegative value") + val ss = + map (fn (name, value) => quote name ^ " set to " ^ value) + options + in + print ("Try again with " ^ + space_implode " " (serial_commas "and" ss) ^ + " to confirm that the " ^ word_model ^ " is genuine.") + end + else + print ("Nitpick is unable to guarantee the authenticity of \ + \the " ^ word_model ^ " in the presence of polymorphic \ + \axioms.") + else + (); + NONE) + else + if not genuine then + (Unsynchronized.inc met_potential; + if check_potential then + let + val status = prove_hol_model scope tac_timeout free_names + sel_names rel_table bounds assms_t + in + (case status of + SOME true => print ("Confirmation by \"auto\": The above " ^ + word_model ^ " is genuine.") + | SOME false => print ("Refutation by \"auto\": The above " ^ + word_model ^ " is spurious.") + | NONE => print "No confirmation by \"auto\"."); + status + end + else + NONE) + else + NONE + end + (* int -> int -> int -> bool -> rich_problem list -> int * int * int *) + fun solve_any_problem max_potential max_genuine donno first_time problems = + let + val max_potential = Int.max (0, max_potential) + val max_genuine = Int.max (0, max_genuine) + (* bool -> int * Kodkod.raw_bound list -> bool option *) + fun print_and_check genuine (j, bounds) = + print_and_check_model genuine bounds (snd (nth problems j)) + val max_solutions = max_potential + max_genuine + |> not need_incremental ? curry Int.min 1 + in + if max_solutions <= 0 then + (0, 0, donno) + else + case Kodkod.solve_any_problem overlord deadline max_threads + max_solutions (map fst problems) of + Kodkod.Normal ([], unsat_js) => + (update_checked_problems problems unsat_js; + (max_potential, max_genuine, donno)) + | Kodkod.Normal (sat_ps, unsat_js) => + let + val (lib_ps, con_ps) = + List.partition (#liberal o snd o nth problems o fst) sat_ps + in + update_checked_problems problems (unsat_js @ map fst lib_ps); + if null con_ps then + let + val num_genuine = Library.take (max_potential, lib_ps) + |> map (print_and_check false) + |> filter (equal (SOME true)) |> length + val max_genuine = max_genuine - num_genuine + val max_potential = max_potential + - (length lib_ps - num_genuine) + in + if max_genuine <= 0 then + (0, 0, donno) + else + let + (* "co_js" is the list of conservative problems whose + liberal pendants couldn't be satisfied and hence that + most probably can't be satisfied themselves. *) + val co_js = + map (fn j => j - 1) unsat_js + |> filter (fn j => + j >= 0 andalso + scopes_equivalent + (#scope (snd (nth problems j))) + (#scope (snd (nth problems (j + 1))))) + val bye_js = sort_distinct int_ord (map fst sat_ps @ + unsat_js @ co_js) + val problems = + problems |> filter_out_indices bye_js + |> max_potential <= 0 + ? filter_out (#liberal o snd) + in + solve_any_problem max_potential max_genuine donno false + problems + end + end + else + let + val _ = Library.take (max_genuine, con_ps) + |> List.app (ignore o print_and_check true) + val max_genuine = max_genuine - length con_ps + in + if max_genuine <= 0 orelse not first_time then + (0, max_genuine, donno) + else + let + val bye_js = sort_distinct int_ord + (map fst sat_ps @ unsat_js) + val problems = + problems |> filter_out_indices bye_js + |> filter_out (#liberal o snd) + in solve_any_problem 0 max_genuine donno false problems end + end + end + | Kodkod.TimedOut unsat_js => + (update_checked_problems problems unsat_js; raise TimeLimit.TimeOut) + | Kodkod.Interrupted NONE => + (checked_problems := NONE; do_interrupted ()) + | Kodkod.Interrupted (SOME unsat_js) => + (update_checked_problems problems unsat_js; do_interrupted ()) + | Kodkod.Error (s, unsat_js) => + (update_checked_problems problems unsat_js; + print_v (K ("Kodkod error: " ^ s ^ ".")); + (max_potential, max_genuine, donno + 1)) + end + + (* int -> int -> scope list -> int * int * int -> int * int * int *) + fun run_batch j n scopes (max_potential, max_genuine, donno) = + let + val _ = + if null scopes then + print_m (K "The scope specification is inconsistent.") + else if verbose then + pprint (Pretty.chunks + [Pretty.blk (0, + pstrs ((if n > 1 then + "Batch " ^ string_of_int (j + 1) ^ " of " ^ + signed_string_of_int n ^ ": " + else + "") ^ + "Trying " ^ string_of_int (length scopes) ^ + " scope" ^ plural_s_for_list scopes ^ ":")), + Pretty.indent indent_size + (Pretty.chunks (map2 + (fn j => fn scope => + Pretty.block ( + (case pretties_for_scope scope true of + [] => [Pretty.str "Empty"] + | pretties => pretties) @ + [Pretty.str (if j = 1 then "." else ";")])) + (length scopes downto 1) scopes))]) + else + () + (* scope * bool -> rich_problem list * bool + -> rich_problem list * bool *) + fun add_problem_for_scope (scope as {datatypes, ...}, liberal) + (problems, donno) = + (check_deadline (); + case problem_for_scope liberal scope of + SOME problem => + (problems + |> (null problems orelse + not (Kodkod.problems_equivalent (fst problem) + (fst (hd problems)))) + ? cons problem, donno) + | NONE => (problems, donno + 1)) + val (problems, donno) = + fold add_problem_for_scope + (map_product pair scopes + ((if max_genuine > 0 then [false] else []) @ + (if max_potential > 0 then [true] else []))) + ([], donno) + val _ = Unsynchronized.change generated_problems (append problems) + val _ = Unsynchronized.change generated_scopes (append scopes) + in + solve_any_problem max_potential max_genuine donno true (rev problems) + end + + (* rich_problem list -> scope -> int *) + fun scope_count (problems : rich_problem list) scope = + length (filter (scopes_equivalent scope o #scope o snd) problems) + (* string -> string *) + fun excipit did_so_and_so = + let + (* rich_problem list -> rich_problem list *) + val do_filter = + if !met_potential = max_potential then filter_out (#liberal o snd) + else I + val total = length (!scopes) + val unsat = + fold (fn scope => + case scope_count (do_filter (!generated_problems)) scope of + 0 => I + | n => + if scope_count (do_filter (these (!checked_problems))) + scope = n then + Integer.add 1 + else + I) (!generated_scopes) 0 + in + "Nitpick " ^ did_so_and_so ^ + (if is_some (!checked_problems) andalso total > 0 then + " after checking " ^ + string_of_int (Int.min (total - 1, unsat)) ^ " of " ^ + string_of_int total ^ " scope" ^ plural_s total + else + "") ^ "." + end + + (* int -> int -> scope list -> int * int * int -> Kodkod.outcome *) + fun run_batches _ _ [] (max_potential, max_genuine, donno) = + if donno > 0 andalso max_genuine > 0 then + (print_m (fn () => excipit "ran out of resources"); "unknown") + else if max_genuine = original_max_genuine then + if max_potential = original_max_potential then + (print_m (K ("Nitpick found no " ^ word_model ^ ".")); "none") + else + (print_m (K ("Nitpick could not find " ^ + (if max_genuine = 1 then "a better " ^ word_model ^ "." + else "any better " ^ word_model ^ "s."))); + "potential") + else + if would_be_genuine then "genuine" else "likely_genuine" + | run_batches j n (batch :: batches) z = + let val (z as (_, max_genuine, _)) = run_batch j n batch z in + run_batches (j + 1) n (if max_genuine > 0 then batches else []) z + end + + val _ = scopes := all_scopes ext_ctxt sym_break cards_assigns maxes_assigns + iters_assigns bisim_depths mono_Ts nonmono_Ts + val batches = batch_list batch_size (!scopes) + val outcome_code = + (run_batches 0 (length batches) batches (max_potential, max_genuine, 0) + handle Exn.Interrupt => do_interrupted ()) + handle TimeLimit.TimeOut => + (print_m (fn () => excipit "ran out of time"); + if !met_potential > 0 then "potential" else "unknown") + | Exn.Interrupt => if auto orelse debug then raise Interrupt + else error (excipit "was interrupted") + val _ = print_v (fn () => "Total time: " ^ + signed_string_of_int (Time.toMilliseconds + (Timer.checkRealTimer timer)) ^ " ms.") + in (outcome_code, !state_ref) end + handle Exn.Interrupt => + if auto orelse #debug params then + raise Interrupt + else + if passed_deadline deadline then + (priority "Nitpick ran out of time."; ("unknown", state)) + else + error "Nitpick was interrupted." + +(* Proof.state -> params -> bool -> term -> string * Proof.state *) +fun pick_nits_in_term state (params as {debug, timeout, expect, ...}) + auto orig_assm_ts orig_t = + let + val deadline = Option.map (curry Time.+ (Time.now ())) timeout + val outcome as (outcome_code, _) = + time_limit (if debug then NONE else timeout) + (pick_them_nits_in_term deadline state params auto orig_assm_ts) + orig_t + in + if expect = "" orelse outcome_code = expect then outcome + else error ("Unexpected outcome: " ^ quote outcome_code ^ ".") + end + +(* Proof.state -> params -> thm -> int -> string * Proof.state *) +fun pick_nits_in_subgoal state params auto subgoal = + let + val ctxt = Proof.context_of state + val t = state |> Proof.get_goal |> snd |> snd |> prop_of + in + if Logic.count_prems t = 0 then + (priority "No subgoal!"; ("none", state)) + else + let + val assms = map term_of (Assumption.all_assms_of ctxt) + val (t, frees) = Logic.goal_params t subgoal + in pick_nits_in_term state params auto assms (subst_bounds (frees, t)) end + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_hol.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,3328 @@ +(* Title: HOL/Nitpick/Tools/nitpick_hol.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Auxiliary HOL-related functions used by Nitpick. +*) + +signature NITPICK_HOL = +sig + type const_table = term list Symtab.table + type special_fun = (styp * int list * term list) * styp + type unrolled = styp * styp + type wf_cache = (styp * (bool * bool)) list + + type extended_context = { + thy: theory, + ctxt: Proof.context, + max_bisim_depth: int, + boxes: (typ option * bool option) list, + wfs: (styp option * bool option) list, + user_axioms: bool option, + debug: bool, + destroy_constrs: bool, + specialize: bool, + skolemize: bool, + star_linear_preds: bool, + uncurry: bool, + fast_descrs: bool, + tac_timeout: Time.time option, + evals: term list, + case_names: (string * int) list, + def_table: const_table, + nondef_table: const_table, + user_nondefs: term list, + simp_table: const_table Unsynchronized.ref, + psimp_table: const_table, + intro_table: const_table, + ground_thm_table: term list Inttab.table, + ersatz_table: (string * string) list, + skolems: (string * string list) list Unsynchronized.ref, + special_funs: special_fun list Unsynchronized.ref, + unrolled_preds: unrolled list Unsynchronized.ref, + wf_cache: wf_cache Unsynchronized.ref} + + val name_sep : string + val numeral_prefix : string + val skolem_prefix : string + val eval_prefix : string + val original_name : string -> string + val unbox_type : typ -> typ + val string_for_type : Proof.context -> typ -> string + val prefix_name : string -> string -> string + val short_name : string -> string + val short_const_name : string -> string + val shorten_const_names_in_term : term -> term + val type_match : theory -> typ * typ -> bool + val const_match : theory -> styp * styp -> bool + val term_match : theory -> term * term -> bool + val is_TFree : typ -> bool + val is_higher_order_type : typ -> bool + val is_fun_type : typ -> bool + val is_set_type : typ -> bool + val is_pair_type : typ -> bool + val is_lfp_iterator_type : typ -> bool + val is_gfp_iterator_type : typ -> bool + val is_fp_iterator_type : typ -> bool + val is_boolean_type : typ -> bool + val is_integer_type : typ -> bool + val is_record_type : typ -> bool + val is_number_type : theory -> typ -> bool + val const_for_iterator_type : typ -> styp + val nth_range_type : int -> typ -> typ + val num_factors_in_type : typ -> int + val num_binder_types : typ -> int + val curried_binder_types : typ -> typ list + val mk_flat_tuple : typ -> term list -> term + val dest_n_tuple : int -> term -> term list + val instantiate_type : theory -> typ -> typ -> typ -> typ + val is_codatatype : theory -> typ -> bool + val is_pure_typedef : theory -> typ -> bool + val is_univ_typedef : theory -> typ -> bool + val is_datatype : theory -> typ -> bool + val is_record_constr : styp -> bool + val is_record_get : theory -> styp -> bool + val is_record_update : theory -> styp -> bool + val is_abs_fun : theory -> styp -> bool + val is_rep_fun : theory -> styp -> bool + val is_constr : theory -> styp -> bool + val is_sel : string -> bool + val discr_for_constr : styp -> styp + val num_sels_for_constr_type : typ -> int + val nth_sel_name_for_constr_name : string -> int -> string + val nth_sel_for_constr : styp -> int -> styp + val boxed_nth_sel_for_constr : extended_context -> styp -> int -> styp + val sel_no_from_name : string -> int + val eta_expand : typ list -> term -> int -> term + val extensionalize : term -> term + val distinctness_formula : typ -> term list -> term + val register_frac_type : string -> (string * string) list -> theory -> theory + val unregister_frac_type : string -> theory -> theory + val register_codatatype : typ -> string -> styp list -> theory -> theory + val unregister_codatatype : typ -> theory -> theory + val datatype_constrs : theory -> typ -> styp list + val boxed_datatype_constrs : extended_context -> typ -> styp list + val num_datatype_constrs : theory -> typ -> int + val constr_name_for_sel_like : string -> string + val boxed_constr_for_sel : extended_context -> styp -> styp + val card_of_type : (typ * int) list -> typ -> int + val bounded_card_of_type : int -> int -> (typ * int) list -> typ -> int + val bounded_precise_card_of_type : + theory -> int -> int -> (typ * int) list -> typ -> int + val is_finite_type : theory -> typ -> bool + val all_axioms_of : theory -> term list * term list * term list + val arity_of_built_in_const : bool -> styp -> int option + val is_built_in_const : bool -> styp -> bool + val case_const_names : theory -> (string * int) list + val const_def_table : Proof.context -> term list -> const_table + val const_nondef_table : term list -> const_table + val const_simp_table : Proof.context -> const_table + val const_psimp_table : Proof.context -> const_table + val inductive_intro_table : Proof.context -> const_table -> const_table + val ground_theorem_table : theory -> term list Inttab.table + val ersatz_table : theory -> (string * string) list + val def_of_const : theory -> const_table -> styp -> term option + val is_inductive_pred : extended_context -> styp -> bool + val is_constr_pattern_lhs : theory -> term -> bool + val is_constr_pattern_formula : theory -> term -> bool + val coalesce_type_vars_in_terms : term list -> term list + val ground_types_in_type : extended_context -> typ -> typ list + val ground_types_in_terms : extended_context -> term list -> typ list + val format_type : int list -> int list -> typ -> typ + val format_term_type : + theory -> const_table -> (term option * int list) list -> term -> typ + val user_friendly_const : + extended_context -> string * string -> (term option * int list) list + -> styp -> term * typ + val assign_operator_for_const : styp -> string + val preprocess_term : + extended_context -> term -> ((term list * term list) * (bool * bool)) * term +end; + +structure NitpickHOL : NITPICK_HOL = +struct + +open NitpickUtil + +type const_table = term list Symtab.table +type special_fun = (styp * int list * term list) * styp +type unrolled = styp * styp +type wf_cache = (styp * (bool * bool)) list + +type extended_context = { + thy: theory, + ctxt: Proof.context, + max_bisim_depth: int, + boxes: (typ option * bool option) list, + wfs: (styp option * bool option) list, + user_axioms: bool option, + debug: bool, + destroy_constrs: bool, + specialize: bool, + skolemize: bool, + star_linear_preds: bool, + uncurry: bool, + fast_descrs: bool, + tac_timeout: Time.time option, + evals: term list, + case_names: (string * int) list, + def_table: const_table, + nondef_table: const_table, + user_nondefs: term list, + simp_table: const_table Unsynchronized.ref, + psimp_table: const_table, + intro_table: const_table, + ground_thm_table: term list Inttab.table, + ersatz_table: (string * string) list, + skolems: (string * string list) list Unsynchronized.ref, + special_funs: special_fun list Unsynchronized.ref, + unrolled_preds: unrolled list Unsynchronized.ref, + wf_cache: wf_cache Unsynchronized.ref} + +structure TheoryData = TheoryDataFun( + type T = {frac_types: (string * (string * string) list) list, + codatatypes: (string * (string * styp list)) list} + val empty = {frac_types = [], codatatypes = []} + val copy = I + val extend = I + fun merge _ ({frac_types = fs1, codatatypes = cs1}, + {frac_types = fs2, codatatypes = cs2}) = + {frac_types = AList.merge (op =) (op =) (fs1, fs2), + codatatypes = AList.merge (op =) (op =) (cs1, cs2)}) + +(* term * term -> term *) +fun s_conj (t1, @{const True}) = t1 + | s_conj (@{const True}, t2) = t2 + | s_conj (t1, t2) = if @{const False} mem [t1, t2] then @{const False} + else HOLogic.mk_conj (t1, t2) +fun s_disj (t1, @{const False}) = t1 + | s_disj (@{const False}, t2) = t2 + | s_disj (t1, t2) = if @{const True} mem [t1, t2] then @{const True} + else HOLogic.mk_disj (t1, t2) +(* term -> term -> term *) +fun mk_exists v t = + HOLogic.exists_const (fastype_of v) $ lambda v (incr_boundvars 1 t) + +(* term -> term -> term list *) +fun strip_connective conn_t (t as (t0 $ t1 $ t2)) = + if t0 = conn_t then strip_connective t0 t2 @ strip_connective t0 t1 else [t] + | strip_connective _ t = [t] +(* term -> term list * term *) +fun strip_any_connective (t as (t0 $ t1 $ t2)) = + if t0 mem [@{const "op &"}, @{const "op |"}] then + (strip_connective t0 t, t0) + else + ([t], @{const Not}) + | strip_any_connective t = ([t], @{const Not}) +(* term -> term list *) +val conjuncts = strip_connective @{const "op &"} +val disjuncts = strip_connective @{const "op |"} + +val name_sep = "$" +val numeral_prefix = nitpick_prefix ^ "num" ^ name_sep +val sel_prefix = nitpick_prefix ^ "sel" +val discr_prefix = nitpick_prefix ^ "is" ^ name_sep +val set_prefix = nitpick_prefix ^ "set" ^ name_sep +val lfp_iterator_prefix = nitpick_prefix ^ "lfpit" ^ name_sep +val gfp_iterator_prefix = nitpick_prefix ^ "gfpit" ^ name_sep +val nwf_prefix = nitpick_prefix ^ "nwf" ^ name_sep +val unrolled_prefix = nitpick_prefix ^ "unroll" ^ name_sep +val base_prefix = nitpick_prefix ^ "base" ^ name_sep +val step_prefix = nitpick_prefix ^ "step" ^ name_sep +val ubfp_prefix = nitpick_prefix ^ "ubfp" ^ name_sep +val lbfp_prefix = nitpick_prefix ^ "lbfp" ^ name_sep +val skolem_prefix = nitpick_prefix ^ "sk" +val special_prefix = nitpick_prefix ^ "sp" +val uncurry_prefix = nitpick_prefix ^ "unc" +val eval_prefix = nitpick_prefix ^ "eval" +val bound_var_prefix = "b" +val cong_var_prefix = "c" +val iter_var_prefix = "i" +val val_var_prefix = nitpick_prefix ^ "v" +val arg_var_prefix = "x" + +(* int -> string *) +fun sel_prefix_for j = sel_prefix ^ string_of_int j ^ name_sep +fun special_prefix_for j = special_prefix ^ string_of_int j ^ name_sep +(* int -> int -> string *) +fun skolem_prefix_for k j = + skolem_prefix ^ string_of_int k ^ "@" ^ string_of_int j ^ name_sep +fun uncurry_prefix_for k j = + uncurry_prefix ^ string_of_int k ^ "@" ^ string_of_int j ^ name_sep + +(* string -> string * string *) +val strip_first_name_sep = + Substring.full #> Substring.position name_sep ##> Substring.triml 1 + #> pairself Substring.string +(* string -> string *) +fun original_name s = + if String.isPrefix nitpick_prefix s then + case strip_first_name_sep s of (s1, "") => s1 | (_, s2) => original_name s2 + else + s +val after_name_sep = snd o strip_first_name_sep + +(* When you add constants to these lists, make sure to handle them in + "NitpickNut.nut_from_term", and perhaps in "NitpickMono.consider_term" as + well. *) +val built_in_consts = + [(@{const_name all}, 1), + (@{const_name "=="}, 2), + (@{const_name "==>"}, 2), + (@{const_name Pure.conjunction}, 2), + (@{const_name Trueprop}, 1), + (@{const_name Not}, 1), + (@{const_name False}, 0), + (@{const_name True}, 0), + (@{const_name All}, 1), + (@{const_name Ex}, 1), + (@{const_name "op ="}, 2), + (@{const_name "op &"}, 2), + (@{const_name "op |"}, 2), + (@{const_name "op -->"}, 2), + (@{const_name If}, 3), + (@{const_name Let}, 2), + (@{const_name Unity}, 0), + (@{const_name Pair}, 2), + (@{const_name fst}, 1), + (@{const_name snd}, 1), + (@{const_name Id}, 0), + (@{const_name insert}, 2), + (@{const_name converse}, 1), + (@{const_name trancl}, 1), + (@{const_name rel_comp}, 2), + (@{const_name image}, 2), + (@{const_name Suc}, 0), + (@{const_name finite}, 1), + (@{const_name nat}, 0), + (@{const_name zero_nat_inst.zero_nat}, 0), + (@{const_name one_nat_inst.one_nat}, 0), + (@{const_name plus_nat_inst.plus_nat}, 0), + (@{const_name minus_nat_inst.minus_nat}, 0), + (@{const_name times_nat_inst.times_nat}, 0), + (@{const_name div_nat_inst.div_nat}, 0), + (@{const_name div_nat_inst.mod_nat}, 0), + (@{const_name ord_nat_inst.less_nat}, 2), + (@{const_name ord_nat_inst.less_eq_nat}, 2), + (@{const_name nat_gcd}, 0), + (@{const_name nat_lcm}, 0), + (@{const_name zero_int_inst.zero_int}, 0), + (@{const_name one_int_inst.one_int}, 0), + (@{const_name plus_int_inst.plus_int}, 0), + (@{const_name minus_int_inst.minus_int}, 0), + (@{const_name times_int_inst.times_int}, 0), + (@{const_name div_int_inst.div_int}, 0), + (@{const_name div_int_inst.mod_int}, 0), + (@{const_name uminus_int_inst.uminus_int}, 0), (* FIXME: needed? *) + (@{const_name ord_int_inst.less_int}, 2), + (@{const_name ord_int_inst.less_eq_int}, 2), + (@{const_name Tha}, 1), + (@{const_name Frac}, 0), + (@{const_name norm_frac}, 0)] +val built_in_descr_consts = + [(@{const_name The}, 1), + (@{const_name Eps}, 1)] +val built_in_typed_consts = + [((@{const_name of_nat}, nat_T --> int_T), 0)] +val built_in_set_consts = + [(@{const_name lower_semilattice_fun_inst.inf_fun}, 2), + (@{const_name upper_semilattice_fun_inst.sup_fun}, 2), + (@{const_name minus_fun_inst.minus_fun}, 2), + (@{const_name ord_fun_inst.less_eq_fun}, 2)] + +(* typ -> typ *) +fun unbox_type (Type (@{type_name fun_box}, Ts)) = + Type ("fun", map unbox_type Ts) + | unbox_type (Type (@{type_name pair_box}, Ts)) = + Type ("*", map unbox_type Ts) + | unbox_type (Type (s, Ts)) = Type (s, map unbox_type Ts) + | unbox_type T = T +(* Proof.context -> typ -> string *) +fun string_for_type ctxt = Syntax.string_of_typ ctxt o unbox_type + +(* string -> string -> string *) +val prefix_name = Long_Name.qualify o Long_Name.base_name +(* string -> string *) +fun short_name s = List.last (space_explode "." s) handle List.Empty => "" +(* string -> term -> term *) +val prefix_abs_vars = Term.map_abs_vars o prefix_name +(* term -> term *) +val shorten_abs_vars = Term.map_abs_vars short_name +(* string -> string *) +fun short_const_name s = + case space_explode name_sep s of + [_] => s |> String.isPrefix nitpick_prefix s ? unprefix nitpick_prefix + | ss => map short_name ss |> space_implode "_" +(* term -> term *) +val shorten_const_names_in_term = + map_aterms (fn Const (s, T) => Const (short_const_name s, T) | t => t) + +(* theory -> typ * typ -> bool *) +fun type_match thy (T1, T2) = + (Sign.typ_match thy (T2, T1) Vartab.empty; true) + handle Type.TYPE_MATCH => false +(* theory -> styp * styp -> bool *) +fun const_match thy ((s1, T1), (s2, T2)) = + s1 = s2 andalso type_match thy (T1, T2) +(* theory -> term * term -> bool *) +fun term_match thy (Const x1, Const x2) = const_match thy (x1, x2) + | term_match thy (Free (s1, T1), Free (s2, T2)) = + const_match thy ((short_name s1, T1), (short_name s2, T2)) + | term_match thy (t1, t2) = t1 aconv t2 + +(* typ -> bool *) +fun is_TFree (TFree _) = true + | is_TFree _ = false +fun is_higher_order_type (Type ("fun", _)) = true + | is_higher_order_type (Type (_, Ts)) = exists is_higher_order_type Ts + | is_higher_order_type _ = false +fun is_fun_type (Type ("fun", _)) = true + | is_fun_type _ = false +fun is_set_type (Type ("fun", [_, @{typ bool}])) = true + | is_set_type _ = false +fun is_pair_type (Type ("*", _)) = true + | is_pair_type _ = false +fun is_lfp_iterator_type (Type (s, _)) = String.isPrefix lfp_iterator_prefix s + | is_lfp_iterator_type _ = false +fun is_gfp_iterator_type (Type (s, _)) = String.isPrefix gfp_iterator_prefix s + | is_gfp_iterator_type _ = false +val is_fp_iterator_type = is_lfp_iterator_type orf is_gfp_iterator_type +val is_boolean_type = equal prop_T orf equal bool_T +val is_integer_type = + member (op =) [nat_T, int_T, @{typ bisim_iterator}] orf is_fp_iterator_type +val is_record_type = not o null o Record.dest_recTs +(* theory -> typ -> bool *) +fun is_frac_type thy (Type (s, [])) = + not (null (these (AList.lookup (op =) (#frac_types (TheoryData.get thy)) + s))) + | is_frac_type _ _ = false +fun is_number_type thy = is_integer_type orf is_frac_type thy + +(* bool -> styp -> typ *) +fun iterator_type_for_const gfp (s, T) = + Type ((if gfp then gfp_iterator_prefix else lfp_iterator_prefix) ^ s, + binder_types T) +(* typ -> styp *) +fun const_for_iterator_type (Type (s, Ts)) = (after_name_sep s, Ts ---> bool_T) + | const_for_iterator_type T = + raise TYPE ("NitpickHOL.const_for_iterator_type", [T], []) + +(* int -> typ -> typ * typ *) +fun strip_n_binders 0 T = ([], T) + | strip_n_binders n (Type ("fun", [T1, T2])) = + strip_n_binders (n - 1) T2 |>> cons T1 + | strip_n_binders n (Type (@{type_name fun_box}, Ts)) = + strip_n_binders n (Type ("fun", Ts)) + | strip_n_binders _ T = raise TYPE ("NitpickHOL.strip_n_binders", [T], []) +(* typ -> typ *) +val nth_range_type = snd oo strip_n_binders + +(* typ -> int *) +fun num_factors_in_type (Type ("*", [T1, T2])) = + fold (Integer.add o num_factors_in_type) [T1, T2] 0 + | num_factors_in_type _ = 1 +fun num_binder_types (Type ("fun", [_, T2])) = 1 + num_binder_types T2 + | num_binder_types _ = 0 +(* typ -> typ list *) +val curried_binder_types = maps HOLogic.flatten_tupleT o binder_types +fun maybe_curried_binder_types T = + (if is_pair_type (body_type T) then binder_types else curried_binder_types) T + +(* typ -> term list -> term *) +fun mk_flat_tuple _ [t] = t + | mk_flat_tuple (Type ("*", [T1, T2])) (t :: ts) = + HOLogic.pair_const T1 T2 $ t $ (mk_flat_tuple T2 ts) + | mk_flat_tuple T ts = raise TYPE ("NitpickHOL.mk_flat_tuple", [T], ts) +(* int -> term -> term list *) +fun dest_n_tuple 1 t = [t] + | dest_n_tuple n t = HOLogic.dest_prod t ||> dest_n_tuple (n - 1) |> op :: + +(* int -> typ -> typ list *) +fun dest_n_tuple_type 1 T = [T] + | dest_n_tuple_type n (Type (_, [T1, T2])) = + T1 :: dest_n_tuple_type (n - 1) T2 + | dest_n_tuple_type _ T = raise TYPE ("NitpickHOL.dest_n_tuple_type", [T], []) + +(* (typ * typ) list -> typ -> typ *) +fun typ_subst [] T = T + | typ_subst ps T = + let + (* typ -> typ *) + fun subst T = + case AList.lookup (op =) ps T of + SOME T' => T' + | NONE => case T of Type (s, Ts) => Type (s, map subst Ts) | _ => T + in subst T end + +(* theory -> typ -> typ -> typ -> typ *) +fun instantiate_type thy T1 T1' T2 = + Same.commit (Envir.subst_type_same + (Sign.typ_match thy (Logic.varifyT T1, T1') Vartab.empty)) + (Logic.varifyT T2) + handle Type.TYPE_MATCH => + raise TYPE ("NitpickHOL.instantiate_type", [T1, T1'], []) + +(* theory -> typ -> typ -> styp *) +fun repair_constr_type thy body_T' T = + instantiate_type thy (body_type T) body_T' T + +(* string -> (string * string) list -> theory -> theory *) +fun register_frac_type frac_s ersaetze thy = + let + val {frac_types, codatatypes} = TheoryData.get thy + val frac_types = AList.update (op =) (frac_s, ersaetze) frac_types + in TheoryData.put {frac_types = frac_types, codatatypes = codatatypes} thy end +(* string -> theory -> theory *) +fun unregister_frac_type frac_s = register_frac_type frac_s [] + +(* typ -> string -> styp list -> theory -> theory *) +fun register_codatatype co_T case_name constr_xs thy = + let + val {frac_types, codatatypes} = TheoryData.get thy + val constr_xs = map (apsnd (repair_constr_type thy co_T)) constr_xs + val (co_s, co_Ts) = dest_Type co_T + val _ = + if forall is_TFree co_Ts andalso not (has_duplicates (op =) co_Ts) then () + else raise TYPE ("NitpickHOL.register_codatatype", [co_T], []) + val codatatypes = AList.update (op =) (co_s, (case_name, constr_xs)) + codatatypes + in TheoryData.put {frac_types = frac_types, codatatypes = codatatypes} thy end +(* typ -> theory -> theory *) +fun unregister_codatatype co_T = register_codatatype co_T "" [] + +type typedef_info = + {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, + set_def: thm option, prop_of_Rep: thm, set_name: string, + Rep_inverse: thm option} + +(* theory -> string -> typedef_info *) +fun typedef_info thy s = + if is_frac_type thy (Type (s, [])) then + SOME {abs_type = Type (s, []), rep_type = @{typ "int * int"}, + Abs_name = @{const_name Abs_Frac}, Rep_name = @{const_name Rep_Frac}, + set_def = NONE, prop_of_Rep = @{prop "Rep_Frac x \ Frac"} + |> Logic.varify, + set_name = @{const_name Frac}, Rep_inverse = NONE} + else case Typedef.get_info thy s of + SOME {abs_type, rep_type, Abs_name, Rep_name, set_def, Rep, Rep_inverse, + ...} => + SOME {abs_type = abs_type, rep_type = rep_type, Abs_name = Abs_name, + Rep_name = Rep_name, set_def = set_def, prop_of_Rep = prop_of Rep, + set_name = set_prefix ^ s, Rep_inverse = SOME Rep_inverse} + | NONE => NONE + +(* string -> bool *) +fun is_basic_datatype s = + s mem [@{type_name "*"}, @{type_name bool}, @{type_name unit}, + @{type_name nat}, @{type_name int}] +(* theory -> string -> bool *) +val is_typedef = is_some oo typedef_info +val is_real_datatype = is_some oo Datatype.get_info +(* theory -> typ -> bool *) +fun is_codatatype thy (T as Type (s, _)) = + not (null (AList.lookup (op =) (#codatatypes (TheoryData.get thy)) s + |> Option.map snd |> these)) + | is_codatatype _ _ = false +fun is_pure_typedef thy (T as Type (s, _)) = + is_typedef thy s andalso + not (is_real_datatype thy s orelse is_codatatype thy T + orelse is_record_type T orelse is_integer_type T) + | is_pure_typedef _ _ = false +fun is_univ_typedef thy (Type (s, _)) = + (case typedef_info thy s of + SOME {set_def, prop_of_Rep, ...} => + (case set_def of + SOME thm => + try (fst o dest_Const o snd o Logic.dest_equals o prop_of) thm + | NONE => + try (fst o dest_Const o snd o HOLogic.dest_mem + o HOLogic.dest_Trueprop) prop_of_Rep) = SOME @{const_name UNIV} + | NONE => false) + | is_univ_typedef _ _ = false +fun is_datatype thy (T as Type (s, _)) = + (is_typedef thy s orelse is_codatatype thy T orelse T = @{typ ind}) + andalso not (is_basic_datatype s) + | is_datatype _ _ = false + +(* theory -> typ -> (string * typ) list * (string * typ) *) +fun all_record_fields thy T = + let val (recs, more) = Record.get_extT_fields thy T in + recs @ more :: all_record_fields thy (snd more) + end + handle TYPE _ => [] +(* styp -> bool *) +fun is_record_constr (x as (s, T)) = + String.isSuffix Record.extN s andalso + let val dataT = body_type T in + is_record_type dataT andalso + s = unsuffix Record.ext_typeN (fst (dest_Type dataT)) ^ Record.extN + end +(* theory -> typ -> int *) +val num_record_fields = Integer.add 1 o length o fst oo Record.get_extT_fields +(* theory -> string -> typ -> int *) +fun no_of_record_field thy s T1 = + find_index (equal s o fst) (Record.get_extT_fields thy T1 ||> single |> op @) +(* theory -> styp -> bool *) +fun is_record_get thy (s, Type ("fun", [T1, _])) = + exists (equal s o fst) (all_record_fields thy T1) + | is_record_get _ _ = false +fun is_record_update thy (s, T) = + String.isSuffix Record.updateN s andalso + exists (equal (unsuffix Record.updateN s) o fst) + (all_record_fields thy (body_type T)) + handle TYPE _ => false +fun is_abs_fun thy (s, Type ("fun", [_, Type (s', _)])) = + (case typedef_info thy s' of + SOME {Abs_name, ...} => s = Abs_name + | NONE => false) + | is_abs_fun _ _ = false +fun is_rep_fun thy (s, Type ("fun", [Type (s', _), _])) = + (case typedef_info thy s' of + SOME {Rep_name, ...} => s = Rep_name + | NONE => false) + | is_rep_fun _ _ = false + +(* theory -> styp -> styp *) +fun mate_of_rep_fun thy (x as (_, Type ("fun", [T1 as Type (s', _), T2]))) = + (case typedef_info thy s' of + SOME {Abs_name, ...} => (Abs_name, Type ("fun", [T2, T1])) + | NONE => raise TERM ("NitpickHOL.mate_of_rep_fun", [Const x])) + | mate_of_rep_fun _ x = raise TERM ("NitpickHOL.mate_of_rep_fun", [Const x]) + +(* theory -> styp -> bool *) +fun is_coconstr thy (s, T) = + let + val {codatatypes, ...} = TheoryData.get thy + val co_T = body_type T + val co_s = dest_Type co_T |> fst + in + exists (fn (s', T') => s = s' andalso repair_constr_type thy co_T T' = T) + (AList.lookup (op =) codatatypes co_s |> Option.map snd |> these) + end + handle TYPE ("dest_Type", _, _) => false +fun is_constr_like thy (s, T) = + s mem [@{const_name FunBox}, @{const_name PairBox}] orelse + let val (x as (s, T)) = (s, unbox_type T) in + Refute.is_IDT_constructor thy x orelse is_record_constr x + orelse (is_abs_fun thy x andalso is_pure_typedef thy (range_type T)) + orelse s mem [@{const_name Zero_Rep}, @{const_name Suc_Rep}] + orelse x = (@{const_name zero_nat_inst.zero_nat}, nat_T) + orelse is_coconstr thy x + end +fun is_constr thy (x as (_, T)) = + is_constr_like thy x + andalso not (is_basic_datatype (fst (dest_Type (body_type T)))) +(* string -> bool *) +val is_sel = String.isPrefix discr_prefix orf String.isPrefix sel_prefix +val is_sel_like_and_no_discr = + String.isPrefix sel_prefix + orf (member (op =) [@{const_name fst}, @{const_name snd}]) + +datatype boxability = + InConstr | InSel | InExpr | InPair | InFunLHS | InFunRHS1 | InFunRHS2 + +(* boxability -> boxability *) +fun in_fun_lhs_for InConstr = InSel + | in_fun_lhs_for _ = InFunLHS +fun in_fun_rhs_for InConstr = InConstr + | in_fun_rhs_for InSel = InSel + | in_fun_rhs_for InFunRHS1 = InFunRHS2 + | in_fun_rhs_for _ = InFunRHS1 + +(* extended_context -> boxability -> typ -> bool *) +fun is_boxing_worth_it (ext_ctxt : extended_context) boxy T = + case T of + Type ("fun", _) => + boxy mem [InPair, InFunLHS] andalso not (is_boolean_type (body_type T)) + | Type ("*", Ts) => + boxy mem [InPair, InFunRHS1, InFunRHS2] + orelse (boxy mem [InExpr, InFunLHS] + andalso exists (is_boxing_worth_it ext_ctxt InPair) + (map (box_type ext_ctxt InPair) Ts)) + | _ => false +(* extended_context -> boxability -> string * typ list -> string *) +and should_box_type (ext_ctxt as {thy, boxes, ...}) boxy (z as (s, Ts)) = + case triple_lookup (type_match thy) boxes (Type z) of + SOME (SOME box_me) => box_me + | _ => is_boxing_worth_it ext_ctxt boxy (Type z) +(* extended_context -> boxability -> typ -> typ *) +and box_type ext_ctxt boxy T = + case T of + Type (z as ("fun", [T1, T2])) => + if not (boxy mem [InConstr, InSel]) + andalso should_box_type ext_ctxt boxy z then + Type (@{type_name fun_box}, + [box_type ext_ctxt InFunLHS T1, box_type ext_ctxt InFunRHS1 T2]) + else + box_type ext_ctxt (in_fun_lhs_for boxy) T1 + --> box_type ext_ctxt (in_fun_rhs_for boxy) T2 + | Type (z as ("*", Ts)) => + if should_box_type ext_ctxt boxy z then + Type (@{type_name pair_box}, map (box_type ext_ctxt InSel) Ts) + else + Type ("*", map (box_type ext_ctxt + (if boxy mem [InConstr, InSel] then boxy + else InPair)) Ts) + | _ => T + +(* styp -> styp *) +fun discr_for_constr (s, T) = (discr_prefix ^ s, body_type T --> bool_T) + +(* typ -> int *) +fun num_sels_for_constr_type T = length (maybe_curried_binder_types T) +(* string -> int -> string *) +fun nth_sel_name_for_constr_name s n = + if s = @{const_name Pair} then + if n = 0 then @{const_name fst} else @{const_name snd} + else + sel_prefix_for n ^ s +(* styp -> int -> styp *) +fun nth_sel_for_constr x ~1 = discr_for_constr x + | nth_sel_for_constr (s, T) n = + (nth_sel_name_for_constr_name s n, + body_type T --> nth (maybe_curried_binder_types T) n) +(* extended_context -> styp -> int -> styp *) +fun boxed_nth_sel_for_constr ext_ctxt = + apsnd (box_type ext_ctxt InSel) oo nth_sel_for_constr + +(* string -> int *) +fun sel_no_from_name s = + if String.isPrefix discr_prefix s then + ~1 + else if String.isPrefix sel_prefix s then + s |> unprefix sel_prefix |> Int.fromString |> the + else if s = @{const_name snd} then + 1 + else + 0 + +(* typ list -> term -> int -> term *) +fun eta_expand _ t 0 = t + | eta_expand Ts (Abs (s, T, t')) n = + Abs (s, T, eta_expand (T :: Ts) t' (n - 1)) + | eta_expand Ts t n = + fold_rev (curry3 Abs ("x\<^isub>\" ^ nat_subscript n)) + (List.take (binder_types (fastype_of1 (Ts, t)), n)) + (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0))) + +(* term -> term *) +fun extensionalize t = + case t of + (t0 as @{const Trueprop}) $ t1 => t0 $ extensionalize t1 + | Const (@{const_name "op ="}, _) $ t1 $ Abs (s, T, t2) => + let val v = Var ((s, maxidx_of_term t + 1), T) in + extensionalize (HOLogic.mk_eq (t1 $ v, subst_bound (v, t2))) + end + | _ => t + +(* typ -> term list -> term *) +fun distinctness_formula T = + all_distinct_unordered_pairs_of + #> map (fn (t1, t2) => @{const Not} $ (HOLogic.eq_const T $ t1 $ t2)) + #> List.foldr (s_conj o swap) @{const True} + +(* typ -> term *) +fun zero_const T = Const (@{const_name zero_nat_inst.zero_nat}, T) +fun suc_const T = Const (@{const_name Suc}, T --> T) + +(* theory -> typ -> styp list *) +fun datatype_constrs thy (T as Type (s, Ts)) = + if is_datatype thy T then + case Datatype.get_info thy s of + SOME {index, descr, ...} => + let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in + map (fn (s', Us) => + (s', map (Refute.typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T)) + constrs + end + | NONE => + case AList.lookup (op =) (#codatatypes (TheoryData.get thy)) s of + SOME (_, xs' as (_ :: _)) => + map (apsnd (repair_constr_type thy T)) xs' + | _ => + if is_record_type T then + let + val s' = unsuffix Record.ext_typeN s ^ Record.extN + val T' = (Record.get_extT_fields thy T + |> apsnd single |> uncurry append |> map snd) ---> T + in [(s', T')] end + else case typedef_info thy s of + SOME {abs_type, rep_type, Abs_name, ...} => + [(Abs_name, instantiate_type thy abs_type T rep_type --> T)] + | NONE => + if T = @{typ ind} then + [dest_Const @{const Zero_Rep}, dest_Const @{const Suc_Rep}] + else + [] + else + [] + | datatype_constrs _ _ = [] +(* extended_context -> typ -> styp list *) +fun boxed_datatype_constrs (ext_ctxt as {thy, ...}) = + map (apsnd (box_type ext_ctxt InConstr)) o datatype_constrs thy +(* theory -> typ -> int *) +val num_datatype_constrs = length oo datatype_constrs + +(* string -> string *) +fun constr_name_for_sel_like @{const_name fst} = @{const_name Pair} + | constr_name_for_sel_like @{const_name snd} = @{const_name Pair} + | constr_name_for_sel_like s' = original_name s' +(* extended_context -> styp -> styp *) +fun boxed_constr_for_sel ext_ctxt (s', T') = + let val s = constr_name_for_sel_like s' in + AList.lookup (op =) (boxed_datatype_constrs ext_ctxt (domain_type T')) s + |> the |> pair s + end +(* theory -> styp -> term *) +fun discr_term_for_constr thy (x as (s, T)) = + let val dataT = body_type T in + if s = @{const_name Suc} then + Abs (Name.uu, dataT, + @{const Not} $ HOLogic.mk_eq (zero_const dataT, Bound 0)) + else if num_datatype_constrs thy dataT >= 2 then + Const (discr_for_constr x) + else + Abs (Name.uu, dataT, @{const True}) + end + +(* theory -> styp -> term -> term *) +fun discriminate_value thy (x as (_, T)) t = + case strip_comb t of + (Const x', args) => + if x = x' then @{const True} + else if is_constr_like thy x' then @{const False} + else betapply (discr_term_for_constr thy x, t) + | _ => betapply (discr_term_for_constr thy x, t) + +(* styp -> term -> term *) +fun nth_arg_sel_term_for_constr (x as (s, T)) n = + let val (arg_Ts, dataT) = strip_type T in + if dataT = nat_T then + @{term "%n::nat. minus_nat_inst.minus_nat n one_nat_inst.one_nat"} + else if is_pair_type dataT then + Const (nth_sel_for_constr x n) + else + let + (* int -> typ -> int * term *) + fun aux m (Type ("*", [T1, T2])) = + let + val (m, t1) = aux m T1 + val (m, t2) = aux m T2 + in (m, HOLogic.mk_prod (t1, t2)) end + | aux m T = + (m + 1, Const (nth_sel_name_for_constr_name s m, dataT --> T) + $ Bound 0) + val m = fold (Integer.add o num_factors_in_type) + (List.take (arg_Ts, n)) 0 + in Abs ("x", dataT, aux m (nth arg_Ts n) |> snd) end + end +(* theory -> styp -> term -> int -> typ -> term *) +fun select_nth_constr_arg thy x t n res_T = + case strip_comb t of + (Const x', args) => + if x = x' then nth args n + else if is_constr_like thy x' then Const (@{const_name unknown}, res_T) + else betapply (nth_arg_sel_term_for_constr x n, t) + | _ => betapply (nth_arg_sel_term_for_constr x n, t) + +(* theory -> styp -> term list -> term *) +fun construct_value _ x [] = Const x + | construct_value thy (x as (s, _)) args = + let val args = map Envir.eta_contract args in + case hd args of + Const (x' as (s', _)) $ t => + if is_sel_like_and_no_discr s' andalso constr_name_for_sel_like s' = s + andalso forall (fn (n, t') => + select_nth_constr_arg thy x t n dummyT = t') + (index_seq 0 (length args) ~~ args) then + t + else + list_comb (Const x, args) + | _ => list_comb (Const x, args) + end + +(* theory -> typ -> term -> term *) +fun constr_expand thy T t = + (case head_of t of + Const x => if is_constr_like thy x then t else raise SAME () + | _ => raise SAME ()) + handle SAME () => + let + val x' as (_, T') = + if is_pair_type T then + let val (T1, T2) = HOLogic.dest_prodT T in + (@{const_name Pair}, [T1, T2] ---> T) + end + else + datatype_constrs thy T |> the_single + val arg_Ts = binder_types T' + in + list_comb (Const x', map2 (select_nth_constr_arg thy x' t) + (index_seq 0 (length arg_Ts)) arg_Ts) + end + +(* (typ * int) list -> typ -> int *) +fun card_of_type asgns (Type ("fun", [T1, T2])) = + reasonable_power (card_of_type asgns T2) (card_of_type asgns T1) + | card_of_type asgns (Type ("*", [T1, T2])) = + card_of_type asgns T1 * card_of_type asgns T2 + | card_of_type _ (Type (@{type_name itself}, _)) = 1 + | card_of_type _ @{typ prop} = 2 + | card_of_type _ @{typ bool} = 2 + | card_of_type _ @{typ unit} = 1 + | card_of_type asgns T = + case AList.lookup (op =) asgns T of + SOME k => k + | NONE => if T = @{typ bisim_iterator} then 0 + else raise TYPE ("NitpickHOL.card_of_type", [T], []) +(* int -> (typ * int) list -> typ -> int *) +fun bounded_card_of_type max default_card asgns (Type ("fun", [T1, T2])) = + let + val k1 = bounded_card_of_type max default_card asgns T1 + val k2 = bounded_card_of_type max default_card asgns T2 + in + if k1 = max orelse k2 = max then max + else Int.min (max, reasonable_power k2 k1) + end + | bounded_card_of_type max default_card asgns (Type ("*", [T1, T2])) = + let + val k1 = bounded_card_of_type max default_card asgns T1 + val k2 = bounded_card_of_type max default_card asgns T2 + in if k1 = max orelse k2 = max then max else Int.min (max, k1 * k2) end + | bounded_card_of_type max default_card asgns T = + Int.min (max, if default_card = ~1 then + card_of_type asgns T + else + card_of_type asgns T + handle TYPE ("NitpickHOL.card_of_type", _, _) => + default_card) +(* theory -> int -> (typ * int) list -> typ -> int *) +fun bounded_precise_card_of_type thy max default_card asgns T = + let + (* typ list -> typ -> int *) + fun aux avoid T = + (if T mem avoid then + 0 + else case T of + Type ("fun", [T1, T2]) => + let + val k1 = aux avoid T1 + val k2 = aux avoid T2 + in + if k1 = 0 orelse k2 = 0 then 0 + else if k1 >= max orelse k2 >= max then max + else Int.min (max, reasonable_power k2 k1) + end + | Type ("*", [T1, T2]) => + let + val k1 = aux avoid T1 + val k2 = aux avoid T2 + in + if k1 = 0 orelse k2 = 0 then 0 + else if k1 >= max orelse k2 >= max then max + else Int.min (max, k1 * k2) + end + | Type (@{type_name itself}, _) => 1 + | @{typ prop} => 2 + | @{typ bool} => 2 + | @{typ unit} => 1 + | Type _ => + (case datatype_constrs thy T of + [] => if is_integer_type T then 0 else raise SAME () + | constrs => + let + val constr_cards = + datatype_constrs thy T + |> map (Integer.prod o map (aux (T :: avoid)) o binder_types + o snd) + in + if exists (equal 0) constr_cards then 0 + else Integer.sum constr_cards + end) + | _ => raise SAME ()) + handle SAME () => AList.lookup (op =) asgns T |> the_default default_card + in Int.min (max, aux [] T) end + +(* theory -> typ -> bool *) +fun is_finite_type thy = not_equal 0 o bounded_precise_card_of_type thy 1 2 [] + +(* term -> bool *) +fun is_ground_term (t1 $ t2) = is_ground_term t1 andalso is_ground_term t2 + | is_ground_term (Const _) = true + | is_ground_term _ = false + +(* term -> word -> word *) +fun hashw_term (t1 $ t2) = Polyhash.hashw (hashw_term t1, hashw_term t2) + | hashw_term (Const (s, _)) = Polyhash.hashw_string (s, 0w0) + | hashw_term _ = 0w0 +(* term -> int *) +val hash_term = Word.toInt o hashw_term + +(* term list -> (indexname * typ) list *) +fun special_bounds ts = + fold Term.add_vars ts [] |> sort (TermOrd.fast_indexname_ord o pairself fst) + +(* indexname * typ -> term -> term *) +fun abs_var ((s, j), T) body = Abs (s, T, abstract_over (Var ((s, j), T), body)) + +(* term -> bool *) +fun is_arity_type_axiom (Const (@{const_name HOL.type_class}, _) + $ Const (@{const_name TYPE}, _)) = true + | is_arity_type_axiom _ = false +(* theory -> bool -> term -> bool *) +fun is_typedef_axiom thy boring (@{const "==>"} $ _ $ t2) = + is_typedef_axiom thy boring t2 + | is_typedef_axiom thy boring + (@{const Trueprop} $ (Const (@{const_name Typedef.type_definition}, _) + $ Const (_, Type ("fun", [Type (s, _), _])) $ Const _ $ _)) = + boring <> (s mem [@{type_name unit}, @{type_name "*"}, @{type_name "+"}] + orelse is_frac_type thy (Type (s, []))) + andalso is_typedef thy s + | is_typedef_axiom _ _ _ = false + +(* Distinguishes between (1) constant definition axioms, (2) type arity and + typedef axioms, and (3) other axioms, and returns the pair ((1), (3)). + Typedef axioms are uninteresting to Nitpick, because it can retrieve them + using "typedef_info". *) +(* theory -> (string * term) list -> string list -> term list * term list *) +fun partition_axioms_by_definitionality thy axioms def_names = + let + val axioms = sort (fast_string_ord o pairself fst) axioms + val defs = OrdList.inter (fast_string_ord o apsnd fst) def_names axioms + val nondefs = + OrdList.subtract (fast_string_ord o apsnd fst) def_names axioms + |> filter_out ((is_arity_type_axiom orf is_typedef_axiom thy true) o snd) + in pairself (map snd) (defs, nondefs) end + +(* Ideally we would check against "Complex_Main", not "Refute", but any theory + will do as long as it contains all the "axioms" and "axiomatization" + commands. *) +(* theory -> bool *) +fun is_built_in_theory thy = Theory.subthy (thy, @{theory Refute}) + +(* term -> bool *) +val is_plain_definition = + let + (* term -> bool *) + fun do_lhs t1 = + case strip_comb t1 of + (Const _, args) => forall is_Var args + andalso not (has_duplicates (op =) args) + | _ => false + fun do_eq (Const (@{const_name "=="}, _) $ t1 $ _) = do_lhs t1 + | do_eq (@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)) = + do_lhs t1 + | do_eq _ = false + in do_eq end + +(* This table is not pretty. A better approach would be to avoid expanding the + operators to their low-level definitions, but this would require dealing with + overloading. *) +val built_in_built_in_defs = + [@{thm div_int_inst.div_int}, @{thm div_int_inst.mod_int}, + @{thm div_nat_inst.div_nat}, @{thm div_nat_inst.mod_nat}, + @{thm lower_semilattice_fun_inst.inf_fun}, @{thm minus_fun_inst.minus_fun}, + @{thm minus_int_inst.minus_int}, @{thm minus_nat_inst.minus_nat}, + @{thm one_int_inst.one_int}, @{thm one_nat_inst.one_nat}, + @{thm ord_fun_inst.less_eq_fun}, @{thm ord_int_inst.less_eq_int}, + @{thm ord_int_inst.less_int}, @{thm ord_nat_inst.less_eq_nat}, + @{thm ord_nat_inst.less_nat}, @{thm plus_int_inst.plus_int}, + @{thm plus_nat_inst.plus_nat}, @{thm times_int_inst.times_int}, + @{thm times_nat_inst.times_nat}, @{thm uminus_int_inst.uminus_int}, + @{thm upper_semilattice_fun_inst.sup_fun}, @{thm zero_int_inst.zero_int}, + @{thm zero_nat_inst.zero_nat}] + |> map prop_of + +(* theory -> term list * term list * term list *) +fun all_axioms_of thy = + let + (* theory list -> term list *) + val axioms_of_thys = maps Thm.axioms_of #> map (apsnd prop_of) + val specs = Defs.all_specifications_of (Theory.defs_of thy) + val def_names = specs |> maps snd |> filter #is_def |> map #name + |> OrdList.make fast_string_ord + val thys = thy :: Theory.ancestors_of thy + val (built_in_thys, user_thys) = List.partition is_built_in_theory thys + val built_in_axioms = axioms_of_thys built_in_thys + val user_axioms = axioms_of_thys user_thys + val (built_in_defs, built_in_nondefs) = + partition_axioms_by_definitionality thy built_in_axioms def_names + ||> filter (is_typedef_axiom thy false) + val (user_defs, user_nondefs) = + partition_axioms_by_definitionality thy user_axioms def_names + val (built_in_nondefs, user_nondefs) = + List.partition (is_typedef_axiom thy false) user_nondefs + |>> append built_in_nondefs + val defs = built_in_built_in_defs @ + (thy |> PureThy.all_thms_of + |> filter (equal Thm.definitionK o Thm.get_kind o snd) + |> map (prop_of o snd) |> filter is_plain_definition) @ + user_defs @ built_in_defs + in (defs, built_in_nondefs, user_nondefs) end + +(* bool -> styp -> int option *) +fun arity_of_built_in_const fast_descrs (s, T) = + if s = @{const_name If} then + if nth_range_type 3 T = @{typ bool} then NONE else SOME 3 + else case AList.lookup (op =) + (built_in_consts + |> fast_descrs ? append built_in_descr_consts) s of + SOME n => SOME n + | NONE => + case AList.lookup (op =) built_in_typed_consts (s, T) of + SOME n => SOME n + | NONE => + if is_fun_type T andalso is_set_type (domain_type T) then + AList.lookup (op =) built_in_set_consts s + else + NONE +(* bool -> styp -> bool *) +val is_built_in_const = is_some oo arity_of_built_in_const + +(* This function is designed to work for both real definition axioms and + simplification rules (equational specifications). *) +(* term -> term *) +fun term_under_def t = + case t of + @{const "==>"} $ _ $ t2 => term_under_def t2 + | Const (@{const_name "=="}, _) $ t1 $ _ => term_under_def t1 + | @{const Trueprop} $ t1 => term_under_def t1 + | Const (@{const_name "op ="}, _) $ t1 $ _ => term_under_def t1 + | Abs (_, _, t') => term_under_def t' + | t1 $ _ => term_under_def t1 + | _ => t + +(* Here we crucially rely on "Refute.specialize_type" performing a preorder + traversal of the term, without which the wrong occurrence of a constant could + be matched in the face of overloading. *) +(* theory -> bool -> const_table -> styp -> term list *) +fun def_props_for_const thy fast_descrs table (x as (s, _)) = + if is_built_in_const fast_descrs x then + [] + else + these (Symtab.lookup table s) + |> map_filter (try (Refute.specialize_type thy x)) + |> filter (equal (Const x) o term_under_def) + +(* term -> term *) +fun normalized_rhs_of thy t = + let + (* term -> term *) + fun aux (v as Var _) t = lambda v t + | aux (c as Const (@{const_name TYPE}, T)) t = lambda c t + | aux _ _ = raise TERM ("NitpickHOL.normalized_rhs_of", [t]) + val (lhs, rhs) = + case t of + Const (@{const_name "=="}, _) $ t1 $ t2 => (t1, t2) + | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ t2) => + (t1, t2) + | _ => raise TERM ("NitpickHOL.normalized_rhs_of", [t]) + val args = strip_comb lhs |> snd + in fold_rev aux args rhs end + +(* theory -> const_table -> styp -> term option *) +fun def_of_const thy table (x as (s, _)) = + if is_built_in_const false x orelse original_name s <> s then + NONE + else + x |> def_props_for_const thy false table |> List.last + |> normalized_rhs_of thy |> prefix_abs_vars s |> SOME + handle List.Empty => NONE + +datatype fixpoint_kind = Lfp | Gfp | NoFp + +(* term -> fixpoint_kind *) +fun fixpoint_kind_of_rhs (Abs (_, _, t)) = fixpoint_kind_of_rhs t + | fixpoint_kind_of_rhs (Const (@{const_name lfp}, _) $ Abs _) = Lfp + | fixpoint_kind_of_rhs (Const (@{const_name gfp}, _) $ Abs _) = Gfp + | fixpoint_kind_of_rhs _ = NoFp + +(* theory -> const_table -> term -> bool *) +fun is_mutually_inductive_pred_def thy table t = + let + (* term -> bool *) + fun is_good_arg (Bound _) = true + | is_good_arg (Const (s, _)) = + s mem [@{const_name True}, @{const_name False}, @{const_name undefined}] + | is_good_arg _ = false + in + case t |> strip_abs_body |> strip_comb of + (Const x, ts as (_ :: _)) => + (case def_of_const thy table x of + SOME t' => fixpoint_kind_of_rhs t' <> NoFp andalso forall is_good_arg ts + | NONE => false) + | _ => false + end +(* theory -> const_table -> term -> term *) +fun unfold_mutually_inductive_preds thy table = + map_aterms (fn t as Const x => + (case def_of_const thy table x of + SOME t' => + let val t' = Envir.eta_contract t' in + if is_mutually_inductive_pred_def thy table t' then t' + else t + end + | NONE => t) + | t => t) + +(* term -> string * term *) +fun pair_for_prop t = + case term_under_def t of + Const (s, _) => (s, t) + | Free _ => raise NOT_SUPPORTED "local definitions" + | t' => raise TERM ("NitpickHOL.pair_for_prop", [t, t']) + +(* (Proof.context -> term list) -> Proof.context -> const_table *) +fun table_for get ctxt = + get ctxt |> map pair_for_prop |> AList.group (op =) |> Symtab.make + +(* theory -> (string * int) list *) +fun case_const_names thy = + Symtab.fold (fn (dtype_s, {index, descr, case_name, ...}) => + if is_basic_datatype dtype_s then + I + else + cons (case_name, AList.lookup (op =) descr index + |> the |> #3 |> length)) + (Datatype.get_all thy) [] @ + map (apsnd length o snd) (#codatatypes (TheoryData.get thy)) + +(* Proof.context -> term list -> const_table *) +fun const_def_table ctxt ts = + table_for (map prop_of o Nitpick_Defs.get) ctxt + |> fold (fn (s, t) => Symtab.map_default (s, []) (cons t)) + (map pair_for_prop ts) +(* term list -> const_table *) +fun const_nondef_table ts = + fold (fn t => append (map (fn s => (s, t)) (Term.add_const_names t []))) ts [] + |> AList.group (op =) |> Symtab.make +(* Proof.context -> const_table *) +val const_simp_table = table_for (map prop_of o Nitpick_Simps.get) +val const_psimp_table = table_for (map prop_of o Nitpick_Psimps.get) +(* Proof.context -> const_table -> const_table *) +fun inductive_intro_table ctxt def_table = + table_for (map (unfold_mutually_inductive_preds (ProofContext.theory_of ctxt) + def_table o prop_of) + o Nitpick_Intros.get) ctxt +(* theory -> term list Inttab.table *) +fun ground_theorem_table thy = + fold ((fn @{const Trueprop} $ t1 => + is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1) + | _ => I) o prop_of o snd) (PureThy.all_thms_of thy) Inttab.empty + +val basic_ersatz_table = + [(@{const_name prod_case}, @{const_name split}), + (@{const_name card}, @{const_name card'}), + (@{const_name setsum}, @{const_name setsum'}), + (@{const_name fold_graph}, @{const_name fold_graph'}), + (@{const_name wf}, @{const_name wf'}), + (@{const_name wf_wfrec}, @{const_name wf_wfrec'}), + (@{const_name wfrec}, @{const_name wfrec'})] + +(* theory -> (string * string) list *) +fun ersatz_table thy = + fold (append o snd) (#frac_types (TheoryData.get thy)) basic_ersatz_table + +(* const_table Unsynchronized.ref -> string -> term list -> unit *) +fun add_simps simp_table s eqs = + Unsynchronized.change simp_table + (Symtab.update (s, eqs @ these (Symtab.lookup (!simp_table) s))) + +(* Similar to "Refute.specialize_type" but returns all matches rather than only + the first (preorder) match. *) +(* theory -> styp -> term -> term list *) +fun multi_specialize_type thy slack (x as (s, T)) t = + let + (* term -> (typ * term) list -> (typ * term) list *) + fun aux (Const (s', T')) ys = + if s = s' then + ys |> (if AList.defined (op =) ys T' then + I + else + cons (T', Refute.monomorphic_term + (Sign.typ_match thy (T', T) Vartab.empty) t) + handle Type.TYPE_MATCH => I + | Refute.REFUTE _ => + if slack then + I + else + raise NOT_SUPPORTED ("too much polymorphism in \ + \axiom involving " ^ quote s)) + else + ys + | aux _ ys = ys + in map snd (fold_aterms aux t []) end + +(* theory -> bool -> const_table -> styp -> term list *) +fun nondef_props_for_const thy slack table (x as (s, _)) = + these (Symtab.lookup table s) |> maps (multi_specialize_type thy slack x) + +(* theory -> styp list -> term list *) +fun optimized_typedef_axioms thy (abs_s, abs_Ts) = + let val abs_T = Type (abs_s, abs_Ts) in + if is_univ_typedef thy abs_T then + [] + else case typedef_info thy abs_s of + SOME {abs_type, rep_type, Abs_name, Rep_name, prop_of_Rep, set_name, + ...} => + let + val rep_T = instantiate_type thy abs_type abs_T rep_type + val rep_t = Const (Rep_name, abs_T --> rep_T) + val set_t = Const (set_name, rep_T --> bool_T) + val set_t' = + prop_of_Rep |> HOLogic.dest_Trueprop + |> Refute.specialize_type thy (dest_Const rep_t) + |> HOLogic.dest_mem |> snd + in + [HOLogic.all_const abs_T + $ Abs (Name.uu, abs_T, set_t $ (rep_t $ Bound 0))] + |> set_t <> set_t' ? cons (HOLogic.mk_eq (set_t, set_t')) + |> map HOLogic.mk_Trueprop + end + | NONE => [] + end +(* theory -> styp -> term *) +fun inverse_axiom_for_rep_fun thy (x as (_, T)) = + typedef_info thy (fst (dest_Type (domain_type T))) + |> the |> #Rep_inverse |> the |> prop_of |> Refute.specialize_type thy x + +(* theory -> int * styp -> term *) +fun constr_case_body thy (j, (x as (_, T))) = + let val arg_Ts = binder_types T in + list_comb (Bound j, map2 (select_nth_constr_arg thy x (Bound 0)) + (index_seq 0 (length arg_Ts)) arg_Ts) + end +(* theory -> typ -> int * styp -> term -> term *) +fun add_constr_case thy res_T (j, x) res_t = + Const (@{const_name If}, [bool_T, res_T, res_T] ---> res_T) + $ discriminate_value thy x (Bound 0) $ constr_case_body thy (j, x) $ res_t +(* theory -> typ -> typ -> term *) +fun optimized_case_def thy dataT res_T = + let + val xs = datatype_constrs thy dataT + val func_Ts = map ((fn T => binder_types T ---> res_T) o snd) xs + val (xs', x) = split_last xs + in + constr_case_body thy (1, x) + |> fold_rev (add_constr_case thy res_T) (length xs downto 2 ~~ xs') + |> fold_rev (curry absdummy) (func_Ts @ [dataT]) + end + +val redefined_in_NitpickDefs_thy = + [@{const_name option_case}, @{const_name nat_case}, @{const_name list_case}, + @{const_name list_size}] + +(* theory -> string -> typ -> typ -> term -> term *) +fun optimized_record_get thy s rec_T res_T t = + let val constr_x = the_single (datatype_constrs thy rec_T) in + case no_of_record_field thy s rec_T of + ~1 => (case rec_T of + Type (_, Ts as _ :: _) => + let + val rec_T' = List.last Ts + val j = num_record_fields thy rec_T - 1 + in + select_nth_constr_arg thy constr_x t j res_T + |> optimized_record_get thy s rec_T' res_T + end + | _ => raise TYPE ("NitpickHOL.optimized_record_get", [rec_T], [])) + | j => select_nth_constr_arg thy constr_x t j res_T + end +(* theory -> string -> typ -> term -> term -> term *) +fun optimized_record_update thy s rec_T fun_t rec_t = + let + val constr_x as (_, constr_T) = the_single (datatype_constrs thy rec_T) + val Ts = binder_types constr_T + val n = length Ts + val special_j = no_of_record_field thy s rec_T + val ts = map2 (fn j => fn T => + let + val t = select_nth_constr_arg thy constr_x rec_t j T + in + if j = special_j then + betapply (fun_t, t) + else if j = n - 1 andalso special_j = ~1 then + optimized_record_update thy s + (rec_T |> dest_Type |> snd |> List.last) fun_t t + else + t + end) (index_seq 0 n) Ts + in list_comb (Const constr_x, ts) end + +(* Constants "c" whose definition is of the form "c == c'", where "c'" is also a + constant, are said to be trivial. For those, we ignore the simplification + rules and use the definition instead, to ensure that built-in symbols like + "ord_nat_inst.less_eq_nat" are picked up correctly. *) +(* theory -> const_table -> styp -> bool *) +fun has_trivial_definition thy table x = + case def_of_const thy table x of SOME (Const _) => true | _ => false + +(* theory -> const_table -> string * typ -> fixpoint_kind *) +fun fixpoint_kind_of_const thy table x = + if is_built_in_const false x then + NoFp + else + fixpoint_kind_of_rhs (the (def_of_const thy table x)) + handle Option.Option => NoFp + +(* extended_context -> styp -> bool *) +fun is_real_inductive_pred ({thy, fast_descrs, def_table, intro_table, ...} + : extended_context) x = + not (null (def_props_for_const thy fast_descrs intro_table x)) + andalso fixpoint_kind_of_const thy def_table x <> NoFp +fun is_real_equational_fun ({thy, fast_descrs, simp_table, psimp_table, ...} + : extended_context) x = + exists (fn table => not (null (def_props_for_const thy fast_descrs table x))) + [!simp_table, psimp_table] +fun is_inductive_pred ext_ctxt = + is_real_inductive_pred ext_ctxt andf (not o is_real_equational_fun ext_ctxt) +fun is_equational_fun (ext_ctxt as {thy, def_table, ...}) = + (is_real_equational_fun ext_ctxt orf is_real_inductive_pred ext_ctxt + orf (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst) + andf (not o has_trivial_definition thy def_table) + andf (not o member (op =) redefined_in_NitpickDefs_thy o fst) + +(* term * term -> term *) +fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t + | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t + | s_betapply p = betapply p +(* term * term list -> term *) +val s_betapplys = Library.foldl s_betapply + +(* term -> term *) +fun lhs_of_equation t = + case t of + Const (@{const_name all}, _) $ Abs (_, _, t1) => lhs_of_equation t1 + | Const (@{const_name "=="}, _) $ t1 $ _ => SOME t1 + | @{const "==>"} $ _ $ t2 => lhs_of_equation t2 + | @{const Trueprop} $ t1 => lhs_of_equation t1 + | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1 + | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1 + | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2 + | _ => NONE +(* theory -> term -> bool *) +fun is_constr_pattern _ (Bound _) = true + | is_constr_pattern thy t = + case strip_comb t of + (Const (x as (s, _)), args) => + is_constr_like thy x andalso forall (is_constr_pattern thy) args + | _ => false +fun is_constr_pattern_lhs thy t = + forall (is_constr_pattern thy) (snd (strip_comb t)) +fun is_constr_pattern_formula thy t = + case lhs_of_equation t of + SOME t' => is_constr_pattern_lhs thy t' + | NONE => false + +val unfold_max_depth = 63 +val axioms_max_depth = 63 + +(* extended_context -> term -> term *) +fun unfold_defs_in_term (ext_ctxt as {thy, destroy_constrs, fast_descrs, + case_names, def_table, ground_thm_table, + ersatz_table, ...}) = + let + (* int -> typ list -> term -> term *) + fun do_term depth Ts t = + case t of + (t0 as Const (@{const_name Int.number_class.number_of}, + Type ("fun", [_, ran_T]))) $ t1 => + ((if is_number_type thy ran_T then + let + val j = t1 |> HOLogic.dest_numeral + |> ran_T <> int_T ? curry Int.max 0 + val s = numeral_prefix ^ signed_string_of_int j + in + if is_integer_type ran_T then + Const (s, ran_T) + else + do_term depth Ts (Const (@{const_name of_int}, int_T --> ran_T) + $ Const (s, int_T)) + end + handle TERM _ => raise SAME () + else + raise SAME ()) + handle SAME () => betapply (do_term depth Ts t0, do_term depth Ts t1)) + | Const (@{const_name refl_on}, T) $ Const (@{const_name UNIV}, _) $ t2 => + do_const depth Ts t (@{const_name refl'}, range_type T) [t2] + | (t0 as Const (x as (@{const_name Sigma}, T))) $ t1 + $ (t2 as Abs (_, _, t2')) => + betapplys (t0 |> loose_bvar1 (t2', 0) ? do_term depth Ts, + map (do_term depth Ts) [t1, t2]) + | Const (x as (@{const_name distinct}, + Type ("fun", [Type (@{type_name list}, [T']), _]))) + $ (t1 as _ $ _) => + (t1 |> HOLogic.dest_list |> distinctness_formula T' + handle TERM _ => do_const depth Ts t x [t1]) + | (t0 as Const (x as (@{const_name If}, _))) $ t1 $ t2 $ t3 => + if is_ground_term t1 + andalso exists (Pattern.matches thy o rpair t1) + (Inttab.lookup_list ground_thm_table + (hash_term t1)) then + do_term depth Ts t2 + else + do_const depth Ts t x [t1, t2, t3] + | Const x $ t1 $ t2 $ t3 => do_const depth Ts t x [t1, t2, t3] + | Const x $ t1 $ t2 => do_const depth Ts t x [t1, t2] + | Const x $ t1 => do_const depth Ts t x [t1] + | Const x => do_const depth Ts t x [] + | t1 $ t2 => betapply (do_term depth Ts t1, do_term depth Ts t2) + | Free _ => t + | Var _ => t + | Bound _ => t + | Abs (s, T, body) => Abs (s, T, do_term depth (T :: Ts) body) + (* int -> typ list -> styp -> term list -> int -> typ -> term * term list *) + and select_nth_constr_arg_with_args _ _ (x as (_, T)) [] n res_T = + (Abs (Name.uu, body_type T, + select_nth_constr_arg thy x (Bound 0) n res_T), []) + | select_nth_constr_arg_with_args depth Ts x (t :: ts) n res_T = + (select_nth_constr_arg thy x (do_term depth Ts t) n res_T, ts) + (* int -> typ list -> term -> styp -> term list -> term *) + and do_const depth Ts t (x as (s, T)) ts = + case AList.lookup (op =) ersatz_table s of + SOME s' => + do_const (depth + 1) Ts (list_comb (Const (s', T), ts)) (s', T) ts + | NONE => + let + val (const, ts) = + if is_built_in_const fast_descrs x then + if s = @{const_name finite} then + if is_finite_type thy (domain_type T) then + (Abs ("A", domain_type T, @{const True}), ts) + else case ts of + [Const (@{const_name UNIV}, _)] => (@{const False}, []) + | _ => (Const x, ts) + else + (Const x, ts) + else case AList.lookup (op =) case_names s of + SOME n => + let + val (dataT, res_T) = nth_range_type n T + |> domain_type pairf range_type + in + (optimized_case_def thy dataT res_T + |> do_term (depth + 1) Ts, ts) + end + | _ => + if is_constr thy x then + (Const x, ts) + else if is_record_get thy x then + case length ts of + 0 => (do_term depth Ts (eta_expand Ts t 1), []) + | _ => (optimized_record_get thy s (domain_type T) + (range_type T) (hd ts), tl ts) + else if is_record_update thy x then + case length ts of + 2 => (optimized_record_update thy (unsuffix Record.updateN s) + (nth_range_type 2 T) + (do_term depth Ts (hd ts)) + (do_term depth Ts (nth ts 1)), + []) + | n => (do_term depth Ts (eta_expand Ts t (2 - n)), []) + else if is_rep_fun thy x then + let val x' = mate_of_rep_fun thy x in + if is_constr thy x' then + select_nth_constr_arg_with_args depth Ts x' ts 0 + (range_type T) + else + (Const x, ts) + end + else if is_equational_fun ext_ctxt x then + (Const x, ts) + else case def_of_const thy def_table x of + SOME def => + if depth > unfold_max_depth then + raise LIMIT ("NitpickHOL.unfold_defs_in_term", + "too many nested definitions (" ^ + string_of_int depth ^ ") while expanding " ^ + quote s) + else if s = @{const_name wfrec'} then + (do_term (depth + 1) Ts (betapplys (def, ts)), []) + else + (do_term (depth + 1) Ts def, ts) + | NONE => (Const x, ts) + in s_betapplys (const, map (do_term depth Ts) ts) |> Envir.beta_norm end + in do_term 0 [] end + +(* theory -> typ -> term list *) +fun codatatype_bisim_axioms thy T = + let + val xs = datatype_constrs thy T + val set_T = T --> bool_T + val iter_T = @{typ bisim_iterator} + val bisim_const = Const (@{const_name bisim}, [iter_T, T, T] ---> bool_T) + val bisim_max = @{const bisim_iterator_max} + val n_var = Var (("n", 0), iter_T) + val n_var_minus_1 = + Const (@{const_name Tha}, (iter_T --> bool_T) --> iter_T) + $ Abs ("m", iter_T, HOLogic.eq_const iter_T + $ (suc_const iter_T $ Bound 0) $ n_var) + val x_var = Var (("x", 0), T) + val y_var = Var (("y", 0), T) + (* styp -> int -> typ -> term *) + fun nth_sub_bisim x n nth_T = + (if is_codatatype thy nth_T then bisim_const $ n_var_minus_1 + else HOLogic.eq_const nth_T) + $ select_nth_constr_arg thy x x_var n nth_T + $ select_nth_constr_arg thy x y_var n nth_T + (* styp -> term *) + fun case_func (x as (_, T)) = + let + val arg_Ts = binder_types T + val core_t = + discriminate_value thy x y_var :: + map2 (nth_sub_bisim x) (index_seq 0 (length arg_Ts)) arg_Ts + |> foldr1 s_conj + in List.foldr absdummy core_t arg_Ts end + in + [HOLogic.eq_const bool_T $ (bisim_const $ n_var $ x_var $ y_var) + $ (@{term "op |"} $ (HOLogic.eq_const iter_T $ n_var $ zero_const iter_T) + $ (betapplys (optimized_case_def thy T bool_T, + map case_func xs @ [x_var]))), + HOLogic.eq_const set_T $ (bisim_const $ bisim_max $ x_var) + $ (Const (@{const_name insert}, [T, set_T] ---> set_T) + $ x_var $ Const (@{const_name bot_fun_inst.bot_fun}, set_T))] + |> map HOLogic.mk_Trueprop + end + +exception NO_TRIPLE of unit + +(* theory -> styp -> term -> term list * term list * term *) +fun triple_for_intro_rule thy x t = + let + val prems = Logic.strip_imp_prems t |> map (ObjectLogic.atomize_term thy) + val concl = Logic.strip_imp_concl t |> ObjectLogic.atomize_term thy + val (main, side) = List.partition (exists_Const (equal x)) prems + (* term -> bool *) + val is_good_head = equal (Const x) o head_of + in + if forall is_good_head main then (side, main, concl) else raise NO_TRIPLE () + end + +(* term -> term *) +val tuple_for_args = HOLogic.mk_tuple o snd o strip_comb + +(* indexname * typ -> term list -> term -> term -> term *) +fun wf_constraint_for rel side concl main = + let + val core = HOLogic.mk_mem (HOLogic.mk_prod (tuple_for_args main, + tuple_for_args concl), Var rel) + val t = List.foldl HOLogic.mk_imp core side + val vars = filter (not_equal rel) (Term.add_vars t []) + in + Library.foldl (fn (t', ((x, j), T)) => + HOLogic.all_const T + $ Abs (x, T, abstract_over (Var ((x, j), T), t'))) + (t, vars) + end + +(* indexname * typ -> term list * term list * term -> term *) +fun wf_constraint_for_triple rel (side, main, concl) = + map (wf_constraint_for rel side concl) main |> foldr1 s_conj + +(* Proof.context -> Time.time option -> thm + -> (Proof.context -> tactic -> tactic) -> bool *) +fun terminates_by ctxt timeout goal tac = + can (SINGLE (Classical.safe_tac (claset_of ctxt)) #> the + #> SINGLE (DETERM_TIMEOUT timeout + (tac ctxt (auto_tac (clasimpset_of ctxt)))) + #> the #> Goal.finish ctxt) goal + +val cached_timeout = Unsynchronized.ref (SOME Time.zeroTime) +val cached_wf_props : (term * bool) list Unsynchronized.ref = + Unsynchronized.ref [] + +val termination_tacs = [Lexicographic_Order.lex_order_tac, + ScnpReconstruct.sizechange_tac] + +(* extended_context -> const_table -> styp -> bool *) +fun is_is_well_founded_inductive_pred + ({thy, ctxt, debug, fast_descrs, tac_timeout, intro_table, ...} + : extended_context) (x as (_, T)) = + case def_props_for_const thy fast_descrs intro_table x of + [] => raise TERM ("NitpickHOL.is_is_well_founded_inductive_pred", [Const x]) + | intro_ts => + (case map (triple_for_intro_rule thy x) intro_ts + |> filter_out (null o #2) of + [] => true + | triples => + let + val binders_T = HOLogic.mk_tupleT (binder_types T) + val rel_T = HOLogic.mk_prodT (binders_T, binders_T) --> bool_T + val j = List.foldl Int.max 0 (map maxidx_of_term intro_ts) + 1 + val rel = (("R", j), rel_T) + val prop = Const (@{const_name wf}, rel_T --> bool_T) $ Var rel :: + map (wf_constraint_for_triple rel) triples + |> foldr1 s_conj |> HOLogic.mk_Trueprop + val _ = if debug then + priority ("Wellfoundedness goal: " ^ + Syntax.string_of_term ctxt prop ^ ".") + else + () + in + if tac_timeout = (!cached_timeout) then () + else (cached_wf_props := []; cached_timeout := tac_timeout); + case AList.lookup (op =) (!cached_wf_props) prop of + SOME wf => wf + | NONE => + let + val goal = prop |> cterm_of thy |> Goal.init + val wf = silence (exists (terminates_by ctxt tac_timeout goal)) + termination_tacs + in Unsynchronized.change cached_wf_props (cons (prop, wf)); wf end + end) + handle List.Empty => false + | NO_TRIPLE () => false + +(* The type constraint below is a workaround for a Poly/ML bug. *) + +(* extended_context -> styp -> bool *) +fun is_well_founded_inductive_pred + (ext_ctxt as {thy, wfs, def_table, wf_cache, ...} : extended_context) + (x as (s, _)) = + case triple_lookup (const_match thy) wfs x of + SOME (SOME b) => b + | _ => s mem [@{const_name Nats}, @{const_name fold_graph'}] + orelse case AList.lookup (op =) (!wf_cache) x of + SOME (_, wf) => wf + | NONE => + let + val gfp = (fixpoint_kind_of_const thy def_table x = Gfp) + val wf = is_is_well_founded_inductive_pred ext_ctxt x + in + Unsynchronized.change wf_cache (cons (x, (gfp, wf))); wf + end + +(* typ list -> typ -> typ -> term -> term *) +fun ap_curry [_] _ _ t = t + | ap_curry arg_Ts tuple_T body_T t = + let val n = length arg_Ts in + list_abs (map (pair "c") arg_Ts, + incr_boundvars n t + $ mk_flat_tuple tuple_T (map Bound (n - 1 downto 0))) + end + +(* int -> term -> int *) +fun num_occs_of_bound_in_term j (t1 $ t2) = + op + (pairself (num_occs_of_bound_in_term j) (t1, t2)) + | num_occs_of_bound_in_term j (Abs (s, T, t')) = + num_occs_of_bound_in_term (j + 1) t' + | num_occs_of_bound_in_term j (Bound j') = if j' = j then 1 else 0 + | num_occs_of_bound_in_term _ _ = 0 + +(* term -> bool *) +val is_linear_inductive_pred_def = + let + (* int -> term -> bool *) + fun do_disjunct j (Const (@{const_name Ex}, _) $ Abs (_, _, t2)) = + do_disjunct (j + 1) t2 + | do_disjunct j t = + case num_occs_of_bound_in_term j t of + 0 => true + | 1 => exists (equal (Bound j) o head_of) (conjuncts t) + | _ => false + (* term -> bool *) + fun do_lfp_def (Const (@{const_name lfp}, _) $ t2) = + let val (xs, body) = strip_abs t2 in + case length xs of + 1 => false + | n => forall (do_disjunct (n - 1)) (disjuncts body) + end + | do_lfp_def _ = false + in do_lfp_def o strip_abs_body end + +(* typ -> typ -> term -> term *) +fun ap_split tuple_T = + HOLogic.mk_psplits (HOLogic.flat_tupleT_paths tuple_T) tuple_T + +(* term -> term * term *) +val linear_pred_base_and_step_rhss = + let + (* term -> term *) + fun aux (Const (@{const_name lfp}, _) $ t2) = + let + val (xs, body) = strip_abs t2 + val arg_Ts = map snd (tl xs) + val tuple_T = HOLogic.mk_tupleT arg_Ts + val j = length arg_Ts + (* int -> term -> term *) + fun repair_rec j (Const (@{const_name Ex}, T1) $ Abs (s2, T2, t2')) = + Const (@{const_name Ex}, T1) + $ Abs (s2, T2, repair_rec (j + 1) t2') + | repair_rec j (@{const "op &"} $ t1 $ t2) = + @{const "op &"} $ repair_rec j t1 $ repair_rec j t2 + | repair_rec j t = + let val (head, args) = strip_comb t in + if head = Bound j then + HOLogic.eq_const tuple_T $ Bound j + $ mk_flat_tuple tuple_T args + else + t + end + val (nonrecs, recs) = + List.partition (equal 0 o num_occs_of_bound_in_term j) + (disjuncts body) + val base_body = nonrecs |> List.foldl s_disj @{const False} + val step_body = recs |> map (repair_rec j) + |> List.foldl s_disj @{const False} + in + (list_abs (tl xs, incr_bv (~1, j, base_body)) + |> ap_split tuple_T bool_T, + Abs ("y", tuple_T, list_abs (tl xs, step_body) + |> ap_split tuple_T bool_T)) + end + | aux t = + raise TERM ("NitpickHOL.linear_pred_base_and_step_rhss.aux", [t]) + in aux end + +(* extended_context -> styp -> term -> term *) +fun closed_linear_pred_const (ext_ctxt as {simp_table, ...}) (x as (s, T)) def = + let + val j = maxidx_of_term def + 1 + val (outer, fp_app) = strip_abs def + val outer_bounds = map Bound (length outer - 1 downto 0) + val outer_vars = map (fn (s, T) => Var ((s, j), T)) outer + val fp_app = subst_bounds (rev outer_vars, fp_app) + val (outer_Ts, rest_T) = strip_n_binders (length outer) T + val tuple_arg_Ts = strip_type rest_T |> fst + val tuple_T = HOLogic.mk_tupleT tuple_arg_Ts + val set_T = tuple_T --> bool_T + val curried_T = tuple_T --> set_T + val uncurried_T = Type ("*", [tuple_T, tuple_T]) --> bool_T + val (base_rhs, step_rhs) = linear_pred_base_and_step_rhss fp_app + val base_x as (base_s, _) = (base_prefix ^ s, outer_Ts ---> set_T) + val base_eq = HOLogic.mk_eq (list_comb (Const base_x, outer_vars), base_rhs) + |> HOLogic.mk_Trueprop + val _ = add_simps simp_table base_s [base_eq] + val step_x as (step_s, _) = (step_prefix ^ s, outer_Ts ---> curried_T) + val step_eq = HOLogic.mk_eq (list_comb (Const step_x, outer_vars), step_rhs) + |> HOLogic.mk_Trueprop + val _ = add_simps simp_table step_s [step_eq] + in + list_abs (outer, + Const (@{const_name Image}, uncurried_T --> set_T --> set_T) + $ (Const (@{const_name rtrancl}, uncurried_T --> uncurried_T) + $ (Const (@{const_name split}, curried_T --> uncurried_T) + $ list_comb (Const step_x, outer_bounds))) + $ list_comb (Const base_x, outer_bounds) + |> ap_curry tuple_arg_Ts tuple_T bool_T) + |> unfold_defs_in_term ext_ctxt + end + +(* extended_context -> bool -> styp -> term *) +fun unrolled_inductive_pred_const (ext_ctxt as {thy, star_linear_preds, + def_table, simp_table, ...}) + gfp (x as (s, T)) = + let + val iter_T = iterator_type_for_const gfp x + val x' as (s', _) = (unrolled_prefix ^ s, iter_T --> T) + val unrolled_const = Const x' $ zero_const iter_T + val def = the (def_of_const thy def_table x) + in + if is_equational_fun ext_ctxt x' then + unrolled_const (* already done *) + else if not gfp andalso is_linear_inductive_pred_def def + andalso star_linear_preds then + closed_linear_pred_const ext_ctxt x def + else + let + val j = maxidx_of_term def + 1 + val (outer, fp_app) = strip_abs def + val outer_bounds = map Bound (length outer - 1 downto 0) + val cur = Var ((iter_var_prefix, j + 1), iter_T) + val next = suc_const iter_T $ cur + val rhs = case fp_app of + Const _ $ t => + betapply (t, list_comb (Const x', next :: outer_bounds)) + | _ => raise TERM ("NitpickHOL.unrolled_inductive_pred_const", + [fp_app]) + val (inner, naked_rhs) = strip_abs rhs + val all = outer @ inner + val bounds = map Bound (length all - 1 downto 0) + val vars = map (fn (s, T) => Var ((s, j), T)) all + val eq = HOLogic.mk_eq (list_comb (Const x', cur :: bounds), naked_rhs) + |> HOLogic.mk_Trueprop |> curry subst_bounds (rev vars) + val _ = add_simps simp_table s' [eq] + in unrolled_const end + end + +(* extended_context -> styp -> term *) +fun raw_inductive_pred_axiom ({thy, def_table, ...} : extended_context) x = + let + val def = the (def_of_const thy def_table x) + val (outer, fp_app) = strip_abs def + val outer_bounds = map Bound (length outer - 1 downto 0) + val rhs = case fp_app of + Const _ $ t => betapply (t, list_comb (Const x, outer_bounds)) + | _ => raise TERM ("NitpickHOL.raw_inductive_pred_axiom", + [fp_app]) + val (inner, naked_rhs) = strip_abs rhs + val all = outer @ inner + val bounds = map Bound (length all - 1 downto 0) + val j = maxidx_of_term def + 1 + val vars = map (fn (s, T) => Var ((s, j), T)) all + in + HOLogic.mk_eq (list_comb (Const x, bounds), naked_rhs) + |> HOLogic.mk_Trueprop |> curry subst_bounds (rev vars) + end +fun inductive_pred_axiom ext_ctxt (x as (s, T)) = + if String.isPrefix ubfp_prefix s orelse String.isPrefix lbfp_prefix s then + let val x' = (after_name_sep s, T) in + raw_inductive_pred_axiom ext_ctxt x' |> subst_atomic [(Const x', Const x)] + end + else + raw_inductive_pred_axiom ext_ctxt x + +(* extended_context -> styp -> term list *) +fun raw_equational_fun_axioms (ext_ctxt as {thy, fast_descrs, simp_table, + psimp_table, ...}) (x as (s, _)) = + if s mem redefined_in_NitpickDefs_thy then + [] + else case def_props_for_const thy fast_descrs (!simp_table) x of + [] => (case def_props_for_const thy fast_descrs psimp_table x of + [] => [inductive_pred_axiom ext_ctxt x] + | psimps => psimps) + | simps => simps + +val equational_fun_axioms = map extensionalize oo raw_equational_fun_axioms + +(* term list -> term list *) +fun coalesce_type_vars_in_terms ts = + let + (* typ -> (sort * string) list -> (sort * string) list *) + fun add_type (TFree (s, S)) table = + (case AList.lookup (op =) table S of + SOME s' => + if string_ord (s', s) = LESS then AList.update (op =) (S, s') table + else table + | NONE => (S, s) :: table) + | add_type _ table = table + val table = fold (fold_types (fold_atyps add_type)) ts [] + (* typ -> typ *) + fun coalesce (TFree (s, S)) = TFree (AList.lookup (op =) table S |> the, S) + | coalesce T = T + in map (map_types (map_atyps coalesce)) ts end + +(* extended_context -> typ -> typ list -> typ list *) +fun add_ground_types ext_ctxt T accum = + case T of + Type ("fun", Ts) => fold (add_ground_types ext_ctxt) Ts accum + | Type ("*", Ts) => fold (add_ground_types ext_ctxt) Ts accum + | Type (@{type_name itself}, [T1]) => add_ground_types ext_ctxt T1 accum + | Type (_, Ts) => + if T mem @{typ prop} :: @{typ bool} :: @{typ unit} :: accum then + accum + else + T :: accum + |> fold (add_ground_types ext_ctxt) + (case boxed_datatype_constrs ext_ctxt T of + [] => Ts + | xs => map snd xs) + | _ => insert (op =) T accum +(* extended_context -> typ -> typ list *) +fun ground_types_in_type ext_ctxt T = add_ground_types ext_ctxt T [] +(* extended_context -> term list -> typ list *) +fun ground_types_in_terms ext_ctxt ts = + fold (fold_types (add_ground_types ext_ctxt)) ts [] + +(* typ list -> int -> term -> bool *) +fun has_heavy_bounds_or_vars Ts level t = + let + (* typ list -> bool *) + fun aux [] = false + | aux [T] = is_fun_type T orelse is_pair_type T + | aux _ = true + in aux (map snd (Term.add_vars t []) @ map (nth Ts) (loose_bnos t)) end + +(* typ list -> int -> int -> int -> term -> term *) +fun fresh_value_var Ts k n j t = + Var ((val_var_prefix ^ nat_subscript (n - j), k), fastype_of1 (Ts, t)) + +(* theory -> typ list -> bool -> int -> int -> term -> term list -> term list + -> term * term list *) +fun pull_out_constr_comb thy Ts relax k level t args seen = + let val t_comb = list_comb (t, args) in + case t of + Const x => + if not relax andalso is_constr thy x + andalso not (is_fun_type (fastype_of1 (Ts, t_comb))) + andalso has_heavy_bounds_or_vars Ts level t_comb + andalso not (loose_bvar (t_comb, level)) then + let + val (j, seen) = case find_index (equal t_comb) seen of + ~1 => (0, t_comb :: seen) + | j => (j, seen) + in (fresh_value_var Ts k (length seen) j t_comb, seen) end + else + (t_comb, seen) + | _ => (t_comb, seen) + end + +(* (term -> term) -> typ list -> int -> term list -> term list *) +fun equations_for_pulled_out_constrs mk_eq Ts k seen = + let val n = length seen in + map2 (fn j => fn t => mk_eq (fresh_value_var Ts k n j t, t)) + (index_seq 0 n) seen + end + +(* theory -> bool -> term -> term *) +fun pull_out_universal_constrs thy def t = + let + val k = maxidx_of_term t + 1 + (* typ list -> bool -> term -> term list -> term list -> term * term list *) + fun do_term Ts def t args seen = + case t of + (t0 as Const (@{const_name "=="}, _)) $ t1 $ t2 => + do_eq_or_imp Ts def t0 t1 t2 seen + | (t0 as @{const "==>"}) $ t1 $ t2 => do_eq_or_imp Ts def t0 t1 t2 seen + | (t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2 => + do_eq_or_imp Ts def t0 t1 t2 seen + | (t0 as @{const "op -->"}) $ t1 $ t2 => do_eq_or_imp Ts def t0 t1 t2 seen + | Abs (s, T, t') => + let val (t', seen) = do_term (T :: Ts) def t' [] seen in + (list_comb (Abs (s, T, t'), args), seen) + end + | t1 $ t2 => + let val (t2, seen) = do_term Ts def t2 [] seen in + do_term Ts def t1 (t2 :: args) seen + end + | _ => pull_out_constr_comb thy Ts def k 0 t args seen + (* typ list -> bool -> term -> term -> term -> term list + -> term * term list *) + and do_eq_or_imp Ts def t0 t1 t2 seen = + let + val (t2, seen) = do_term Ts def t2 [] seen + val (t1, seen) = do_term Ts false t1 [] seen + in (t0 $ t1 $ t2, seen) end + val (concl, seen) = do_term [] def t [] [] + in + Logic.list_implies (equations_for_pulled_out_constrs Logic.mk_equals [] k + seen, concl) + end + +(* theory -> bool -> term -> term *) +fun destroy_pulled_out_constrs thy axiom t = + let + (* styp -> int *) + val num_occs_of_var = + fold_aterms (fn Var z => (fn f => fn z' => f z' |> z = z' ? Integer.add 1) + | _ => I) t (K 0) + (* bool -> term -> term *) + fun aux careful ((t0 as Const (@{const_name "=="}, _)) $ t1 $ t2) = + aux_eq careful true t0 t1 t2 + | aux careful ((t0 as @{const "==>"}) $ t1 $ t2) = + t0 $ aux false t1 $ aux careful t2 + | aux careful ((t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2) = + aux_eq careful true t0 t1 t2 + | aux careful ((t0 as @{const "op -->"}) $ t1 $ t2) = + t0 $ aux false t1 $ aux careful t2 + | aux careful (Abs (s, T, t')) = Abs (s, T, aux careful t') + | aux careful (t1 $ t2) = aux careful t1 $ aux careful t2 + | aux _ t = t + (* bool -> bool -> term -> term -> term -> term *) + and aux_eq careful pass1 t0 t1 t2 = + (if careful then + raise SAME () + else if axiom andalso is_Var t2 + andalso num_occs_of_var (dest_Var t2) = 1 then + @{const True} + else case strip_comb t2 of + (Const (x as (s, T)), args) => + let val arg_Ts = binder_types T in + if length arg_Ts = length args + andalso (is_constr thy x orelse s mem [@{const_name Pair}] + orelse x = dest_Const @{const Suc}) + andalso (not careful orelse not (is_Var t1) + orelse String.isPrefix val_var_prefix + (fst (fst (dest_Var t1)))) then + discriminate_value thy x t1 :: + map3 (sel_eq x t1) (index_seq 0 (length args)) arg_Ts args + |> foldr1 s_conj + |> body_type (type_of t0) = prop_T ? HOLogic.mk_Trueprop + else + raise SAME () + end + | _ => raise SAME ()) + handle SAME () => if pass1 then aux_eq careful false t0 t2 t1 + else t0 $ aux false t2 $ aux false t1 + (* styp -> term -> int -> typ -> term -> term *) + and sel_eq x t n nth_T nth_t = + HOLogic.eq_const nth_T $ nth_t $ select_nth_constr_arg thy x t n nth_T + |> aux false + in aux axiom t end + +(* theory -> term -> term *) +fun simplify_constrs_and_sels thy t = + let + (* term -> int -> term *) + fun is_nth_sel_on t' n (Const (s, _) $ t) = + (t = t' andalso is_sel_like_and_no_discr s + andalso sel_no_from_name s = n) + | is_nth_sel_on _ _ _ = false + (* term -> term list -> term *) + fun do_term (Const (@{const_name Rep_Frac}, _) + $ (Const (@{const_name Abs_Frac}, _) $ t1)) [] = do_term t1 [] + | do_term (Const (@{const_name Abs_Frac}, _) + $ (Const (@{const_name Rep_Frac}, _) $ t1)) [] = do_term t1 [] + | do_term (t1 $ t2) args = do_term t1 (do_term t2 [] :: args) + | do_term (t as Const (x as (s, T))) (args as _ :: _) = + ((if is_constr_like thy x then + if length args = num_binder_types T then + case hd args of + Const (x' as (_, T')) $ t' => + if domain_type T' = body_type T + andalso forall (uncurry (is_nth_sel_on t')) + (index_seq 0 (length args) ~~ args) then + t' + else + raise SAME () + | _ => raise SAME () + else + raise SAME () + else if is_sel_like_and_no_discr s then + case strip_comb (hd args) of + (Const (x' as (s', T')), ts') => + if is_constr_like thy x' + andalso constr_name_for_sel_like s = s' + andalso not (exists is_pair_type (binder_types T')) then + list_comb (nth ts' (sel_no_from_name s), tl args) + else + raise SAME () + | _ => raise SAME () + else + raise SAME ()) + handle SAME () => betapplys (t, args)) + | do_term (Abs (s, T, t')) args = + betapplys (Abs (s, T, do_term t' []), args) + | do_term t args = betapplys (t, args) + in do_term t [] end + +(* term -> term *) +fun curry_assms (@{const "==>"} $ (@{const Trueprop} + $ (@{const "op &"} $ t1 $ t2)) $ t3) = + curry_assms (Logic.list_implies ([t1, t2] |> map HOLogic.mk_Trueprop, t3)) + | curry_assms (@{const "==>"} $ t1 $ t2) = + @{const "==>"} $ curry_assms t1 $ curry_assms t2 + | curry_assms t = t + +(* term -> term *) +val destroy_universal_equalities = + let + (* term list -> (indexname * typ) list -> term -> term *) + fun aux prems zs t = + case t of + @{const "==>"} $ t1 $ t2 => aux_implies prems zs t1 t2 + | _ => Logic.list_implies (rev prems, t) + (* term list -> (indexname * typ) list -> term -> term -> term *) + and aux_implies prems zs t1 t2 = + case t1 of + Const (@{const_name "=="}, _) $ Var z $ t' => aux_eq prems zs z t' t1 t2 + | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ Var z $ t') => + aux_eq prems zs z t' t1 t2 + | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t' $ Var z) => + aux_eq prems zs z t' t1 t2 + | _ => aux (t1 :: prems) (Term.add_vars t1 zs) t2 + (* term list -> (indexname * typ) list -> indexname * typ -> term -> term + -> term -> term *) + and aux_eq prems zs z t' t1 t2 = + if not (z mem zs) andalso not (exists_subterm (equal (Var z)) t') then + aux prems zs (subst_free [(Var z, t')] t2) + else + aux (t1 :: prems) (Term.add_vars t1 zs) t2 + in aux [] [] end + +(* theory -> term -> term *) +fun pull_out_existential_constrs thy t = + let + val k = maxidx_of_term t + 1 + (* typ list -> int -> term -> term list -> term list -> term * term list *) + fun aux Ts num_exists t args seen = + case t of + (t0 as Const (@{const_name Ex}, _)) $ Abs (s1, T1, t1) => + let + val (t1, seen') = aux (T1 :: Ts) (num_exists + 1) t1 [] [] + val n = length seen' + (* unit -> term list *) + fun vars () = map2 (fresh_value_var Ts k n) (index_seq 0 n) seen' + in + (equations_for_pulled_out_constrs HOLogic.mk_eq Ts k seen' + |> List.foldl s_conj t1 |> fold mk_exists (vars ()) + |> curry3 Abs s1 T1 |> curry (op $) t0, seen) + end + | t1 $ t2 => + let val (t2, seen) = aux Ts num_exists t2 [] seen in + aux Ts num_exists t1 (t2 :: args) seen + end + | Abs (s, T, t') => + let + val (t', seen) = aux (T :: Ts) 0 t' [] (map (incr_boundvars 1) seen) + in (list_comb (Abs (s, T, t'), args), map (incr_boundvars ~1) seen) end + | _ => + if num_exists > 0 then + pull_out_constr_comb thy Ts false k num_exists t args seen + else + (list_comb (t, args), seen) + in aux [] 0 t [] [] |> fst end + +(* theory -> int -> term list -> term list -> (term * term list) option *) +fun find_bound_assign _ _ _ [] = NONE + | find_bound_assign thy j seen (t :: ts) = + let + (* bool -> term -> term -> (term * term list) option *) + fun aux pass1 t1 t2 = + (if loose_bvar1 (t2, j) then + if pass1 then aux false t2 t1 else raise SAME () + else case t1 of + Bound j' => if j' = j then SOME (t2, ts @ seen) else raise SAME () + | Const (s, Type ("fun", [T1, T2])) $ Bound j' => + if j' = j andalso s = sel_prefix_for 0 ^ @{const_name FunBox} then + SOME (construct_value thy (@{const_name FunBox}, T2 --> T1) [t2], + ts @ seen) + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => find_bound_assign thy j (t :: seen) ts + in + case t of + Const (@{const_name "op ="}, _) $ t1 $ t2 => aux true t1 t2 + | _ => find_bound_assign thy j (t :: seen) ts + end + +(* int -> term -> term -> term *) +fun subst_one_bound j arg t = + let + fun aux (Bound i, lev) = + if i < lev then raise SAME () + else if i = lev then incr_boundvars (lev - j) arg + else Bound (i - 1) + | aux (Abs (a, T, body), lev) = Abs (a, T, aux (body, lev + 1)) + | aux (f $ t, lev) = + (aux (f, lev) $ (aux (t, lev) handle SAME () => t) + handle SAME () => f $ aux (t, lev)) + | aux _ = raise SAME () + in aux (t, j) handle SAME () => t end + +(* theory -> term -> term *) +fun destroy_existential_equalities thy = + let + (* string list -> typ list -> term list -> term *) + fun kill [] [] ts = foldr1 s_conj ts + | kill (s :: ss) (T :: Ts) ts = + (case find_bound_assign thy (length ss) [] ts of + SOME (_, []) => @{const True} + | SOME (arg_t, ts) => + kill ss Ts (map (subst_one_bound (length ss) + (incr_bv (~1, length ss + 1, arg_t))) ts) + | NONE => + Const (@{const_name Ex}, (T --> bool_T) --> bool_T) + $ Abs (s, T, kill ss Ts ts)) + | kill _ _ _ = raise UnequalLengths + (* string list -> typ list -> term -> term *) + fun gather ss Ts ((t0 as Const (@{const_name Ex}, _)) $ Abs (s1, T1, t1)) = + gather (ss @ [s1]) (Ts @ [T1]) t1 + | gather [] [] (Abs (s, T, t1)) = Abs (s, T, gather [] [] t1) + | gather [] [] (t1 $ t2) = gather [] [] t1 $ gather [] [] t2 + | gather [] [] t = t + | gather ss Ts t = kill ss Ts (conjuncts (gather [] [] t)) + in gather [] [] end + +(* term -> term *) +fun distribute_quantifiers t = + case t of + (t0 as Const (@{const_name All}, T0)) $ Abs (s, T1, t1) => + (case t1 of + (t10 as @{const "op &"}) $ t11 $ t12 => + t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11)) + $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) + | (t10 as @{const Not}) $ t11 => + t10 $ distribute_quantifiers (Const (@{const_name Ex}, T0) + $ Abs (s, T1, t11)) + | t1 => + if not (loose_bvar1 (t1, 0)) then + distribute_quantifiers (incr_boundvars ~1 t1) + else + t0 $ Abs (s, T1, distribute_quantifiers t1)) + | (t0 as Const (@{const_name Ex}, T0)) $ Abs (s, T1, t1) => + (case distribute_quantifiers t1 of + (t10 as @{const "op |"}) $ t11 $ t12 => + t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11)) + $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) + | (t10 as @{const "op -->"}) $ t11 $ t12 => + t10 $ distribute_quantifiers (Const (@{const_name All}, T0) + $ Abs (s, T1, t11)) + $ distribute_quantifiers (t0 $ Abs (s, T1, t12)) + | (t10 as @{const Not}) $ t11 => + t10 $ distribute_quantifiers (Const (@{const_name All}, T0) + $ Abs (s, T1, t11)) + | t1 => + if not (loose_bvar1 (t1, 0)) then + distribute_quantifiers (incr_boundvars ~1 t1) + else + t0 $ Abs (s, T1, distribute_quantifiers t1)) + | t1 $ t2 => distribute_quantifiers t1 $ distribute_quantifiers t2 + | Abs (s, T, t') => Abs (s, T, distribute_quantifiers t') + | _ => t + +(* int -> int -> (int -> int) -> term -> term *) +fun renumber_bounds j n f t = + case t of + t1 $ t2 => renumber_bounds j n f t1 $ renumber_bounds j n f t2 + | Abs (s, T, t') => Abs (s, T, renumber_bounds (j + 1) n f t') + | Bound j' => + Bound (if j' >= j andalso j' < j + n then f (j' - j) + j else j') + | _ => t + +val quantifier_cluster_max_size = 8 + +(* theory -> term -> term *) +fun push_quantifiers_inward thy = + let + (* string -> string list -> typ list -> term -> term *) + fun aux quant_s ss Ts t = + (case t of + (t0 as Const (s0, _)) $ Abs (s1, T1, t1 as _ $ _) => + if s0 = quant_s andalso length Ts < quantifier_cluster_max_size then + aux s0 (s1 :: ss) (T1 :: Ts) t1 + else if quant_s = "" + andalso s0 mem [@{const_name All}, @{const_name Ex}] then + aux s0 [s1] [T1] t1 + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => + case t of + t1 $ t2 => + if quant_s = "" then + aux "" [] [] t1 $ aux "" [] [] t2 + else + let + val typical_card = 4 + (* ('a -> ''b list) -> 'a list -> ''b list *) + fun big_union proj ps = + fold (fold (insert (op =)) o proj) ps [] + val (ts, connective) = strip_any_connective t + val T_costs = + map (bounded_card_of_type 65536 typical_card []) Ts + val t_costs = map size_of_term ts + val num_Ts = length Ts + (* int -> int *) + val flip = curry (op -) (num_Ts - 1) + val t_boundss = map (map flip o loose_bnos) ts + (* (int list * int) list -> int list -> int *) + fun cost boundss_cum_costs [] = + map snd boundss_cum_costs |> Integer.sum + | cost boundss_cum_costs (j :: js) = + let + val (yeas, nays) = + List.partition (fn (bounds, _) => j mem bounds) + boundss_cum_costs + val yeas_bounds = big_union fst yeas + val yeas_cost = Integer.sum (map snd yeas) + * nth T_costs j + in cost ((yeas_bounds, yeas_cost) :: nays) js end + val js = all_permutations (index_seq 0 num_Ts) + |> map (`(cost (t_boundss ~~ t_costs))) + |> sort (int_ord o pairself fst) |> hd |> snd + val back_js = map (fn j => find_index (equal j) js) + (index_seq 0 num_Ts) + val ts = map (renumber_bounds 0 num_Ts (nth back_js o flip)) + ts + (* (term * int list) list -> term *) + fun mk_connection [] = + raise ARG ("NitpickHOL.push_quantifiers_inward.aux.\ + \mk_connection", "") + | mk_connection ts_cum_bounds = + ts_cum_bounds |> map fst + |> foldr1 (fn (t1, t2) => connective $ t1 $ t2) + (* (term * int list) list -> int list -> term *) + fun build ts_cum_bounds [] = ts_cum_bounds |> mk_connection + | build ts_cum_bounds (j :: js) = + let + val (yeas, nays) = + List.partition (fn (_, bounds) => j mem bounds) + ts_cum_bounds + ||> map (apfst (incr_boundvars ~1)) + in + if null yeas then + build nays js + else + let val T = nth Ts (flip j) in + build ((Const (quant_s, (T --> bool_T) --> bool_T) + $ Abs (nth ss (flip j), T, + mk_connection yeas), + big_union snd yeas) :: nays) js + end + end + in build (ts ~~ t_boundss) js end + | Abs (s, T, t') => Abs (s, T, aux "" [] [] t') + | _ => t + in aux "" [] [] end + +(* polarity -> string -> bool *) +fun is_positive_existential polar quant_s = + (polar = Pos andalso quant_s = @{const_name Ex}) + orelse (polar = Neg andalso quant_s <> @{const_name Ex}) + +(* extended_context -> int -> term -> term *) +fun skolemize_term_and_more (ext_ctxt as {thy, def_table, skolems, ...}) + skolem_depth = + let + (* int list -> int list *) + val incrs = map (Integer.add 1) + (* string list -> typ list -> int list -> int -> polarity -> term -> term *) + fun aux ss Ts js depth polar t = + let + (* string -> typ -> string -> typ -> term -> term *) + fun do_quantifier quant_s quant_T abs_s abs_T t = + if not (loose_bvar1 (t, 0)) then + aux ss Ts js depth polar (incr_boundvars ~1 t) + else if depth <= skolem_depth + andalso is_positive_existential polar quant_s then + let + val j = length (!skolems) + 1 + val sko_s = skolem_prefix_for (length js) j ^ abs_s + val _ = Unsynchronized.change skolems (cons (sko_s, ss)) + val sko_t = list_comb (Const (sko_s, rev Ts ---> abs_T), + map Bound (rev js)) + val abs_t = Abs (abs_s, abs_T, aux ss Ts (incrs js) depth polar t) + in + if null js then betapply (abs_t, sko_t) + else Const (@{const_name Let}, abs_T --> quant_T) $ sko_t $ abs_t + end + else + Const (quant_s, quant_T) + $ Abs (abs_s, abs_T, + if is_higher_order_type abs_T then + t + else + aux (abs_s :: ss) (abs_T :: Ts) (0 :: incrs js) + (depth + 1) polar t) + in + case t of + Const (s0 as @{const_name all}, T0) $ Abs (s1, T1, t1) => + do_quantifier s0 T0 s1 T1 t1 + | @{const "==>"} $ t1 $ t2 => + @{const "==>"} $ aux ss Ts js depth (flip_polarity polar) t1 + $ aux ss Ts js depth polar t2 + | @{const Pure.conjunction} $ t1 $ t2 => + @{const Pure.conjunction} $ aux ss Ts js depth polar t1 + $ aux ss Ts js depth polar t2 + | @{const Trueprop} $ t1 => + @{const Trueprop} $ aux ss Ts js depth polar t1 + | @{const Not} $ t1 => + @{const Not} $ aux ss Ts js depth (flip_polarity polar) t1 + | Const (s0 as @{const_name All}, T0) $ Abs (s1, T1, t1) => + do_quantifier s0 T0 s1 T1 t1 + | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) => + do_quantifier s0 T0 s1 T1 t1 + | @{const "op &"} $ t1 $ t2 => + @{const "op &"} $ aux ss Ts js depth polar t1 + $ aux ss Ts js depth polar t2 + | @{const "op |"} $ t1 $ t2 => + @{const "op |"} $ aux ss Ts js depth polar t1 + $ aux ss Ts js depth polar t2 + | @{const "op -->"} $ t1 $ t2 => + @{const "op -->"} $ aux ss Ts js depth (flip_polarity polar) t1 + $ aux ss Ts js depth polar t2 + | (t0 as Const (@{const_name Let}, T0)) $ t1 $ t2 => + t0 $ t1 $ aux ss Ts js depth polar t2 + | Const (x as (s, T)) => + if is_inductive_pred ext_ctxt x + andalso not (is_well_founded_inductive_pred ext_ctxt x) then + let + val gfp = (fixpoint_kind_of_const thy def_table x = Gfp) + val (pref, connective, set_oper) = + if gfp then + (lbfp_prefix, + @{const "op |"}, + @{const_name upper_semilattice_fun_inst.sup_fun}) + else + (ubfp_prefix, + @{const "op &"}, + @{const_name lower_semilattice_fun_inst.inf_fun}) + (* unit -> term *) + fun pos () = unrolled_inductive_pred_const ext_ctxt gfp x + |> aux ss Ts js depth polar + fun neg () = Const (pref ^ s, T) + in + (case polar |> gfp ? flip_polarity of + Pos => pos () + | Neg => neg () + | Neut => + if is_fun_type T then + let + val ((trunk_arg_Ts, rump_arg_T), body_T) = + T |> strip_type |>> split_last + val set_T = rump_arg_T --> body_T + (* (unit -> term) -> term *) + fun app f = + list_comb (f (), + map Bound (length trunk_arg_Ts - 1 downto 0)) + in + List.foldl absdummy + (Const (set_oper, [set_T, set_T] ---> set_T) + $ app pos $ app neg) trunk_arg_Ts + end + else + connective $ pos () $ neg ()) + end + else + Const x + | t1 $ t2 => + betapply (aux ss Ts [] (skolem_depth + 1) polar t1, + aux ss Ts [] depth Neut t2) + | Abs (s, T, t1) => Abs (s, T, aux ss Ts (incrs js) depth polar t1) + | _ => t + end + in aux [] [] [] 0 Pos end + +(* extended_context -> styp -> (int * term option) list *) +fun static_args_in_term ({ersatz_table, ...} : extended_context) x t = + let + (* term -> term list -> term list -> term list list *) + fun fun_calls (Abs (_, _, t)) _ = fun_calls t [] + | fun_calls (t1 $ t2) args = fun_calls t2 [] #> fun_calls t1 (t2 :: args) + | fun_calls t args = + (case t of + Const (x' as (s', T')) => + x = x' orelse (case AList.lookup (op =) ersatz_table s' of + SOME s'' => x = (s'', T') + | NONE => false) + | _ => false) ? cons args + (* term list list -> term list list -> term list -> term list list *) + fun call_sets [] [] vs = [vs] + | call_sets [] uss vs = vs :: call_sets uss [] [] + | call_sets ([] :: _) _ _ = [] + | call_sets ((t :: ts) :: tss) uss vs = + OrdList.insert TermOrd.term_ord t vs |> call_sets tss (ts :: uss) + val sets = call_sets (fun_calls t [] []) [] [] + val indexed_sets = sets ~~ (index_seq 0 (length sets)) + in + fold_rev (fn (set, j) => + case set of + [Var _] => AList.lookup (op =) indexed_sets set = SOME j + ? cons (j, NONE) + | [t as Const _] => cons (j, SOME t) + | [t as Free _] => cons (j, SOME t) + | _ => I) indexed_sets [] + end +(* extended_context -> styp -> term list -> (int * term option) list *) +fun static_args_in_terms ext_ctxt x = + map (static_args_in_term ext_ctxt x) + #> fold1 (OrdList.inter (prod_ord int_ord (option_ord TermOrd.term_ord))) + +(* term -> term list *) +fun params_in_equation (@{const "==>"} $ _ $ t2) = params_in_equation t2 + | params_in_equation (@{const Trueprop} $ t1) = params_in_equation t1 + | params_in_equation (Const (@{const_name "op ="}, _) $ t1 $ _) = + snd (strip_comb t1) + | params_in_equation _ = [] + +(* styp -> styp -> int list -> term list -> term list -> term -> term *) +fun specialize_fun_axiom x x' fixed_js fixed_args extra_args t = + let + val k = fold Integer.max (map maxidx_of_term (fixed_args @ extra_args)) 0 + + 1 + val t = map_aterms (fn Var ((s, i), T) => Var ((s, k + i), T) | t' => t') t + val fixed_params = filter_indices fixed_js (params_in_equation t) + (* term list -> term -> term *) + fun aux args (Abs (s, T, t)) = list_comb (Abs (s, T, aux [] t), args) + | aux args (t1 $ t2) = aux (aux [] t2 :: args) t1 + | aux args t = + if t = Const x then + list_comb (Const x', extra_args @ filter_out_indices fixed_js args) + else + let val j = find_index (equal t) fixed_params in + list_comb (if j >= 0 then nth fixed_args j else t, args) + end + in aux [] t end + +(* typ list -> term -> bool *) +fun is_eligible_arg Ts t = + let val bad_Ts = map snd (Term.add_vars t []) @ map (nth Ts) (loose_bnos t) in + null bad_Ts + orelse (is_higher_order_type (fastype_of1 (Ts, t)) + andalso forall (not o is_higher_order_type) bad_Ts) + end + +(* (int * term option) list -> (int * term) list -> int list *) +fun overlapping_indices [] _ = [] + | overlapping_indices _ [] = [] + | overlapping_indices (ps1 as (j1, t1) :: ps1') (ps2 as (j2, t2) :: ps2') = + if j1 < j2 then overlapping_indices ps1' ps2 + else if j1 > j2 then overlapping_indices ps1 ps2' + else overlapping_indices ps1' ps2' |> the_default t2 t1 = t2 ? cons j1 + +val special_depth = 20 + +(* extended_context -> int -> term -> term *) +fun specialize_consts_in_term (ext_ctxt as {thy, specialize, simp_table, + special_funs, ...}) depth t = + if not specialize orelse depth > special_depth then + t + else + let + (* FIXME: strong enough in the face of user-defined axioms? *) + val blacklist = if depth = 0 then [] + else case term_under_def t of Const x => [x] | _ => [] + (* term list -> typ list -> term -> term *) + fun aux args Ts (Const (x as (s, T))) = + ((if not (x mem blacklist) andalso not (null args) + andalso not (String.isPrefix special_prefix s) + andalso is_equational_fun ext_ctxt x then + let + val eligible_args = filter (is_eligible_arg Ts o snd) + (index_seq 0 (length args) ~~ args) + val _ = not (null eligible_args) orelse raise SAME () + val old_axs = equational_fun_axioms ext_ctxt x + |> map (destroy_existential_equalities thy) + val static_params = static_args_in_terms ext_ctxt x old_axs + val fixed_js = overlapping_indices static_params eligible_args + val _ = not (null fixed_js) orelse raise SAME () + val fixed_args = filter_indices fixed_js args + val vars = fold Term.add_vars fixed_args [] + |> sort (TermOrd.fast_indexname_ord o pairself fst) + val bound_js = fold (fn t => fn js => add_loose_bnos (t, 0, js)) + fixed_args [] + |> sort int_ord + val live_args = filter_out_indices fixed_js args + val extra_args = map Var vars @ map Bound bound_js @ live_args + val extra_Ts = map snd vars @ filter_indices bound_js Ts + val k = maxidx_of_term t + 1 + (* int -> term *) + fun var_for_bound_no j = + Var ((bound_var_prefix ^ + nat_subscript (find_index (equal j) bound_js + 1), k), + nth Ts j) + val fixed_args_in_axiom = + map (curry subst_bounds + (map var_for_bound_no (index_seq 0 (length Ts)))) + fixed_args + in + case AList.lookup (op =) (!special_funs) + (x, fixed_js, fixed_args_in_axiom) of + SOME x' => list_comb (Const x', extra_args) + | NONE => + let + val extra_args_in_axiom = + map Var vars @ map var_for_bound_no bound_js + val x' as (s', _) = + (special_prefix_for (length (!special_funs) + 1) ^ s, + extra_Ts @ filter_out_indices fixed_js (binder_types T) + ---> body_type T) + val new_axs = + map (specialize_fun_axiom x x' fixed_js + fixed_args_in_axiom extra_args_in_axiom) old_axs + val _ = + Unsynchronized.change special_funs + (cons ((x, fixed_js, fixed_args_in_axiom), x')) + val _ = add_simps simp_table s' new_axs + in list_comb (Const x', extra_args) end + end + else + raise SAME ()) + handle SAME () => list_comb (Const x, args)) + | aux args Ts (Abs (s, T, t)) = + list_comb (Abs (s, T, aux [] (T :: Ts) t), args) + | aux args Ts (t1 $ t2) = aux (aux [] Ts t2 :: args) Ts t1 + | aux args _ t = list_comb (t, args) + in aux [] [] t end + +(* theory -> term -> int Termtab.tab -> int Termtab.tab *) +fun add_to_uncurry_table thy t = + let + (* term -> term list -> int Termtab.tab -> int Termtab.tab *) + fun aux (t1 $ t2) args table = + let val table = aux t2 [] table in aux t1 (t2 :: args) table end + | aux (Abs (_, _, t')) _ table = aux t' [] table + | aux (t as Const (x as (s, _))) args table = + if is_built_in_const false x orelse is_constr_like thy x orelse is_sel s + orelse s = @{const_name Sigma} then + table + else + Termtab.map_default (t, 65536) (curry Int.min (length args)) table + | aux _ _ table = table + in aux t [] end + +(* int Termtab.tab term -> term *) +fun uncurry_term table t = + let + (* term -> term list -> term *) + fun aux (t1 $ t2) args = aux t1 (aux t2 [] :: args) + | aux (Abs (s, T, t')) args = betapplys (Abs (s, T, aux t' []), args) + | aux (t as Const (s, T)) args = + (case Termtab.lookup table t of + SOME n => + if n >= 2 then + let + val (arg_Ts, rest_T) = strip_n_binders n T + val j = + if hd arg_Ts = @{typ bisim_iterator} + orelse is_fp_iterator_type (hd arg_Ts) then + 1 + else case find_index (not_equal bool_T) arg_Ts of + ~1 => n + | j => j + val ((before_args, tuple_args), after_args) = + args |> chop n |>> chop j + val ((before_arg_Ts, tuple_arg_Ts), rest_T) = + T |> strip_n_binders n |>> chop j + val tuple_T = HOLogic.mk_tupleT tuple_arg_Ts + in + if n - j < 2 then + betapplys (t, args) + else + betapplys (Const (uncurry_prefix_for (n - j) j ^ s, + before_arg_Ts ---> tuple_T --> rest_T), + before_args @ [mk_flat_tuple tuple_T tuple_args] @ + after_args) + end + else + betapplys (t, args) + | NONE => betapplys (t, args)) + | aux t args = betapplys (t, args) + in aux t [] end + +(* (term -> term) -> int -> term -> term *) +fun coerce_bound_no f j t = + case t of + t1 $ t2 => coerce_bound_no f j t1 $ coerce_bound_no f j t2 + | Abs (s, T, t') => Abs (s, T, coerce_bound_no f (j + 1) t') + | Bound j' => if j' = j then f t else t + | _ => t + +(* extended_context -> bool -> term -> term *) +fun box_fun_and_pair_in_term (ext_ctxt as {thy, fast_descrs, ...}) def orig_t = + let + (* typ -> typ *) + fun box_relational_operator_type (Type ("fun", Ts)) = + Type ("fun", map box_relational_operator_type Ts) + | box_relational_operator_type (Type ("*", Ts)) = + Type ("*", map (box_type ext_ctxt InPair) Ts) + | box_relational_operator_type T = T + (* typ -> typ -> term -> term *) + fun coerce_bound_0_in_term new_T old_T = + old_T <> new_T ? coerce_bound_no (coerce_term [new_T] old_T new_T) 0 + (* typ list -> typ -> term -> term *) + and coerce_term Ts new_T old_T t = + if old_T = new_T then + t + else + case (new_T, old_T) of + (Type (new_s, new_Ts as [new_T1, new_T2]), + Type ("fun", [old_T1, old_T2])) => + (case eta_expand Ts t 1 of + Abs (s, _, t') => + Abs (s, new_T1, + t' |> coerce_bound_0_in_term new_T1 old_T1 + |> coerce_term (new_T1 :: Ts) new_T2 old_T2) + |> Envir.eta_contract + |> new_s <> "fun" + ? construct_value thy (@{const_name FunBox}, + Type ("fun", new_Ts) --> new_T) o single + | t' => raise TERM ("NitpickHOL.box_fun_and_pair_in_term.\ + \coerce_term", [t'])) + | (Type (new_s, new_Ts as [new_T1, new_T2]), + Type (old_s, old_Ts as [old_T1, old_T2])) => + if old_s mem [@{type_name fun_box}, @{type_name pair_box}, "*"] then + case constr_expand thy old_T t of + Const (@{const_name FunBox}, _) $ t1 => + if new_s = "fun" then + coerce_term Ts new_T (Type ("fun", old_Ts)) t1 + else + construct_value thy + (@{const_name FunBox}, Type ("fun", new_Ts) --> new_T) + [coerce_term Ts (Type ("fun", new_Ts)) + (Type ("fun", old_Ts)) t1] + | Const _ $ t1 $ t2 => + construct_value thy + (if new_s = "*" then @{const_name Pair} + else @{const_name PairBox}, new_Ts ---> new_T) + [coerce_term Ts new_T1 old_T1 t1, + coerce_term Ts new_T2 old_T2 t2] + | t' => raise TERM ("NitpickHOL.box_fun_and_pair_in_term.\ + \coerce_term", [t']) + else + raise TYPE ("coerce_term", [new_T, old_T], [t]) + | _ => raise TYPE ("coerce_term", [new_T, old_T], [t]) + (* indexname * typ -> typ * term -> typ option list -> typ option list *) + fun add_boxed_types_for_var (z as (_, T)) (T', t') = + case t' of + Var z' => z' = z ? insert (op =) T' + | Const (@{const_name Pair}, _) $ t1 $ t2 => + (case T' of + Type (_, [T1, T2]) => + fold (add_boxed_types_for_var z) [(T1, t1), (T2, t2)] + | _ => raise TYPE ("NitpickHOL.box_fun_and_pair_in_term.\ + \add_boxed_types_for_var", [T'], [])) + | _ => exists_subterm (equal (Var z)) t' ? insert (op =) T + (* typ list -> typ list -> term -> indexname * typ -> typ *) + fun box_var_in_def new_Ts old_Ts t (z as (_, T)) = + case t of + @{const Trueprop} $ t1 => box_var_in_def new_Ts old_Ts t1 z + | Const (s0, _) $ t1 $ _ => + if s0 mem [@{const_name "=="}, @{const_name "op ="}] then + let + val (t', args) = strip_comb t1 + val T' = fastype_of1 (new_Ts, do_term new_Ts old_Ts Neut t') + in + case fold (add_boxed_types_for_var z) + (fst (strip_n_binders (length args) T') ~~ args) [] of + [T''] => T'' + | _ => T + end + else + T + | _ => T + (* typ list -> typ list -> polarity -> string -> typ -> string -> typ + -> term -> term *) + and do_quantifier new_Ts old_Ts polar quant_s quant_T abs_s abs_T t = + let + val abs_T' = + if polar = Neut orelse is_positive_existential polar quant_s then + box_type ext_ctxt InFunLHS abs_T + else + abs_T + val body_T = body_type quant_T + in + Const (quant_s, (abs_T' --> body_T) --> body_T) + $ Abs (abs_s, abs_T', + t |> do_term (abs_T' :: new_Ts) (abs_T :: old_Ts) polar) + end + (* typ list -> typ list -> string -> typ -> term -> term -> term *) + and do_equals new_Ts old_Ts s0 T0 t1 t2 = + let + val (t1, t2) = pairself (do_term new_Ts old_Ts Neut) (t1, t2) + val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2) + val T = [T1, T2] |> sort TermOrd.typ_ord |> List.last + in + list_comb (Const (s0, [T, T] ---> body_type T0), + map2 (coerce_term new_Ts T) [T1, T2] [t1, t2]) + end + (* string -> typ -> term *) + and do_description_operator s T = + let val T1 = box_type ext_ctxt InFunLHS (range_type T) in + Const (s, (T1 --> bool_T) --> T1) + end + (* typ list -> typ list -> polarity -> term -> term *) + and do_term new_Ts old_Ts polar t = + case t of + Const (s0 as @{const_name all}, T0) $ Abs (s1, T1, t1) => + do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 + | Const (s0 as @{const_name "=="}, T0) $ t1 $ t2 => + do_equals new_Ts old_Ts s0 T0 t1 t2 + | @{const "==>"} $ t1 $ t2 => + @{const "==>"} $ do_term new_Ts old_Ts (flip_polarity polar) t1 + $ do_term new_Ts old_Ts polar t2 + | @{const Pure.conjunction} $ t1 $ t2 => + @{const Pure.conjunction} $ do_term new_Ts old_Ts polar t1 + $ do_term new_Ts old_Ts polar t2 + | @{const Trueprop} $ t1 => + @{const Trueprop} $ do_term new_Ts old_Ts polar t1 + | @{const Not} $ t1 => + @{const Not} $ do_term new_Ts old_Ts (flip_polarity polar) t1 + | Const (s0 as @{const_name All}, T0) $ Abs (s1, T1, t1) => + do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 + | Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) => + do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1 + | Const (s0 as @{const_name "op ="}, T0) $ t1 $ t2 => + do_equals new_Ts old_Ts s0 T0 t1 t2 + | @{const "op &"} $ t1 $ t2 => + @{const "op &"} $ do_term new_Ts old_Ts polar t1 + $ do_term new_Ts old_Ts polar t2 + | @{const "op |"} $ t1 $ t2 => + @{const "op |"} $ do_term new_Ts old_Ts polar t1 + $ do_term new_Ts old_Ts polar t2 + | @{const "op -->"} $ t1 $ t2 => + @{const "op -->"} $ do_term new_Ts old_Ts (flip_polarity polar) t1 + $ do_term new_Ts old_Ts polar t2 + | Const (s as @{const_name The}, T) => do_description_operator s T + | Const (s as @{const_name Eps}, T) => do_description_operator s T + | Const (s as @{const_name Tha}, T) => do_description_operator s T + | Const (x as (s, T)) => + Const (s, if s mem [@{const_name converse}, @{const_name trancl}] then + box_relational_operator_type T + else if is_built_in_const fast_descrs x + orelse s = @{const_name Sigma} then + T + else if is_constr_like thy x then + box_type ext_ctxt InConstr T + else if is_sel s orelse is_rep_fun thy x then + box_type ext_ctxt InSel T + else + box_type ext_ctxt InExpr T) + | t1 $ Abs (s, T, t2') => + let + val t1 = do_term new_Ts old_Ts Neut t1 + val T1 = fastype_of1 (new_Ts, t1) + val (s1, Ts1) = dest_Type T1 + val T' = hd (snd (dest_Type (hd Ts1))) + val t2 = Abs (s, T', do_term (T' :: new_Ts) (T :: old_Ts) Neut t2') + val T2 = fastype_of1 (new_Ts, t2) + val t2 = coerce_term new_Ts (hd Ts1) T2 t2 + in + betapply (if s1 = "fun" then + t1 + else + select_nth_constr_arg thy + (@{const_name FunBox}, Type ("fun", Ts1) --> T1) t1 0 + (Type ("fun", Ts1)), t2) + end + | t1 $ t2 => + let + val t1 = do_term new_Ts old_Ts Neut t1 + val T1 = fastype_of1 (new_Ts, t1) + val (s1, Ts1) = dest_Type T1 + val t2 = do_term new_Ts old_Ts Neut t2 + val T2 = fastype_of1 (new_Ts, t2) + val t2 = coerce_term new_Ts (hd Ts1) T2 t2 + in + betapply (if s1 = "fun" then + t1 + else + select_nth_constr_arg thy + (@{const_name FunBox}, Type ("fun", Ts1) --> T1) t1 0 + (Type ("fun", Ts1)), t2) + end + | Free (s, T) => Free (s, box_type ext_ctxt InExpr T) + | Var (z as (x, T)) => + Var (x, if def then box_var_in_def new_Ts old_Ts orig_t z + else box_type ext_ctxt InExpr T) + | Bound _ => t + | Abs (s, T, t') => + Abs (s, T, do_term (T :: new_Ts) (T :: old_Ts) Neut t') + in do_term [] [] Pos orig_t end + +(* int -> term -> term *) +fun eval_axiom_for_term j t = + Logic.mk_equals (Const (eval_prefix ^ string_of_int j, fastype_of t), t) + +(* extended_context -> styp -> bool *) +fun is_equational_fun_surely_complete ext_ctxt x = + case raw_equational_fun_axioms ext_ctxt x of + [@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)] => + strip_comb t1 |> snd |> forall is_Var + | _ => false + +type special = int list * term list * styp + +(* styp -> special -> special -> term *) +fun special_congruence_axiom (s, T) (js1, ts1, x1) (js2, ts2, x2) = + let + val (bounds1, bounds2) = pairself (map Var o special_bounds) (ts1, ts2) + val Ts = binder_types T + val max_j = fold (fold (curry Int.max)) [js1, js2] ~1 + val (eqs, (args1, args2)) = + fold (fn j => case pairself (fn ps => AList.lookup (op =) ps j) + (js1 ~~ ts1, js2 ~~ ts2) of + (SOME t1, SOME t2) => apfst (cons (t1, t2)) + | (SOME t1, NONE) => apsnd (apsnd (cons t1)) + | (NONE, SOME t2) => apsnd (apfst (cons t2)) + | (NONE, NONE) => + let val v = Var ((cong_var_prefix ^ nat_subscript j, 0), + nth Ts j) in + apsnd (pairself (cons v)) + end) (max_j downto 0) ([], ([], [])) + in + Logic.list_implies (eqs |> filter_out (op =) |> distinct (op =) + |> map Logic.mk_equals, + Logic.mk_equals (list_comb (Const x1, bounds1 @ args1), + list_comb (Const x2, bounds2 @ args2))) + |> Refute.close_form + end + +(* extended_context -> styp list -> term list *) +fun special_congruence_axioms (ext_ctxt as {special_funs, ...}) xs = + let + val groups = + !special_funs + |> map (fn ((x, js, ts), x') => (x, (js, ts, x'))) + |> AList.group (op =) + |> filter_out (is_equational_fun_surely_complete ext_ctxt o fst) + |> map (fn (x, zs) => (x, zs |> (x mem xs) ? cons ([], [], x))) + (* special -> int *) + fun generality (js, _, _) = ~(length js) + (* special -> special -> bool *) + fun is_more_specific (j1, t1, x1) (j2, t2, x2) = + x1 <> x2 andalso OrdList.subset (prod_ord int_ord TermOrd.term_ord) + (j2 ~~ t2, j1 ~~ t1) + (* styp -> special list -> special list -> special list -> term list + -> term list *) + fun do_pass_1 _ [] [_] [_] = I + | do_pass_1 x skipped _ [] = do_pass_2 x skipped + | do_pass_1 x skipped all (z :: zs) = + case filter (is_more_specific z) all + |> sort (int_ord o pairself generality) of + [] => do_pass_1 x (z :: skipped) all zs + | (z' :: _) => cons (special_congruence_axiom x z z') + #> do_pass_1 x skipped all zs + (* styp -> special list -> term list -> term list *) + and do_pass_2 _ [] = I + | do_pass_2 x (z :: zs) = + fold (cons o special_congruence_axiom x z) zs #> do_pass_2 x zs + in fold (fn (x, zs) => do_pass_1 x [] zs zs) groups [] end + +(* term -> bool *) +val is_trivial_equation = the_default false o try (op aconv o Logic.dest_equals) + +(* 'a Symtab.table -> 'a list *) +fun all_table_entries table = Symtab.fold (append o snd) table [] +(* const_table -> string -> const_table *) +fun extra_table table s = Symtab.make [(s, all_table_entries table)] + +(* extended_context -> term -> (term list * term list) * (bool * bool) *) +fun axioms_for_term + (ext_ctxt as {thy, max_bisim_depth, user_axioms, fast_descrs, evals, + def_table, nondef_table, user_nondefs, ...}) t = + let + type accumulator = styp list * (term list * term list) + (* (term list * term list -> term list) + -> ((term list -> term list) -> term list * term list + -> term list * term list) + -> int -> term -> accumulator -> accumulator *) + fun add_axiom get app depth t (accum as (xs, axs)) = + let + val t = t |> unfold_defs_in_term ext_ctxt + |> skolemize_term_and_more ext_ctxt ~1 + in + if is_trivial_equation t then + accum + else + let val t' = t |> specialize_consts_in_term ext_ctxt depth in + if exists (member (op aconv) (get axs)) [t, t'] then accum + else add_axioms_for_term (depth + 1) t' (xs, app (cons t') axs) + end + end + (* int -> term -> accumulator -> accumulator *) + and add_nondef_axiom depth = add_axiom snd apsnd depth + and add_def_axiom depth t = + (if head_of t = @{const "==>"} then add_nondef_axiom + else add_axiom fst apfst) depth t + (* int -> term -> accumulator -> accumulator *) + and add_axioms_for_term depth t (accum as (xs, axs)) = + case t of + t1 $ t2 => accum |> fold (add_axioms_for_term depth) [t1, t2] + | Const (x as (s, T)) => + (if x mem xs orelse is_built_in_const fast_descrs x then + accum + else + let val accum as (xs, _) = (x :: xs, axs) in + if depth > axioms_max_depth then + raise LIMIT ("NitpickHOL.axioms_for_term.add_axioms_for_term", + "too many nested axioms (" ^ string_of_int depth ^ + ")") + else if Refute.is_const_of_class thy x then + let + val class = Logic.class_of_const s + val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), + class) + val ax1 = try (Refute.specialize_type thy x) of_class + val ax2 = Option.map (Refute.specialize_type thy x o snd) + (Refute.get_classdef thy class) + in fold (add_def_axiom depth) (map_filter I [ax1, ax2]) accum end + else if is_constr thy x then + accum + else if is_equational_fun ext_ctxt x then + fold (add_def_axiom depth) (equational_fun_axioms ext_ctxt x) + accum + else if is_abs_fun thy x then + accum |> fold (add_nondef_axiom depth) + (nondef_props_for_const thy false nondef_table x) + |> fold (add_def_axiom depth) + (nondef_props_for_const thy true + (extra_table def_table s) x) + else if is_rep_fun thy x then + accum |> fold (add_nondef_axiom depth) + (nondef_props_for_const thy false nondef_table x) + |> fold (add_def_axiom depth) + (nondef_props_for_const thy true + (extra_table def_table s) x) + |> add_axioms_for_term depth + (Const (mate_of_rep_fun thy x)) + |> add_def_axiom depth (inverse_axiom_for_rep_fun thy x) + else + accum |> user_axioms <> SOME false + ? fold (add_nondef_axiom depth) + (nondef_props_for_const thy false nondef_table x) + end) + |> add_axioms_for_type depth T + | Free (_, T) => add_axioms_for_type depth T accum + | Var (_, T) => add_axioms_for_type depth T accum + | Bound _ => accum + | Abs (_, T, t) => accum |> add_axioms_for_term depth t + |> add_axioms_for_type depth T + (* int -> typ -> accumulator -> accumulator *) + and add_axioms_for_type depth T = + case T of + Type ("fun", Ts) => fold (add_axioms_for_type depth) Ts + | Type ("*", Ts) => fold (add_axioms_for_type depth) Ts + | @{typ prop} => I + | @{typ bool} => I + | @{typ unit} => I + | TFree (_, S) => add_axioms_for_sort depth T S + | TVar (_, S) => add_axioms_for_sort depth T S + | Type (z as (_, Ts)) => + fold (add_axioms_for_type depth) Ts + #> (if is_pure_typedef thy T then + fold (add_def_axiom depth) (optimized_typedef_axioms thy z) + else if max_bisim_depth >= 0 andalso is_codatatype thy T then + fold (add_def_axiom depth) (codatatype_bisim_axioms thy T) + else + I) + (* int -> typ -> sort -> accumulator -> accumulator *) + and add_axioms_for_sort depth T S = + let + val supers = Sign.complete_sort thy S + val class_axioms = + maps (fn class => map prop_of (AxClass.get_info thy class |> #axioms + handle ERROR _ => [])) supers + val monomorphic_class_axioms = + map (fn t => case Term.add_tvars t [] of + [] => t + | [(x, S)] => + Refute.monomorphic_term (Vartab.make [(x, (S, T))]) t + | _ => raise TERM ("NitpickHOL.axioms_for_term.\ + \add_axioms_for_sort", [t])) + class_axioms + in fold (add_nondef_axiom depth) monomorphic_class_axioms end + val (mono_user_nondefs, poly_user_nondefs) = + List.partition (null o Term.hidden_polymorphism) user_nondefs + val eval_axioms = map2 eval_axiom_for_term (index_seq 0 (length evals)) + evals + val (xs, (defs, nondefs)) = + ([], ([], [])) |> add_axioms_for_term 1 t + |> fold_rev (add_def_axiom 1) eval_axioms + |> user_axioms = SOME true + ? fold (add_nondef_axiom 1) mono_user_nondefs + val defs = defs @ special_congruence_axioms ext_ctxt xs + in + ((defs, nondefs), (user_axioms = SOME true orelse null mono_user_nondefs, + null poly_user_nondefs)) + end + +(* theory -> const_table -> styp -> int list *) +fun const_format thy def_table (x as (s, T)) = + if String.isPrefix unrolled_prefix s then + const_format thy def_table (original_name s, range_type T) + else if String.isPrefix skolem_prefix s then + let + val k = unprefix skolem_prefix s + |> strip_first_name_sep |> fst |> space_explode "@" + |> hd |> Int.fromString |> the + in [k, num_binder_types T - k] end + else if original_name s <> s then + [num_binder_types T] + else case def_of_const thy def_table x of + SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then + let val k = length (strip_abs_vars t') in + [k, num_binder_types T - k] + end + else + [num_binder_types T] + | NONE => [num_binder_types T] +(* int list -> int list -> int list *) +fun intersect_formats _ [] = [] + | intersect_formats [] _ = [] + | intersect_formats ks1 ks2 = + let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in + intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else [])) + (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @ + [Int.min (k1, k2)] + end + +(* theory -> const_table -> (term option * int list) list -> term -> int list *) +fun lookup_format thy def_table formats t = + case AList.lookup (fn (SOME x, SOME y) => + (term_match thy) (x, y) | _ => false) + formats (SOME t) of + SOME format => format + | NONE => let val format = the (AList.lookup (op =) formats NONE) in + case t of + Const x => intersect_formats format + (const_format thy def_table x) + | _ => format + end + +(* int list -> int list -> typ -> typ *) +fun format_type default_format format T = + let + val T = unbox_type T + val format = format |> filter (curry (op <) 0) + in + if forall (equal 1) format then + T + else + let + val (binder_Ts, body_T) = strip_type T + val batched = + binder_Ts + |> map (format_type default_format default_format) + |> rev |> chunk_list_unevenly (rev format) + |> map (HOLogic.mk_tupleT o rev) + in List.foldl (op -->) body_T batched end + end +(* theory -> const_table -> (term option * int list) list -> term -> typ *) +fun format_term_type thy def_table formats t = + format_type (the (AList.lookup (op =) formats NONE)) + (lookup_format thy def_table formats t) (fastype_of t) + +(* int list -> int -> int list -> int list *) +fun repair_special_format js m format = + m - 1 downto 0 |> chunk_list_unevenly (rev format) + |> map (rev o filter_out (member (op =) js)) + |> filter_out null |> map length |> rev + +(* extended_context -> string * string -> (term option * int list) list + -> styp -> term * typ *) +fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...} + : extended_context) (base_name, step_name) formats = + let + val default_format = the (AList.lookup (op =) formats NONE) + (* styp -> term * typ *) + fun do_const (x as (s, T)) = + (if String.isPrefix special_prefix s then + let + (* term -> term *) + val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t') + val (x' as (_, T'), js, ts) = + AList.find (op =) (!special_funs) (s, unbox_type T) |> the_single + val max_j = List.last js + val Ts = List.take (binder_types T', max_j + 1) + val missing_js = filter_out (member (op =) js) (0 upto max_j) + val missing_Ts = filter_indices missing_js Ts + (* int -> indexname *) + fun nth_missing_var n = + ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n) + val missing_vars = map nth_missing_var (0 upto length missing_js - 1) + val vars = special_bounds ts @ missing_vars + val ts' = map2 (fn T => fn j => + case AList.lookup (op =) (js ~~ ts) j of + SOME t => do_term t + | NONE => + Var (nth missing_vars + (find_index (equal j) missing_js))) + Ts (0 upto max_j) + val t = do_const x' |> fst + val format = + case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2) + | _ => false) formats (SOME t) of + SOME format => + repair_special_format js (num_binder_types T') format + | NONE => + const_format thy def_table x' + |> repair_special_format js (num_binder_types T') + |> intersect_formats default_format + in + (list_comb (t, ts') |> fold_rev abs_var vars, + format_type default_format format T) + end + else if String.isPrefix uncurry_prefix s then + let + val (ss, s') = unprefix uncurry_prefix s + |> strip_first_name_sep |>> space_explode "@" + in + if String.isPrefix step_prefix s' then + do_const (s', T) + else + let + val k = the (Int.fromString (hd ss)) + val j = the (Int.fromString (List.last ss)) + val (before_Ts, (tuple_T, rest_T)) = + strip_n_binders j T ||> (strip_n_binders 1 #>> hd) + val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T + in do_const (s', T') end + end + else if String.isPrefix unrolled_prefix s then + let val t = Const (original_name s, range_type T) in + (lambda (Free (iter_var_prefix, nat_T)) t, + format_type default_format + (lookup_format thy def_table formats t) T) + end + else if String.isPrefix base_prefix s then + (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T), + format_type default_format default_format T) + else if String.isPrefix step_prefix s then + (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T), + format_type default_format default_format T) + else if String.isPrefix skolem_prefix s then + let + val ss = the (AList.lookup (op =) (!skolems) s) + val (Ts, Ts') = chop (length ss) (binder_types T) + val frees = map Free (ss ~~ Ts) + val s' = original_name s + in + (fold lambda frees (Const (s', Ts' ---> T)), + format_type default_format + (lookup_format thy def_table formats (Const x)) T) + end + else if String.isPrefix eval_prefix s then + let + val t = nth evals (the (Int.fromString (unprefix eval_prefix s))) + in (t, format_term_type thy def_table formats t) end + else if s = @{const_name undefined_fast_The} then + (Const (nitpick_prefix ^ "The fallback", T), + format_type default_format + (lookup_format thy def_table formats + (Const (@{const_name The}, (T --> bool_T) --> T))) T) + else if s = @{const_name undefined_fast_Eps} then + (Const (nitpick_prefix ^ "Eps fallback", T), + format_type default_format + (lookup_format thy def_table formats + (Const (@{const_name Eps}, (T --> bool_T) --> T))) T) + else + let val t = Const (original_name s, T) in + (t, format_term_type thy def_table formats t) + end) + |>> map_types (typ_subst [(@{typ bisim_iterator}, nat_T)] o unbox_type) + |>> shorten_const_names_in_term |>> shorten_abs_vars + in do_const end + +(* styp -> string *) +fun assign_operator_for_const (s, T) = + if String.isPrefix ubfp_prefix s then + if is_fun_type T then "\" else "\" + else if String.isPrefix lbfp_prefix s then + if is_fun_type T then "\" else "\" + else if original_name s <> s then + assign_operator_for_const (after_name_sep s, T) + else + "=" + +(* extended_context -> term + -> ((term list * term list) * (bool * bool)) * term *) +fun preprocess_term (ext_ctxt as {thy, destroy_constrs, boxes, skolemize, + uncurry, ...}) t = + let + val skolem_depth = if skolemize then 4 else ~1 + val (((def_ts, nondef_ts), (got_all_mono_user_axioms, no_poly_user_axioms)), + core_t) = t |> unfold_defs_in_term ext_ctxt + |> Refute.close_form + |> skolemize_term_and_more ext_ctxt skolem_depth + |> specialize_consts_in_term ext_ctxt 0 + |> `(axioms_for_term ext_ctxt) + val maybe_box = exists (not_equal (SOME false) o snd) boxes + val table = + Termtab.empty |> uncurry + ? fold (add_to_uncurry_table thy) (core_t :: def_ts @ nondef_ts) + (* bool -> bool -> term -> term *) + fun do_rest def core = + uncurry ? uncurry_term table + #> maybe_box ? box_fun_and_pair_in_term ext_ctxt def + #> destroy_constrs ? (pull_out_universal_constrs thy def + #> pull_out_existential_constrs thy + #> destroy_pulled_out_constrs thy def) + #> curry_assms + #> destroy_universal_equalities + #> destroy_existential_equalities thy + #> simplify_constrs_and_sels thy + #> distribute_quantifiers + #> push_quantifiers_inward thy + #> not core ? Refute.close_form + #> shorten_abs_vars + in + (((map (do_rest true false) def_ts, map (do_rest false false) nondef_ts), + (got_all_mono_user_axioms, no_poly_user_axioms)), + do_rest false true core_t) + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_isar.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,496 @@ +(* Title: HOL/Nitpick/Tools/nitpick_isar.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Adds the "nitpick" and "nitpick_params" commands to Isabelle/Isar's outer +syntax. +*) + +signature NITPICK_ISAR = +sig + type params = Nitpick.params + + val default_params : theory -> (string * string) list -> params +end + +structure NitpickIsar : NITPICK_ISAR = +struct + +open NitpickUtil +open NitpickHOL +open NitpickRep +open NitpickNut +open Nitpick + +type raw_param = string * string list + +val default_default_params = + [("card", ["1\8"]), + ("iter", ["0,1,2,4,8,12,16,24"]), + ("bisim_depth", ["7"]), + ("box", ["smart"]), + ("mono", ["smart"]), + ("wf", ["smart"]), + ("sat_solver", ["smart"]), + ("batch_size", ["smart"]), + ("auto", ["false"]), + ("blocking", ["true"]), + ("falsify", ["true"]), + ("user_axioms", ["smart"]), + ("assms", ["true"]), + ("coalesce_type_vars", ["false"]), + ("destroy_constrs", ["true"]), + ("specialize", ["true"]), + ("skolemize", ["true"]), + ("star_linear_preds", ["true"]), + ("uncurry", ["true"]), + ("fast_descrs", ["true"]), + ("peephole_optim", ["true"]), + ("timeout", ["30 s"]), + ("auto_timeout", ["5 s"]), + ("tac_timeout", ["500 ms"]), + ("sym_break", ["20"]), + ("sharing_depth", ["3"]), + ("flatten_props", ["false"]), + ("max_threads", ["0"]), + ("verbose", ["false"]), + ("debug", ["false"]), + ("overlord", ["false"]), + ("show_all", ["false"]), + ("show_skolems", ["true"]), + ("show_datatypes", ["false"]), + ("show_consts", ["false"]), + ("format", ["1"]), + ("max_potential", ["1"]), + ("max_genuine", ["1"]), + ("check_potential", ["false"]), + ("check_genuine", ["false"])] + +val negated_params = + [("dont_box", "box"), + ("non_mono", "mono"), + ("non_wf", "wf"), + ("no_auto", "auto"), + ("non_blocking", "blocking"), + ("satisfy", "falsify"), + ("no_user_axioms", "user_axioms"), + ("no_assms", "assms"), + ("dont_coalesce_type_vars", "coalesce_type_vars"), + ("dont_destroy_constrs", "destroy_constrs"), + ("dont_specialize", "specialize"), + ("dont_skolemize", "skolemize"), + ("dont_star_linear_preds", "star_linear_preds"), + ("dont_uncurry", "uncurry"), + ("full_descrs", "fast_descrs"), + ("no_peephole_optim", "peephole_optim"), + ("dont_flatten_props", "flatten_props"), + ("quiet", "verbose"), + ("no_debug", "debug"), + ("no_overlord", "overlord"), + ("dont_show_all", "show_all"), + ("hide_skolems", "show_skolems"), + ("hide_datatypes", "show_datatypes"), + ("hide_consts", "show_consts"), + ("trust_potential", "check_potential"), + ("trust_genuine", "check_genuine")] + +(* string -> bool *) +fun is_known_raw_param s = + AList.defined (op =) default_default_params s + orelse AList.defined (op =) negated_params s + orelse s mem ["max", "eval", "expect"] + orelse exists (fn p => String.isPrefix (p ^ " ") s) + ["card", "max", "iter", "box", "dont_box", "mono", "non_mono", + "wf", "non_wf", "format"] + +(* string * 'a -> unit *) +fun check_raw_param (s, _) = + if is_known_raw_param s then () + else error ("Unknown parameter " ^ quote s ^ ".") + +(* string -> string option *) +fun unnegate_param_name name = + case AList.lookup (op =) negated_params name of + NONE => if String.isPrefix "dont_" name then SOME (unprefix "dont_" name) + else if String.isPrefix "non_" name then SOME (unprefix "non_" name) + else NONE + | some_name => some_name +(* raw_param -> raw_param *) +fun unnegate_raw_param (name, value) = + case unnegate_param_name name of + SOME name' => (name', case value of + ["false"] => ["true"] + | ["true"] => ["false"] + | [] => ["false"] + | _ => value) + | NONE => (name, value) + +structure TheoryData = TheoryDataFun( + type T = {params: raw_param list, registered_auto: bool} + val empty = {params = rev default_default_params, registered_auto = false} + val copy = I + val extend = I + fun merge _ ({params = ps1, registered_auto = a1}, + {params = ps2, registered_auto = a2}) = + {params = AList.merge (op =) (op =) (ps1, ps2), + registered_auto = a1 orelse a2}) + +(* raw_param -> theory -> theory *) +fun set_default_raw_param param thy = + let val {params, registered_auto} = TheoryData.get thy in + TheoryData.put + {params = AList.update (op =) (unnegate_raw_param param) params, + registered_auto = registered_auto} thy + end +(* theory -> raw_param list *) +val default_raw_params = #params o TheoryData.get + +(* theory -> theory *) +fun set_registered_auto thy = + TheoryData.put {params = default_raw_params thy, registered_auto = true} thy +(* theory -> bool *) +val is_registered_auto = #registered_auto o TheoryData.get + +(* string -> bool *) +fun is_punctuation s = (s = "," orelse s = "-" orelse s = "\") + +(* string list -> string *) +fun stringify_raw_param_value [] = "" + | stringify_raw_param_value [s] = s + | stringify_raw_param_value (s1 :: s2 :: ss) = + s1 ^ (if is_punctuation s1 orelse is_punctuation s2 then "" else " ") ^ + stringify_raw_param_value (s2 :: ss) + +(* bool -> string -> string -> bool option *) +fun bool_option_from_string option name s = + (case s of + "smart" => if option then NONE else raise Option + | "false" => SOME false + | "true" => SOME true + | "" => SOME true + | s => raise Option) + handle Option.Option => + let val ss = map quote ((option ? cons "smart") ["true", "false"]) in + error ("Parameter " ^ quote name ^ " must be assigned " ^ + space_implode " " (serial_commas "or" ss) ^ ".") + end +(* bool -> raw_param list -> bool option -> string -> bool option *) +fun general_lookup_bool option raw_params default_value name = + case AList.lookup (op =) raw_params name of + SOME s => s |> stringify_raw_param_value + |> bool_option_from_string option name + | NONE => default_value + +(* int -> string -> int *) +fun maxed_int_from_string min_int s = Int.max (min_int, the (Int.fromString s)) + +(* Proof.context -> bool -> raw_param list -> raw_param list -> params *) +fun extract_params ctxt auto default_params override_params = + let + val override_params = map unnegate_raw_param override_params + val raw_params = rev override_params @ rev default_params + val lookup = + Option.map stringify_raw_param_value o AList.lookup (op =) raw_params + (* string -> string *) + fun lookup_string name = the_default "" (lookup name) + (* string -> bool *) + val lookup_bool = the o general_lookup_bool false raw_params (SOME false) + (* string -> bool option *) + val lookup_bool_option = general_lookup_bool true raw_params NONE + (* string -> string option -> int *) + fun do_int name value = + case value of + SOME s => (case Int.fromString s of + SOME i => i + | NONE => error ("Parameter " ^ quote name ^ + " must be assigned an integer value.")) + | NONE => 0 + (* string -> int *) + fun lookup_int name = do_int name (lookup name) + (* string -> int option *) + fun lookup_int_option name = + case lookup name of + SOME "smart" => NONE + | value => SOME (do_int name value) + (* string -> int -> string -> int list *) + fun int_range_from_string name min_int s = + let + val (k1, k2) = + (case space_explode "-" s of + [s] => the_default (s, s) (first_field "\" s) + | ["", s2] => ("-" ^ s2, "-" ^ s2) + | [s1, s2] => (s1, s2) + | _ => raise Option) + |> pairself (maxed_int_from_string min_int) + in if k1 <= k2 then k1 upto k2 else k1 downto k2 end + handle Option.Option => + error ("Parameter " ^ quote name ^ + " must be assigned a sequence of integers.") + (* string -> int -> string -> int list *) + fun int_seq_from_string name min_int s = + maps (int_range_from_string name min_int) (space_explode "," s) + (* string -> int -> int list *) + fun lookup_int_seq name min_int = + case lookup name of + SOME s => (case int_seq_from_string name min_int s of + [] => [min_int] + | value => value) + | NONE => [min_int] + (* (string -> 'a) -> int -> string -> ('a option * int list) list *) + fun lookup_ints_assigns read prefix min_int = + (NONE, lookup_int_seq prefix min_int) + :: map (fn (name, value) => + (SOME (read (String.extract (name, size prefix + 1, NONE))), + value |> stringify_raw_param_value + |> int_seq_from_string name min_int)) + (filter (String.isPrefix (prefix ^ " ") o fst) raw_params) + (* (string -> 'a) -> string -> ('a option * bool option) list *) + fun lookup_bool_option_assigns read prefix = + (NONE, lookup_bool_option prefix) + :: map (fn (name, value) => + (SOME (read (String.extract (name, size prefix + 1, NONE))), + value |> stringify_raw_param_value + |> bool_option_from_string true name)) + (filter (String.isPrefix (prefix ^ " ") o fst) raw_params) + (* string -> Time.time option *) + fun lookup_time name = + case lookup name of + NONE => NONE + | SOME "none" => NONE + | SOME s => + let + val msecs = + case space_explode " " s of + [s1, "min"] => 60000 * the (Int.fromString s1) + | [s1, "s"] => 1000 * the (Int.fromString s1) + | [s1, "ms"] => the (Int.fromString s1) + | _ => 0 + in + if msecs <= 0 then + error ("Parameter " ^ quote name ^ " must be assigned a positive \ + \time value (e.g., \"60 s\", \"200 ms\") or \"none\".") + else + SOME (Time.fromMilliseconds msecs) + end + (* string -> term list *) + val lookup_term_list = + AList.lookup (op =) raw_params #> these #> Syntax.read_terms ctxt + val read_type_polymorphic = + Syntax.read_typ ctxt #> Logic.mk_type + #> singleton (Variable.polymorphic ctxt) #> Logic.dest_type + (* string -> term *) + val read_term_polymorphic = + Syntax.read_term ctxt #> singleton (Variable.polymorphic ctxt) + (* string -> styp *) + val read_const_polymorphic = read_term_polymorphic #> dest_Const + val cards_assigns = lookup_ints_assigns read_type_polymorphic "card" 1 + val maxes_assigns = lookup_ints_assigns read_const_polymorphic "max" ~1 + val iters_assigns = lookup_ints_assigns read_const_polymorphic "iter" 0 + val bisim_depths = lookup_int_seq "bisim_depth" ~1 + val boxes = + lookup_bool_option_assigns read_type_polymorphic "box" @ + map_filter (fn (SOME T, _) => + if is_fun_type T orelse is_pair_type T then + SOME (SOME T, SOME true) + else + NONE + | (NONE, _) => NONE) cards_assigns + val monos = lookup_bool_option_assigns read_type_polymorphic "mono" + val wfs = lookup_bool_option_assigns read_const_polymorphic "wf" + val sat_solver = lookup_string "sat_solver" + val blocking = not auto andalso lookup_bool "blocking" + val falsify = lookup_bool "falsify" + val debug = not auto andalso lookup_bool "debug" + val verbose = debug orelse (not auto andalso lookup_bool "verbose") + val overlord = lookup_bool "overlord" + val user_axioms = lookup_bool_option "user_axioms" + val assms = lookup_bool "assms" + val coalesce_type_vars = lookup_bool "coalesce_type_vars" + val destroy_constrs = lookup_bool "destroy_constrs" + val specialize = lookup_bool "specialize" + val skolemize = lookup_bool "skolemize" + val star_linear_preds = lookup_bool "star_linear_preds" + val uncurry = lookup_bool "uncurry" + val fast_descrs = lookup_bool "fast_descrs" + val peephole_optim = lookup_bool "peephole_optim" + val timeout = if auto then lookup_time "auto_timeout" + else lookup_time "timeout" + val tac_timeout = lookup_time "tac_timeout" + val sym_break = Int.max (0, lookup_int "sym_break") + val sharing_depth = Int.max (1, lookup_int "sharing_depth") + val flatten_props = lookup_bool "flatten_props" + val max_threads = Int.max (0, lookup_int "max_threads") + val show_all = debug orelse lookup_bool "show_all" + val show_skolems = show_all orelse lookup_bool "show_skolems" + val show_datatypes = show_all orelse lookup_bool "show_datatypes" + val show_consts = show_all orelse lookup_bool "show_consts" + val formats = lookup_ints_assigns read_term_polymorphic "format" 0 + val evals = lookup_term_list "eval" + val max_potential = if auto then 0 + else Int.max (0, lookup_int "max_potential") + val max_genuine = Int.max (0, lookup_int "max_genuine") + val check_potential = lookup_bool "check_potential" + val check_genuine = lookup_bool "check_genuine" + val batch_size = case lookup_int_option "batch_size" of + SOME n => Int.max (1, n) + | NONE => if debug then 1 else 64 + val expect = lookup_string "expect" + in + {cards_assigns = cards_assigns, maxes_assigns = maxes_assigns, + iters_assigns = iters_assigns, bisim_depths = bisim_depths, boxes = boxes, + monos = monos, wfs = wfs, sat_solver = sat_solver, blocking = blocking, + falsify = falsify, debug = debug, verbose = verbose, overlord = overlord, + user_axioms = user_axioms, assms = assms, + coalesce_type_vars = coalesce_type_vars, destroy_constrs = destroy_constrs, + specialize = specialize, skolemize = skolemize, + star_linear_preds = star_linear_preds, uncurry = uncurry, + fast_descrs = fast_descrs, peephole_optim = peephole_optim, + timeout = timeout, tac_timeout = tac_timeout, sym_break = sym_break, + sharing_depth = sharing_depth, flatten_props = flatten_props, + max_threads = max_threads, show_skolems = show_skolems, + show_datatypes = show_datatypes, show_consts = show_consts, + formats = formats, evals = evals, max_potential = max_potential, + max_genuine = max_genuine, check_potential = check_potential, + check_genuine = check_genuine, batch_size = batch_size, expect = expect} + end + +(* theory -> (string * string) list -> params *) +fun default_params thy = + extract_params (ProofContext.init thy) false (default_raw_params thy) + o map (apsnd single) + +(* OuterParse.token list -> string * OuterParse.token list *) +val scan_key = Scan.repeat1 OuterParse.typ_group >> space_implode " " + +(* OuterParse.token list -> string list * OuterParse.token list *) +val scan_value = + Scan.repeat1 (OuterParse.minus >> single + || Scan.repeat1 (Scan.unless OuterParse.minus OuterParse.name) + || OuterParse.$$$ "," |-- OuterParse.number >> prefix "," + >> single) >> flat + +(* OuterParse.token list -> raw_param * OuterParse.token list *) +val scan_param = + scan_key -- (Scan.option (OuterParse.$$$ "=" |-- scan_value) >> these) +(* OuterParse.token list -> raw_param list option * OuterParse.token list *) +val scan_params = Scan.option (OuterParse.$$$ "[" |-- OuterParse.list scan_param + --| OuterParse.$$$ "]") + +(* Proof.context -> ('a -> 'a) -> 'a -> 'a *) +fun handle_exceptions ctxt f x = + f x + handle ARG (loc, details) => + error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ ".") + | BAD (loc, details) => + error ("Internal error (" ^ quote loc ^ "): " ^ details ^ ".") + | LIMIT (_, details) => + (warning ("Limit reached: " ^ details ^ "."); x) + | NOT_SUPPORTED details => + (warning ("Unsupported case: " ^ details ^ "."); x) + | NUT (loc, us) => + error ("Invalid intermediate term" ^ plural_s_for_list us ^ + " (" ^ quote loc ^ "): " ^ + commas (map (string_for_nut ctxt) us) ^ ".") + | REP (loc, Rs) => + error ("Invalid representation" ^ plural_s_for_list Rs ^ + " (" ^ quote loc ^ "): " ^ commas (map string_for_rep Rs) ^ ".") + | TERM (loc, ts) => + error ("Invalid term" ^ plural_s_for_list ts ^ + " (" ^ quote loc ^ "): " ^ + commas (map (Syntax.string_of_term ctxt) ts) ^ ".") + | TYPE (loc, Ts, ts) => + error ("Invalid type" ^ plural_s_for_list Ts ^ + (if null ts then + "" + else + " for term" ^ plural_s_for_list ts ^ " " ^ + commas (map (quote o Syntax.string_of_term ctxt) ts)) ^ + " (" ^ quote loc ^ "): " ^ + commas (map (Syntax.string_of_typ ctxt) Ts) ^ ".") + | Kodkod.SYNTAX (_, details) => + (warning ("Ill-formed Kodkodi output: " ^ details ^ "."); x) + | Refute.REFUTE (loc, details) => + error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ ".") + +(* raw_param list -> bool -> int -> Proof.state -> Proof.state *) +fun pick_nits override_params auto subgoal state = + let + val thy = Proof.theory_of state + val ctxt = Proof.context_of state + val thm = snd (snd (Proof.get_goal state)) + val _ = List.app check_raw_param override_params + val params as {blocking, debug, ...} = + extract_params ctxt auto (default_raw_params thy) override_params + (* unit -> Proof.state *) + fun go () = + (if auto then perhaps o try + else if debug then fn f => fn x => f x + else handle_exceptions ctxt) + (fn state => pick_nits_in_subgoal state params auto subgoal |> snd) + state + in + if auto orelse blocking then + go () + else + (SimpleThread.fork true (fn () => (go (); ()) handle Exn.Interrupt => ()); + state) + end + +(* (TableFun().key * string list) list option * int option + -> Toplevel.transition -> Toplevel.transition *) +fun nitpick_trans (opt_params, opt_subgoal) = + Toplevel.keep (K () + o pick_nits (these opt_params) false (the_default 1 opt_subgoal) + o Toplevel.proof_of) + +(* raw_param -> string *) +fun string_for_raw_param (name, value) = + name ^ " = " ^ stringify_raw_param_value value + +(* bool -> Proof.state -> Proof.state *) +fun pick_nits_auto interactive state = + let val thy = Proof.theory_of state in + ((interactive andalso not (!Toplevel.quiet) + andalso the (general_lookup_bool false (default_raw_params thy) + (SOME false) "auto")) + ? pick_nits [] true 0) state + end + +(* theory -> theory *) +fun register_auto thy = + (not (is_registered_auto thy) + ? (set_registered_auto + #> Context.theory_map (Specification.add_theorem_hook pick_nits_auto))) + thy + +(* (TableFun().key * string) list option -> Toplevel.transition + -> Toplevel.transition *) +fun nitpick_params_trans opt_params = + Toplevel.theory + (fn thy => + let val thy = fold set_default_raw_param (these opt_params) thy in + writeln ("Default parameters for Nitpick:\n" ^ + (case rev (default_raw_params thy) of + [] => "none" + | params => + (map check_raw_param params; + params |> map string_for_raw_param |> sort_strings + |> cat_lines))); + register_auto thy + end) + +(* OuterParse.token list + -> (Toplevel.transition -> Toplevel.transition) * OuterParse.token list *) +fun scan_nitpick_command tokens = + (scan_params -- Scan.option OuterParse.nat) tokens |>> nitpick_trans +fun scan_nitpick_params_command tokens = + scan_params tokens |>> nitpick_params_trans + +val _ = OuterSyntax.improper_command "nitpick" + "try to find a counterexample for a given subgoal using Kodkod" + OuterKeyword.diag scan_nitpick_command +val _ = OuterSyntax.command "nitpick_params" + "set and display the default parameters for Nitpick" + OuterKeyword.thy_decl scan_nitpick_params_command + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_kodkod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,1737 @@ +(* Title: HOL/Nitpick/Tools/nitpick_kodkod.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Kodkod problem generator part of Kodkod. +*) + +signature NITPICK_KODKOD = +sig + type extended_context = NitpickHOL.extended_context + type dtype_spec = NitpickScope.dtype_spec + type kodkod_constrs = NitpickPeephole.kodkod_constrs + type nut = NitpickNut.nut + type nfa_transition = Kodkod.rel_expr * typ + type nfa_entry = typ * nfa_transition list + type nfa_table = nfa_entry list + + structure NameTable : TABLE + + val univ_card : + int -> int -> int -> Kodkod.bound list -> Kodkod.formula -> int + val check_arity : int -> int -> unit + val kk_tuple : bool -> int -> int list -> Kodkod.tuple + val tuple_set_from_atom_schema : (int * int) list -> Kodkod.tuple_set + val sequential_int_bounds : int -> Kodkod.int_bound list + val bounds_for_built_in_rels_in_formula : + bool -> int -> int -> int -> int -> Kodkod.formula -> Kodkod.bound list + val bound_for_plain_rel : Proof.context -> bool -> nut -> Kodkod.bound + val bound_for_sel_rel : + Proof.context -> bool -> dtype_spec list -> nut -> Kodkod.bound + val merge_bounds : Kodkod.bound list -> Kodkod.bound list + val declarative_axiom_for_plain_rel : kodkod_constrs -> nut -> Kodkod.formula + val declarative_axioms_for_datatypes : + extended_context -> int Typtab.table -> kodkod_constrs + -> nut NameTable.table -> dtype_spec list -> Kodkod.formula list + val kodkod_formula_from_nut : + int Typtab.table -> bool -> kodkod_constrs -> nut -> Kodkod.formula +end; + +structure NitpickKodkod : NITPICK_KODKOD = +struct + +open NitpickUtil +open NitpickHOL +open NitpickScope +open NitpickPeephole +open NitpickRep +open NitpickNut + +type nfa_transition = Kodkod.rel_expr * typ +type nfa_entry = typ * nfa_transition list +type nfa_table = nfa_entry list + +structure NfaGraph = Graph(type key = typ val ord = TermOrd.typ_ord) + +(* int -> Kodkod.int_expr list *) +fun flip_nums n = index_seq 1 n @ [0] |> map Kodkod.Num + +(* int -> int -> int -> Kodkod.bound list -> Kodkod.formula -> int *) +fun univ_card nat_card int_card main_j0 bounds formula = + let + (* Kodkod.rel_expr -> int -> int *) + fun rel_expr_func r k = + Int.max (k, case r of + Kodkod.Atom j => j + 1 + | Kodkod.AtomSeq (k', j0) => j0 + k' + | _ => 0) + (* Kodkod.tuple -> int -> int *) + fun tuple_func t k = + case t of + Kodkod.Tuple js => fold Integer.max (map (Integer.add 1) js) k + | _ => k + (* Kodkod.tuple_set -> int -> int *) + fun tuple_set_func ts k = + Int.max (k, case ts of Kodkod.TupleAtomSeq (k', j0) => j0 + k' | _ => 0) + val expr_F = {formula_func = K I, rel_expr_func = rel_expr_func, + int_expr_func = K I} + val tuple_F = {tuple_func = tuple_func, tuple_set_func = tuple_set_func} + val card = fold (Kodkod.fold_bound expr_F tuple_F) bounds 1 + |> Kodkod.fold_formula expr_F formula + in Int.max (main_j0 + fold Integer.max [2, nat_card, int_card] 0, card) end + +(* Proof.context -> bool -> string -> typ -> rep -> string *) +fun bound_comment ctxt debug nick T R = + short_const_name nick ^ + (if debug then " :: " ^ plain_string_from_yxml (Syntax.string_of_typ ctxt T) + else "") ^ " : " ^ string_for_rep R + +(* int -> int -> unit *) +fun check_arity univ_card n = + if n > Kodkod.max_arity univ_card then + raise LIMIT ("NitpickKodkod.check_arity", + "arity " ^ string_of_int n ^ " too large for universe of \ + \cardinality " ^ string_of_int univ_card) + else + () + +(* bool -> int -> int list -> Kodkod.tuple *) +fun kk_tuple debug univ_card js = + if debug then + Kodkod.Tuple js + else + Kodkod.TupleIndex (length js, + fold (fn j => fn accum => accum * univ_card + j) js 0) + +(* (int * int) list -> Kodkod.tuple_set *) +val tuple_set_from_atom_schema = + foldl1 Kodkod.TupleProduct o map Kodkod.TupleAtomSeq +(* rep -> Kodkod.tuple_set *) +val upper_bound_for_rep = tuple_set_from_atom_schema o atom_schema_of_rep + +(* int -> Kodkod.int_bound list *) +fun sequential_int_bounds n = + [(NONE, map (Kodkod.TupleSet o single o Kodkod.Tuple o single) + (index_seq 0 n))] + +(* Kodkod.formula -> Kodkod.n_ary_index list *) +fun built_in_rels_in_formula formula = + let + (* Kodkod.rel_expr -> Kodkod.n_ary_index list -> Kodkod.n_ary_index list *) + fun rel_expr_func (Kodkod.Rel (n, j)) rels = + (case AList.lookup (op =) (#rels initial_pool) n of + SOME k => (j < k ? insert (op =) (n, j)) rels + | NONE => rels) + | rel_expr_func _ rels = rels + val expr_F = {formula_func = K I, rel_expr_func = rel_expr_func, + int_expr_func = K I} + in Kodkod.fold_formula expr_F formula [] end + +val max_table_size = 65536 + +(* int -> unit *) +fun check_table_size k = + if k > max_table_size then + raise LIMIT ("NitpickKodkod.check_table_size", + "precomputed table too large (" ^ string_of_int k ^ ")") + else + () + +(* bool -> int -> int * int -> (int -> int) -> Kodkod.tuple list *) +fun tabulate_func1 debug univ_card (k, j0) f = + (check_table_size k; + map_filter (fn j1 => let val j2 = f j1 in + if j2 >= 0 then + SOME (kk_tuple debug univ_card [j1 + j0, j2 + j0]) + else + NONE + end) (index_seq 0 k)) +(* bool -> int -> int * int -> int -> (int * int -> int) -> Kodkod.tuple list *) +fun tabulate_op2 debug univ_card (k, j0) res_j0 f = + (check_table_size (k * k); + map_filter (fn j => let + val j1 = j div k + val j2 = j - j1 * k + val j3 = f (j1, j2) + in + if j3 >= 0 then + SOME (kk_tuple debug univ_card + [j1 + j0, j2 + j0, j3 + res_j0]) + else + NONE + end) (index_seq 0 (k * k))) +(* bool -> int -> int * int -> int -> (int * int -> int * int) + -> Kodkod.tuple list *) +fun tabulate_op2_2 debug univ_card (k, j0) res_j0 f = + (check_table_size (k * k); + map_filter (fn j => let + val j1 = j div k + val j2 = j - j1 * k + val (j3, j4) = f (j1, j2) + in + if j3 >= 0 andalso j4 >= 0 then + SOME (kk_tuple debug univ_card + [j1 + j0, j2 + j0, j3 + res_j0, + j4 + res_j0]) + else + NONE + end) (index_seq 0 (k * k))) +(* bool -> int -> int * int -> (int * int -> int) -> Kodkod.tuple list *) +fun tabulate_nat_op2 debug univ_card (k, j0) f = + tabulate_op2 debug univ_card (k, j0) j0 (atom_for_nat (k, 0) o f) +fun tabulate_int_op2 debug univ_card (k, j0) f = + tabulate_op2 debug univ_card (k, j0) j0 + (atom_for_int (k, 0) o f o pairself (int_for_atom (k, 0))) +(* bool -> int -> int * int -> (int * int -> int * int) -> Kodkod.tuple list *) +fun tabulate_int_op2_2 debug univ_card (k, j0) f = + tabulate_op2_2 debug univ_card (k, j0) j0 + (pairself (atom_for_int (k, 0)) o f + o pairself (int_for_atom (k, 0))) + +(* int * int -> int *) +fun isa_div (m, n) = m div n handle General.Div => 0 +fun isa_mod (m, n) = m mod n handle General.Div => m +fun isa_gcd (m, 0) = m + | isa_gcd (m, n) = isa_gcd (n, isa_mod (m, n)) +fun isa_lcm (m, n) = isa_div (m * n, isa_gcd (m, n)) +val isa_zgcd = isa_gcd o pairself abs +(* int * int -> int * int *) +fun isa_norm_frac (m, n) = + if n < 0 then isa_norm_frac (~m, ~n) + else if m = 0 orelse n = 0 then (0, 1) + else let val p = isa_zgcd (m, n) in (isa_div (m, p), isa_div (n, p)) end + +(* bool -> int -> int -> int -> int -> int * int + -> string * bool * Kodkod.tuple list *) +fun tabulate_built_in_rel debug univ_card nat_card int_card j0 (x as (n, _)) = + (check_arity univ_card n; + if Kodkod.Rel x = not3_rel then + ("not3", tabulate_func1 debug univ_card (2, j0) (curry (op -) 1)) + else if Kodkod.Rel x = suc_rel then + ("suc", tabulate_func1 debug univ_card (univ_card - j0 - 1, j0) + (Integer.add 1)) + else if Kodkod.Rel x = nat_add_rel then + ("nat_add", tabulate_nat_op2 debug univ_card (nat_card, j0) (op +)) + else if Kodkod.Rel x = int_add_rel then + ("int_add", tabulate_int_op2 debug univ_card (int_card, j0) (op +)) + else if Kodkod.Rel x = nat_subtract_rel then + ("nat_subtract", + tabulate_op2 debug univ_card (nat_card, j0) j0 (op nat_minus)) + else if Kodkod.Rel x = int_subtract_rel then + ("int_subtract", tabulate_int_op2 debug univ_card (int_card, j0) (op -)) + else if Kodkod.Rel x = nat_multiply_rel then + ("nat_multiply", tabulate_nat_op2 debug univ_card (nat_card, j0) (op * )) + else if Kodkod.Rel x = int_multiply_rel then + ("int_multiply", tabulate_int_op2 debug univ_card (int_card, j0) (op * )) + else if Kodkod.Rel x = nat_divide_rel then + ("nat_divide", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_div) + else if Kodkod.Rel x = int_divide_rel then + ("int_divide", tabulate_int_op2 debug univ_card (int_card, j0) isa_div) + else if Kodkod.Rel x = nat_modulo_rel then + ("nat_modulo", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_mod) + else if Kodkod.Rel x = int_modulo_rel then + ("int_modulo", tabulate_int_op2 debug univ_card (int_card, j0) isa_mod) + else if Kodkod.Rel x = nat_less_rel then + ("nat_less", tabulate_nat_op2 debug univ_card (nat_card, j0) + (int_for_bool o op <)) + else if Kodkod.Rel x = int_less_rel then + ("int_less", tabulate_int_op2 debug univ_card (int_card, j0) + (int_for_bool o op <)) + else if Kodkod.Rel x = gcd_rel then + ("gcd", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_gcd) + else if Kodkod.Rel x = lcm_rel then + ("lcm", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_lcm) + else if Kodkod.Rel x = norm_frac_rel then + ("norm_frac", tabulate_int_op2_2 debug univ_card (int_card, j0) + isa_norm_frac) + else + raise ARG ("NitpickKodkod.tabulate_built_in_rel", "unknown relation")) + +(* bool -> int -> int -> int -> int -> int * int -> Kodkod.rel_expr + -> Kodkod.bound *) +fun bound_for_built_in_rel debug univ_card nat_card int_card j0 x = + let + val (nick, ts) = tabulate_built_in_rel debug univ_card nat_card int_card + j0 x + in ([(x, nick)], [Kodkod.TupleSet ts]) end + +(* bool -> int -> int -> int -> int -> Kodkod.formula -> Kodkod.bound list *) +fun bounds_for_built_in_rels_in_formula debug univ_card nat_card int_card j0 = + map (bound_for_built_in_rel debug univ_card nat_card int_card j0) + o built_in_rels_in_formula + +(* Proof.context -> bool -> nut -> Kodkod.bound *) +fun bound_for_plain_rel ctxt debug (u as FreeRel (x, T, R, nick)) = + ([(x, bound_comment ctxt debug nick T R)], + if nick = @{const_name bisim_iterator_max} then + case R of + Atom (k, j0) => [Kodkod.TupleSet [Kodkod.Tuple [k - 1 + j0]]] + | _ => raise NUT ("NitpickKodkod.bound_for_plain_rel", [u]) + else + [Kodkod.TupleSet [], upper_bound_for_rep R]) + | bound_for_plain_rel _ _ u = + raise NUT ("NitpickKodkod.bound_for_plain_rel", [u]) + +(* Proof.context -> bool -> dtype_spec list -> nut -> Kodkod.bound *) +fun bound_for_sel_rel ctxt debug dtypes + (FreeRel (x, T as Type ("fun", [T1, T2]), R as Func (Atom (_, j0), R2), + nick)) = + let + val constr as {delta, epsilon, exclusive, explicit_max, ...} = + constr_spec dtypes (original_name nick, T1) + in + ([(x, bound_comment ctxt debug nick T R)], + if explicit_max = 0 then + [Kodkod.TupleSet []] + else + let val ts = Kodkod.TupleAtomSeq (epsilon - delta, delta + j0) in + if R2 = Formula Neut then + [ts] |> not exclusive ? cons (Kodkod.TupleSet []) + else + [Kodkod.TupleSet [], + Kodkod.TupleProduct (ts, upper_bound_for_rep R2)] + end) + end + | bound_for_sel_rel _ _ _ u = + raise NUT ("NitpickKodkod.bound_for_sel_rel", [u]) + +(* Kodkod.bound list -> Kodkod.bound list *) +fun merge_bounds bs = + let + (* Kodkod.bound -> int *) + fun arity (zs, _) = fst (fst (hd zs)) + (* Kodkod.bound list -> Kodkod.bound -> Kodkod.bound list + -> Kodkod.bound list *) + fun add_bound ds b [] = List.revAppend (ds, [b]) + | add_bound ds b (c :: cs) = + if arity b = arity c andalso snd b = snd c then + List.revAppend (ds, (fst c @ fst b, snd c) :: cs) + else + add_bound (c :: ds) b cs + in fold (add_bound []) bs [] end + +(* int -> int -> Kodkod.rel_expr list *) +fun unary_var_seq j0 n = map (curry Kodkod.Var 1) (index_seq j0 n) + +(* int list -> Kodkod.rel_expr *) +val singleton_from_combination = foldl1 Kodkod.Product o map Kodkod.Atom +(* rep -> Kodkod.rel_expr list *) +fun all_singletons_for_rep R = + if is_lone_rep R then + all_combinations_for_rep R |> map singleton_from_combination + else + raise REP ("NitpickKodkod.all_singletons_for_rep", [R]) + +(* Kodkod.rel_expr -> Kodkod.rel_expr list *) +fun unpack_products (Kodkod.Product (r1, r2)) = + unpack_products r1 @ unpack_products r2 + | unpack_products r = [r] +fun unpack_joins (Kodkod.Join (r1, r2)) = unpack_joins r1 @ unpack_joins r2 + | unpack_joins r = [r] + +(* rep -> Kodkod.rel_expr *) +val empty_rel_for_rep = empty_n_ary_rel o arity_of_rep +fun full_rel_for_rep R = + case atom_schema_of_rep R of + [] => raise REP ("NitpickKodkod.full_rel_for_rep", [R]) + | schema => foldl1 Kodkod.Product (map Kodkod.AtomSeq schema) + +(* int -> int list -> Kodkod.decl list *) +fun decls_for_atom_schema j0 schema = + map2 (fn j => fn x => Kodkod.DeclOne ((1, j), Kodkod.AtomSeq x)) + (index_seq j0 (length schema)) schema + +(* The type constraint below is a workaround for a Poly/ML bug. *) + +(* FIXME: clean up *) +(* kodkod_constrs -> rep -> Kodkod.rel_expr -> Kodkod.formula *) +fun d_n_ary_function ({kk_all, kk_join, kk_lone, kk_one, ...} : kodkod_constrs) + R r = + let val body_R = body_rep R in + if is_lone_rep body_R then + let + val binder_schema = atom_schema_of_reps (binder_reps R) + val body_schema = atom_schema_of_rep body_R + val one = is_one_rep body_R + val opt_x = case r of Kodkod.Rel x => SOME x | _ => NONE + in + if opt_x <> NONE andalso length binder_schema = 1 + andalso length body_schema = 1 then + (if one then Kodkod.Function else Kodkod.Functional) + (the opt_x, Kodkod.AtomSeq (hd binder_schema), + Kodkod.AtomSeq (hd body_schema)) + else + let + val decls = decls_for_atom_schema ~1 binder_schema + val vars = unary_var_seq ~1 (length binder_schema) + val kk_xone = if one then kk_one else kk_lone + in kk_all decls (kk_xone (fold kk_join vars r)) end + end + else + Kodkod.True + end +fun kk_n_ary_function kk R (r as Kodkod.Rel _) = + (* FIXME: weird test *) + if not (is_opt_rep R) then + if r = suc_rel then + Kodkod.False + else if r = nat_add_rel then + formula_for_bool (card_of_rep (body_rep R) = 1) + else if r = nat_multiply_rel then + formula_for_bool (card_of_rep (body_rep R) <= 2) + else + d_n_ary_function kk R r + else if r = nat_subtract_rel then + Kodkod.True + else + d_n_ary_function kk R r + | kk_n_ary_function kk R r = d_n_ary_function kk R r + +(* kodkod_constrs -> Kodkod.rel_expr list -> Kodkod.formula *) +fun kk_disjoint_sets _ [] = Kodkod.True + | kk_disjoint_sets (kk as {kk_and, kk_no, kk_intersect, ...} : kodkod_constrs) + (r :: rs) = + fold (kk_and o kk_no o kk_intersect r) rs (kk_disjoint_sets kk rs) + +(* int -> kodkod_constrs -> (Kodkod.rel_expr -> Kodkod.rel_expr) + -> Kodkod.rel_expr -> Kodkod.rel_expr *) +fun basic_rel_let j ({kk_rel_let, ...} : kodkod_constrs) f r = + if inline_rel_expr r then + f r + else + let val x = (Kodkod.arity_of_rel_expr r, j) in + kk_rel_let [Kodkod.AssignRelReg (x, r)] (f (Kodkod.RelReg x)) + end + +(* kodkod_constrs -> (Kodkod.rel_expr -> Kodkod.rel_expr) -> Kodkod.rel_expr + -> Kodkod.rel_expr *) +val single_rel_let = basic_rel_let 0 +(* kodkod_constrs -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr) + -> Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr *) +fun double_rel_let kk f r1 r2 = + single_rel_let kk (fn r1 => basic_rel_let 1 kk (f r1) r2) r1 +(* kodkod_constrs + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr) + -> Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr + -> Kodkod.rel_expr *) +fun triple_rel_let kk f r1 r2 r3 = + double_rel_let kk (fn r1 => fn r2 => basic_rel_let 2 kk (f r1 r2) r3) r1 r2 + +(* kodkod_constrs -> int -> Kodkod.formula -> Kodkod.rel_expr *) +fun atom_from_formula ({kk_rel_if, ...} : kodkod_constrs) j0 f = + kk_rel_if f (Kodkod.Atom (j0 + 1)) (Kodkod.Atom j0) +(* kodkod_constrs -> rep -> Kodkod.formula -> Kodkod.rel_expr *) +fun rel_expr_from_formula kk R f = + case unopt_rep R of + Atom (2, j0) => atom_from_formula kk j0 f + | _ => raise REP ("NitpickKodkod.rel_expr_from_formula", [R]) + +(* kodkod_cotrs -> int -> int -> Kodkod.rel_expr -> Kodkod.rel_expr list *) +fun unpack_vect_in_chunks ({kk_project_seq, ...} : kodkod_constrs) chunk_arity + num_chunks r = + List.tabulate (num_chunks, fn j => kk_project_seq r (j * chunk_arity) + chunk_arity) + +(* kodkod_constrs -> bool -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr + -> Kodkod.rel_expr *) +fun kk_n_fold_join + (kk as {kk_intersect, kk_product, kk_join, kk_project_seq, ...}) one R1 + res_R r1 r2 = + case arity_of_rep R1 of + 1 => kk_join r1 r2 + | arity1 => + let + val unpacked_rs1 = + if inline_rel_expr r1 then unpack_vect_in_chunks kk 1 arity1 r1 + else unpack_products r1 + in + if one andalso length unpacked_rs1 = arity1 then + fold kk_join unpacked_rs1 r2 + else + kk_project_seq + (kk_intersect (kk_product r1 (full_rel_for_rep res_R)) r2) + arity1 (arity_of_rep res_R) + end + +(* kodkod_constrs -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr list + -> Kodkod.rel_expr list -> Kodkod.rel_expr *) +fun kk_case_switch (kk as {kk_union, kk_product, ...}) R1 R2 r rs1 rs2 = + if rs1 = rs2 then r + else kk_n_fold_join kk true R1 R2 r (fold1 kk_union (map2 kk_product rs1 rs2)) + +val lone_rep_fallback_max_card = 4096 +val some_j0 = 0 + +(* kodkod_constrs -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +fun lone_rep_fallback kk new_R old_R r = + if old_R = new_R then + r + else + let val card = card_of_rep old_R in + if is_lone_rep old_R andalso is_lone_rep new_R + andalso card = card_of_rep new_R then + if card >= lone_rep_fallback_max_card then + raise LIMIT ("NitpickKodkod.lone_rep_fallback", + "too high cardinality (" ^ string_of_int card ^ ")") + else + kk_case_switch kk old_R new_R r (all_singletons_for_rep old_R) + (all_singletons_for_rep new_R) + else + raise REP ("NitpickKodkod.lone_rep_fallback", [old_R, new_R]) + end +(* kodkod_constrs -> int * int -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and atom_from_rel_expr kk (x as (k, j0)) old_R r = + case old_R of + Func (R1, R2) => + let + val dom_card = card_of_rep R1 + val R2' = case R2 of Atom _ => R2 | _ => Atom (card_of_rep R2, some_j0) + in + atom_from_rel_expr kk x (Vect (dom_card, R2')) + (vect_from_rel_expr kk dom_card R2' old_R r) + end + | Opt _ => raise REP ("NitpickKodkod.atom_from_rel_expr", [old_R]) + | _ => lone_rep_fallback kk (Atom x) old_R r +(* kodkod_constrs -> rep list -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and struct_from_rel_expr kk Rs old_R r = + case old_R of + Atom _ => lone_rep_fallback kk (Struct Rs) old_R r + | Struct Rs' => + let + val Rs = filter (not_equal Unit) Rs + val Rs' = filter (not_equal Unit) Rs' + in + if Rs' = Rs then + r + else if map card_of_rep Rs' = map card_of_rep Rs then + let + val old_arities = map arity_of_rep Rs' + val old_offsets = offset_list old_arities + val old_rs = map2 (#kk_project_seq kk r) old_offsets old_arities + in + fold1 (#kk_product kk) + (map3 (rel_expr_from_rel_expr kk) Rs Rs' old_rs) + end + else + lone_rep_fallback kk (Struct Rs) old_R r + end + | _ => raise REP ("NitpickKodkod.struct_from_rel_expr", [old_R]) +(* kodkod_constrs -> int -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and vect_from_rel_expr kk k R old_R r = + case old_R of + Atom _ => lone_rep_fallback kk (Vect (k, R)) old_R r + | Vect (k', R') => + if k = k' andalso R = R' then r + else lone_rep_fallback kk (Vect (k, R)) old_R r + | Func (R1, Formula Neut) => + if k = card_of_rep R1 then + fold1 (#kk_product kk) + (map (fn arg_r => + rel_expr_from_formula kk R (#kk_subset kk arg_r r)) + (all_singletons_for_rep R1)) + else + raise REP ("NitpickKodkod.vect_from_rel_expr", [old_R]) + | Func (Unit, R2) => rel_expr_from_rel_expr kk R R2 r + | Func (R1, R2) => + fold1 (#kk_product kk) + (map (fn arg_r => + rel_expr_from_rel_expr kk R R2 + (kk_n_fold_join kk true R1 R2 arg_r r)) + (all_singletons_for_rep R1)) + | _ => raise REP ("NitpickKodkod.vect_from_rel_expr", [old_R]) +(* kodkod_constrs -> rep -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and func_from_no_opt_rel_expr kk R1 R2 (Atom x) r = + let + val dom_card = card_of_rep R1 + val R2' = case R2 of Atom _ => R2 | _ => Atom (card_of_rep R2, some_j0) + in + func_from_no_opt_rel_expr kk R1 R2 (Vect (dom_card, R2')) + (vect_from_rel_expr kk dom_card R2' (Atom x) r) + end + | func_from_no_opt_rel_expr kk Unit R2 old_R r = + (case old_R of + Vect (k, R') => rel_expr_from_rel_expr kk R2 R' r + | Func (Unit, R2') => rel_expr_from_rel_expr kk R2 R2' r + | Func (Atom (1, _), Formula Neut) => + (case unopt_rep R2 of + Atom (2, j0) => atom_from_formula kk j0 (#kk_some kk r) + | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr", + [old_R, Func (Unit, R2)])) + | Func (R1', R2') => + rel_expr_from_rel_expr kk R2 R2' (#kk_project_seq kk r (arity_of_rep R1') + (arity_of_rep R2')) + | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr", + [old_R, Func (Unit, R2)])) + | func_from_no_opt_rel_expr kk R1 (Formula Neut) old_R r = + (case old_R of + Vect (k, Atom (2, j0)) => + let + val args_rs = all_singletons_for_rep R1 + val vals_rs = unpack_vect_in_chunks kk 1 k r + (* Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr *) + fun empty_or_singleton_set_for arg_r val_r = + #kk_join kk val_r (#kk_product kk (Kodkod.Atom (j0 + 1)) arg_r) + in + fold1 (#kk_union kk) (map2 empty_or_singleton_set_for args_rs vals_rs) + end + | Func (R1', Formula Neut) => + if R1 = R1' then + r + else + let + val schema = atom_schema_of_rep R1 + val r1 = fold1 (#kk_product kk) (unary_var_seq ~1 (length schema)) + |> rel_expr_from_rel_expr kk R1' R1 + in + #kk_comprehension kk (decls_for_atom_schema ~1 schema) + (#kk_subset kk r1 r) + end + | Func (Unit, (Atom (2, j0))) => + #kk_rel_if kk (#kk_rel_eq kk r (Kodkod.Atom (j0 + 1))) + (full_rel_for_rep R1) (empty_rel_for_rep R1) + | Func (R1', Atom (2, j0)) => + func_from_no_opt_rel_expr kk R1 (Formula Neut) + (Func (R1', Formula Neut)) (#kk_join kk r (Kodkod.Atom (j0 + 1))) + | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr", + [old_R, Func (R1, Formula Neut)])) + | func_from_no_opt_rel_expr kk R1 R2 old_R r = + case old_R of + Vect (k, R) => + let + val args_rs = all_singletons_for_rep R1 + val vals_rs = unpack_vect_in_chunks kk (arity_of_rep R) k r + |> map (rel_expr_from_rel_expr kk R2 R) + in fold1 (#kk_union kk) (map2 (#kk_product kk) args_rs vals_rs) end + | Func (R1', Formula Neut) => + (case R2 of + Atom (x as (2, j0)) => + let val schema = atom_schema_of_rep R1 in + if length schema = 1 then + #kk_override kk (#kk_product kk (Kodkod.AtomSeq (hd schema)) + (Kodkod.Atom j0)) + (#kk_product kk r (Kodkod.Atom (j0 + 1))) + else + let + val r1 = fold1 (#kk_product kk) (unary_var_seq ~1 (length schema)) + |> rel_expr_from_rel_expr kk R1' R1 + val r2 = Kodkod.Var (1, ~(length schema) - 1) + val r3 = atom_from_formula kk j0 (#kk_subset kk r1 r) + in + #kk_comprehension kk (decls_for_atom_schema ~1 (schema @ [x])) + (#kk_rel_eq kk r2 r3) + end + end + | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr", + [old_R, Func (R1, R2)])) + | Func (Unit, R2') => + let val j0 = some_j0 in + func_from_no_opt_rel_expr kk R1 R2 (Func (Atom (1, j0), R2')) + (#kk_product kk (Kodkod.Atom j0) r) + end + | Func (R1', R2') => + if R1 = R1' andalso R2 = R2' then + r + else + let + val dom_schema = atom_schema_of_rep R1 + val ran_schema = atom_schema_of_rep R2 + val dom_prod = fold1 (#kk_product kk) + (unary_var_seq ~1 (length dom_schema)) + |> rel_expr_from_rel_expr kk R1' R1 + val ran_prod = fold1 (#kk_product kk) + (unary_var_seq (~(length dom_schema) - 1) + (length ran_schema)) + |> rel_expr_from_rel_expr kk R2' R2 + val app = kk_n_fold_join kk true R1' R2' dom_prod r + in + #kk_comprehension kk (decls_for_atom_schema ~1 + (dom_schema @ ran_schema)) + (#kk_subset kk ran_prod app) + end + | _ => raise REP ("NitpickKodkod.func_from_no_opt_rel_expr", + [old_R, Func (R1, R2)]) +(* kodkod_constrs -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and rel_expr_from_rel_expr kk new_R old_R r = + let + val unopt_old_R = unopt_rep old_R + val unopt_new_R = unopt_rep new_R + in + if unopt_old_R <> old_R andalso unopt_new_R = new_R then + raise REP ("NitpickKodkod.rel_expr_from_rel_expr", [old_R, new_R]) + else if unopt_new_R = unopt_old_R then + r + else + (case unopt_new_R of + Atom x => atom_from_rel_expr kk x + | Struct Rs => struct_from_rel_expr kk Rs + | Vect (k, R') => vect_from_rel_expr kk k R' + | Func (R1, R2) => func_from_no_opt_rel_expr kk R1 R2 + | _ => raise REP ("NitpickKodkod.rel_expr_from_rel_expr", + [old_R, new_R])) + unopt_old_R r + end +(* kodkod_constrs -> rep -> rep -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) +and rel_expr_to_func kk R1 R2 = rel_expr_from_rel_expr kk (Func (R1, R2)) + +(* kodkod_constrs -> nut -> Kodkod.formula *) +fun declarative_axiom_for_plain_rel kk (FreeRel (x, _, R as Func _, nick)) = + kk_n_ary_function kk (R |> nick = @{const_name List.set} ? unopt_rep) + (Kodkod.Rel x) + | declarative_axiom_for_plain_rel ({kk_lone, kk_one, ...} : kodkod_constrs) + (FreeRel (x, _, R, _)) = + if is_one_rep R then kk_one (Kodkod.Rel x) + else if is_lone_rep R andalso card_of_rep R > 1 then kk_lone (Kodkod.Rel x) + else Kodkod.True + | declarative_axiom_for_plain_rel _ u = + raise NUT ("NitpickKodkod.declarative_axiom_for_plain_rel", [u]) + +(* nut NameTable.table -> styp -> Kodkod.rel_expr * rep * int *) +fun const_triple rel_table (x as (s, T)) = + case the_name rel_table (ConstName (s, T, Any)) of + FreeRel ((n, j), _, R, _) => (Kodkod.Rel (n, j), R, n) + | _ => raise TERM ("NitpickKodkod.const_triple", [Const x]) + +(* nut NameTable.table -> styp -> Kodkod.rel_expr *) +fun discr_rel_expr rel_table = #1 o const_triple rel_table o discr_for_constr + +(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list + -> styp -> int -> nfa_transition list *) +fun nfa_transitions_for_sel ext_ctxt ({kk_project, ...} : kodkod_constrs) + rel_table (dtypes : dtype_spec list) constr_x n = + let + val x as (_, T) = boxed_nth_sel_for_constr ext_ctxt constr_x n + val (r, R, arity) = const_triple rel_table x + val type_schema = type_schema_of_rep T R + in + map_filter (fn (j, T) => + if forall (not_equal T o #typ) dtypes then NONE + else SOME (kk_project r (map Kodkod.Num [0, j]), T)) + (index_seq 1 (arity - 1) ~~ tl type_schema) + end +(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list + -> styp -> nfa_transition list *) +fun nfa_transitions_for_constr ext_ctxt kk rel_table dtypes (x as (_, T)) = + maps (nfa_transitions_for_sel ext_ctxt kk rel_table dtypes x) + (index_seq 0 (num_sels_for_constr_type T)) +(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list + -> dtype_spec -> nfa_entry option *) +fun nfa_entry_for_datatype _ _ _ _ ({co = true, ...} : dtype_spec) = NONE + | nfa_entry_for_datatype ext_ctxt kk rel_table dtypes + ({typ, constrs, ...} : dtype_spec) = + SOME (typ, maps (nfa_transitions_for_constr ext_ctxt kk rel_table dtypes + o #const) constrs) + +val empty_rel = Kodkod.Product (Kodkod.None, Kodkod.None) + +(* nfa_table -> typ -> typ -> Kodkod.rel_expr list *) +fun direct_path_rel_exprs nfa start final = + case AList.lookup (op =) nfa final of + SOME trans => map fst (filter (equal start o snd) trans) + | NONE => [] +(* kodkod_constrs -> nfa_table -> typ list -> typ -> typ -> Kodkod.rel_expr *) +and any_path_rel_expr ({kk_union, ...} : kodkod_constrs) nfa [] start final = + fold kk_union (direct_path_rel_exprs nfa start final) + (if start = final then Kodkod.Iden else empty_rel) + | any_path_rel_expr (kk as {kk_union, ...}) nfa (q :: qs) start final = + kk_union (any_path_rel_expr kk nfa qs start final) + (knot_path_rel_expr kk nfa qs start q final) +(* kodkod_constrs -> nfa_table -> typ list -> typ -> typ -> typ + -> Kodkod.rel_expr *) +and knot_path_rel_expr (kk as {kk_join, kk_reflexive_closure, ...}) nfa qs start + knot final = + kk_join (kk_join (any_path_rel_expr kk nfa qs knot final) + (kk_reflexive_closure (loop_path_rel_expr kk nfa qs knot))) + (any_path_rel_expr kk nfa qs start knot) +(* kodkod_constrs -> nfa_table -> typ list -> typ -> Kodkod.rel_expr *) +and loop_path_rel_expr ({kk_union, ...} : kodkod_constrs) nfa [] start = + fold kk_union (direct_path_rel_exprs nfa start start) empty_rel + | loop_path_rel_expr (kk as {kk_union, kk_closure, ...}) nfa (q :: qs) start = + if start = q then + kk_closure (loop_path_rel_expr kk nfa qs start) + else + kk_union (loop_path_rel_expr kk nfa qs start) + (knot_path_rel_expr kk nfa qs start q start) + +(* nfa_table -> unit NfaGraph.T *) +fun graph_for_nfa nfa = + let + (* typ -> unit NfaGraph.T -> unit NfaGraph.T *) + fun new_node q = perhaps (try (NfaGraph.new_node (q, ()))) + (* nfa_table -> unit NfaGraph.T -> unit NfaGraph.T *) + fun add_nfa [] = I + | add_nfa ((_, []) :: nfa) = add_nfa nfa + | add_nfa ((q, ((_, q') :: transitions)) :: nfa) = + add_nfa ((q, transitions) :: nfa) o NfaGraph.add_edge (q, q') o + new_node q' o new_node q + in add_nfa nfa NfaGraph.empty end + +(* nfa_table -> nfa_table list *) +fun strongly_connected_sub_nfas nfa = + nfa |> graph_for_nfa |> NfaGraph.strong_conn + |> map (fn keys => filter (member (op =) keys o fst) nfa) + +(* dtype_spec list -> kodkod_constrs -> nfa_table -> typ -> Kodkod.formula *) +fun acyclicity_axiom_for_datatype dtypes kk nfa start = + #kk_no kk (#kk_intersect kk + (loop_path_rel_expr kk nfa (map fst nfa) start) Kodkod.Iden) +(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list + -> Kodkod.formula list *) +fun acyclicity_axioms_for_datatypes ext_ctxt kk rel_table dtypes = + map_filter (nfa_entry_for_datatype ext_ctxt kk rel_table dtypes) dtypes + |> strongly_connected_sub_nfas + |> maps (fn nfa => map (acyclicity_axiom_for_datatype dtypes kk nfa o fst) + nfa) + +(* extended_context -> int -> kodkod_constrs -> nut NameTable.table + -> Kodkod.rel_expr -> constr_spec -> int -> Kodkod.formula *) +fun sel_axiom_for_sel ext_ctxt j0 + (kk as {kk_all, kk_implies, kk_formula_if, kk_subset, kk_rel_eq, kk_no, + kk_join, kk_project, ...}) rel_table dom_r + ({const, delta, epsilon, exclusive, explicit_max, ...} : constr_spec) + n = + let + val x as (_, T) = boxed_nth_sel_for_constr ext_ctxt const n + val (r, R, arity) = const_triple rel_table x + val R2 = dest_Func R |> snd + val z = (epsilon - delta, delta + j0) + in + if exclusive then + kk_n_ary_function kk (Func (Atom z, R2)) r + else + let val r' = kk_join (Kodkod.Var (1, 0)) r in + kk_all [Kodkod.DeclOne ((1, 0), Kodkod.AtomSeq z)] + (kk_formula_if (kk_subset (Kodkod.Var (1, 0)) dom_r) + (kk_n_ary_function kk R2 r') + (kk_no r')) + end + end +(* extended_context -> int -> kodkod_constrs -> nut NameTable.table + -> constr_spec -> Kodkod.formula list *) +fun sel_axioms_for_constr ext_ctxt j0 kk rel_table + (constr as {const, delta, epsilon, explicit_max, ...}) = + let + val honors_explicit_max = + explicit_max < 0 orelse epsilon - delta <= explicit_max + in + if explicit_max = 0 then + [formula_for_bool honors_explicit_max] + else + let + val ran_r = discr_rel_expr rel_table const + val max_axiom = + if honors_explicit_max then Kodkod.True + else Kodkod.LE (Kodkod.Cardinality ran_r, Kodkod.Num explicit_max) + in + max_axiom :: + map (sel_axiom_for_sel ext_ctxt j0 kk rel_table ran_r constr) + (index_seq 0 (num_sels_for_constr_type (snd const))) + end + end +(* extended_context -> int -> kodkod_constrs -> nut NameTable.table + -> dtype_spec -> Kodkod.formula list *) +fun sel_axioms_for_datatype ext_ctxt j0 kk rel_table + ({constrs, ...} : dtype_spec) = + maps (sel_axioms_for_constr ext_ctxt j0 kk rel_table) constrs + +(* extended_context -> kodkod_constrs -> nut NameTable.table -> constr_spec + -> Kodkod.formula list *) +fun uniqueness_axiom_for_constr ext_ctxt + ({kk_all, kk_implies, kk_and, kk_rel_eq, kk_lone, kk_join, ...} + : kodkod_constrs) rel_table ({const, ...} : constr_spec) = + let + (* Kodkod.rel_expr -> Kodkod.formula *) + fun conjunct_for_sel r = + kk_rel_eq (kk_join (Kodkod.Var (1, 0)) r) + (kk_join (Kodkod.Var (1, 1)) r) + val num_sels = num_sels_for_constr_type (snd const) + val triples = map (const_triple rel_table + o boxed_nth_sel_for_constr ext_ctxt const) + (~1 upto num_sels - 1) + val j0 = case triples |> hd |> #2 of + Func (Atom (_, j0), _) => j0 + | R => raise REP ("NitpickKodkod.uniqueness_axiom_for_constr", [R]) + val set_r = triples |> hd |> #1 + in + if num_sels = 0 then + kk_lone set_r + else + kk_all (map (Kodkod.DeclOne o rpair set_r o pair 1) [0, 1]) + (kk_implies + (fold1 kk_and (map (conjunct_for_sel o #1) (tl triples))) + (kk_rel_eq (Kodkod.Var (1, 0)) (Kodkod.Var (1, 1)))) + end +(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec + -> Kodkod.formula list *) +fun uniqueness_axioms_for_datatype ext_ctxt kk rel_table + ({constrs, ...} : dtype_spec) = + map (uniqueness_axiom_for_constr ext_ctxt kk rel_table) constrs + +(* constr_spec -> int *) +fun effective_constr_max ({delta, epsilon, ...} : constr_spec) = epsilon - delta +(* int -> kodkod_constrs -> nut NameTable.table -> dtype_spec + -> Kodkod.formula list *) +fun partition_axioms_for_datatype j0 (kk as {kk_rel_eq, kk_union, ...}) + rel_table + ({card, constrs, ...} : dtype_spec) = + if forall #exclusive constrs then + [Integer.sum (map effective_constr_max constrs) = card |> formula_for_bool] + else + let val rs = map (discr_rel_expr rel_table o #const) constrs in + [kk_rel_eq (fold1 kk_union rs) (Kodkod.AtomSeq (card, j0)), + kk_disjoint_sets kk rs] + end + +(* extended_context -> int Typtab.table -> kodkod_constrs -> nut NameTable.table + -> dtype_spec -> Kodkod.formula list *) +fun other_axioms_for_datatype ext_ctxt ofs kk rel_table (dtype as {typ, ...}) = + let val j0 = offset_of_type ofs typ in + sel_axioms_for_datatype ext_ctxt j0 kk rel_table dtype @ + uniqueness_axioms_for_datatype ext_ctxt kk rel_table dtype @ + partition_axioms_for_datatype j0 kk rel_table dtype + end + +(* extended_context -> int Typtab.table -> kodkod_constrs -> nut NameTable.table + -> dtype_spec list -> Kodkod.formula list *) +fun declarative_axioms_for_datatypes ext_ctxt ofs kk rel_table dtypes = + acyclicity_axioms_for_datatypes ext_ctxt kk rel_table dtypes @ + maps (other_axioms_for_datatype ext_ctxt ofs kk rel_table) dtypes + +(* int Typtab.table -> bool -> kodkod_constrs -> nut -> Kodkod.formula *) +fun kodkod_formula_from_nut ofs liberal + (kk as {kk_all, kk_exist, kk_formula_let, kk_formula_if, kk_or, kk_not, + kk_iff, kk_implies, kk_and, kk_subset, kk_rel_eq, kk_no, kk_one, + kk_some, kk_rel_let, kk_rel_if, kk_union, kk_difference, + kk_intersect, kk_product, kk_join, kk_closure, kk_comprehension, + kk_project, kk_project_seq, kk_not3, kk_nat_less, kk_int_less, + ...}) u = + let + val main_j0 = offset_of_type ofs bool_T + val bool_j0 = main_j0 + val bool_atom_R = Atom (2, main_j0) + val false_atom = Kodkod.Atom bool_j0 + val true_atom = Kodkod.Atom (bool_j0 + 1) + + (* polarity -> int -> Kodkod.rel_expr -> Kodkod.formula *) + fun formula_from_opt_atom polar j0 r = + case polar of + Neg => kk_not (kk_rel_eq r (Kodkod.Atom j0)) + | _ => kk_rel_eq r (Kodkod.Atom (j0 + 1)) + (* int -> Kodkod.rel_expr -> Kodkod.formula *) + val formula_from_atom = formula_from_opt_atom Pos + + (* Kodkod.formula -> Kodkod.formula -> Kodkod.formula *) + fun kk_notimplies f1 f2 = kk_and f1 (kk_not f2) + (* Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr *) + val kk_or3 = + double_rel_let kk + (fn r1 => fn r2 => + kk_rel_if (kk_subset true_atom (kk_union r1 r2)) true_atom + (kk_intersect r1 r2)) + val kk_and3 = + double_rel_let kk + (fn r1 => fn r2 => + kk_rel_if (kk_subset false_atom (kk_union r1 r2)) false_atom + (kk_intersect r1 r2)) + fun kk_notimplies3 r1 r2 = kk_and3 r1 (kk_not3 r2) + + (* int -> Kodkod.rel_expr -> Kodkod.formula list *) + val unpack_formulas = + map (formula_from_atom bool_j0) oo unpack_vect_in_chunks kk 1 + (* (Kodkod.formula -> Kodkod.formula -> Kodkod.formula) -> int + -> Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr *) + fun kk_vect_set_op connective k r1 r2 = + fold1 kk_product (map2 (atom_from_formula kk bool_j0 oo connective) + (unpack_formulas k r1) (unpack_formulas k r2)) + (* (Kodkod.formula -> Kodkod.formula -> Kodkod.formula) -> int + -> Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.formula *) + fun kk_vect_set_bool_op connective k r1 r2 = + fold1 kk_and (map2 connective (unpack_formulas k r1) + (unpack_formulas k r2)) + + (* nut -> Kodkod.formula *) + fun to_f u = + case rep_of u of + Formula polar => + (case u of + Cst (False, _, _) => Kodkod.False + | Cst (True, _, _) => Kodkod.True + | Op1 (Not, _, _, u1) => kk_not (to_f u1) + | Op1 (Finite, _, _, u1) => + let val opt1 = is_opt_rep (rep_of u1) in + case polar of + Neut => if opt1 then + raise NUT ("NitpickKodkod.to_f (Finite)", [u]) + else + Kodkod.True + | Pos => formula_for_bool (not opt1) + | Neg => Kodkod.True + end + | Op1 (Cast, _, _, u1) => to_f_with_polarity polar u1 + | Op2 (All, _, _, u1, u2) => kk_all (untuple to_decl u1) (to_f u2) + | Op2 (Exist, _, _, u1, u2) => kk_exist (untuple to_decl u1) (to_f u2) + | Op2 (Or, _, _, u1, u2) => kk_or (to_f u1) (to_f u2) + | Op2 (And, _, _, u1, u2) => kk_and (to_f u1) (to_f u2) + | Op2 (Less, T, Formula polar, u1, u2) => + formula_from_opt_atom polar bool_j0 + (to_r (Op2 (Less, T, Opt bool_atom_R, u1, u2))) + | Op2 (Subset, _, _, u1, u2) => + let + val dom_T = domain_type (type_of u1) + val R1 = rep_of u1 + val R2 = rep_of u2 + val (dom_R, ran_R) = + case min_rep R1 R2 of + Func (Unit, R') => + (Atom (1, offset_of_type ofs dom_T), R') + | Func Rp => Rp + | R => (Atom (card_of_domain_from_rep 2 R, + offset_of_type ofs dom_T), + if is_opt_rep R then Opt bool_atom_R else Formula Neut) + val set_R = Func (dom_R, ran_R) + in + if not (is_opt_rep ran_R) then + to_set_bool_op kk_implies kk_subset u1 u2 + else if polar = Neut then + raise NUT ("NitpickKodkod.to_f (Subset)", [u]) + else + let + (* bool -> nut -> Kodkod.rel_expr *) + fun set_to_r widen u = + if widen then + kk_difference (full_rel_for_rep dom_R) + (kk_join (to_rep set_R u) false_atom) + else + kk_join (to_rep set_R u) true_atom + val widen1 = (polar = Pos andalso is_opt_rep R1) + val widen2 = (polar = Neg andalso is_opt_rep R2) + in kk_subset (set_to_r widen1 u1) (set_to_r widen2 u2) end + end + | Op2 (DefEq, _, _, u1, u2) => + (case min_rep (rep_of u1) (rep_of u2) of + Unit => Kodkod.True + | Formula polar => + kk_iff (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) + | min_R => + let + (* nut -> nut list *) + fun args (Op2 (Apply, _, _, u1, u2)) = u2 :: args u1 + | args (Tuple (_, _, us)) = us + | args _ = [] + val opt_arg_us = filter (is_opt_rep o rep_of) (args u1) + in + if null opt_arg_us orelse not (is_Opt min_R) + orelse is_eval_name u1 then + fold (kk_or o (kk_no o to_r)) opt_arg_us + (kk_rel_eq (to_rep min_R u1) (to_rep min_R u2)) + else + kk_no (kk_difference (to_rep min_R u1) (to_rep min_R u2)) + end) + | Op2 (Eq, T, R, u1, u2) => + (case min_rep (rep_of u1) (rep_of u2) of + Unit => Kodkod.True + | Formula polar => + kk_iff (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) + | min_R => + if is_opt_rep min_R then + if polar = Neut then + (* continuation of hackish optimization *) + kk_rel_eq (to_rep min_R u1) (to_rep min_R u2) + else if is_Cst Unrep u1 then + to_could_be_unrep (polar = Neg) u2 + else if is_Cst Unrep u2 then + to_could_be_unrep (polar = Neg) u1 + else + let + val r1 = to_rep min_R u1 + val r2 = to_rep min_R u2 + val both_opt = forall (is_opt_rep o rep_of) [u1, u2] + in + (if polar = Pos then + if not both_opt then + kk_rel_eq r1 r2 + else if is_lone_rep min_R + andalso arity_of_rep min_R = 1 then + kk_some (kk_intersect r1 r2) + else + raise SAME () + else + if is_lone_rep min_R then + if arity_of_rep min_R = 1 then + kk_subset (kk_product r1 r2) Kodkod.Iden + else if not both_opt then + (r1, r2) |> is_opt_rep (rep_of u2) ? swap + |> uncurry kk_difference |> kk_no + else + raise SAME () + else + raise SAME ()) + handle SAME () => + formula_from_opt_atom polar bool_j0 + (to_guard [u1, u2] bool_atom_R + (rel_expr_from_formula kk bool_atom_R + (kk_rel_eq r1 r2))) + end + else + let + val r1 = to_rep min_R u1 + val r2 = to_rep min_R u2 + in + if is_one_rep min_R then + let + val rs1 = unpack_products r1 + val rs2 = unpack_products r2 + in + if length rs1 = length rs2 + andalso map Kodkod.arity_of_rel_expr rs1 + = map Kodkod.arity_of_rel_expr rs2 then + fold1 kk_and (map2 kk_subset rs1 rs2) + else + kk_subset r1 r2 + end + else + kk_rel_eq r1 r2 + end) + | Op2 (Apply, T, _, u1, u2) => + (case (polar, rep_of u1) of + (Neg, Func (R, Formula Neut)) => kk_subset (to_opt R u2) (to_r u1) + | _ => + to_f_with_polarity polar + (Op2 (Apply, T, Opt (Atom (2, offset_of_type ofs T)), u1, u2))) + | Op3 (Let, _, _, u1, u2, u3) => + kk_formula_let [to_expr_assign u1 u2] (to_f u3) + | Op3 (If, _, _, u1, u2, u3) => + kk_formula_if (to_f u1) (to_f u2) (to_f u3) + | FormulaReg (j, _, _) => Kodkod.FormulaReg j + | _ => raise NUT ("NitpickKodkod.to_f", [u])) + | Atom (2, j0) => formula_from_atom j0 (to_r u) + | _ => raise NUT ("NitpickKodkod.to_f", [u]) + (* polarity -> nut -> Kodkod.formula *) + and to_f_with_polarity polar u = + case rep_of u of + Formula _ => to_f u + | Atom (2, j0) => formula_from_atom j0 (to_r u) + | Opt (Atom (2, j0)) => formula_from_opt_atom polar j0 (to_r u) + | _ => raise NUT ("NitpickKodkod.to_f_with_polarity", [u]) + (* nut -> Kodkod.rel_expr *) + and to_r u = + case u of + Cst (False, _, Atom _) => false_atom + | Cst (True, _, Atom _) => true_atom + | Cst (Iden, T, Func (Struct [R1, R2], Formula Neut)) => + if R1 = R2 andalso arity_of_rep R1 = 1 then + kk_intersect Kodkod.Iden (kk_product (full_rel_for_rep R1) + Kodkod.Univ) + else + let + val schema1 = atom_schema_of_rep R1 + val schema2 = atom_schema_of_rep R2 + val arity1 = length schema1 + val arity2 = length schema2 + val r1 = fold1 kk_product (unary_var_seq 0 arity1) + val r2 = fold1 kk_product (unary_var_seq arity1 arity2) + val min_R = min_rep R1 R2 + in + kk_comprehension + (decls_for_atom_schema 0 (schema1 @ schema2)) + (kk_rel_eq (rel_expr_from_rel_expr kk min_R R1 r1) + (rel_expr_from_rel_expr kk min_R R2 r2)) + end + | Cst (Iden, T, Func (Atom (1, j0), Formula Neut)) => Kodkod.Atom j0 + | Cst (Iden, T as Type ("fun", [T1, _]), R as Func (R1, _)) => + to_rep R (Cst (Iden, T, Func (one_rep ofs T1 R1, Formula Neut))) + | Cst (Num j, @{typ int}, R) => + (case atom_for_int (card_of_rep R, offset_of_type ofs int_T) j of + ~1 => if is_opt_rep R then Kodkod.None + else raise NUT ("NitpickKodkod.to_r (Num)", [u]) + | j' => Kodkod.Atom j') + | Cst (Num j, T, R) => + if j < card_of_rep R then Kodkod.Atom (j + offset_of_type ofs T) + else if is_opt_rep R then Kodkod.None + else raise NUT ("NitpickKodkod.to_r", [u]) + | Cst (Unknown, _, R) => empty_rel_for_rep R + | Cst (Unrep, _, R) => empty_rel_for_rep R + | Cst (Suc, T, Func (Atom x, _)) => + if domain_type T <> nat_T then suc_rel + else kk_intersect suc_rel (kk_product Kodkod.Univ (Kodkod.AtomSeq x)) + | Cst (Add, Type ("fun", [@{typ nat}, _]), _) => nat_add_rel + | Cst (Add, Type ("fun", [@{typ int}, _]), _) => int_add_rel + | Cst (Subtract, Type ("fun", [@{typ nat}, _]), _) => nat_subtract_rel + | Cst (Subtract, Type ("fun", [@{typ int}, _]), _) => int_subtract_rel + | Cst (Multiply, Type ("fun", [@{typ nat}, _]), _) => nat_multiply_rel + | Cst (Multiply, Type ("fun", [@{typ int}, _]), _) => int_multiply_rel + | Cst (Divide, Type ("fun", [@{typ nat}, _]), _) => nat_divide_rel + | Cst (Divide, Type ("fun", [@{typ int}, _]), _) => int_divide_rel + | Cst (Modulo, Type ("fun", [@{typ nat}, _]), _) => nat_modulo_rel + | Cst (Modulo, Type ("fun", [@{typ int}, _]), _) => int_modulo_rel + | Cst (Gcd, _, _) => gcd_rel + | Cst (Lcm, _, _) => lcm_rel + | Cst (Fracs, _, Func (Atom (1, _), _)) => Kodkod.None + | Cst (Fracs, _, Func (Struct _, _)) => + kk_project_seq norm_frac_rel 2 2 + | Cst (NormFrac, _, _) => norm_frac_rel + | Cst (NatToInt, _, Func (Atom _, Atom _)) => Kodkod.Iden + | Cst (NatToInt, _, + Func (Atom (nat_k, nat_j0), Opt (Atom (int_k, int_j0)))) => + if nat_j0 = int_j0 then + kk_intersect Kodkod.Iden + (kk_product (Kodkod.AtomSeq (max_int_for_card int_k + 1, nat_j0)) + Kodkod.Univ) + else + raise BAD ("NitpickKodkod.to_r (NatToInt)", "\"nat_j0 <> int_j0\"") + | Cst (IntToNat, _, Func (Atom (int_k, int_j0), nat_R)) => + let + val abs_card = max_int_for_card int_k + 1 + val (nat_k, nat_j0) = the_single (atom_schema_of_rep nat_R) + val overlap = Int.min (nat_k, abs_card) + in + if nat_j0 = int_j0 then + kk_union (kk_product (Kodkod.AtomSeq (int_k - abs_card, + int_j0 + abs_card)) + (Kodkod.Atom nat_j0)) + (kk_intersect Kodkod.Iden + (kk_product (Kodkod.AtomSeq (overlap, int_j0)) + Kodkod.Univ)) + else + raise BAD ("NitpickKodkod.to_r (IntToNat)", "\"nat_j0 <> int_j0\"") + end + | Op1 (Not, _, R, u1) => kk_not3 (to_rep R u1) + | Op1 (Finite, _, Opt (Atom _), _) => Kodkod.None + | Op1 (Converse, T, R, u1) => + let + val (b_T, a_T) = HOLogic.dest_prodT (domain_type T) + val (b_R, a_R) = + case R of + Func (Struct [R1, R2], _) => (R1, R2) + | Func (R1, _) => + if card_of_rep R1 <> 1 then + raise REP ("NitpickKodkod.to_r (Converse)", [R]) + else + pairself (Atom o pair 1 o offset_of_type ofs) (b_T, a_T) + | _ => raise REP ("NitpickKodkod.to_r (Converse)", [R]) + val body_R = body_rep R + val a_arity = arity_of_rep a_R + val b_arity = arity_of_rep b_R + val ab_arity = a_arity + b_arity + val body_arity = arity_of_rep body_R + in + kk_project (to_rep (Func (Struct [a_R, b_R], body_R)) u1) + (map Kodkod.Num (index_seq a_arity b_arity @ + index_seq 0 a_arity @ + index_seq ab_arity body_arity)) + |> rel_expr_from_rel_expr kk R (Func (Struct [b_R, a_R], body_R)) + end + | Op1 (Closure, _, R, u1) => + if is_opt_rep R then + let + val T1 = type_of u1 + val R' = rep_to_binary_rel_rep ofs T1 (unopt_rep (rep_of u1)) + val R'' = opt_rep ofs T1 R' + in + single_rel_let kk + (fn r => + let + val true_r = kk_closure (kk_join r true_atom) + val full_r = full_rel_for_rep R' + val false_r = kk_difference full_r + (kk_closure (kk_difference full_r + (kk_join r false_atom))) + in + rel_expr_from_rel_expr kk R R'' + (kk_union (kk_product true_r true_atom) + (kk_product false_r false_atom)) + end) (to_rep R'' u1) + end + else + let val R' = rep_to_binary_rel_rep ofs (type_of u1) (rep_of u1) in + rel_expr_from_rel_expr kk R R' (kk_closure (to_rep R' u1)) + end + | Op1 (SingletonSet, _, Func (R1, Opt _), Cst (Unrep, _, _)) => + (if R1 = Unit then I else kk_product (full_rel_for_rep R1)) false_atom + | Op1 (SingletonSet, _, R, u1) => + (case R of + Func (R1, Formula Neut) => to_rep R1 u1 + | Func (Unit, Opt R) => to_guard [u1] R true_atom + | Func (R1, R2 as Opt _) => + single_rel_let kk + (fn r => kk_rel_if (kk_no r) (empty_rel_for_rep R) + (rel_expr_to_func kk R1 bool_atom_R + (Func (R1, Formula Neut)) r)) + (to_opt R1 u1) + | _ => raise NUT ("NitpickKodkod.to_r (SingletonSet)", [u])) + | Op1 (Tha, T, R, u1) => + if is_opt_rep R then + kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom + else + to_rep (Func (R, Formula Neut)) u1 + | Op1 (First, T, R, u1) => to_nth_pair_sel 0 T R u1 + | Op1 (Second, T, R, u1) => to_nth_pair_sel 1 T R u1 + | Op1 (Cast, _, R, u1) => + ((case rep_of u1 of + Formula _ => + (case unopt_rep R of + Atom (2, j0) => atom_from_formula kk j0 (to_f u1) + | _ => raise SAME ()) + | _ => raise SAME ()) + handle SAME () => rel_expr_from_rel_expr kk R (rep_of u1) (to_r u1)) + | Op2 (All, T, R as Opt _, u1, u2) => + to_r (Op1 (Not, T, R, + Op2 (Exist, T, R, u1, Op1 (Not, T, rep_of u2, u2)))) + | Op2 (Exist, T, Opt _, u1, u2) => + let val rs1 = untuple to_decl u1 in + if not (is_opt_rep (rep_of u2)) then + kk_rel_if (kk_exist rs1 (to_f u2)) true_atom Kodkod.None + else + let val r2 = to_r u2 in + kk_union (kk_rel_if (kk_exist rs1 (kk_rel_eq r2 true_atom)) + true_atom Kodkod.None) + (kk_rel_if (kk_all rs1 (kk_rel_eq r2 false_atom)) + false_atom Kodkod.None) + end + end + | Op2 (Or, _, _, u1, u2) => + if is_opt_rep (rep_of u1) then kk_rel_if (to_f u2) true_atom (to_r u1) + else kk_rel_if (to_f u1) true_atom (to_r u2) + | Op2 (And, _, _, u1, u2) => + if is_opt_rep (rep_of u1) then kk_rel_if (to_f u2) (to_r u1) false_atom + else kk_rel_if (to_f u1) (to_r u2) false_atom + | Op2 (Less, _, _, u1, u2) => + if type_of u1 = nat_T then + if is_Cst Unrep u1 then to_compare_with_unrep u2 false_atom + else if is_Cst Unrep u2 then to_compare_with_unrep u1 true_atom + else kk_nat_less (to_integer u1) (to_integer u2) + else + kk_int_less (to_integer u1) (to_integer u2) + | Op2 (The, T, R, u1, u2) => + if is_opt_rep R then + let val r1 = to_opt (Func (unopt_rep R, bool_atom_R)) u1 in + kk_rel_if (kk_one (kk_join r1 true_atom)) (kk_join r1 true_atom) + (kk_rel_if (kk_or (kk_some (kk_join r1 true_atom)) + (kk_subset (full_rel_for_rep R) + (kk_join r1 false_atom))) + (to_rep R u2) Kodkod.None) + end + else + let val r1 = to_rep (Func (R, Formula Neut)) u1 in + kk_rel_if (kk_one r1) r1 (to_rep R u2) + end + | Op2 (Eps, T, R, u1, u2) => + if is_opt_rep (rep_of u1) then + let + val r1 = to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1 + val r2 = to_rep R u2 + in + kk_union (kk_rel_if (kk_one (kk_join r1 true_atom)) + (kk_join r1 true_atom) Kodkod.None) + (kk_rel_if (kk_or (kk_subset r2 (kk_join r1 true_atom)) + (kk_subset (full_rel_for_rep R) + (kk_join r1 false_atom))) + r2 Kodkod.None) + end + else + let + val r1 = to_rep (Func (unopt_rep R, Formula Neut)) u1 + val r2 = to_rep R u2 + in + kk_union (kk_rel_if (kk_one r1) r1 Kodkod.None) + (kk_rel_if (kk_or (kk_no r1) (kk_subset r2 r1)) + r2 Kodkod.None) + end + | Op2 (Triad, T, Opt (Atom (2, j0)), u1, u2) => + let + val f1 = to_f u1 + val f2 = to_f u2 + in + if f1 = f2 then + atom_from_formula kk j0 f1 + else + kk_union (kk_rel_if f1 true_atom Kodkod.None) + (kk_rel_if f2 Kodkod.None false_atom) + end + | Op2 (Union, _, R, u1, u2) => + to_set_op kk_or kk_or3 kk_union kk_union kk_intersect false R u1 u2 + | Op2 (SetDifference, _, R, u1, u2) => + to_set_op kk_notimplies kk_notimplies3 kk_difference kk_intersect + kk_union true R u1 u2 + | Op2 (Intersect, _, R, u1, u2) => + to_set_op kk_and kk_and3 kk_intersect kk_intersect kk_union false R + u1 u2 + | Op2 (Composition, _, R, u1, u2) => + let + val (a_T, b_T) = HOLogic.dest_prodT (domain_type (type_of u2)) + val (_, c_T) = HOLogic.dest_prodT (domain_type (type_of u1)) + val ab_k = card_of_domain_from_rep 2 (rep_of u2) + val bc_k = card_of_domain_from_rep 2 (rep_of u1) + val ac_k = card_of_domain_from_rep 2 R + val a_k = exact_root 2 (ac_k * ab_k div bc_k) + val b_k = exact_root 2 (ab_k * bc_k div ac_k) + val c_k = exact_root 2 (bc_k * ac_k div ab_k) + val a_R = Atom (a_k, offset_of_type ofs a_T) + val b_R = Atom (b_k, offset_of_type ofs b_T) + val c_R = Atom (c_k, offset_of_type ofs c_T) + val body_R = body_rep R + in + (case body_R of + Formula Neut => + kk_join (to_rep (Func (Struct [a_R, b_R], Formula Neut)) u2) + (to_rep (Func (Struct [b_R, c_R], Formula Neut)) u1) + | Opt (Atom (2, _)) => + let + (* Kodkod.rel_expr -> rep -> rep -> nut -> Kodkod.rel_expr *) + fun do_nut r R1 R2 u = + kk_join (to_rep (Func (Struct [R1, R2], body_R)) u) r + (* Kodkod.rel_expr -> Kodkod.rel_expr *) + fun do_term r = + kk_product (kk_join (do_nut r a_R b_R u2) + (do_nut r b_R c_R u1)) r + in kk_union (do_term true_atom) (do_term false_atom) end + | _ => raise NUT ("NitpickKodkod.to_r (Composition)", [u])) + |> rel_expr_from_rel_expr kk R (Func (Struct [a_R, c_R], body_R)) + end + | Op2 (Product, T, R, u1, u2) => + let + val (a_T, b_T) = HOLogic.dest_prodT (domain_type T) + val a_k = card_of_domain_from_rep 2 (rep_of u1) + val b_k = card_of_domain_from_rep 2 (rep_of u2) + val a_R = Atom (a_k, offset_of_type ofs a_T) + val b_R = Atom (b_k, offset_of_type ofs b_T) + val body_R = body_rep R + in + (case body_R of + Formula Neut => + kk_product (to_rep (Func (a_R, Formula Neut)) u1) + (to_rep (Func (b_R, Formula Neut)) u2) + | Opt (Atom (2, _)) => + let + (* Kodkod.rel_expr -> rep -> nut -> Kodkod.rel_expr *) + fun do_nut r R u = kk_join (to_rep (Func (R, body_R)) u) r + (* Kodkod.rel_expr -> Kodkod.rel_expr *) + fun do_term r = + kk_product (kk_product (do_nut r a_R u1) (do_nut r b_R u2)) r + in kk_union (do_term true_atom) (do_term false_atom) end + | _ => raise NUT ("NitpickKodkod.to_r (Product)", [u])) + |> rel_expr_from_rel_expr kk R (Func (Struct [a_R, b_R], body_R)) + end + | Op2 (Image, T, R, u1, u2) => + (case (rep_of u1, rep_of u2) of + (Func (R11, R12), Func (R21, Formula Neut)) => + if R21 = R11 andalso is_lone_rep R12 then + let + (* Kodkod.rel_expr -> Kodkod.rel_expr *) + fun big_join r = kk_n_fold_join kk false R21 R12 r (to_r u1) + val core_r = big_join (to_r u2) + val core_R = Func (R12, Formula Neut) + in + if is_opt_rep R12 then + let + val schema = atom_schema_of_rep R21 + val decls = decls_for_atom_schema ~1 schema + val vars = unary_var_seq ~1 (length decls) + val f = kk_some (big_join (fold1 kk_product vars)) + in + kk_rel_if (kk_all decls f) + (rel_expr_from_rel_expr kk R core_R core_r) + (rel_expr_from_rel_expr kk R (opt_rep ofs T core_R) + (kk_product core_r true_atom)) + end + else + rel_expr_from_rel_expr kk R core_R core_r + end + else + raise NUT ("NitpickKodkod.to_r (Image)", [u1, u2]) + | _ => raise NUT ("NitpickKodkod.to_r (Image)", [u1, u2])) + | Op2 (Apply, @{typ nat}, _, + Op2 (Apply, _, _, Cst (Subtract, _, _), u1), u2) => + if is_Cst Unrep u2 andalso not (is_opt_rep (rep_of u1)) then + Kodkod.Atom (offset_of_type ofs nat_T) + else + fold kk_join [to_integer u1, to_integer u2] nat_subtract_rel + | Op2 (Apply, _, R, u1, u2) => + if is_Cst Unrep u2 andalso is_set_type (type_of u1) + andalso not (is_opt_rep (rep_of u1)) then + false_atom + else + to_apply R u1 u2 + | Op2 (Lambda, T, R as Opt (Atom (1, j0)), u1, u2) => + to_guard [u1, u2] R (Kodkod.Atom j0) + | Op2 (Lambda, T, Func (_, Formula Neut), u1, u2) => + kk_comprehension (untuple to_decl u1) (to_f u2) + | Op2 (Lambda, T, Func (_, R2), u1, u2) => + let + val dom_decls = untuple to_decl u1 + val ran_schema = atom_schema_of_rep R2 + val ran_decls = decls_for_atom_schema ~1 ran_schema + val ran_vars = unary_var_seq ~1 (length ran_decls) + in + kk_comprehension (dom_decls @ ran_decls) + (kk_subset (fold1 kk_product ran_vars) + (to_rep R2 u2)) + end + | Op3 (Let, _, R, u1, u2, u3) => + kk_rel_let [to_expr_assign u1 u2] (to_rep R u3) + | Op3 (If, _, R, u1, u2, u3) => + if is_opt_rep (rep_of u1) then + triple_rel_let kk + (fn r1 => fn r2 => fn r3 => + let val empty_r = empty_rel_for_rep R in + fold1 kk_union + [kk_rel_if (kk_rel_eq r1 true_atom) r2 empty_r, + kk_rel_if (kk_rel_eq r1 false_atom) r3 empty_r, + kk_rel_if (kk_rel_eq r2 r3) + (if inline_rel_expr r2 then r2 else r3) empty_r] + end) + (to_r u1) (to_rep R u2) (to_rep R u3) + else + kk_rel_if (to_f u1) (to_rep R u2) (to_rep R u3) + | Tuple (_, R, us) => + (case unopt_rep R of + Struct Rs => to_product Rs us + | Vect (k, R) => to_product (replicate k R) us + | Atom (1, j0) => + (case filter (not_equal Unit o rep_of) us of + [] => Kodkod.Atom j0 + | us' => + kk_rel_if (kk_some (fold1 kk_product (map to_r us'))) + (Kodkod.Atom j0) Kodkod.None) + | _ => raise NUT ("NitpickKodkod.to_r (Tuple)", [u])) + | Construct ([u'], _, _, []) => to_r u' + | Construct (_ :: sel_us, T, R, arg_us) => + let + val set_rs = + map2 (fn sel_u => fn arg_u => + let + val (R1, R2) = dest_Func (rep_of sel_u) + val sel_r = to_r sel_u + val arg_r = to_opt R2 arg_u + in + if is_one_rep R2 then + kk_n_fold_join kk true R2 R1 arg_r + (kk_project sel_r (flip_nums (arity_of_rep R2))) + else + kk_comprehension + (decls_for_atom_schema ~1 (atom_schema_of_rep R1)) + (kk_rel_eq (kk_join (Kodkod.Var (1, ~1)) sel_r) + arg_r) + end) sel_us arg_us + in fold1 kk_intersect set_rs end + | BoundRel (x, _, _, _) => Kodkod.Var x + | FreeRel (x, _, _, _) => Kodkod.Rel x + | RelReg (j, _, R) => Kodkod.RelReg (arity_of_rep R, j) + | u => raise NUT ("NitpickKodkod.to_r", [u]) + (* nut -> Kodkod.decl *) + and to_decl (BoundRel (x, _, R, _)) = + Kodkod.DeclOne (x, Kodkod.AtomSeq (the_single (atom_schema_of_rep R))) + | to_decl u = raise NUT ("NitpickKodkod.to_decl", [u]) + (* nut -> Kodkod.expr_assign *) + and to_expr_assign (FormulaReg (j, _, R)) u = + Kodkod.AssignFormulaReg (j, to_f u) + | to_expr_assign (RelReg (j, _, R)) u = + Kodkod.AssignRelReg ((arity_of_rep R, j), to_r u) + | to_expr_assign u1 _ = raise NUT ("NitpickKodkod.to_expr_assign", [u1]) + (* int * int -> nut -> Kodkod.rel_expr *) + and to_atom (x as (k, j0)) u = + case rep_of u of + Formula _ => atom_from_formula kk j0 (to_f u) + | Unit => if k = 1 then Kodkod.Atom j0 + else raise NUT ("NitpickKodkod.to_atom", [u]) + | R => atom_from_rel_expr kk x R (to_r u) + (* rep list -> nut -> Kodkod.rel_expr *) + and to_struct Rs u = + case rep_of u of + Unit => full_rel_for_rep (Struct Rs) + | R' => struct_from_rel_expr kk Rs R' (to_r u) + (* int -> rep -> nut -> Kodkod.rel_expr *) + and to_vect k R u = + case rep_of u of + Unit => full_rel_for_rep (Vect (k, R)) + | R' => vect_from_rel_expr kk k R R' (to_r u) + (* rep -> rep -> nut -> Kodkod.rel_expr *) + and to_func R1 R2 u = + case rep_of u of + Unit => full_rel_for_rep (Func (R1, R2)) + | R' => rel_expr_to_func kk R1 R2 R' (to_r u) + (* rep -> nut -> Kodkod.rel_expr *) + and to_opt R u = + let val old_R = rep_of u in + if is_opt_rep old_R then + rel_expr_from_rel_expr kk (Opt R) old_R (to_r u) + else + to_rep R u + end + (* rep -> nut -> Kodkod.rel_expr *) + and to_rep (Atom x) u = to_atom x u + | to_rep (Struct Rs) u = to_struct Rs u + | to_rep (Vect (k, R)) u = to_vect k R u + | to_rep (Func (R1, R2)) u = to_func R1 R2 u + | to_rep (Opt R) u = to_opt R u + | to_rep R _ = raise REP ("NitpickKodkod.to_rep", [R]) + (* nut -> Kodkod.rel_expr *) + and to_integer u = to_opt (one_rep ofs (type_of u) (rep_of u)) u + (* nut list -> rep -> Kodkod.rel_expr -> Kodkod.rel_expr *) + and to_guard guard_us R r = + let + val unpacked_rs = unpack_joins r + val plain_guard_rs = + map to_r (filter (is_Opt o rep_of) guard_us) + |> filter_out (member (op =) unpacked_rs) + val func_guard_us = + filter ((is_Func andf is_opt_rep) o rep_of) guard_us + val func_guard_rs = map to_r func_guard_us + val guard_fs = + map kk_no plain_guard_rs @ + map2 (kk_not oo kk_n_ary_function kk) + (map (unopt_rep o rep_of) func_guard_us) func_guard_rs + in + if null guard_fs then + r + else + kk_rel_if (fold1 kk_or guard_fs) (empty_rel_for_rep R) r + end + (* rep -> rep -> Kodkod.rel_expr -> int -> Kodkod.rel_expr *) + and to_project new_R old_R r j0 = + rel_expr_from_rel_expr kk new_R old_R + (kk_project_seq r j0 (arity_of_rep old_R)) + (* rep list -> nut list -> Kodkod.rel_expr *) + and to_product Rs us = + case map (uncurry to_opt) (filter (not_equal Unit o fst) (Rs ~~ us)) of + [] => raise REP ("NitpickKodkod.to_product", Rs) + | rs => fold1 kk_product rs + (* int -> typ -> rep -> nut -> Kodkod.rel_expr *) + and to_nth_pair_sel n res_T res_R u = + case u of + Tuple (_, _, us) => to_rep res_R (nth us n) + | _ => let + val R = rep_of u + val (a_T, b_T) = HOLogic.dest_prodT (type_of u) + val Rs = + case unopt_rep R of + Struct (Rs as [_, _]) => Rs + | _ => + let + val res_card = card_of_rep res_R + val other_card = card_of_rep R div res_card + val (a_card, b_card) = (res_card, other_card) + |> n = 1 ? swap + in + [Atom (a_card, offset_of_type ofs a_T), + Atom (b_card, offset_of_type ofs b_T)] + end + val nth_R = nth Rs n + val j0 = if n = 0 then 0 else arity_of_rep (hd Rs) + in + case arity_of_rep nth_R of + 0 => to_guard [u] res_R + (to_rep res_R (Cst (Unity, res_T, Unit))) + | arity => to_project res_R nth_R (to_rep (Opt (Struct Rs)) u) j0 + end + (* (Kodkod.formula -> Kodkod.formula -> Kodkod.formula) + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.formula) -> nut -> nut + -> Kodkod.formula *) + and to_set_bool_op connective set_oper u1 u2 = + let + val min_R = min_rep (rep_of u1) (rep_of u2) + val r1 = to_rep min_R u1 + val r2 = to_rep min_R u2 + in + case min_R of + Vect (k, Atom _) => kk_vect_set_bool_op connective k r1 r2 + | Func (R1, Formula Neut) => set_oper r1 r2 + | Func (Unit, Atom (2, j0)) => + connective (formula_from_atom j0 r1) (formula_from_atom j0 r2) + | Func (R1, Atom _) => set_oper (kk_join r1 true_atom) + (kk_join r2 true_atom) + | _ => raise REP ("NitpickKodkod.to_set_bool_op", [min_R]) + end + (* (Kodkod.formula -> Kodkod.formula -> Kodkod.formula) + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.rel_expr) + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.formula) + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.formula) + -> (Kodkod.rel_expr -> Kodkod.rel_expr -> Kodkod.formula) -> bool -> rep + -> nut -> nut -> Kodkod.rel_expr *) + and to_set_op connective connective3 set_oper true_set_oper false_set_oper + neg_second R u1 u2 = + let + val min_R = min_rep (rep_of u1) (rep_of u2) + val r1 = to_rep min_R u1 + val r2 = to_rep min_R u2 + val unopt_R = unopt_rep R + in + rel_expr_from_rel_expr kk unopt_R (unopt_rep min_R) + (case min_R of + Opt (Vect (k, Atom _)) => kk_vect_set_op connective k r1 r2 + | Vect (k, Atom _) => kk_vect_set_op connective k r1 r2 + | Func (_, Formula Neut) => set_oper r1 r2 + | Func (Unit, _) => connective3 r1 r2 + | Func (R1, _) => + double_rel_let kk + (fn r1 => fn r2 => + kk_union + (kk_product + (true_set_oper (kk_join r1 true_atom) + (kk_join r2 (atom_for_bool bool_j0 + (not neg_second)))) + true_atom) + (kk_product + (false_set_oper (kk_join r1 false_atom) + (kk_join r2 (atom_for_bool bool_j0 + neg_second))) + false_atom)) + r1 r2 + | _ => raise REP ("NitpickKodkod.to_set_op", [min_R])) + end + (* rep -> rep -> Kodkod.rel_expr -> nut -> Kodkod.rel_expr *) + and to_apply res_R func_u arg_u = + case unopt_rep (rep_of func_u) of + Unit => + let val j0 = offset_of_type ofs (type_of func_u) in + to_guard [arg_u] res_R + (rel_expr_from_rel_expr kk res_R (Atom (1, j0)) + (Kodkod.Atom j0)) + end + | Atom (1, j0) => + to_guard [arg_u] res_R + (rel_expr_from_rel_expr kk res_R (Atom (1, j0)) (to_r func_u)) + | Atom (k, j0) => + let + val dom_card = card_of_rep (rep_of arg_u) + val ran_R = Atom (exact_root dom_card k, + offset_of_type ofs (range_type (type_of func_u))) + in + to_apply_vect dom_card ran_R res_R (to_vect dom_card ran_R func_u) + arg_u + end + | Vect (1, R') => + to_guard [arg_u] res_R + (rel_expr_from_rel_expr kk res_R R' (to_r func_u)) + | Vect (k, R') => to_apply_vect k R' res_R (to_r func_u) arg_u + | Func (R, Formula Neut) => + to_guard [arg_u] res_R (rel_expr_from_formula kk res_R + (kk_subset (to_opt R arg_u) (to_r func_u))) + | Func (Unit, R2) => + to_guard [arg_u] res_R + (rel_expr_from_rel_expr kk res_R R2 (to_r func_u)) + | Func (R1, R2) => + rel_expr_from_rel_expr kk res_R R2 + (kk_n_fold_join kk true R1 R2 (to_opt R1 arg_u) (to_r func_u)) + |> body_rep R2 = Formula Neut ? to_guard [arg_u] res_R + | _ => raise NUT ("NitpickKodkod.to_apply", [func_u]) + (* int -> rep -> rep -> Kodkod.rel_expr -> nut *) + and to_apply_vect k R' res_R func_r arg_u = + let + val arg_R = one_rep ofs (type_of arg_u) (unopt_rep (rep_of arg_u)) + val vect_r = vect_from_rel_expr kk k res_R (Vect (k, R')) func_r + val vect_rs = unpack_vect_in_chunks kk (arity_of_rep res_R) k vect_r + in + kk_case_switch kk arg_R res_R (to_opt arg_R arg_u) + (all_singletons_for_rep arg_R) vect_rs + end + (* bool -> nut -> Kodkod.formula *) + and to_could_be_unrep neg u = + if neg andalso is_opt_rep (rep_of u) then kk_no (to_r u) + else Kodkod.False + (* nut -> Kodkod.rel_expr -> Kodkod.rel_expr *) + and to_compare_with_unrep u r = + if is_opt_rep (rep_of u) then kk_rel_if (kk_some (to_r u)) r Kodkod.None + else r + in to_f_with_polarity Pos u end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_model.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,703 @@ +(* Title: HOL/Nitpick/Tools/nitpick_model.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Model reconstruction for Nitpick. +*) + +signature NITPICK_MODEL = +sig + type scope = NitpickScope.scope + type rep = NitpickRep.rep + type nut = NitpickNut.nut + + type params = { + show_skolems: bool, + show_datatypes: bool, + show_consts: bool} + + structure NameTable : TABLE + + val tuple_list_for_name : + nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list + val reconstruct_hol_model : + params -> scope -> (term option * int list) list -> styp list -> nut list + -> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list + -> Pretty.T * bool + val prove_hol_model : + scope -> Time.time option -> nut list -> nut list -> nut NameTable.table + -> Kodkod.raw_bound list -> term -> bool option +end; + +structure NitpickModel : NITPICK_MODEL = +struct + +open NitpickUtil +open NitpickHOL +open NitpickScope +open NitpickPeephole +open NitpickRep +open NitpickNut + +type params = { + show_skolems: bool, + show_datatypes: bool, + show_consts: bool} + +val unknown = "?" +val unrep = "\" +val maybe_mixfix = "_\<^sup>?" +val base_mixfix = "_\<^bsub>base\<^esub>" +val step_mixfix = "_\<^bsub>step\<^esub>" +val abs_mixfix = "\_\" +val non_opt_name = nitpick_prefix ^ "non_opt" + +(* string -> typ -> int -> string *) +fun atom_name prefix (T as Type (s, _)) j = + prefix ^ substring (short_name s, 0, 1) ^ nat_subscript (j + 1) + | atom_name prefix (T as TFree (s, _)) j = + prefix ^ perhaps (try (unprefix "'")) s ^ nat_subscript (j + 1) + | atom_name _ T _ = raise TYPE ("NitpickModel.atom_name", [T], []) +(* bool -> typ -> int -> term *) +fun atom for_auto T j = + if for_auto then + Free (atom_name (hd (space_explode "." nitpick_prefix)) T j, T) + else + Const (atom_name "" T j, T) + +(* nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list *) +fun tuple_list_for_name rel_table bounds name = + the (AList.lookup (op =) bounds (the_rel rel_table name)) handle NUT _ => [[]] + +(* term -> term *) +fun unbox_term (Const (@{const_name FunBox}, _) $ t1) = unbox_term t1 + | unbox_term (Const (@{const_name PairBox}, + Type ("fun", [T1, Type ("fun", [T2, T3])])) $ t1 $ t2) = + let val Ts = map unbox_type [T1, T2] in + Const (@{const_name Pair}, Ts ---> Type ("*", Ts)) + $ unbox_term t1 $ unbox_term t2 + end + | unbox_term (Const (s, T)) = Const (s, unbox_type T) + | unbox_term (t1 $ t2) = unbox_term t1 $ unbox_term t2 + | unbox_term (Free (s, T)) = Free (s, unbox_type T) + | unbox_term (Var (x, T)) = Var (x, unbox_type T) + | unbox_term (Bound j) = Bound j + | unbox_term (Abs (s, T, t')) = Abs (s, unbox_type T, unbox_term t') + +(* typ -> typ -> (typ * typ) * (typ * typ) *) +fun factor_out_types (T1 as Type ("*", [T11, T12])) + (T2 as Type ("*", [T21, T22])) = + let val (n1, n2) = pairself num_factors_in_type (T11, T21) in + if n1 = n2 then + let + val ((T11', opt_T12'), (T21', opt_T22')) = factor_out_types T12 T22 + in + ((Type ("*", [T11, T11']), opt_T12'), + (Type ("*", [T21, T21']), opt_T22')) + end + else if n1 < n2 then + case factor_out_types T1 T21 of + (p1, (T21', NONE)) => (p1, (T21', SOME T22)) + | (p1, (T21', SOME T22')) => + (p1, (T21', SOME (Type ("*", [T22', T22])))) + else + swap (factor_out_types T2 T1) + end + | factor_out_types (Type ("*", [T11, T12])) T2 = ((T11, SOME T12), (T2, NONE)) + | factor_out_types T1 (Type ("*", [T21, T22])) = ((T1, NONE), (T21, SOME T22)) + | factor_out_types T1 T2 = ((T1, NONE), (T2, NONE)) + +(* bool -> typ -> typ -> (term * term) list -> term *) +fun make_plain_fun maybe_opt T1 T2 = + let + (* typ -> typ -> (term * term) list -> term *) + fun aux T1 T2 [] = + Const (if maybe_opt orelse T2 <> bool_T then @{const_name undefined} + else non_opt_name, T1 --> T2) + | aux T1 T2 ((t1, t2) :: ps) = + Const (@{const_name fun_upd}, [T1 --> T2, T1, T2] ---> T1 --> T2) + $ aux T1 T2 ps $ t1 $ t2 + in aux T1 T2 o rev end +(* term -> bool *) +fun is_plain_fun (Const (s, _)) = s mem [@{const_name undefined}, non_opt_name] + | is_plain_fun (Const (@{const_name fun_upd}, _) $ t0 $ _ $ _) = + is_plain_fun t0 + | is_plain_fun _ = false +(* term -> bool * (term list * term list) *) +val dest_plain_fun = + let + (* term -> term list * term list *) + fun aux (Const (s, _)) = (s <> non_opt_name, ([], [])) + | aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = + let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end + | aux t = raise TERM ("NitpickModel.dest_plain_fun", [t]) + in apsnd (pairself rev) o aux end + +(* typ -> term -> term list * term *) +fun break_in_two (Type ("*", [T1, T2])) + (Const (@{const_name Pair}, _) $ t1 $ t2) = + break_in_two T2 t2 |>> cons t1 + | break_in_two _ (Const (@{const_name Pair}, _) $ t1 $ t2) = ([t1], t2) + | break_in_two _ t = raise TERM ("NitpickModel.break_in_two", [t]) +(* typ -> term -> term -> term *) +fun pair_up (Type ("*", [T1', T2'])) + (t1 as Const (@{const_name Pair}, + Type ("fun", [_, Type ("fun", [_, T1])])) $ t11 $ t12) + t2 = + if T1 = T1' then HOLogic.mk_prod (t1, t2) + else HOLogic.mk_prod (t11, pair_up T2' t12 t2) + | pair_up _ t1 t2 = HOLogic.mk_prod (t1, t2) +(* typ -> term -> term list * term list -> (term * term) list*) +fun multi_pair_up T1 t1 (ts2, ts3) = map2 (pair o pair_up T1 t1) ts2 ts3 + +(* typ -> typ -> typ -> term -> term *) +fun typecast_fun (Type ("fun", [T1', T2'])) T1 T2 t = + let + (* typ -> typ -> typ -> term -> term *) + fun do_curry T1a T1b T2 t = + let + val (maybe_opt, ps) = dest_plain_fun t + val ps = + ps |>> map (break_in_two T1a #>> mk_flat_tuple T1a) + |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2)))) + |> AList.coalesce (op =) + |> map (apsnd (make_plain_fun maybe_opt T1b T2)) + in make_plain_fun maybe_opt T1a (T1b --> T2) ps end + (* typ -> typ -> term -> term *) + and do_uncurry T1 T2 t = + let + val (maybe_opt, tsp) = dest_plain_fun t + val ps = + tsp |> op ~~ + |> maps (fn (t1, t2) => + multi_pair_up T1 t1 (snd (dest_plain_fun t2))) + in make_plain_fun maybe_opt T1 T2 ps end + (* typ -> typ -> typ -> typ -> term -> term *) + and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2') + | do_arrow T1' T2' T1 T2 + (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = + Const (@{const_name fun_upd}, + [T1' --> T2', T1', T2'] ---> T1' --> T2') + $ do_arrow T1' T2' T1 T2 t0 $ do_term T1' T1 t1 $ do_term T2' T2 t2 + | do_arrow _ _ _ _ t = + raise TERM ("NitpickModel.typecast_fun.do_arrow", [t]) + and do_fun T1' T2' T1 T2 t = + case factor_out_types T1' T1 of + ((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2 + | ((_, NONE), (T1a, SOME T1b)) => + t |> do_curry T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2) + | ((T1a', SOME T1b'), (_, NONE)) => + t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2' + | _ => raise TYPE ("NitpickModel.typecast_fun.do_fun", [T1, T1'], []) + (* typ -> typ -> term -> term *) + and do_term (Type ("fun", [T1', T2'])) (Type ("fun", [T1, T2])) t = + do_fun T1' T2' T1 T2 t + | do_term (T' as Type ("*", Ts' as [T1', T2'])) (Type ("*", [T1, T2])) + (Const (@{const_name Pair}, _) $ t1 $ t2) = + Const (@{const_name Pair}, Ts' ---> T') + $ do_term T1' T1 t1 $ do_term T2' T2 t2 + | do_term T' T t = + if T = T' then t + else raise TYPE ("NitpickModel.typecast_fun.do_term", [T, T'], []) + in if T1' = T1 andalso T2' = T2 then t else do_fun T1' T2' T1 T2 t end + | typecast_fun T' _ _ _ = raise TYPE ("NitpickModel.typecast_fun", [T'], []) + +(* term -> string *) +fun truth_const_sort_key @{const True} = "0" + | truth_const_sort_key @{const False} = "2" + | truth_const_sort_key _ = "1" + +(* typ -> term list -> term *) +fun mk_tuple (Type ("*", [T1, T2])) ts = + HOLogic.mk_prod (mk_tuple T1 ts, + mk_tuple T2 (List.drop (ts, length (HOLogic.flatten_tupleT T1)))) + | mk_tuple _ (t :: _) = t + +(* string * string * string * string -> scope -> nut list -> nut list + -> nut list -> nut NameTable.table -> Kodkod.raw_bound list -> typ -> typ + -> rep -> int list list -> term *) +fun reconstruct_term (maybe_name, base_name, step_name, abs_name) + ({ext_ctxt as {thy, ctxt, ...}, card_assigns, datatypes, ofs, ...} + : scope) sel_names rel_table bounds = + let + val for_auto = (maybe_name = "") + (* bool -> typ -> typ -> (term * term) list -> term *) + fun make_set maybe_opt T1 T2 = + let + val empty_const = Const (@{const_name Set.empty}, T1 --> T2) + val insert_const = Const (@{const_name insert}, + [T1, T1 --> T2] ---> T1 --> T2) + (* (term * term) list -> term *) + fun aux [] = + if maybe_opt andalso not (is_precise_type datatypes T1) then + insert_const $ Const (unrep, T1) $ empty_const + else + empty_const + | aux ((t1, t2) :: zs) = + aux zs |> t2 <> @{const False} + ? curry (op $) (insert_const + $ (t1 |> t2 <> @{const True} + ? curry (op $) + (Const (maybe_name, + T1 --> T1)))) + in aux end + (* typ -> typ -> typ -> (term * term) list -> term *) + fun make_map T1 T2 T2' = + let + val update_const = Const (@{const_name fun_upd}, + [T1 --> T2, T1, T2] ---> T1 --> T2) + (* (term * term) list -> term *) + fun aux' [] = Const (@{const_name Map.empty}, T1 --> T2) + | aux' ((t1, t2) :: ps) = + (case t2 of + Const (@{const_name None}, _) => aux' ps + | _ => update_const $ aux' ps $ t1 $ t2) + fun aux ps = + if not (is_precise_type datatypes T1) then + update_const $ aux' ps $ Const (unrep, T1) + $ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2')) + else + aux' ps + in aux end + (* typ list -> term -> term *) + fun setify_mapify_funs Ts t = + (case fastype_of1 (Ts, t) of + Type ("fun", [T1, T2]) => + if is_plain_fun t then + case T2 of + @{typ bool} => + let + val (maybe_opt, ts_pair) = + dest_plain_fun t ||> pairself (map (setify_mapify_funs Ts)) + in + make_set maybe_opt T1 T2 + (sort_wrt (truth_const_sort_key o snd) (op ~~ ts_pair)) + end + | Type (@{type_name option}, [T2']) => + let + val ts_pair = snd (dest_plain_fun t) + |> pairself (map (setify_mapify_funs Ts)) + in make_map T1 T2 T2' (rev (op ~~ ts_pair)) end + | _ => raise SAME () + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => + case t of + t1 $ t2 => setify_mapify_funs Ts t1 $ setify_mapify_funs Ts t2 + | Abs (s, T, t') => Abs (s, T, setify_mapify_funs (T :: Ts) t') + | _ => t + (* bool -> typ -> typ -> typ -> term list -> term list -> term *) + fun make_fun maybe_opt T1 T2 T' ts1 ts2 = + ts1 ~~ ts2 |> T1 = @{typ bisim_iterator} ? rev + |> make_plain_fun (maybe_opt andalso not for_auto) T1 T2 + |> unbox_term + |> typecast_fun (unbox_type T') (unbox_type T1) (unbox_type T2) + (* (typ * int) list -> typ -> typ -> int -> term *) + fun term_for_atom seen (T as Type ("fun", [T1, T2])) T' j = + let + val k1 = card_of_type card_assigns T1 + val k2 = card_of_type card_assigns T2 + in + term_for_rep seen T T' (Vect (k1, Atom (k2, 0))) + [nth_combination (replicate k1 (k2, 0)) j] + handle General.Subscript => + raise ARG ("NitpickModel.reconstruct_term.term_for_atom", + signed_string_of_int j ^ " for " ^ + string_for_rep (Vect (k1, Atom (k2, 0)))) + end + | term_for_atom seen (Type ("*", [T1, T2])) _ j = + let val k1 = card_of_type card_assigns T1 in + list_comb (HOLogic.pair_const T1 T2, + map2 (fn T => term_for_atom seen T T) [T1, T2] + [j div k1, j mod k1]) + end + | term_for_atom seen @{typ prop} _ j = + HOLogic.mk_Trueprop (term_for_atom seen bool_T bool_T j) + | term_for_atom _ @{typ bool} _ j = + if j = 0 then @{const False} else @{const True} + | term_for_atom _ @{typ unit} _ _ = @{const Unity} + | term_for_atom seen T _ j = + if T = nat_T then + HOLogic.mk_number nat_T j + else if T = int_T then + HOLogic.mk_number int_T + (int_for_atom (card_of_type card_assigns int_T, 0) j) + else if is_fp_iterator_type T then + HOLogic.mk_number nat_T (card_of_type card_assigns T - j - 1) + else if T = @{typ bisim_iterator} then + HOLogic.mk_number nat_T j + else case datatype_spec datatypes T of + NONE => atom for_auto T j + | SOME {constrs, co, ...} => + let + (* styp -> int list *) + fun tuples_for_const (s, T) = + tuple_list_for_name rel_table bounds (ConstName (s, T, Any)) + (* unit -> indexname * typ *) + fun var () = ((atom_name "" T j, 0), T) + val discr_jsss = map (tuples_for_const o discr_for_constr o #const) + constrs + val real_j = j + offset_of_type ofs T + val constr_x as (constr_s, constr_T) = + get_first (fn (jss, {const, ...}) => + if [real_j] mem jss then SOME const else NONE) + (discr_jsss ~~ constrs) |> the + val arg_Ts = curried_binder_types constr_T + val sel_xs = map (boxed_nth_sel_for_constr ext_ctxt constr_x) + (index_seq 0 (length arg_Ts)) + val sel_Rs = + map (fn x => get_first + (fn ConstName (s', T', R) => + if (s', T') = x then SOME R else NONE + | u => raise NUT ("NitpickModel.reconstruct_\ + \term.term_for_atom", [u])) + sel_names |> the) sel_xs + val arg_Rs = map (snd o dest_Func) sel_Rs + val sel_jsss = map tuples_for_const sel_xs + val arg_jsss = + map (map_filter (fn js => if hd js = real_j then SOME (tl js) + else NONE)) sel_jsss + val uncur_arg_Ts = binder_types constr_T + in + if co andalso (T, j) mem seen then + Var (var ()) + else + let + val seen = seen |> co ? cons (T, j) + val ts = + if length arg_Ts = 0 then + [] + else + map3 (fn Ts => term_for_rep seen Ts Ts) arg_Ts arg_Rs + arg_jsss + |> mk_tuple (HOLogic.mk_tupleT uncur_arg_Ts) + |> dest_n_tuple (length uncur_arg_Ts) + val t = + if constr_s = @{const_name Abs_Frac} then + let + val num_T = body_type T + (* int -> term *) + val mk_num = HOLogic.mk_number num_T + in + case ts of + [Const (@{const_name Pair}, _) $ t1 $ t2] => + (case snd (HOLogic.dest_number t1) of + 0 => mk_num 0 + | n1 => case HOLogic.dest_number t2 |> snd of + 1 => mk_num n1 + | n2 => Const (@{const_name HOL.divide}, + [num_T, num_T] ---> num_T) + $ mk_num n1 $ mk_num n2) + | _ => raise TERM ("NitpickModel.reconstruct_term.term_\ + \for_atom (Abs_Frac)", ts) + end + else if not for_auto andalso is_abs_fun thy constr_x then + Const (abs_name, constr_T) $ the_single ts + else + list_comb (Const constr_x, ts) + in + if co then + let val var = var () in + if exists_subterm (equal (Var var)) t then + Const (@{const_name The}, (T --> bool_T) --> T) + $ Abs ("\", T, + Const (@{const_name "op ="}, [T, T] ---> bool_T) + $ Bound 0 $ abstract_over (Var var, t)) + else + t + end + else + t + end + end + (* (typ * int) list -> int -> rep -> typ -> typ -> typ -> int list + -> term *) + and term_for_vect seen k R T1 T2 T' js = + make_fun true T1 T2 T' (map (term_for_atom seen T1 T1) (index_seq 0 k)) + (map (term_for_rep seen T2 T2 R o single) + (batch_list (arity_of_rep R) js)) + (* (typ * int) list -> typ -> typ -> rep -> int list list -> term *) + and term_for_rep seen T T' Unit [[]] = term_for_atom seen T T' 0 + | term_for_rep seen T T' (R as Atom (k, j0)) [[j]] = + if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0) + else raise REP ("NitpickModel.reconstruct_term.term_for_rep", [R]) + | term_for_rep seen (Type ("*", [T1, T2])) _ (Struct [R1, R2]) [js] = + let + val arity1 = arity_of_rep R1 + val (js1, js2) = chop arity1 js + in + list_comb (HOLogic.pair_const T1 T2, + map3 (fn T => term_for_rep seen T T) [T1, T2] [R1, R2] + [[js1], [js2]]) + end + | term_for_rep seen (Type ("fun", [T1, T2])) T' (R as Vect (k, R')) [js] = + term_for_vect seen k R' T1 T2 T' js + | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, Formula Neut)) + jss = + let + val jss1 = all_combinations_for_rep R1 + val ts1 = map (term_for_rep seen T1 T1 R1 o single) jss1 + val ts2 = + map (fn js => term_for_rep seen T2 T2 (Atom (2, 0)) + [[int_for_bool (js mem jss)]]) jss1 + in make_fun false T1 T2 T' ts1 ts2 end + | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, R2)) jss = + let + val arity1 = arity_of_rep R1 + val jss1 = all_combinations_for_rep R1 + val ts1 = map (term_for_rep seen T1 T1 R1 o single) jss1 + val grouped_jss2 = AList.group (op =) (map (chop arity1) jss) + val ts2 = map (term_for_rep seen T2 T2 R2 o the_default [] + o AList.lookup (op =) grouped_jss2) jss1 + in make_fun true T1 T2 T' ts1 ts2 end + | term_for_rep seen T T' (Opt R) jss = + if null jss then Const (unknown, T) else term_for_rep seen T T' R jss + | term_for_rep seen T _ R jss = + raise ARG ("NitpickModel.reconstruct_term.term_for_rep", + Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^ + string_of_int (length jss)) + in + (not for_auto ? setify_mapify_funs []) o unbox_term oooo term_for_rep [] + end + +(* scope -> nut list -> nut NameTable.table -> Kodkod.raw_bound list -> nut + -> term *) +fun term_for_name scope sel_names rel_table bounds name = + let val T = type_of name in + tuple_list_for_name rel_table bounds name + |> reconstruct_term ("", "", "", "") scope sel_names rel_table bounds T T + (rep_of name) + end + +(* Proof.context + -> (string * string * string * string * string) * Proof.context *) +fun add_wacky_syntax ctxt = + let + (* term -> string *) + val name_of = fst o dest_Const + val thy = ProofContext.theory_of ctxt |> Context.reject_draft + val (maybe_t, thy) = + Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}), + Mixfix (maybe_mixfix, [1000], 1000)) thy + val (base_t, thy) = + Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}), + Mixfix (base_mixfix, [1000], 1000)) thy + val (step_t, thy) = + Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}), + Mixfix (step_mixfix, [1000], 1000)) thy + val (abs_t, thy) = + Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}), + Mixfix (abs_mixfix, [40], 40)) thy + in + ((name_of maybe_t, name_of base_t, name_of step_t, name_of abs_t), + ProofContext.transfer_syntax thy ctxt) + end + +(* term -> term *) +fun unfold_outer_the_binders (t as Const (@{const_name The}, _) + $ Abs (s, T, Const (@{const_name "op ="}, _) + $ Bound 0 $ t')) = + betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders + | unfold_outer_the_binders t = t +(* typ list -> int -> term * term -> bool *) +fun bisimilar_values _ 0 _ = true + | bisimilar_values coTs max_depth (t1, t2) = + let val T = fastype_of t1 in + if exists_subtype (member (op =) coTs) T then + let + val ((head1, args1), (head2, args2)) = + pairself (strip_comb o unfold_outer_the_binders) (t1, t2) + val max_depth = max_depth - (if T mem coTs then 1 else 0) + in + head1 = head2 + andalso forall (bisimilar_values coTs max_depth) (args1 ~~ args2) + end + else + t1 = t2 + end + +(* params -> scope -> (term option * int list) list -> styp list -> nut list + -> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list + -> Pretty.T * bool *) +fun reconstruct_hol_model {show_skolems, show_datatypes, show_consts} + ({ext_ctxt as {thy, ctxt, max_bisim_depth, boxes, user_axioms, debug, + wfs, destroy_constrs, specialize, skolemize, + star_linear_preds, uncurry, fast_descrs, tac_timeout, + evals, case_names, def_table, nondef_table, user_nondefs, + simp_table, psimp_table, intro_table, ground_thm_table, + ersatz_table, skolems, special_funs, unrolled_preds, + wf_cache}, + card_assigns, bisim_depth, datatypes, ofs} : scope) formats all_frees + free_names sel_names nonsel_names rel_table bounds = + let + val (wacky_names as (_, base_name, step_name, _), ctxt) = + add_wacky_syntax ctxt + val ext_ctxt = + {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes, + wfs = wfs, user_axioms = user_axioms, debug = debug, + destroy_constrs = destroy_constrs, specialize = specialize, + skolemize = skolemize, star_linear_preds = star_linear_preds, + uncurry = uncurry, fast_descrs = fast_descrs, tac_timeout = tac_timeout, + evals = evals, case_names = case_names, def_table = def_table, + nondef_table = nondef_table, user_nondefs = user_nondefs, + simp_table = simp_table, psimp_table = psimp_table, + intro_table = intro_table, ground_thm_table = ground_thm_table, + ersatz_table = ersatz_table, skolems = skolems, + special_funs = special_funs, unrolled_preds = unrolled_preds, + wf_cache = wf_cache} + val scope = {ext_ctxt = ext_ctxt, card_assigns = card_assigns, + bisim_depth = bisim_depth, datatypes = datatypes, ofs = ofs} + (* typ -> typ -> rep -> int list list -> term *) + val term_for_rep = reconstruct_term wacky_names scope sel_names rel_table + bounds + (* typ -> typ -> typ *) + fun nth_value_of_type T card n = term_for_rep T T (Atom (card, 0)) [[n]] + (* dtype_spec list -> dtype_spec -> bool *) + fun is_codatatype_wellformed (cos : dtype_spec list) + ({typ, card, ...} : dtype_spec) = + let + val ts = map (nth_value_of_type typ card) (index_seq 0 card) + val max_depth = Integer.sum (map #card cos) + in + forall (not o bisimilar_values (map #typ cos) max_depth) + (all_distinct_unordered_pairs_of ts) + end + (* string -> Pretty.T *) + fun pretty_for_assign name = + let + val (oper, (t1, T'), T) = + case name of + FreeName (s, T, _) => + let val t = Free (s, unbox_type T) in + ("=", (t, format_term_type thy def_table formats t), T) + end + | ConstName (s, T, _) => + (assign_operator_for_const (s, T), + user_friendly_const ext_ctxt (base_name, step_name) formats (s, T), + T) + | _ => raise NUT ("NitpickModel.reconstruct_hol_model.\ + \pretty_for_assign", [name]) + val t2 = if rep_of name = Any then + Const (@{const_name undefined}, T') + else + tuple_list_for_name rel_table bounds name + |> term_for_rep T T' (rep_of name) + in + Pretty.block (Pretty.breaks + [(setmp_CRITICAL show_question_marks false o setmp_show_all_types) + (Syntax.pretty_term ctxt) t1, + Pretty.str oper, Syntax.pretty_term ctxt t2]) + end + (* dtype_spec -> Pretty.T *) + fun pretty_for_datatype ({typ, card, precise, ...} : dtype_spec) = + Pretty.block (Pretty.breaks + [Syntax.pretty_typ ctxt (unbox_type typ), Pretty.str "=", + Pretty.enum "," "{" "}" + (map (Syntax.pretty_term ctxt o nth_value_of_type typ card) + (index_seq 0 card) @ + (if precise then [] else [Pretty.str unrep]))]) + (* typ -> dtype_spec list *) + fun integer_datatype T = + [{typ = T, card = card_of_type card_assigns T, co = false, + precise = false, constrs = []}] + handle TYPE ("NitpickHOL.card_of_type", _, _) => [] + val (codatatypes, datatypes) = + List.partition #co datatypes + ||> append (integer_datatype nat_T @ integer_datatype int_T) + val block_of_datatypes = + if show_datatypes andalso not (null datatypes) then + [Pretty.big_list ("Datatype" ^ plural_s_for_list datatypes ^ ":") + (map pretty_for_datatype datatypes)] + else + [] + val block_of_codatatypes = + if show_datatypes andalso not (null codatatypes) then + [Pretty.big_list ("Codatatype" ^ plural_s_for_list codatatypes ^ ":") + (map pretty_for_datatype codatatypes)] + else + [] + (* bool -> string -> nut list -> Pretty.T list *) + fun block_of_names show title names = + if show andalso not (null names) then + Pretty.str (title ^ plural_s_for_list names ^ ":") + :: map (Pretty.indent indent_size o pretty_for_assign) + (sort_wrt (original_name o nickname_of) names) + else + [] + val (skolem_names, nonskolem_nonsel_names) = + List.partition is_skolem_name nonsel_names + val (eval_names, noneval_nonskolem_nonsel_names) = + List.partition (String.isPrefix eval_prefix o nickname_of) + nonskolem_nonsel_names + ||> filter_out (equal @{const_name bisim_iterator_max} o nickname_of) + val free_names = + map (fn x as (s, T) => + case filter (equal x o nickname_of pairf (unbox_type o type_of)) + free_names of + [name] => name + | [] => FreeName (s, T, Any) + | _ => raise TERM ("NitpickModel.reconstruct_hol_model", + [Const x])) all_frees + val chunks = block_of_names true "Free variable" free_names @ + block_of_names show_skolems "Skolem constant" skolem_names @ + block_of_names true "Evaluated term" eval_names @ + block_of_datatypes @ block_of_codatatypes @ + block_of_names show_consts "Constant" + noneval_nonskolem_nonsel_names + in + (Pretty.chunks (if null chunks then [Pretty.str "Empty assignment"] + else chunks), + bisim_depth >= 0 + orelse forall (is_codatatype_wellformed codatatypes) codatatypes) + end + +(* scope -> Time.time option -> nut list -> nut list -> nut NameTable.table + -> Kodkod.raw_bound list -> term -> bool option *) +fun prove_hol_model (scope as {ext_ctxt as {thy, ctxt, ...}, card_assigns, ...}) + auto_timeout free_names sel_names rel_table bounds prop = + let + (* typ * int -> term *) + fun free_type_assm (T, k) = + let + (* int -> term *) + val atom = atom true T + fun equation_for_atom j = HOLogic.eq_const T $ Bound 0 $ atom j + val eqs = map equation_for_atom (index_seq 0 k) + val compreh_assm = + Const (@{const_name All}, (T --> bool_T) --> bool_T) + $ Abs ("x", T, foldl1 HOLogic.mk_disj eqs) + val distinct_assm = distinctness_formula T (map atom (index_seq 0 k)) + in HOLogic.mk_conj (compreh_assm, distinct_assm) end + (* nut -> term *) + fun free_name_assm name = + HOLogic.mk_eq (Free (nickname_of name, type_of name), + term_for_name scope sel_names rel_table bounds name) + val freeT_assms = map free_type_assm (filter (is_TFree o fst) card_assigns) + val model_assms = map free_name_assm free_names + val assm = List.foldr HOLogic.mk_conj @{const True} + (freeT_assms @ model_assms) + (* bool -> bool *) + fun try_out negate = + let + val concl = (negate ? curry (op $) @{const Not}) + (ObjectLogic.atomize_term thy prop) + val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (assm, concl)) + |> map_types (map_type_tfree + (fn (s, []) => TFree (s, HOLogic.typeS) + | x => TFree x)) + |> cterm_of thy |> Goal.init + in + (goal |> SINGLE (DETERM_TIMEOUT auto_timeout + (auto_tac (clasimpset_of ctxt))) + |> the |> Goal.finish ctxt; true) + handle THM _ => false + | TimeLimit.TimeOut => false + end + in + if silence try_out false then SOME true + else if silence try_out true then SOME false + else NONE + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_mono.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,943 @@ +(* Title: HOL/Nitpick/Tools/nitpick_mono.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2009 + +Monotonicity predicate for higher-order logic. +*) + +signature NITPICK_MONO = +sig + type extended_context = NitpickHOL.extended_context + + val formulas_monotonic : + extended_context -> typ -> term list -> term list -> term -> bool +end; + +structure NitpickMono : NITPICK_MONO = +struct + +open NitpickUtil +open NitpickHOL + +type var = int + +datatype sign = Pos | Neg +datatype sign_atom = S of sign | V of var + +type literal = var * sign + +datatype ctype = + CAlpha | + CFun of ctype * sign_atom * ctype | + CPair of ctype * ctype | + CType of string * ctype list | + CRec of string * typ list + +type cdata = + {ext_ctxt: extended_context, + alpha_T: typ, + max_fresh: int Unsynchronized.ref, + datatype_cache: ((string * typ list) * ctype) list Unsynchronized.ref, + constr_cache: (styp * ctype) list Unsynchronized.ref} + +exception CTYPE of string * ctype list + +(* string -> unit *) +fun print_g (s : string) = () + +(* var -> string *) +val string_for_var = signed_string_of_int +(* string -> var list -> string *) +fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>" + | string_for_vars sep xs = space_implode sep (map string_for_var xs) +fun subscript_string_for_vars sep xs = + if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>" + +(* sign -> string *) +fun string_for_sign Pos = "+" + | string_for_sign Neg = "-" + +(* sign -> sign -> sign *) +fun xor sn1 sn2 = if sn1 = sn2 then Pos else Neg +(* sign -> sign *) +val negate = xor Neg + +(* sign_atom -> string *) +fun string_for_sign_atom (S sn) = string_for_sign sn + | string_for_sign_atom (V j) = string_for_var j + +(* literal -> string *) +fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn + +val bool_C = CType (@{type_name bool}, []) + +(* ctype -> bool *) +fun is_CRec (CRec _) = true + | is_CRec _ = false + +val no_prec = 100 +val prec_CFun = 1 +val prec_CPair = 2 + +(* tuple_set -> int *) +fun precedence_of_ctype (CFun _) = prec_CFun + | precedence_of_ctype (CPair _) = prec_CPair + | precedence_of_ctype _ = no_prec + +(* ctype -> string *) +val string_for_ctype = + let + (* int -> ctype -> string *) + fun aux outer_prec C = + let + val prec = precedence_of_ctype C + val need_parens = (prec < outer_prec) + in + (if need_parens then "(" else "") ^ + (case C of + CAlpha => "\" + | CFun (C1, a, C2) => + aux (prec + 1) C1 ^ " \\<^bsup>" ^ + string_for_sign_atom a ^ "\<^esup> " ^ aux prec C2 + | CPair (C1, C2) => aux (prec + 1) C1 ^ " \ " ^ aux prec C2 + | CType (s, []) => + if s mem [@{type_name prop}, @{type_name bool}] then "o" else s + | CType (s, Cs) => "(" ^ commas (map (aux 0) Cs) ^ ") " ^ s + | CRec (s, _) => "[" ^ s ^ "]") ^ + (if need_parens then ")" else "") + end + in aux 0 end + +(* ctype -> ctype list *) +fun flatten_ctype (CPair (C1, C2)) = maps flatten_ctype [C1, C2] + | flatten_ctype (CType (_, Cs)) = maps flatten_ctype Cs + | flatten_ctype C = [C] + +(* extended_context -> typ -> cdata *) +fun initial_cdata ext_ctxt alpha_T = + ({ext_ctxt = ext_ctxt, alpha_T = alpha_T, max_fresh = Unsynchronized.ref 0, + datatype_cache = Unsynchronized.ref [], + constr_cache = Unsynchronized.ref []} : cdata) + +(* typ -> typ -> bool *) +fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) = + T = alpha_T orelse (not (is_fp_iterator_type T) + andalso exists (could_exist_alpha_subtype alpha_T) Ts) + | could_exist_alpha_subtype alpha_T T = (T = alpha_T) +(* theory -> typ -> typ -> bool *) +fun could_exist_alpha_sub_ctype _ (alpha_T as TFree _) = + could_exist_alpha_subtype alpha_T + | could_exist_alpha_sub_ctype thy alpha_T = equal alpha_T orf is_datatype thy + +(* ctype -> bool *) +fun exists_alpha_sub_ctype CAlpha = true + | exists_alpha_sub_ctype (CFun (C1, _, C2)) = + exists exists_alpha_sub_ctype [C1, C2] + | exists_alpha_sub_ctype (CPair (C1, C2)) = + exists exists_alpha_sub_ctype [C1, C2] + | exists_alpha_sub_ctype (CType (_, Cs)) = exists exists_alpha_sub_ctype Cs + | exists_alpha_sub_ctype (CRec _) = true + +(* ctype -> bool *) +fun exists_alpha_sub_ctype_fresh CAlpha = true + | exists_alpha_sub_ctype_fresh (CFun (_, V _, _)) = true + | exists_alpha_sub_ctype_fresh (CFun (_, _, C2)) = + exists_alpha_sub_ctype_fresh C2 + | exists_alpha_sub_ctype_fresh (CPair (C1, C2)) = + exists exists_alpha_sub_ctype_fresh [C1, C2] + | exists_alpha_sub_ctype_fresh (CType (_, Cs)) = + exists exists_alpha_sub_ctype_fresh Cs + | exists_alpha_sub_ctype_fresh (CRec _) = true + +(* string * typ list -> ctype list -> ctype *) +fun constr_ctype_for_binders z Cs = + fold_rev (fn C => curry3 CFun C (S Neg)) Cs (CRec z) + +(* ((string * typ list) * ctype) list -> ctype list -> ctype -> ctype *) +fun repair_ctype _ _ CAlpha = CAlpha + | repair_ctype cache seen (CFun (C1, a, C2)) = + CFun (repair_ctype cache seen C1, a, repair_ctype cache seen C2) + | repair_ctype cache seen (CPair Cp) = + CPair (pairself (repair_ctype cache seen) Cp) + | repair_ctype cache seen (CType (s, Cs)) = + CType (s, maps (flatten_ctype o repair_ctype cache seen) Cs) + | repair_ctype cache seen (CRec (z as (s, _))) = + case AList.lookup (op =) cache z |> the of + CRec _ => CType (s, []) + | C => if C mem seen then CType (s, []) + else repair_ctype cache (C :: seen) C + +(* ((string * typ list) * ctype) list Unsynchronized.ref -> unit *) +fun repair_datatype_cache cache = + let + (* (string * typ list) * ctype -> unit *) + fun repair_one (z, C) = + Unsynchronized.change cache + (AList.update (op =) (z, repair_ctype (!cache) [] C)) + in List.app repair_one (rev (!cache)) end + +(* (typ * ctype) list -> (styp * ctype) list Unsynchronized.ref -> unit *) +fun repair_constr_cache dtype_cache constr_cache = + let + (* styp * ctype -> unit *) + fun repair_one (x, C) = + Unsynchronized.change constr_cache + (AList.update (op =) (x, repair_ctype dtype_cache [] C)) + in List.app repair_one (!constr_cache) end + +(* cdata -> typ -> ctype *) +fun fresh_ctype_for_type ({ext_ctxt as {thy, ...}, alpha_T, max_fresh, + datatype_cache, constr_cache, ...} : cdata) = + let + (* typ -> typ -> ctype *) + fun do_fun T1 T2 = + let + val C1 = do_type T1 + val C2 = do_type T2 + val a = if is_boolean_type (body_type T2) + andalso exists_alpha_sub_ctype_fresh C1 then + V (Unsynchronized.inc max_fresh) + else + S Neg + in CFun (C1, a, C2) end + (* typ -> ctype *) + and do_type T = + if T = alpha_T then + CAlpha + else case T of + Type ("fun", [T1, T2]) => do_fun T1 T2 + | Type (@{type_name fun_box}, [T1, T2]) => do_fun T1 T2 + | Type ("*", [T1, T2]) => CPair (pairself do_type (T1, T2)) + | Type (z as (s, _)) => + if could_exist_alpha_sub_ctype thy alpha_T T then + case AList.lookup (op =) (!datatype_cache) z of + SOME C => C + | NONE => + let + val _ = Unsynchronized.change datatype_cache (cons (z, CRec z)) + val xs = datatype_constrs thy T + val (all_Cs, constr_Cs) = + fold_rev (fn (_, T') => fn (all_Cs, constr_Cs) => + let + val binder_Cs = map do_type (binder_types T') + val new_Cs = filter exists_alpha_sub_ctype_fresh + binder_Cs + val constr_C = constr_ctype_for_binders z + binder_Cs + in + (union (op =) new_Cs all_Cs, + constr_C :: constr_Cs) + end) + xs ([], []) + val C = CType (s, all_Cs) + val _ = Unsynchronized.change datatype_cache + (AList.update (op =) (z, C)) + val _ = Unsynchronized.change constr_cache + (append (xs ~~ constr_Cs)) + in + if forall (not o is_CRec o snd) (!datatype_cache) then + (repair_datatype_cache datatype_cache; + repair_constr_cache (!datatype_cache) constr_cache; + AList.lookup (op =) (!datatype_cache) z |> the) + else + C + end + else + CType (s, []) + | _ => CType (Refute.string_of_typ T, []) + in do_type end + +(* ctype -> ctype list *) +fun prodC_factors (CPair (C1, C2)) = maps prodC_factors [C1, C2] + | prodC_factors C = [C] +(* ctype -> ctype list * ctype *) +fun curried_strip_ctype (CFun (C1, S Neg, C2)) = + curried_strip_ctype C2 |>> append (prodC_factors C1) + | curried_strip_ctype C = ([], C) +(* string -> ctype -> ctype *) +fun sel_ctype_from_constr_ctype s C = + let val (arg_Cs, dataC) = curried_strip_ctype C in + CFun (dataC, S Neg, + case sel_no_from_name s of ~1 => bool_C | n => nth arg_Cs n) + end + +(* cdata -> styp -> ctype *) +fun ctype_for_constr (cdata as {ext_ctxt as {thy, ...}, alpha_T, constr_cache, + ...}) (x as (_, T)) = + if could_exist_alpha_sub_ctype thy alpha_T T then + case AList.lookup (op =) (!constr_cache) x of + SOME C => C + | NONE => (fresh_ctype_for_type cdata (body_type T); + AList.lookup (op =) (!constr_cache) x |> the) + else + fresh_ctype_for_type cdata T +fun ctype_for_sel (cdata as {ext_ctxt, ...}) (x as (s, _)) = + x |> boxed_constr_for_sel ext_ctxt |> ctype_for_constr cdata + |> sel_ctype_from_constr_ctype s + +(* literal list -> ctype -> ctype *) +fun instantiate_ctype lits = + let + (* ctype -> ctype *) + fun aux CAlpha = CAlpha + | aux (CFun (C1, V x, C2)) = + let + val a = case AList.lookup (op =) lits x of + SOME sn => S sn + | NONE => V x + in CFun (aux C1, a, aux C2) end + | aux (CFun (C1, a, C2)) = CFun (aux C1, a, aux C2) + | aux (CPair Cp) = CPair (pairself aux Cp) + | aux (CType (s, Cs)) = CType (s, map aux Cs) + | aux (CRec z) = CRec z + in aux end + +datatype comp_op = Eq | Leq + +type comp = sign_atom * sign_atom * comp_op * var list +type sign_expr = literal list + +datatype constraint_set = + UnsolvableCSet | + CSet of literal list * comp list * sign_expr list + +(* comp_op -> string *) +fun string_for_comp_op Eq = "=" + | string_for_comp_op Leq = "\" + +(* sign_expr -> string *) +fun string_for_sign_expr [] = "\" + | string_for_sign_expr lits = + space_implode " \ " (map string_for_literal lits) + +(* constraint_set *) +val slack = CSet ([], [], []) + +(* literal -> literal list option -> literal list option *) +fun do_literal _ NONE = NONE + | do_literal (x, sn) (SOME lits) = + case AList.lookup (op =) lits x of + SOME sn' => if sn = sn' then SOME lits else NONE + | NONE => SOME ((x, sn) :: lits) + +(* comp_op -> var list -> sign_atom -> sign_atom -> literal list * comp list + -> (literal list * comp list) option *) +fun do_sign_atom_comp Eq [] a1 a2 (accum as (lits, comps)) = + (case (a1, a2) of + (S sn1, S sn2) => if sn1 = sn2 then SOME accum else NONE + | (V x1, S sn2) => + Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits)) + | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps) + | _ => do_sign_atom_comp Eq [] a2 a1 accum) + | do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) = + (case (a1, a2) of + (_, S Neg) => SOME accum + | (S Pos, _) => SOME accum + | (S Neg, S Pos) => NONE + | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps) + | _ => do_sign_atom_comp Eq [] a1 a2 accum) + | do_sign_atom_comp cmp xs a1 a2 (accum as (lits, comps)) = + SOME (lits, insert (op =) (a1, a2, cmp, xs) comps) + +(* comp -> var list -> ctype -> ctype -> (literal list * comp list) option + -> (literal list * comp list) option *) +fun do_ctype_comp _ _ _ _ NONE = NONE + | do_ctype_comp _ _ CAlpha CAlpha accum = accum + | do_ctype_comp Eq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22)) + (SOME accum) = + accum |> do_sign_atom_comp Eq xs a1 a2 |> do_ctype_comp Eq xs C11 C21 + |> do_ctype_comp Eq xs C12 C22 + | do_ctype_comp Leq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22)) + (SOME accum) = + (if exists_alpha_sub_ctype C11 then + accum |> do_sign_atom_comp Leq xs a1 a2 + |> do_ctype_comp Leq xs C21 C11 + |> (case a2 of + S Neg => I + | S Pos => do_ctype_comp Leq xs C11 C21 + | V x => do_ctype_comp Leq (x :: xs) C11 C21) + else + SOME accum) + |> do_ctype_comp Leq xs C12 C22 + | do_ctype_comp cmp xs (C1 as CPair (C11, C12)) (C2 as CPair (C21, C22)) + accum = + (accum |> fold (uncurry (do_ctype_comp cmp xs)) [(C11, C21), (C12, C22)] + handle Library.UnequalLengths => + raise CTYPE ("NitpickMono.do_ctype_comp", [C1, C2])) + | do_ctype_comp cmp xs (CType _) (CType _) accum = + accum (* no need to compare them thanks to the cache *) + | do_ctype_comp _ _ C1 C2 _ = + raise CTYPE ("NitpickMono.do_ctype_comp", [C1, C2]) + +(* comp_op -> ctype -> ctype -> constraint_set -> constraint_set *) +fun add_ctype_comp _ _ _ UnsolvableCSet = UnsolvableCSet + | add_ctype_comp cmp C1 C2 (CSet (lits, comps, sexps)) = + (print_g ("*** Add " ^ string_for_ctype C1 ^ " " ^ string_for_comp_op cmp ^ + " " ^ string_for_ctype C2); + case do_ctype_comp cmp [] C1 C2 (SOME (lits, comps)) of + NONE => (print_g "**** Unsolvable"; UnsolvableCSet) + | SOME (lits, comps) => CSet (lits, comps, sexps)) + +(* ctype -> ctype -> constraint_set -> constraint_set *) +val add_ctypes_equal = add_ctype_comp Eq +val add_is_sub_ctype = add_ctype_comp Leq + +(* sign -> sign_expr -> ctype -> (literal list * sign_expr list) option + -> (literal list * sign_expr list) option *) +fun do_notin_ctype_fv _ _ _ NONE = NONE + | do_notin_ctype_fv Neg _ CAlpha accum = accum + | do_notin_ctype_fv Pos [] CAlpha _ = NONE + | do_notin_ctype_fv Pos [(x, sn)] CAlpha (SOME (lits, sexps)) = + SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps) + | do_notin_ctype_fv Pos sexp CAlpha (SOME (lits, sexps)) = + SOME (lits, insert (op =) sexp sexps) + | do_notin_ctype_fv sn sexp (CFun (C1, S sn', C2)) accum = + accum |> (if sn' = Pos andalso sn = Pos then do_notin_ctype_fv Pos sexp C1 + else I) + |> (if sn' = Neg orelse sn = Pos then do_notin_ctype_fv Neg sexp C1 + else I) + |> do_notin_ctype_fv sn sexp C2 + | do_notin_ctype_fv Pos sexp (CFun (C1, V x, C2)) accum = + accum |> (case do_literal (x, Neg) (SOME sexp) of + NONE => I + | SOME sexp' => do_notin_ctype_fv Pos sexp' C1) + |> do_notin_ctype_fv Neg sexp C1 + |> do_notin_ctype_fv Pos sexp C2 + | do_notin_ctype_fv Neg sexp (CFun (C1, V x, C2)) accum = + accum |> (case do_literal (x, Pos) (SOME sexp) of + NONE => I + | SOME sexp' => do_notin_ctype_fv Pos sexp' C1) + |> do_notin_ctype_fv Neg sexp C2 + | do_notin_ctype_fv sn sexp (CPair (C1, C2)) accum = + accum |> fold (do_notin_ctype_fv sn sexp) [C1, C2] + | do_notin_ctype_fv sn sexp (CType (_, Cs)) accum = + accum |> fold (do_notin_ctype_fv sn sexp) Cs + | do_notin_ctype_fv _ _ C _ = + raise CTYPE ("NitpickMono.do_notin_ctype_fv", [C]) + +(* sign -> ctype -> constraint_set -> constraint_set *) +fun add_notin_ctype_fv _ _ UnsolvableCSet = UnsolvableCSet + | add_notin_ctype_fv sn C (CSet (lits, comps, sexps)) = + (print_g ("*** Add " ^ string_for_ctype C ^ " is right-" ^ + (case sn of Neg => "unique" | Pos => "total") ^ "."); + case do_notin_ctype_fv sn [] C (SOME (lits, sexps)) of + NONE => (print_g "**** Unsolvable"; UnsolvableCSet) + | SOME (lits, sexps) => CSet (lits, comps, sexps)) + +(* ctype -> constraint_set -> constraint_set *) +val add_ctype_is_right_unique = add_notin_ctype_fv Neg +val add_ctype_is_right_total = add_notin_ctype_fv Pos + +(* constraint_set -> constraint_set -> constraint_set *) +fun unite (CSet (lits1, comps1, sexps1)) (CSet (lits2, comps2, sexps2)) = + (case SOME lits1 |> fold do_literal lits2 of + NONE => (print_g "**** Unsolvable"; UnsolvableCSet) + | SOME lits => CSet (lits, comps1 @ comps2, sexps1 @ sexps2)) + | unite _ _ = UnsolvableCSet + +(* sign -> bool *) +fun bool_from_sign Pos = false + | bool_from_sign Neg = true +(* bool -> sign *) +fun sign_from_bool false = Pos + | sign_from_bool true = Neg + +(* literal -> PropLogic.prop_formula *) +fun prop_for_literal (x, sn) = + (not (bool_from_sign sn) ? PropLogic.Not) (PropLogic.BoolVar x) +(* sign_atom -> PropLogic.prop_formula *) +fun prop_for_sign_atom_eq (S sn', sn) = + if sn = sn' then PropLogic.True else PropLogic.False + | prop_for_sign_atom_eq (V x, sn) = prop_for_literal (x, sn) +(* sign_expr -> PropLogic.prop_formula *) +fun prop_for_sign_expr xs = PropLogic.exists (map prop_for_literal xs) +(* var list -> sign -> PropLogic.prop_formula *) +fun prop_for_exists_eq xs sn = + PropLogic.exists (map (fn x => prop_for_literal (x, sn)) xs) +(* comp -> PropLogic.prop_formula *) +fun prop_for_comp (a1, a2, Eq, []) = + PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []), + prop_for_comp (a2, a1, Leq, [])) + | prop_for_comp (a1, a2, Leq, []) = + PropLogic.SOr (prop_for_sign_atom_eq (a1, Pos), + prop_for_sign_atom_eq (a2, Neg)) + | prop_for_comp (a1, a2, cmp, xs) = + PropLogic.SOr (prop_for_exists_eq xs Neg, prop_for_comp (a1, a2, cmp, [])) + +(* var -> (int -> bool option) -> literal list -> literal list *) +fun literals_from_assignments max_var asgns lits = + fold (fn x => fn accum => + if AList.defined (op =) lits x then + accum + else case asgns x of + SOME b => (x, sign_from_bool b) :: accum + | NONE => accum) (max_var downto 1) lits + +(* literal list -> sign_atom -> sign option *) +fun lookup_sign_atom _ (S sn) = SOME sn + | lookup_sign_atom lit (V x) = AList.lookup (op =) lit x + +(* comp -> string *) +fun string_for_comp (a1, a2, cmp, xs) = + string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^ + subscript_string_for_vars " \ " xs ^ " " ^ string_for_sign_atom a2 + +(* literal list -> comp list -> sign_expr list -> unit *) +fun print_problem lits comps sexps = + print_g ("*** Problem:\n" ^ cat_lines (map string_for_literal lits @ + map string_for_comp comps @ + map string_for_sign_expr sexps)) + +(* literal list -> unit *) +fun print_solution lits = + let val (pos, neg) = List.partition (equal Pos o snd) lits in + print_g ("*** Solution:\n" ^ + "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^ + "-: " ^ commas (map (string_for_var o fst) neg)) + end + +(* var -> constraint_set -> literal list list option *) +fun solve _ UnsolvableCSet = (print_g "*** Problem: Unsolvable"; NONE) + | solve max_var (CSet (lits, comps, sexps)) = + let + val _ = print_problem lits comps sexps + val prop = PropLogic.all (map prop_for_literal lits @ + map prop_for_comp comps @ + map prop_for_sign_expr sexps) + in + case silence (SatSolver.invoke_solver "dpll") prop of + SatSolver.SATISFIABLE asgns => + SOME (literals_from_assignments max_var asgns lits + |> tap print_solution) + | _ => NONE + end + +(* var -> constraint_set -> bool *) +val is_solvable = is_some oo solve + +type ctype_schema = ctype * constraint_set +type ctype_context = + {bounds: ctype list, + frees: (styp * ctype) list, + consts: (styp * ctype_schema) list} + +type accumulator = ctype_context * constraint_set + +val initial_gamma = {bounds = [], frees = [], consts = []} +val unsolvable_accum = (initial_gamma, UnsolvableCSet) + +(* ctype -> ctype_context -> ctype_context *) +fun push_bound C {bounds, frees, consts} = + {bounds = C :: bounds, frees = frees, consts = consts} +(* ctype_context -> ctype_context *) +fun pop_bound {bounds, frees, consts} = + {bounds = tl bounds, frees = frees, consts = consts} + handle List.Empty => initial_gamma + +(* cdata -> term -> accumulator -> ctype * accumulator *) +fun consider_term (cdata as {ext_ctxt as {ctxt, thy, def_table, ...}, alpha_T, + max_fresh, ...}) = + let + (* typ -> ctype *) + val ctype_for = fresh_ctype_for_type cdata + (* ctype -> ctype *) + fun pos_set_ctype_for_dom C = + CFun (C, S (if exists_alpha_sub_ctype C then Pos else Neg), bool_C) + (* typ -> accumulator -> ctype * accumulator *) + fun do_quantifier T (gamma, cset) = + let + val abs_C = ctype_for (domain_type (domain_type T)) + val body_C = ctype_for (range_type T) + in + (CFun (CFun (abs_C, S Neg, body_C), S Neg, body_C), + (gamma, cset |> add_ctype_is_right_total abs_C)) + end + fun do_equals T (gamma, cset) = + let val C = ctype_for (domain_type T) in + (CFun (C, S Neg, CFun (C, S Neg, ctype_for (nth_range_type 2 T))), + (gamma, cset |> add_ctype_is_right_unique C)) + end + fun do_robust_set_operation T (gamma, cset) = + let + val set_T = domain_type T + val C1 = ctype_for set_T + val C2 = ctype_for set_T + val C3 = ctype_for set_T + in + (CFun (C1, S Neg, CFun (C2, S Neg, C3)), + (gamma, cset |> add_is_sub_ctype C1 C3 |> add_is_sub_ctype C2 C3)) + end + fun do_fragile_set_operation T (gamma, cset) = + let + val set_T = domain_type T + val set_C = ctype_for set_T + (* typ -> ctype *) + fun custom_ctype_for (T as Type ("fun", [T1, T2])) = + if T = set_T then set_C + else CFun (custom_ctype_for T1, S Neg, custom_ctype_for T2) + | custom_ctype_for T = ctype_for T + in + (custom_ctype_for T, (gamma, cset |> add_ctype_is_right_unique set_C)) + end + (* typ -> accumulator -> ctype * accumulator *) + fun do_pair_constr T accum = + case ctype_for (nth_range_type 2 T) of + C as CPair (a_C, b_C) => + (CFun (a_C, S Neg, CFun (b_C, S Neg, C)), accum) + | C => raise CTYPE ("NitpickMono.consider_term.do_pair_constr", [C]) + (* int -> typ -> accumulator -> ctype * accumulator *) + fun do_nth_pair_sel n T = + case ctype_for (domain_type T) of + C as CPair (a_C, b_C) => + pair (CFun (C, S Neg, if n = 0 then a_C else b_C)) + | C => raise CTYPE ("NitpickMono.consider_term.do_nth_pair_sel", [C]) + val unsolvable = (CType ("unsolvable", []), unsolvable_accum) + (* typ -> term -> accumulator -> ctype * accumulator *) + fun do_bounded_quantifier abs_T bound_t body_t accum = + let + val abs_C = ctype_for abs_T + val (bound_C, accum) = accum |>> push_bound abs_C |> do_term bound_t + val expected_bound_C = pos_set_ctype_for_dom abs_C + in + accum ||> add_ctypes_equal expected_bound_C bound_C |> do_term body_t + ||> apfst pop_bound + end + (* term -> accumulator -> ctype * accumulator *) + and do_term _ (_, UnsolvableCSet) = unsolvable + | do_term t (accum as (gamma as {bounds, frees, consts}, cset)) = + (case t of + Const (x as (s, T)) => + (case AList.lookup (op =) consts x of + SOME (C, cset') => (C, (gamma, cset |> unite cset')) + | NONE => + if not (could_exist_alpha_subtype alpha_T T) then + (ctype_for T, accum) + else case s of + @{const_name all} => do_quantifier T accum + | @{const_name "=="} => do_equals T accum + | @{const_name All} => do_quantifier T accum + | @{const_name Ex} => do_quantifier T accum + | @{const_name "op ="} => do_equals T accum + | @{const_name The} => (print_g "*** The"; unsolvable) + | @{const_name Eps} => (print_g "*** Eps"; unsolvable) + | @{const_name If} => + do_robust_set_operation (range_type T) accum + |>> curry3 CFun bool_C (S Neg) + | @{const_name Pair} => do_pair_constr T accum + | @{const_name fst} => do_nth_pair_sel 0 T accum + | @{const_name snd} => do_nth_pair_sel 1 T accum + | @{const_name Id} => + (CFun (ctype_for (domain_type T), S Neg, bool_C), accum) + | @{const_name insert} => + let + val set_T = domain_type (range_type T) + val C1 = ctype_for (domain_type set_T) + val C1' = pos_set_ctype_for_dom C1 + val C2 = ctype_for set_T + val C3 = ctype_for set_T + in + (CFun (C1, S Neg, CFun (C2, S Neg, C3)), + (gamma, cset |> add_ctype_is_right_unique C1 + |> add_is_sub_ctype C1' C3 + |> add_is_sub_ctype C2 C3)) + end + | @{const_name converse} => + let + val x = Unsynchronized.inc max_fresh + (* typ -> ctype *) + fun ctype_for_set T = + CFun (ctype_for (domain_type T), V x, bool_C) + val ab_set_C = domain_type T |> ctype_for_set + val ba_set_C = range_type T |> ctype_for_set + in (CFun (ab_set_C, S Neg, ba_set_C), accum) end + | @{const_name trancl} => do_fragile_set_operation T accum + | @{const_name rtrancl} => (print_g "*** rtrancl"; unsolvable) + | @{const_name lower_semilattice_fun_inst.inf_fun} => + do_robust_set_operation T accum + | @{const_name upper_semilattice_fun_inst.sup_fun} => + do_robust_set_operation T accum + | @{const_name finite} => + let val C1 = ctype_for (domain_type (domain_type T)) in + (CFun (pos_set_ctype_for_dom C1, S Neg, bool_C), accum) + end + | @{const_name rel_comp} => + let + val x = Unsynchronized.inc max_fresh + (* typ -> ctype *) + fun ctype_for_set T = + CFun (ctype_for (domain_type T), V x, bool_C) + val bc_set_C = domain_type T |> ctype_for_set + val ab_set_C = domain_type (range_type T) |> ctype_for_set + val ac_set_C = nth_range_type 2 T |> ctype_for_set + in + (CFun (bc_set_C, S Neg, CFun (ab_set_C, S Neg, ac_set_C)), + accum) + end + | @{const_name image} => + let + val a_C = ctype_for (domain_type (domain_type T)) + val b_C = ctype_for (range_type (domain_type T)) + in + (CFun (CFun (a_C, S Neg, b_C), S Neg, + CFun (pos_set_ctype_for_dom a_C, S Neg, + pos_set_ctype_for_dom b_C)), accum) + end + | @{const_name Sigma} => + let + val x = Unsynchronized.inc max_fresh + (* typ -> ctype *) + fun ctype_for_set T = + CFun (ctype_for (domain_type T), V x, bool_C) + val a_set_T = domain_type T + val a_C = ctype_for (domain_type a_set_T) + val b_set_C = ctype_for_set (range_type (domain_type + (range_type T))) + val a_set_C = ctype_for_set a_set_T + val a_to_b_set_C = CFun (a_C, S Neg, b_set_C) + val ab_set_C = ctype_for_set (nth_range_type 2 T) + in + (CFun (a_set_C, S Neg, CFun (a_to_b_set_C, S Neg, ab_set_C)), + accum) + end + | @{const_name minus_fun_inst.minus_fun} => + let + val set_T = domain_type T + val left_set_C = ctype_for set_T + val right_set_C = ctype_for set_T + in + (CFun (left_set_C, S Neg, + CFun (right_set_C, S Neg, left_set_C)), + (gamma, cset |> add_ctype_is_right_unique right_set_C + (* FIXME: |> add_is_sub_ctype right_set_C left_set_C *))) + end + | @{const_name ord_fun_inst.less_eq_fun} => + do_fragile_set_operation T accum + | @{const_name Tha} => + let + val a_C = ctype_for (domain_type (domain_type T)) + val a_set_C = pos_set_ctype_for_dom a_C + in (CFun (a_set_C, S Neg, a_C), accum) end + | @{const_name FunBox} => + let val dom_C = ctype_for (domain_type T) in + (CFun (dom_C, S Neg, dom_C), accum) + end + | _ => if is_sel s then + if constr_name_for_sel_like s = @{const_name FunBox} then + let val dom_C = ctype_for (domain_type T) in + (CFun (dom_C, S Neg, dom_C), accum) + end + else + (ctype_for_sel cdata x, accum) + else if is_constr thy x then + (ctype_for_constr cdata x, accum) + else if is_built_in_const true x then + case def_of_const thy def_table x of + SOME t' => do_term t' accum + | NONE => (print_g ("*** built-in " ^ s); unsolvable) + else + (ctype_for T, accum)) + | Free (x as (_, T)) => + (case AList.lookup (op =) frees x of + SOME C => (C, accum) + | NONE => + let val C = ctype_for T in + (C, ({bounds = bounds, frees = (x, C) :: frees, + consts = consts}, cset)) + end) + | Var _ => (print_g "*** Var"; unsolvable) + | Bound j => (nth bounds j, accum) + | Abs (_, T, @{const False}) => (ctype_for (T --> bool_T), accum) + | Abs (s, T, t') => + let + val C = ctype_for T + val (C', accum) = do_term t' (accum |>> push_bound C) + in (CFun (C, S Neg, C'), accum |>> pop_bound) end + | Const (@{const_name All}, _) + $ Abs (_, T', @{const "op -->"} $ (t1 $ Bound 0) $ t2) => + do_bounded_quantifier T' t1 t2 accum + | Const (@{const_name Ex}, _) + $ Abs (_, T', @{const "op &"} $ (t1 $ Bound 0) $ t2) => + do_bounded_quantifier T' t1 t2 accum + | Const (@{const_name Let}, _) $ t1 $ t2 => + do_term (betapply (t2, t1)) accum + | t1 $ t2 => + let + val (C1, accum) = do_term t1 accum + val (C2, accum) = do_term t2 accum + in + case accum of + (_, UnsolvableCSet) => unsolvable + | _ => case C1 of + CFun (C11, _, C12) => + (C12, accum ||> add_is_sub_ctype C2 C11) + | _ => raise CTYPE ("NitpickMono.consider_term.do_term \ + \(op $)", [C1]) + end) + |> tap (fn (C, _) => + print_g (" \ \ " ^ + Syntax.string_of_term ctxt t ^ " : " ^ + string_for_ctype C)) + in do_term end + +(* cdata -> sign -> term -> accumulator -> accumulator *) +fun consider_general_formula (cdata as {ext_ctxt as {ctxt, ...}, ...}) = + let + (* typ -> ctype *) + val ctype_for = fresh_ctype_for_type cdata + (* term -> accumulator -> ctype * accumulator *) + val do_term = consider_term cdata + (* term -> accumulator -> accumulator *) + val do_boolean_term = snd oo do_term + (* sign -> term -> accumulator -> accumulator *) + fun do_formula _ _ (_, UnsolvableCSet) = unsolvable_accum + | do_formula sn t (accum as (gamma as {bounds, frees, consts}, cset)) = + let + (* term -> accumulator -> accumulator *) + val do_co_formula = do_formula sn + val do_contra_formula = do_formula (negate sn) + (* string -> typ -> term -> accumulator *) + fun do_quantifier quant_s abs_T body_t = + let + val abs_C = ctype_for abs_T + val side_cond = ((sn = Neg) = (quant_s = @{const_name Ex})) + val cset = cset |> side_cond ? add_ctype_is_right_total abs_C + in + (gamma |> push_bound abs_C, cset) |> do_co_formula body_t + |>> pop_bound + end + (* typ -> term -> accumulator *) + fun do_bounded_quantifier abs_T body_t = + accum |>> push_bound (ctype_for abs_T) |> do_co_formula body_t + |>> pop_bound + (* term -> term -> accumulator *) + fun do_equals t1 t2 = + case sn of + Pos => do_boolean_term t accum + | Neg => let + val (C1, accum) = do_term t1 accum + val (C2, accum) = do_term t2 accum + in accum (* FIXME: ||> add_ctypes_equal C1 C2 *) end + in + case t of + Const (s0 as @{const_name all}, _) $ Abs (_, T1, t1) => + do_quantifier s0 T1 t1 + | Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2 + | @{const "==>"} $ t1 $ t2 => + accum |> do_contra_formula t1 |> do_co_formula t2 + | @{const Trueprop} $ t1 => do_co_formula t1 accum + | @{const Not} $ t1 => do_contra_formula t1 accum + | Const (@{const_name All}, _) + $ Abs (_, T1, t1 as @{const "op -->"} $ (_ $ Bound 0) $ _) => + do_bounded_quantifier T1 t1 + | Const (s0 as @{const_name All}, _) $ Abs (_, T1, t1) => + do_quantifier s0 T1 t1 + | Const (@{const_name Ex}, _) + $ Abs (_, T1, t1 as @{const "op &"} $ (_ $ Bound 0) $ _) => + do_bounded_quantifier T1 t1 + | Const (s0 as @{const_name Ex}, _) $ Abs (_, T1, t1) => + do_quantifier s0 T1 t1 + | Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 + | @{const "op &"} $ t1 $ t2 => + accum |> do_co_formula t1 |> do_co_formula t2 + | @{const "op |"} $ t1 $ t2 => + accum |> do_co_formula t1 |> do_co_formula t2 + | @{const "op -->"} $ t1 $ t2 => + accum |> do_contra_formula t1 |> do_co_formula t2 + | Const (@{const_name If}, _) $ t1 $ t2 $ t3 => + accum |> do_boolean_term t1 |> do_co_formula t2 |> do_co_formula t3 + | Const (@{const_name Let}, _) $ t1 $ t2 => + do_co_formula (betapply (t2, t1)) accum + | _ => do_boolean_term t accum + end + |> tap (fn _ => print_g ("\ \ " ^ + Syntax.string_of_term ctxt t ^ + " : o\<^sup>" ^ string_for_sign sn)) + in do_formula end + +(* The harmless axiom optimization below is somewhat too aggressive in the face + of (rather peculiar) user-defined axioms. *) +val harmless_consts = + [@{const_name ord_class.less}, @{const_name ord_class.less_eq}] +val bounteous_consts = [@{const_name bisim}] + +(* term -> bool *) +fun is_harmless_axiom t = + Term.add_consts t [] |> filter_out (is_built_in_const true) + |> (forall (member (op =) harmless_consts o original_name o fst) + orf exists (member (op =) bounteous_consts o fst)) + +(* cdata -> sign -> term -> accumulator -> accumulator *) +fun consider_nondefinitional_axiom cdata sn t = + not (is_harmless_axiom t) ? consider_general_formula cdata sn t + +(* cdata -> term -> accumulator -> accumulator *) +fun consider_definitional_axiom (cdata as {ext_ctxt as {thy, ...}, ...}) t = + if not (is_constr_pattern_formula thy t) then + consider_nondefinitional_axiom cdata Pos t + else if is_harmless_axiom t then + I + else + let + (* term -> accumulator -> accumulator *) + val do_term = consider_term cdata + (* typ -> term -> accumulator -> accumulator *) + fun do_all abs_T body_t accum = + let val abs_C = fresh_ctype_for_type cdata abs_T in + accum |>> push_bound abs_C |> do_formula body_t |>> pop_bound + end + (* term -> term -> accumulator -> accumulator *) + and do_implies t1 t2 = do_term t1 #> snd #> do_formula t2 + and do_equals t1 t2 accum = + let + val (C1, accum) = do_term t1 accum + val (C2, accum) = do_term t2 accum + in accum ||> add_ctypes_equal C1 C2 end + (* term -> accumulator -> accumulator *) + and do_formula _ (_, UnsolvableCSet) = unsolvable_accum + | do_formula t accum = + case t of + Const (@{const_name all}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum + | @{const Trueprop} $ t1 => do_formula t1 accum + | Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2 accum + | @{const "==>"} $ t1 $ t2 => do_implies t1 t2 accum + | @{const Pure.conjunction} $ t1 $ t2 => + accum |> do_formula t1 |> do_formula t2 + | Const (@{const_name All}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum + | Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 accum + | @{const "op &"} $ t1 $ t2 => accum |> do_formula t1 |> do_formula t2 + | @{const "op -->"} $ t1 $ t2 => do_implies t1 t2 accum + | _ => raise TERM ("NitpickMono.consider_definitional_axiom.\ + \do_formula", [t]) + in do_formula t end + +(* Proof.context -> literal list -> term -> ctype -> string *) +fun string_for_ctype_of_term ctxt lits t C = + Syntax.string_of_term ctxt t ^ " : " ^ + string_for_ctype (instantiate_ctype lits C) + +(* theory -> literal list -> ctype_context -> unit *) +fun print_ctype_context ctxt lits ({frees, consts, ...} : ctype_context) = + map (fn (x, C) => string_for_ctype_of_term ctxt lits (Free x) C) frees @ + map (fn (x, (C, _)) => string_for_ctype_of_term ctxt lits (Const x) C) consts + |> cat_lines |> print_g + +(* extended_context -> typ -> term list -> term list -> term -> bool *) +fun formulas_monotonic (ext_ctxt as {ctxt, ...}) alpha_T def_ts nondef_ts + core_t = + let + val _ = print_g ("****** " ^ string_for_ctype CAlpha ^ " is " ^ + Syntax.string_of_typ ctxt alpha_T) + val cdata as {max_fresh, ...} = initial_cdata ext_ctxt alpha_T + val (gamma, cset) = + (initial_gamma, slack) + |> fold (consider_definitional_axiom cdata) def_ts + |> fold (consider_nondefinitional_axiom cdata Pos) nondef_ts + |> consider_general_formula cdata Pos core_t + in + case solve (!max_fresh) cset of + SOME lits => (print_ctype_context ctxt lits gamma; true) + | _ => false + end + handle CTYPE (loc, Cs) => raise BAD (loc, commas (map string_for_ctype Cs)) + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_nut.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,1363 @@ +(* Title: HOL/Nitpick/Tools/nitpick_nut.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Nitpick underlying terms (nuts). +*) + +signature NITPICK_NUT = +sig + type special_fun = NitpickHOL.special_fun + type extended_context = NitpickHOL.extended_context + type scope = NitpickScope.scope + type name_pool = NitpickPeephole.name_pool + type rep = NitpickRep.rep + + datatype cst = + Unity | + False | + True | + Iden | + Num of int | + Unknown | + Unrep | + Suc | + Add | + Subtract | + Multiply | + Divide | + Modulo | + Gcd | + Lcm | + Fracs | + NormFrac | + NatToInt | + IntToNat + + datatype op1 = + Not | + Finite | + Converse | + Closure | + SingletonSet | + Tha | + First | + Second | + Cast + + datatype op2 = + All | + Exist | + Or | + And | + Less | + Subset | + DefEq | + Eq | + The | + Eps | + Triad | + Union | + SetDifference | + Intersect | + Composition | + Product | + Image | + Apply | + Lambda + + datatype op3 = + Let | + If + + datatype nut = + Cst of cst * typ * rep | + Op1 of op1 * typ * rep * nut | + Op2 of op2 * typ * rep * nut * nut | + Op3 of op3 * typ * rep * nut * nut * nut | + Tuple of typ * rep * nut list | + Construct of nut list * typ * rep * nut list | + BoundName of int * typ * rep * string | + FreeName of string * typ * rep | + ConstName of string * typ * rep | + BoundRel of Kodkod.n_ary_index * typ * rep * string | + FreeRel of Kodkod.n_ary_index * typ * rep * string | + RelReg of int * typ * rep | + FormulaReg of int * typ * rep + + structure NameTable : TABLE + + exception NUT of string * nut list + + val string_for_nut : Proof.context -> nut -> string + val inline_nut : nut -> bool + val type_of : nut -> typ + val rep_of : nut -> rep + val nickname_of : nut -> string + val is_skolem_name : nut -> bool + val is_eval_name : nut -> bool + val is_Cst : cst -> nut -> bool + val fold_nut : (nut -> 'a -> 'a) -> nut -> 'a -> 'a + val map_nut : (nut -> nut) -> nut -> nut + val untuple : (nut -> 'a) -> nut -> 'a list + val add_free_and_const_names : + nut -> nut list * nut list -> nut list * nut list + val name_ord : (nut * nut) -> order + val the_name : 'a NameTable.table -> nut -> 'a + val the_rel : nut NameTable.table -> nut -> Kodkod.n_ary_index + val nut_from_term : theory -> bool -> special_fun list -> op2 -> term -> nut + val choose_reps_for_free_vars : + scope -> nut list -> rep NameTable.table -> nut list * rep NameTable.table + val choose_reps_for_consts : + scope -> bool -> nut list -> rep NameTable.table + -> nut list * rep NameTable.table + val choose_reps_for_all_sels : + scope -> rep NameTable.table -> nut list * rep NameTable.table + val choose_reps_in_nut : + scope -> bool -> rep NameTable.table -> bool -> nut -> nut + val rename_free_vars : + nut list -> name_pool -> nut NameTable.table + -> nut list * name_pool * nut NameTable.table + val rename_vars_in_nut : name_pool -> nut NameTable.table -> nut -> nut +end; + +structure NitpickNut : NITPICK_NUT = +struct + +open NitpickUtil +open NitpickHOL +open NitpickScope +open NitpickPeephole +open NitpickRep + +datatype cst = + Unity | + False | + True | + Iden | + Num of int | + Unknown | + Unrep | + Suc | + Add | + Subtract | + Multiply | + Divide | + Modulo | + Gcd | + Lcm | + Fracs | + NormFrac | + NatToInt | + IntToNat + +datatype op1 = + Not | + Finite | + Converse | + Closure | + SingletonSet | + Tha | + First | + Second | + Cast + +datatype op2 = + All | + Exist | + Or | + And | + Less | + Subset | + DefEq | + Eq | + The | + Eps | + Triad | + Union | + SetDifference | + Intersect | + Composition | + Product | + Image | + Apply | + Lambda + +datatype op3 = + Let | + If + +datatype nut = + Cst of cst * typ * rep | + Op1 of op1 * typ * rep * nut | + Op2 of op2 * typ * rep * nut * nut | + Op3 of op3 * typ * rep * nut * nut * nut | + Tuple of typ * rep * nut list | + Construct of nut list * typ * rep * nut list | + BoundName of int * typ * rep * string | + FreeName of string * typ * rep | + ConstName of string * typ * rep | + BoundRel of Kodkod.n_ary_index * typ * rep * string | + FreeRel of Kodkod.n_ary_index * typ * rep * string | + RelReg of int * typ * rep | + FormulaReg of int * typ * rep + +exception NUT of string * nut list + +(* cst -> string *) +fun string_for_cst Unity = "Unity" + | string_for_cst False = "False" + | string_for_cst True = "True" + | string_for_cst Iden = "Iden" + | string_for_cst (Num j) = "Num " ^ signed_string_of_int j + | string_for_cst Unknown = "Unknown" + | string_for_cst Unrep = "Unrep" + | string_for_cst Suc = "Suc" + | string_for_cst Add = "Add" + | string_for_cst Subtract = "Subtract" + | string_for_cst Multiply = "Multiply" + | string_for_cst Divide = "Divide" + | string_for_cst Modulo = "Modulo" + | string_for_cst Gcd = "Gcd" + | string_for_cst Lcm = "Lcm" + | string_for_cst Fracs = "Fracs" + | string_for_cst NormFrac = "NormFrac" + | string_for_cst NatToInt = "NatToInt" + | string_for_cst IntToNat = "IntToNat" + +(* op1 -> string *) +fun string_for_op1 Not = "Not" + | string_for_op1 Finite = "Finite" + | string_for_op1 Converse = "Converse" + | string_for_op1 Closure = "Closure" + | string_for_op1 SingletonSet = "SingletonSet" + | string_for_op1 Tha = "Tha" + | string_for_op1 First = "First" + | string_for_op1 Second = "Second" + | string_for_op1 Cast = "Cast" + +(* op2 -> string *) +fun string_for_op2 All = "All" + | string_for_op2 Exist = "Exist" + | string_for_op2 Or = "Or" + | string_for_op2 And = "And" + | string_for_op2 Less = "Less" + | string_for_op2 Subset = "Subset" + | string_for_op2 DefEq = "DefEq" + | string_for_op2 Eq = "Eq" + | string_for_op2 The = "The" + | string_for_op2 Eps = "Eps" + | string_for_op2 Triad = "Triad" + | string_for_op2 Union = "Union" + | string_for_op2 SetDifference = "SetDifference" + | string_for_op2 Intersect = "Intersect" + | string_for_op2 Composition = "Composition" + | string_for_op2 Product = "Product" + | string_for_op2 Image = "Image" + | string_for_op2 Apply = "Apply" + | string_for_op2 Lambda = "Lambda" + +(* op3 -> string *) +fun string_for_op3 Let = "Let" + | string_for_op3 If = "If" + +(* int -> Proof.context -> nut -> string *) +fun basic_string_for_nut indent ctxt u = + let + (* nut -> string *) + val sub = basic_string_for_nut (indent + 1) ctxt + in + (if indent = 0 then "" else "\n" ^ implode (replicate (2 * indent) " ")) ^ + "(" ^ + (case u of + Cst (c, T, R) => + "Cst " ^ string_for_cst c ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R + | Op1 (oper, T, R, u1) => + "Op1 " ^ string_for_op1 oper ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R ^ " " ^ sub u1 + | Op2 (oper, T, R, u1, u2) => + "Op2 " ^ string_for_op2 oper ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R ^ " " ^ sub u1 ^ " " ^ sub u2 + | Op3 (oper, T, R, u1, u2, u3) => + "Op3 " ^ string_for_op3 oper ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R ^ " " ^ sub u1 ^ " " ^ sub u2 ^ " " ^ sub u3 + | Tuple (T, R, us) => + "Tuple " ^ Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ + implode (map sub us) + | Construct (us', T, R, us) => + "Construct " ^ implode (map sub us') ^ Syntax.string_of_typ ctxt T ^ + " " ^ string_for_rep R ^ " " ^ implode (map sub us) + | BoundName (j, T, R, nick) => + "BoundName " ^ signed_string_of_int j ^ " " ^ + Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ " " ^ nick + | FreeName (s, T, R) => + "FreeName " ^ s ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R + | ConstName (s, T, R) => + "ConstName " ^ s ^ " " ^ Syntax.string_of_typ ctxt T ^ " " ^ + string_for_rep R + | BoundRel ((n, j), T, R, nick) => + "BoundRel " ^ string_of_int n ^ "." ^ signed_string_of_int j ^ " " ^ + Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ " " ^ nick + | FreeRel ((n, j), T, R, nick) => + "FreeRel " ^ string_of_int n ^ "." ^ signed_string_of_int j ^ " " ^ + Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ " " ^ nick + | RelReg (j, T, R) => + "RelReg " ^ signed_string_of_int j ^ " " ^ Syntax.string_of_typ ctxt T ^ + " " ^ string_for_rep R + | FormulaReg (j, T, R) => + "FormulaReg " ^ signed_string_of_int j ^ " " ^ + Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R) ^ + ")" + end +(* Proof.context -> nut -> string *) +val string_for_nut = basic_string_for_nut 0 + +(* nut -> bool *) +fun inline_nut (Op1 _) = false + | inline_nut (Op2 _) = false + | inline_nut (Op3 _) = false + | inline_nut (Tuple (_, _, us)) = forall inline_nut us + | inline_nut _ = true + +(* nut -> typ *) +fun type_of (Cst (_, T, _)) = T + | type_of (Op1 (_, T, _, _)) = T + | type_of (Op2 (_, T, _, _, _)) = T + | type_of (Op3 (_, T, _, _, _, _)) = T + | type_of (Tuple (T, _, _)) = T + | type_of (Construct (_, T, _, _)) = T + | type_of (BoundName (_, T, _, _)) = T + | type_of (FreeName (_, T, _)) = T + | type_of (ConstName (_, T, _)) = T + | type_of (BoundRel (_, T, _, _)) = T + | type_of (FreeRel (_, T, _, _)) = T + | type_of (RelReg (_, T, _)) = T + | type_of (FormulaReg (_, T, _)) = T + +(* nut -> rep *) +fun rep_of (Cst (_, _, R)) = R + | rep_of (Op1 (_, _, R, _)) = R + | rep_of (Op2 (_, _, R, _, _)) = R + | rep_of (Op3 (_, _, R, _, _, _)) = R + | rep_of (Tuple (_, R, _)) = R + | rep_of (Construct (_, _, R, _)) = R + | rep_of (BoundName (_, _, R, _)) = R + | rep_of (FreeName (_, _, R)) = R + | rep_of (ConstName (_, _, R)) = R + | rep_of (BoundRel (_, _, R, _)) = R + | rep_of (FreeRel (_, _, R, _)) = R + | rep_of (RelReg (_, _, R)) = R + | rep_of (FormulaReg (_, _, R)) = R + +(* nut -> string *) +fun nickname_of (BoundName (_, _, _, nick)) = nick + | nickname_of (FreeName (s, _, _)) = s + | nickname_of (ConstName (s, _, _)) = s + | nickname_of (BoundRel (_, _, _, nick)) = nick + | nickname_of (FreeRel (_, _, _, nick)) = nick + | nickname_of u = raise NUT ("NitpickNut.nickname_of", [u]) + +(* nut -> bool *) +fun is_skolem_name u = + space_explode name_sep (nickname_of u) + |> exists (String.isPrefix skolem_prefix) + handle NUT ("NitpickNut.nickname_of", _) => false +fun is_eval_name u = + String.isPrefix eval_prefix (nickname_of u) + handle NUT ("NitpickNut.nickname_of", _) => false +(* cst -> nut -> bool *) +fun is_Cst cst (Cst (cst', _, _)) = (cst = cst') + | is_Cst _ _ = false + +(* (nut -> 'a -> 'a) -> nut -> 'a -> 'a *) +fun fold_nut f u = + case u of + Op1 (_, _, _, u1) => fold_nut f u1 + | Op2 (_, _, _, u1, u2) => fold_nut f u1 #> fold_nut f u2 + | Op3 (_, _, _, u1, u2, u3) => fold_nut f u1 #> fold_nut f u2 #> fold_nut f u3 + | Tuple (_, _, us) => fold (fold_nut f) us + | Construct (us', _, _, us) => fold (fold_nut f) us #> fold (fold_nut f) us' + | _ => f u +(* (nut -> nut) -> nut -> nut *) +fun map_nut f u = + case u of + Op1 (oper, T, R, u1) => Op1 (oper, T, R, map_nut f u1) + | Op2 (oper, T, R, u1, u2) => Op2 (oper, T, R, map_nut f u1, map_nut f u2) + | Op3 (oper, T, R, u1, u2, u3) => + Op3 (oper, T, R, map_nut f u1, map_nut f u2, map_nut f u3) + | Tuple (T, R, us) => Tuple (T, R, map (map_nut f) us) + | Construct (us', T, R, us) => + Construct (map (map_nut f) us', T, R, map (map_nut f) us) + | _ => f u + +(* nut * nut -> order *) +fun name_ord (BoundName (j1, _, _, _), BoundName (j2, _, _, _)) = + int_ord (j1, j2) + | name_ord (BoundName _, _) = LESS + | name_ord (_, BoundName _) = GREATER + | name_ord (FreeName (s1, T1, _), FreeName (s2, T2, _)) = + (case fast_string_ord (s1, s2) of + EQUAL => TermOrd.typ_ord (T1, T2) + | ord => ord) + | name_ord (FreeName _, _) = LESS + | name_ord (_, FreeName _) = GREATER + | name_ord (ConstName (s1, T1, _), ConstName (s2, T2, _)) = + (case fast_string_ord (s1, s2) of + EQUAL => TermOrd.typ_ord (T1, T2) + | ord => ord) + | name_ord (u1, u2) = raise NUT ("NitpickNut.name_ord", [u1, u2]) + +(* nut -> nut -> int *) +fun num_occs_in_nut needle_u stack_u = + fold_nut (fn u => if u = needle_u then Integer.add 1 else I) stack_u 0 +(* nut -> nut -> bool *) +val is_subterm_of = not_equal 0 oo num_occs_in_nut + +(* nut -> nut -> nut -> nut *) +fun substitute_in_nut needle_u needle_u' = + map_nut (fn u => if u = needle_u then needle_u' else u) + +(* nut -> nut list * nut list -> nut list * nut list *) +val add_free_and_const_names = + fold_nut (fn u => case u of + FreeName _ => apfst (insert (op =) u) + | ConstName _ => apsnd (insert (op =) u) + | _ => I) + +(* nut -> rep -> nut *) +fun modify_name_rep (BoundName (j, T, _, nick)) R = BoundName (j, T, R, nick) + | modify_name_rep (FreeName (s, T, _)) R = FreeName (s, T, R) + | modify_name_rep (ConstName (s, T, _)) R = ConstName (s, T, R) + | modify_name_rep u _ = raise NUT ("NitpickNut.modify_name_rep", [u]) + +structure NameTable = Table(type key = nut val ord = name_ord) + +(* 'a NameTable.table -> nut -> 'a *) +fun the_name table name = + case NameTable.lookup table name of + SOME u => u + | NONE => raise NUT ("NitpickNut.the_name", [name]) +(* nut NameTable.table -> nut -> Kodkod.n_ary_index *) +fun the_rel table name = + case the_name table name of + FreeRel (x, _, _, _) => x + | u => raise NUT ("NitpickNut.the_rel", [u]) + +(* typ * term -> typ * term *) +fun mk_fst (_, Const (@{const_name Pair}, T) $ t1 $ _) = (domain_type T, t1) + | mk_fst (T, t) = + let val res_T = fst (HOLogic.dest_prodT T) in + (res_T, Const (@{const_name fst}, T --> res_T) $ t) + end +fun mk_snd (_, Const (@{const_name Pair}, T) $ _ $ t2) = + (domain_type (range_type T), t2) + | mk_snd (T, t) = + let val res_T = snd (HOLogic.dest_prodT T) in + (res_T, Const (@{const_name snd}, T --> res_T) $ t) + end +(* typ * term -> (typ * term) list *) +fun factorize (z as (Type ("*", _), _)) = maps factorize [mk_fst z, mk_snd z] + | factorize z = [z] + +(* theory -> bool -> special_fun list -> op2 -> term -> nut *) +fun nut_from_term thy fast_descrs special_funs eq = + let + (* string list -> typ list -> term -> nut *) + fun aux eq ss Ts t = + let + (* term -> nut *) + val sub = aux Eq ss Ts + val sub' = aux eq ss Ts + (* string -> typ -> term -> nut *) + fun sub_abs s T = aux eq (s :: ss) (T :: Ts) + (* typ -> term -> term -> nut *) + fun sub_equals T t1 t2 = + let + val (binder_Ts, body_T) = strip_type (domain_type T) + val n = length binder_Ts + in + if eq = Eq andalso n > 0 then + let + val t1 = incr_boundvars n t1 + val t2 = incr_boundvars n t2 + val xs = map Bound (n - 1 downto 0) + val equation = Const (@{const_name "op ="}, + body_T --> body_T --> bool_T) + $ betapplys (t1, xs) $ betapplys (t2, xs) + val t = + fold_rev (fn T => fn (t, j) => + (Const (@{const_name All}, T --> bool_T) + $ Abs ("x" ^ nat_subscript j, T, t), j - 1)) + binder_Ts (equation, n) |> fst + in sub' t end + else + Op2 (eq, bool_T, Any, aux Eq ss Ts t1, aux Eq ss Ts t2) + end + (* op2 -> string -> typ -> term -> nut *) + fun do_quantifier quant s T t1 = + let + val bound_u = BoundName (length Ts, T, Any, s) + val body_u = sub_abs s T t1 + in + if is_subterm_of bound_u body_u then + Op2 (quant, bool_T, Any, bound_u, body_u) + else + body_u + end + (* term -> term list -> nut *) + fun do_apply t0 ts = + let + val (ts', t2) = split_last ts + val t1 = list_comb (t0, ts') + val T1 = fastype_of1 (Ts, t1) + in Op2 (Apply, range_type T1, Any, sub t1, sub t2) end + in + case strip_comb t of + (Const (@{const_name all}, _), [Abs (s, T, t1)]) => + do_quantifier All s T t1 + | (t0 as Const (@{const_name all}, T), [t1]) => + sub' (t0 $ eta_expand Ts t1 1) + | (Const (@{const_name "=="}, T), [t1, t2]) => sub_equals T t1 t2 + | (Const (@{const_name "==>"}, _), [t1, t2]) => + Op2 (Or, prop_T, Any, Op1 (Not, prop_T, Any, sub t1), sub' t2) + | (Const (@{const_name Pure.conjunction}, _), [t1, t2]) => + Op2 (And, prop_T, Any, sub' t1, sub' t2) + | (Const (@{const_name Trueprop}, _), [t1]) => sub' t1 + | (Const (@{const_name Not}, _), [t1]) => + (case sub t1 of + Op1 (Not, _, _, u11) => u11 + | u1 => Op1 (Not, bool_T, Any, u1)) + | (Const (@{const_name False}, T), []) => Cst (False, T, Any) + | (Const (@{const_name True}, T), []) => Cst (True, T, Any) + | (Const (@{const_name All}, _), [Abs (s, T, t1)]) => + do_quantifier All s T t1 + | (t0 as Const (@{const_name All}, T), [t1]) => + sub' (t0 $ eta_expand Ts t1 1) + | (Const (@{const_name Ex}, _), [Abs (s, T, t1)]) => + do_quantifier Exist s T t1 + | (t0 as Const (@{const_name Ex}, T), [t1]) => + sub' (t0 $ eta_expand Ts t1 1) + | (t0 as Const (@{const_name The}, T), [t1]) => + if fast_descrs then + Op2 (The, range_type T, Any, sub t1, + sub (Const (@{const_name undefined_fast_The}, range_type T))) + else + do_apply t0 [t1] + | (t0 as Const (@{const_name Eps}, T), [t1]) => + if fast_descrs then + Op2 (Eps, range_type T, Any, sub t1, + sub (Const (@{const_name undefined_fast_Eps}, range_type T))) + else + do_apply t0 [t1] + | (Const (@{const_name "op ="}, T), [t1, t2]) => sub_equals T t1 t2 + | (Const (@{const_name "op &"}, _), [t1, t2]) => + Op2 (And, bool_T, Any, sub' t1, sub' t2) + | (Const (@{const_name "op |"}, _), [t1, t2]) => + Op2 (Or, bool_T, Any, sub t1, sub t2) + | (Const (@{const_name "op -->"}, _), [t1, t2]) => + Op2 (Or, bool_T, Any, Op1 (Not, bool_T, Any, sub t1), sub' t2) + | (Const (@{const_name If}, T), [t1, t2, t3]) => + Op3 (If, nth_range_type 3 T, Any, sub t1, sub t2, sub t3) + | (Const (@{const_name Let}, T), [t1, Abs (s, T', t2)]) => + Op3 (Let, nth_range_type 2 T, Any, BoundName (length Ts, T', Any, s), + sub t1, sub_abs s T' t2) + | (t0 as Const (@{const_name Let}, T), [t1, t2]) => + sub (t0 $ t1 $ eta_expand Ts t2 1) + | (Const (@{const_name unknown}, T), []) => Cst (Unknown, T, Any) + | (@{const Unity}, []) => Cst (Unity, @{typ unit}, Any) + | (Const (@{const_name Pair}, T), [t1, t2]) => + Tuple (nth_range_type 2 T, Any, map sub [t1, t2]) + | (Const (@{const_name fst}, T), [t1]) => + Op1 (First, range_type T, Any, sub t1) + | (Const (@{const_name snd}, T), [t1]) => + Op1 (Second, range_type T, Any, sub t1) + | (Const (@{const_name Id}, T), []) => Cst (Iden, T, Any) + | (Const (@{const_name insert}, T), [t1, t2]) => + (case t2 of + Abs (_, _, @{const False}) => + Op1 (SingletonSet, nth_range_type 2 T, Any, sub t1) + | _ => + Op2 (Union, nth_range_type 2 T, Any, + Op1 (SingletonSet, nth_range_type 2 T, Any, sub t1), sub t2)) + | (Const (@{const_name converse}, T), [t1]) => + Op1 (Converse, range_type T, Any, sub t1) + | (Const (@{const_name trancl}, T), [t1]) => + Op1 (Closure, range_type T, Any, sub t1) + | (Const (@{const_name rel_comp}, T), [t1, t2]) => + Op2 (Composition, nth_range_type 2 T, Any, sub t1, sub t2) + | (Const (@{const_name Sigma}, T), [t1, Abs (s, T', t2')]) => + Op2 (Product, nth_range_type 2 T, Any, sub t1, sub_abs s T' t2') + | (Const (@{const_name image}, T), [t1, t2]) => + Op2 (Image, nth_range_type 2 T, Any, sub t1, sub t2) + | (Const (@{const_name Suc}, T), []) => Cst (Suc, T, Any) + | (Const (@{const_name finite}, T), [t1]) => + Op1 (Finite, bool_T, Any, sub t1) + | (Const (@{const_name nat}, T), []) => Cst (IntToNat, T, Any) + | (Const (@{const_name zero_nat_inst.zero_nat}, T), []) => + Cst (Num 0, T, Any) + | (Const (@{const_name one_nat_inst.one_nat}, T), []) => + Cst (Num 1, T, Any) + | (Const (@{const_name plus_nat_inst.plus_nat}, T), []) => + Cst (Add, T, Any) + | (Const (@{const_name minus_nat_inst.minus_nat}, T), []) => + Cst (Subtract, T, Any) + | (Const (@{const_name times_nat_inst.times_nat}, T), []) => + Cst (Multiply, T, Any) + | (Const (@{const_name div_nat_inst.div_nat}, T), []) => + Cst (Divide, T, Any) + | (Const (@{const_name div_nat_inst.mod_nat}, T), []) => + Cst (Modulo, T, Any) + | (Const (@{const_name ord_nat_inst.less_nat}, T), [t1, t2]) => + Op2 (Less, bool_T, Any, sub t1, sub t2) + | (Const (@{const_name ord_nat_inst.less_eq_nat}, T), [t1, t2]) => + Op1 (Not, bool_T, Any, Op2 (Less, bool_T, Any, sub t2, sub t1)) + | (Const (@{const_name nat_gcd}, T), []) => Cst (Gcd, T, Any) + | (Const (@{const_name nat_lcm}, T), []) => Cst (Lcm, T, Any) + | (Const (@{const_name zero_int_inst.zero_int}, T), []) => + Cst (Num 0, T, Any) + | (Const (@{const_name one_int_inst.one_int}, T), []) => + Cst (Num 1, T, Any) + | (Const (@{const_name plus_int_inst.plus_int}, T), []) => + Cst (Add, T, Any) + | (Const (@{const_name minus_int_inst.minus_int}, T), []) => + Cst (Subtract, T, Any) + | (Const (@{const_name times_int_inst.times_int}, T), []) => + Cst (Multiply, T, Any) + | (Const (@{const_name div_int_inst.div_int}, T), []) => + Cst (Divide, T, Any) + | (Const (@{const_name div_int_inst.mod_int}, T), []) => + Cst (Modulo, T, Any) + | (Const (@{const_name uminus_int_inst.uminus_int}, T), []) => + Op2 (Apply, int_T --> int_T, Any, + Cst (Subtract, [int_T, int_T] ---> int_T, Any), + Cst (Num 0, int_T, Any)) + | (Const (@{const_name ord_int_inst.less_int}, T), [t1, t2]) => + Op2 (Less, bool_T, Any, sub t1, sub t2) + | (Const (@{const_name ord_int_inst.less_eq_int}, T), [t1, t2]) => + Op1 (Not, bool_T, Any, Op2 (Less, bool_T, Any, sub t2, sub t1)) + | (Const (@{const_name Tha}, Type ("fun", [_, T2])), [t1]) => + Op1 (Tha, T2, Any, sub t1) + | (Const (@{const_name Frac}, T), []) => Cst (Fracs, T, Any) + | (Const (@{const_name norm_frac}, T), []) => Cst (NormFrac, T, Any) + | (Const (@{const_name of_nat}, T as @{typ "nat => int"}), []) => + Cst (NatToInt, T, Any) + | (Const (@{const_name lower_semilattice_fun_inst.inf_fun}, T), + [t1, t2]) => + Op2 (Intersect, nth_range_type 2 T, Any, sub t1, sub t2) + | (Const (@{const_name upper_semilattice_fun_inst.sup_fun}, T), + [t1, t2]) => + Op2 (Union, nth_range_type 2 T, Any, sub t1, sub t2) + | (t0 as Const (@{const_name minus_fun_inst.minus_fun}, T), [t1, t2]) => + Op2 (SetDifference, nth_range_type 2 T, Any, sub t1, sub t2) + | (t0 as Const (@{const_name ord_fun_inst.less_eq_fun}, T), [t1, t2]) => + Op2 (Subset, bool_T, Any, sub t1, sub t2) + | (t0 as Const (x as (s, T)), ts) => + if is_constr thy x then + case num_binder_types T - length ts of + 0 => Construct (map ((fn (s, T) => ConstName (s, T, Any)) + o nth_sel_for_constr x) + (~1 upto num_sels_for_constr_type T - 1), + body_type T, Any, + ts |> map (`(curry fastype_of1 Ts)) + |> maps factorize |> map (sub o snd)) + | k => sub (eta_expand Ts t k) + else if String.isPrefix numeral_prefix s then + Cst (Num (the (Int.fromString (unprefix numeral_prefix s))), T, Any) + else + (case arity_of_built_in_const fast_descrs x of + SOME n => + (case n - length ts of + 0 => raise TERM ("NitpickNut.nut_from_term.aux", [t]) + | k => if k > 0 then sub (eta_expand Ts t k) + else do_apply t0 ts) + | NONE => if null ts then ConstName (s, T, Any) + else do_apply t0 ts) + | (Free (s, T), []) => FreeName (s, T, Any) + | (Var _, []) => raise TERM ("NitpickNut.nut_from_term", [t]) + | (Bound j, []) => + BoundName (length Ts - j - 1, nth Ts j, Any, nth ss j) + | (Abs (s, T, t1), []) => + Op2 (Lambda, T --> fastype_of1 (T :: Ts, t1), Any, + BoundName (length Ts, T, Any, s), sub_abs s T t1) + | (t0, ts) => do_apply t0 ts + end + in aux eq [] [] end + +(* scope -> typ -> rep *) +fun rep_for_abs_fun scope T = + let val (R1, R2) = best_non_opt_symmetric_reps_for_fun_type scope T in + Func (R1, (card_of_rep R1 <> card_of_rep R2 ? Opt) R2) + end + +(* scope -> nut -> nut list * rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_rep_for_free_var scope v (vs, table) = + let + val R = best_non_opt_set_rep_for_type scope (type_of v) + val v = modify_name_rep v R + in (v :: vs, NameTable.update (v, R) table) end +(* scope -> bool -> nut -> nut list * rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_rep_for_const (scope as {ext_ctxt as {thy, ctxt, ...}, datatypes, + ofs, ...}) all_precise v (vs, table) = + let + val x as (s, T) = (nickname_of v, type_of v) + val R = (if is_abs_fun thy x then + rep_for_abs_fun + else if is_rep_fun thy x then + Func oo best_non_opt_symmetric_reps_for_fun_type + else if all_precise orelse is_skolem_name v + orelse s mem [@{const_name undefined_fast_The}, + @{const_name undefined_fast_Eps}, + @{const_name bisim}] then + best_non_opt_set_rep_for_type + else if original_name s + mem [@{const_name set}, @{const_name distinct}, + @{const_name ord_class.less}, + @{const_name ord_class.less_eq}, + @{const_name bisim_iterator_max}] then + best_set_rep_for_type + else + best_opt_set_rep_for_type) scope T + val v = modify_name_rep v R + in (v :: vs, NameTable.update (v, R) table) end + +(* scope -> nut list -> rep NameTable.table -> nut list * rep NameTable.table *) +fun choose_reps_for_free_vars scope vs table = + fold (choose_rep_for_free_var scope) vs ([], table) +(* scope -> bool -> nut list -> rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_reps_for_consts scope all_precise vs table = + fold (choose_rep_for_const scope all_precise) vs ([], table) + +(* scope -> styp -> int -> nut list * rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_rep_for_nth_sel_for_constr (scope as {ext_ctxt, ...}) x n + (vs, table) = + let + val (s', T') = boxed_nth_sel_for_constr ext_ctxt x n + val R' = if n = ~1 then best_non_opt_set_rep_for_type scope T' + else best_opt_set_rep_for_type scope T' |> unopt_rep + val v = ConstName (s', T', R') + in (v :: vs, NameTable.update (v, R') table) end +(* scope -> styp -> nut list * rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_rep_for_sels_for_constr scope (x as (_, T)) = + fold_rev (choose_rep_for_nth_sel_for_constr scope x) + (~1 upto num_sels_for_constr_type T - 1) +(* scope -> dtype_spec -> nut list * rep NameTable.table + -> nut list * rep NameTable.table *) +fun choose_rep_for_sels_of_datatype scope ({constrs, ...} : dtype_spec) = + fold_rev (choose_rep_for_sels_for_constr scope o #const) constrs +(* scope -> rep NameTable.table -> nut list * rep NameTable.table *) +fun choose_reps_for_all_sels (scope as {datatypes, ...}) = + fold (choose_rep_for_sels_of_datatype scope) datatypes o pair [] + +(* scope -> nut -> rep NameTable.table -> rep NameTable.table *) +fun choose_rep_for_bound_var scope v table = + let val R = best_one_rep_for_type scope (type_of v) in + NameTable.update (v, R) table + end + +(* A nut is said to be constructive if whenever it evaluates to unknown in our + three-valued logic, it would evaluate to a unrepresentable value ("unrep") + according to the HOL semantics. For example, "Suc n" is + constructive if "n" is representable or "Unrep", because unknown implies + unrep. *) +(* nut -> bool *) +fun is_constructive u = + is_Cst Unrep u orelse + (not (is_fun_type (type_of u)) andalso not (is_opt_rep (rep_of u))) orelse + case u of + Cst (Num _, _, _) => true + | Cst (cst, T, _) => cst = Suc orelse (body_type T = nat_T andalso cst = Add) + | Op2 (Apply, _, _, u1, u2) => forall is_constructive [u1, u2] + | Op3 (If, _, _, u1, u2, u3) => + not (is_opt_rep (rep_of u1)) andalso forall is_constructive [u2, u3] + | Tuple (_, _, us) => forall is_constructive us + | Construct (_, _, _, us) => forall is_constructive us + | _ => false + +(* nut -> nut *) +fun optimize_unit u = + if rep_of u = Unit then Cst (Unity, type_of u, Unit) else u +(* typ -> rep -> nut *) +fun unknown_boolean T R = + Cst (case R of + Formula Pos => False + | Formula Neg => True + | _ => Unknown, T, R) + +(* op1 -> typ -> rep -> nut -> nut *) +fun s_op1 oper T R u1 = + ((if oper = Not then + if is_Cst True u1 then Cst (False, T, R) + else if is_Cst False u1 then Cst (True, T, R) + else raise SAME () + else + raise SAME ()) + handle SAME () => Op1 (oper, T, R, u1)) + |> optimize_unit +(* op2 -> typ -> rep -> nut -> nut -> nut *) +fun s_op2 oper T R u1 u2 = + ((case oper of + Or => + if exists (is_Cst True) [u1, u2] then Cst (True, T, unopt_rep R) + else if is_Cst False u1 then u2 + else if is_Cst False u2 then u1 + else raise SAME () + | And => + if exists (is_Cst False) [u1, u2] then Cst (False, T, unopt_rep R) + else if is_Cst True u1 then u2 + else if is_Cst True u2 then u1 + else raise SAME () + | Eq => + (case pairself (is_Cst Unrep) (u1, u2) of + (true, true) => unknown_boolean T R + | (false, false) => raise SAME () + | _ => if forall (is_opt_rep o rep_of) [u1, u2] then raise SAME () + else Cst (False, T, Formula Neut)) + | Triad => + if is_Cst True u1 then u1 + else if is_Cst False u2 then u2 + else raise SAME () + | Apply => + if is_Cst Unrep u1 then + Cst (Unrep, T, R) + else if is_Cst Unrep u2 then + if is_constructive u1 then + Cst (Unrep, T, R) + else if is_boolean_type T andalso not (is_opt_rep (rep_of u1)) then + (* Selectors are an unfortunate exception to the rule that non-"Opt" + predicates return "False" for unrepresentable domain values. *) + case u1 of + ConstName (s, _, _) => if is_sel s then unknown_boolean T R + else Cst (False, T, Formula Neut) + | _ => Cst (False, T, Formula Neut) + else case u1 of + Op2 (Apply, _, _, ConstName (@{const_name List.append}, _, _), _) => + Cst (Unrep, T, R) + | _ => raise SAME () + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => Op2 (oper, T, R, u1, u2)) + |> optimize_unit +(* op3 -> typ -> rep -> nut -> nut -> nut -> nut *) +fun s_op3 oper T R u1 u2 u3 = + ((case oper of + Let => + if inline_nut u2 orelse num_occs_in_nut u1 u3 < 2 then + substitute_in_nut u1 u2 u3 + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => Op3 (oper, T, R, u1, u2, u3)) + |> optimize_unit +(* typ -> rep -> nut list -> nut *) +fun s_tuple T R us = + (if exists (is_Cst Unrep) us then Cst (Unrep, T, R) else Tuple (T, R, us)) + |> optimize_unit + +(* theory -> nut -> nut *) +fun optimize_nut u = + case u of + Op1 (oper, T, R, u1) => s_op1 oper T R (optimize_nut u1) + | Op2 (oper, T, R, u1, u2) => + s_op2 oper T R (optimize_nut u1) (optimize_nut u2) + | Op3 (oper, T, R, u1, u2, u3) => + s_op3 oper T R (optimize_nut u1) (optimize_nut u2) (optimize_nut u3) + | Tuple (T, R, us) => s_tuple T R (map optimize_nut us) + | Construct (us', T, R, us) => Construct (us', T, R, map optimize_nut us) + | _ => optimize_unit u + +(* (nut -> 'a) -> nut -> 'a list *) +fun untuple f (Tuple (_, _, us)) = maps (untuple f) us + | untuple f u = if rep_of u = Unit then [] else [f u] + +(* scope -> bool -> rep NameTable.table -> bool -> nut -> nut *) +fun choose_reps_in_nut (scope as {ext_ctxt as {thy, ctxt, ...}, card_assigns, + datatypes, ofs, ...}) liberal table def = + let + val bool_atom_R = Atom (2, offset_of_type ofs bool_T) + (* polarity -> bool -> rep *) + fun bool_rep polar opt = + if polar = Neut andalso opt then Opt bool_atom_R else Formula polar + (* nut -> nut -> nut *) + fun triad u1 u2 = s_op2 Triad (type_of u1) (Opt bool_atom_R) u1 u2 + (* (polarity -> nut) -> nut *) + fun triad_fn f = triad (f Pos) (f Neg) + (* rep NameTable.table -> bool -> polarity -> nut -> nut -> nut *) + fun unrepify_nut_in_nut table def polar needle_u = + let val needle_T = type_of needle_u in + substitute_in_nut needle_u (Cst (if is_fun_type needle_T then Unknown + else Unrep, needle_T, Any)) + #> aux table def polar + end + (* rep NameTable.table -> bool -> polarity -> nut -> nut *) + and aux table def polar u = + let + (* bool -> polarity -> nut -> nut *) + val gsub = aux table + (* nut -> nut *) + val sub = gsub false Neut + in + case u of + Cst (False, T, _) => Cst (False, T, Formula Neut) + | Cst (True, T, _) => Cst (True, T, Formula Neut) + | Cst (Num j, T, _) => + (case spec_of_type scope T of + (1, j0) => if j = 0 then Cst (Unity, T, Unit) + else Cst (Unrep, T, Opt (Atom (1, j0))) + | (k, j0) => + let + val ok = (if T = int_T then atom_for_int (k, j0) j <> ~1 + else j < k) + in + if ok then Cst (Num j, T, Atom (k, j0)) + else Cst (Unrep, T, Opt (Atom (k, j0))) + end) + | Cst (Suc, T as Type ("fun", [T1, _]), _) => + let val R = Atom (spec_of_type scope T1) in + Cst (Suc, T, Func (R, Opt R)) + end + | Cst (Fracs, T, _) => + Cst (Fracs, T, best_non_opt_set_rep_for_type scope T) + | Cst (NormFrac, T, _) => + let val R1 = Atom (spec_of_type scope (domain_type T)) in + Cst (NormFrac, T, Func (R1, Func (R1, Opt (Struct [R1, R1])))) + end + | Cst (cst, T, _) => + if cst mem [Unknown, Unrep] then + case (is_boolean_type T, polar) of + (true, Pos) => Cst (False, T, Formula Pos) + | (true, Neg) => Cst (True, T, Formula Neg) + | _ => Cst (cst, T, best_opt_set_rep_for_type scope T) + else if cst mem [Add, Subtract, Multiply, Divide, Modulo, Gcd, + Lcm] then + let + val T1 = domain_type T + val R1 = Atom (spec_of_type scope T1) + val total = + T1 = nat_T andalso cst mem [Subtract, Divide, Modulo, Gcd] + in Cst (cst, T, Func (R1, Func (R1, (not total ? Opt) R1))) end + else if cst mem [NatToInt, IntToNat] then + let + val (nat_card, nat_j0) = spec_of_type scope nat_T + val (int_card, int_j0) = spec_of_type scope int_T + in + if cst = NatToInt then + let val total = (max_int_for_card int_card >= nat_card + 1) in + Cst (cst, T, + Func (Atom (nat_card, nat_j0), + (not total ? Opt) (Atom (int_card, int_j0)))) + end + else + let val total = (max_int_for_card int_card < nat_card) in + Cst (cst, T, Func (Atom (int_card, int_j0), + Atom (nat_card, nat_j0)) |> not total ? Opt) + end + end + else + Cst (cst, T, best_set_rep_for_type scope T) + | Op1 (Not, T, _, u1) => + (case gsub def (flip_polarity polar) u1 of + Op2 (Triad, T, R, u11, u12) => + triad (s_op1 Not T (Formula Pos) u12) + (s_op1 Not T (Formula Neg) u11) + | u1' => s_op1 Not T (flip_rep_polarity (rep_of u1')) u1') + | Op1 (oper, T, _, u1) => + let + val u1 = sub u1 + val R1 = rep_of u1 + val R = case oper of + Finite => bool_rep polar (is_opt_rep R1) + | _ => (if is_opt_rep R1 then best_opt_set_rep_for_type + else best_non_opt_set_rep_for_type) scope T + in s_op1 oper T R u1 end + | Op2 (Less, T, _, u1, u2) => + let + val u1 = sub u1 + val u2 = sub u2 + val R = bool_rep polar (exists (is_opt_rep o rep_of) [u1, u2]) + in s_op2 Less T R u1 u2 end + | Op2 (Subset, T, _, u1, u2) => + let + val u1 = sub u1 + val u2 = sub u2 + val opt = exists (is_opt_rep o rep_of) [u1, u2] + val R = bool_rep polar opt + in + if is_opt_rep R then + triad_fn (fn polar => s_op2 Subset T (Formula polar) u1 u2) + else if opt andalso polar = Pos andalso + not (is_fully_comparable_type datatypes (type_of u1)) then + Cst (False, T, Formula Pos) + else + s_op2 Subset T R u1 u2 + end + | Op2 (DefEq, T, _, u1, u2) => + s_op2 DefEq T (Formula Neut) (sub u1) (sub u2) + | Op2 (Eq, T, _, u1, u2) => + let + val u1' = sub u1 + val u2' = sub u2 + (* unit -> nut *) + fun non_opt_case () = s_op2 Eq T (Formula polar) u1' u2' + (* unit -> nut *) + fun opt_opt_case () = + if polar = Neut then + triad_fn (fn polar => s_op2 Eq T (Formula polar) u1' u2') + else + non_opt_case () + (* nut -> nut *) + fun hybrid_case u = + (* hackish optimization *) + if is_constructive u then s_op2 Eq T (Formula Neut) u1' u2' + else opt_opt_case () + in + if liberal orelse polar = Neg + orelse is_fully_comparable_type datatypes (type_of u1) then + case (is_opt_rep (rep_of u1'), is_opt_rep (rep_of u2')) of + (true, true) => opt_opt_case () + | (true, false) => hybrid_case u1' + | (false, true) => hybrid_case u2' + | (false, false) => non_opt_case () + else + Cst (False, T, Formula Pos) + |> polar = Neut ? (fn pos_u => triad pos_u (gsub def Neg u)) + end + | Op2 (Image, T, _, u1, u2) => + let + val u1' = sub u1 + val u2' = sub u2 + in + (case (rep_of u1', rep_of u2') of + (Func (R11, R12), Func (R21, Formula Neut)) => + if R21 = R11 andalso is_lone_rep R12 then + let + val R = + best_non_opt_set_rep_for_type scope T + |> exists (is_opt_rep o rep_of) [u1', u2'] ? opt_rep ofs T + in s_op2 Image T R u1' u2' end + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => + let + val T1 = type_of u1 + val dom_T = domain_type T1 + val ran_T = range_type T1 + val x_u = BoundName (~1, dom_T, Any, "image.x") + val y_u = BoundName (~2, ran_T, Any, "image.y") + in + Op2 (Lambda, T, Any, y_u, + Op2 (Exist, bool_T, Any, x_u, + Op2 (And, bool_T, Any, + case u2 of + Op2 (Lambda, _, _, u21, u22) => + if num_occs_in_nut u21 u22 = 0 then (* FIXME: move to s_op2 *) + u22 + else + Op2 (Apply, bool_T, Any, u2, x_u) + | _ => Op2 (Apply, bool_T, Any, u2, x_u), + Op2 (Eq, bool_T, Any, y_u, + Op2 (Apply, ran_T, Any, u1, x_u))))) + |> sub + end + end + | Op2 (Apply, T, _, u1, u2) => + let + val u1 = sub u1 + val u2 = sub u2 + val T1 = type_of u1 + val R1 = rep_of u1 + val R2 = rep_of u2 + val opt = + case (u1, is_opt_rep R2) of + (ConstName (@{const_name set}, _, _), false) => false + | _ => exists is_opt_rep [R1, R2] + val ran_R = + if is_boolean_type T then + bool_rep polar opt + else + smart_range_rep ofs T1 (fn () => card_of_type card_assigns T) + (unopt_rep R1) + |> opt ? opt_rep ofs T + in s_op2 Apply T ran_R u1 u2 end + | Op2 (Lambda, T, _, u1, u2) => + (case best_set_rep_for_type scope T of + Unit => Cst (Unity, T, Unit) + | R as Func (R1, _) => + let + val table' = NameTable.update (u1, R1) table + val u1' = aux table' false Neut u1 + val u2' = aux table' false Neut u2 + val R = + if is_opt_rep (rep_of u2') + orelse (range_type T = bool_T andalso + not (is_Cst False + (unrepify_nut_in_nut table false Neut + u1 u2 + |> optimize_nut))) then + opt_rep ofs T R + else + unopt_rep R + in s_op2 Lambda T R u1' u2' end + | R => raise NUT ("NitpickNut.aux.choose_reps_in_nut", [u])) + | Op2 (oper, T, _, u1, u2) => + if oper mem [All, Exist] then + let + val table' = fold (choose_rep_for_bound_var scope) (untuple I u1) + table + val u1' = aux table' def polar u1 + val u2' = aux table' def polar u2 + in + if polar = Neut andalso is_opt_rep (rep_of u2') then + triad_fn (fn polar => gsub def polar u) + else + let val quant_u = s_op2 oper T (Formula polar) u1' u2' in + if def + orelse (liberal andalso (polar = Pos) = (oper = All)) + orelse is_precise_type datatypes (type_of u1) then + quant_u + else + let + val connective = if oper = All then And else Or + val unrepified_u = unrepify_nut_in_nut table def polar + u1 u2 + in + s_op2 connective T + (min_rep (rep_of quant_u) (rep_of unrepified_u)) + quant_u unrepified_u + end + end + end + else if oper mem [Or, And] then + let + val u1' = gsub def polar u1 + val u2' = gsub def polar u2 + in + (if polar = Neut then + case (is_opt_rep (rep_of u1'), is_opt_rep (rep_of u2')) of + (true, true) => triad_fn (fn polar => gsub def polar u) + | (true, false) => + s_op2 oper T (Opt bool_atom_R) + (triad_fn (fn polar => gsub def polar u1)) u2' + | (false, true) => + s_op2 oper T (Opt bool_atom_R) + u1' (triad_fn (fn polar => gsub def polar u2)) + | (false, false) => raise SAME () + else + raise SAME ()) + handle SAME () => s_op2 oper T (Formula polar) u1' u2' + end + else if oper mem [The, Eps] then + let + val u1' = sub u1 + val opt1 = is_opt_rep (rep_of u1') + val unopt_R = best_one_rep_for_type scope T |> optable_rep ofs T + val R = unopt_R |> (oper = Eps orelse opt1) ? opt_rep ofs T + val u = Op2 (oper, T, R, u1', sub u2) + in + if is_precise_type datatypes T orelse not opt1 then + u + else + let + val x_u = BoundName (~1, T, unopt_R, "descr.x") + val R = R |> opt_rep ofs T + in + Op3 (If, T, R, + Op2 (Exist, bool_T, Formula Pos, x_u, + s_op2 Apply bool_T (Formula Pos) (gsub false Pos u1) + x_u), u, Cst (Unknown, T, R)) + end + end + else + let + val u1 = sub u1 + val u2 = sub u2 + val R = + best_non_opt_set_rep_for_type scope T + |> exists (is_opt_rep o rep_of) [u1, u2] ? opt_rep ofs T + in s_op2 oper T R u1 u2 end + | Op3 (Let, T, _, u1, u2, u3) => + let + val u2 = sub u2 + val R2 = rep_of u2 + val table' = NameTable.update (u1, R2) table + val u1 = modify_name_rep u1 R2 + val u3 = aux table' false polar u3 + in s_op3 Let T (rep_of u3) u1 u2 u3 end + | Op3 (If, T, _, u1, u2, u3) => + let + val u1 = sub u1 + val u2 = gsub def polar u2 + val u3 = gsub def polar u3 + val min_R = min_rep (rep_of u2) (rep_of u3) + val R = min_R |> is_opt_rep (rep_of u1) ? opt_rep ofs T + in s_op3 If T R u1 u2 u3 end + | Tuple (T, _, us) => + let + val Rs = map (best_one_rep_for_type scope o type_of) us + val us = map sub us + val R = if forall (equal Unit) Rs then Unit else Struct Rs + val R' = (exists (is_opt_rep o rep_of) us ? opt_rep ofs T) R + in s_tuple T R' us end + | Construct (us', T, _, us) => + let + val us = map sub us + val Rs = map rep_of us + val R = best_one_rep_for_type scope T + val {total, ...} = + constr_spec datatypes (original_name (nickname_of (hd us')), T) + val opt = exists is_opt_rep Rs orelse not total + in Construct (map sub us', T, R |> opt ? Opt, us) end + | _ => + let val u = modify_name_rep u (the_name table u) in + if polar = Neut orelse not (is_boolean_type (type_of u)) + orelse not (is_opt_rep (rep_of u)) then + u + else + s_op1 Cast (type_of u) (Formula polar) u + end + end + |> optimize_unit + in aux table def Pos end + +(* int -> Kodkod.n_ary_index list -> Kodkod.n_ary_index list + -> int * Kodkod.n_ary_index list *) +fun fresh_n_ary_index n [] ys = (0, (n, 1) :: ys) + | fresh_n_ary_index n ((m, j) :: xs) ys = + if m = n then (j, ys @ ((m, j + 1) :: xs)) + else fresh_n_ary_index n xs ((m, j) :: ys) +(* int -> name_pool -> int * name_pool *) +fun fresh_rel n {rels, vars, formula_reg, rel_reg} = + let val (j, rels') = fresh_n_ary_index n rels [] in + (j, {rels = rels', vars = vars, formula_reg = formula_reg, + rel_reg = rel_reg}) + end +(* int -> name_pool -> int * name_pool *) +fun fresh_var n {rels, vars, formula_reg, rel_reg} = + let val (j, vars') = fresh_n_ary_index n vars [] in + (j, {rels = rels, vars = vars', formula_reg = formula_reg, + rel_reg = rel_reg}) + end +(* int -> name_pool -> int * name_pool *) +fun fresh_formula_reg {rels, vars, formula_reg, rel_reg} = + (formula_reg, {rels = rels, vars = vars, formula_reg = formula_reg + 1, + rel_reg = rel_reg}) +(* int -> name_pool -> int * name_pool *) +fun fresh_rel_reg {rels, vars, formula_reg, rel_reg} = + (rel_reg, {rels = rels, vars = vars, formula_reg = formula_reg, + rel_reg = rel_reg + 1}) + +(* nut -> nut list * name_pool * nut NameTable.table + -> nut list * name_pool * nut NameTable.table *) +fun rename_plain_var v (ws, pool, table) = + let + val is_formula = (rep_of v = Formula Neut) + val fresh = if is_formula then fresh_formula_reg else fresh_rel_reg + val (j, pool) = fresh pool + val constr = if is_formula then FormulaReg else RelReg + val w = constr (j, type_of v, rep_of v) + in (w :: ws, pool, NameTable.update (v, w) table) end + +(* typ -> rep -> nut list -> nut *) +fun shape_tuple (T as Type ("*", [T1, T2])) (R as Struct [R1, R2]) us = + let val arity1 = arity_of_rep R1 in + Tuple (T, R, [shape_tuple T1 R1 (List.take (us, arity1)), + shape_tuple T2 R2 (List.drop (us, arity1))]) + end + | shape_tuple (T as Type ("fun", [_, T2])) (R as Vect (k, R')) us = + Tuple (T, R, map (shape_tuple T2 R') (batch_list (length us div k) us)) + | shape_tuple T R [u] = + if type_of u = T then u else raise NUT ("NitpickNut.shape_tuple", [u]) + | shape_tuple T Unit [] = Cst (Unity, T, Unit) + | shape_tuple _ _ us = raise NUT ("NitpickNut.shape_tuple", us) + +(* bool -> nut -> nut list * name_pool * nut NameTable.table + -> nut list * name_pool * nut NameTable.table *) +fun rename_n_ary_var rename_free v (ws, pool, table) = + let + val T = type_of v + val R = rep_of v + val arity = arity_of_rep R + val nick = nickname_of v + val (constr, fresh) = if rename_free then (FreeRel, fresh_rel) + else (BoundRel, fresh_var) + in + if not rename_free andalso arity > 1 then + let + val atom_schema = atom_schema_of_rep R + val type_schema = type_schema_of_rep T R + val (js, pool) = funpow arity (fn (js, pool) => + let val (j, pool) = fresh 1 pool in + (j :: js, pool) + end) + ([], pool) + val ws' = map3 (fn j => fn x => fn T => + constr ((1, j), T, Atom x, + nick ^ " [" ^ string_of_int j ^ "]")) + (rev js) atom_schema type_schema + in (ws' @ ws, pool, NameTable.update (v, shape_tuple T R ws') table) end + else + let + val (j, pool) = fresh arity pool + val w = constr ((arity, j), T, R, nick) + in (w :: ws, pool, NameTable.update (v, w) table) end + end + +(* nut list -> name_pool -> nut NameTable.table + -> nut list * name_pool * nut NameTable.table *) +fun rename_free_vars vs pool table = + let + val vs = filter (not_equal Unit o rep_of) vs + val (vs, pool, table) = fold (rename_n_ary_var true) vs ([], pool, table) + in (rev vs, pool, table) end + +(* name_pool -> nut NameTable.table -> nut -> nut *) +fun rename_vars_in_nut pool table u = + case u of + Cst _ => u + | Op1 (oper, T, R, u1) => Op1 (oper, T, R, rename_vars_in_nut pool table u1) + | Op2 (oper, T, R, u1, u2) => + if oper mem [All, Exist, Lambda] then + let + val (_, pool, table) = fold (rename_n_ary_var false) (untuple I u1) + ([], pool, table) + in + Op2 (oper, T, R, rename_vars_in_nut pool table u1, + rename_vars_in_nut pool table u2) + end + else + Op2 (oper, T, R, rename_vars_in_nut pool table u1, + rename_vars_in_nut pool table u2) + | Op3 (Let, T, R, u1, u2, u3) => + if rep_of u2 = Unit orelse inline_nut u2 then + let + val u2 = rename_vars_in_nut pool table u2 + val table = NameTable.update (u1, u2) table + in rename_vars_in_nut pool table u3 end + else + let + val bs = untuple I u1 + val (_, pool, table') = fold rename_plain_var bs ([], pool, table) + val u11 = rename_vars_in_nut pool table' u1 + in + Op3 (Let, T, R, rename_vars_in_nut pool table' u1, + rename_vars_in_nut pool table u2, + rename_vars_in_nut pool table' u3) + end + | Op3 (oper, T, R, u1, u2, u3) => + Op3 (oper, T, R, rename_vars_in_nut pool table u1, + rename_vars_in_nut pool table u2, rename_vars_in_nut pool table u3) + | Tuple (T, R, us) => Tuple (T, R, map (rename_vars_in_nut pool table) us) + | Construct (us', T, R, us) => + Construct (map (rename_vars_in_nut pool table) us', T, R, + map (rename_vars_in_nut pool table) us) + | _ => the_name table u + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_peephole.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_peephole.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,643 @@ +(* Title: HOL/Nitpick/Tools/nitpick_peephole.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Peephole optimizer for Nitpick. +*) + +signature NITPICK_PEEPHOLE = +sig + type formula = Kodkod.formula + type int_expr = Kodkod.int_expr + type rel_expr = Kodkod.rel_expr + type decl = Kodkod.decl + type expr_assign = Kodkod.expr_assign + + type name_pool = { + rels: Kodkod.n_ary_index list, + vars: Kodkod.n_ary_index list, + formula_reg: int, + rel_reg: int} + + val initial_pool : name_pool + val not3_rel : rel_expr + val suc_rel : rel_expr + val nat_add_rel : rel_expr + val int_add_rel : rel_expr + val nat_subtract_rel : rel_expr + val int_subtract_rel : rel_expr + val nat_multiply_rel : rel_expr + val int_multiply_rel : rel_expr + val nat_divide_rel : rel_expr + val int_divide_rel : rel_expr + val nat_modulo_rel : rel_expr + val int_modulo_rel : rel_expr + val nat_less_rel : rel_expr + val int_less_rel : rel_expr + val gcd_rel : rel_expr + val lcm_rel : rel_expr + val norm_frac_rel : rel_expr + val atom_for_bool : int -> bool -> rel_expr + val formula_for_bool : bool -> formula + val atom_for_nat : int * int -> int -> int + val min_int_for_card : int -> int + val max_int_for_card : int -> int + val int_for_atom : int * int -> int -> int + val atom_for_int : int * int -> int -> int + val inline_rel_expr : rel_expr -> bool + val empty_n_ary_rel : int -> rel_expr + val num_seq : int -> int -> int_expr list + val s_and : formula -> formula -> formula + + type kodkod_constrs = { + kk_all: decl list -> formula -> formula, + kk_exist: decl list -> formula -> formula, + kk_formula_let: expr_assign list -> formula -> formula, + kk_formula_if: formula -> formula -> formula -> formula, + kk_or: formula -> formula -> formula, + kk_not: formula -> formula, + kk_iff: formula -> formula -> formula, + kk_implies: formula -> formula -> formula, + kk_and: formula -> formula -> formula, + kk_subset: rel_expr -> rel_expr -> formula, + kk_rel_eq: rel_expr -> rel_expr -> formula, + kk_no: rel_expr -> formula, + kk_lone: rel_expr -> formula, + kk_one: rel_expr -> formula, + kk_some: rel_expr -> formula, + kk_rel_let: expr_assign list -> rel_expr -> rel_expr, + kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr, + kk_union: rel_expr -> rel_expr -> rel_expr, + kk_difference: rel_expr -> rel_expr -> rel_expr, + kk_override: rel_expr -> rel_expr -> rel_expr, + kk_intersect: rel_expr -> rel_expr -> rel_expr, + kk_product: rel_expr -> rel_expr -> rel_expr, + kk_join: rel_expr -> rel_expr -> rel_expr, + kk_closure: rel_expr -> rel_expr, + kk_reflexive_closure: rel_expr -> rel_expr, + kk_comprehension: decl list -> formula -> rel_expr, + kk_project: rel_expr -> int_expr list -> rel_expr, + kk_project_seq: rel_expr -> int -> int -> rel_expr, + kk_not3: rel_expr -> rel_expr, + kk_nat_less: rel_expr -> rel_expr -> rel_expr, + kk_int_less: rel_expr -> rel_expr -> rel_expr + } + + val kodkod_constrs : bool -> int -> int -> int -> kodkod_constrs +end; + +structure NitpickPeephole : NITPICK_PEEPHOLE = +struct + +open Kodkod +open NitpickUtil + +type name_pool = { + rels: n_ary_index list, + vars: n_ary_index list, + formula_reg: int, + rel_reg: int} + +(* If you add new built-in relations, make sure to increment the counters here + as well to avoid name clashes (which fortunately would be detected by + Kodkodi). *) +val initial_pool = + {rels = [(2, 10), (3, 20), (4, 10)], vars = [], formula_reg = 10, + rel_reg = 10} + +val not3_rel = Rel (2, 0) +val suc_rel = Rel (2, 1) +val nat_add_rel = Rel (3, 0) +val int_add_rel = Rel (3, 1) +val nat_subtract_rel = Rel (3, 2) +val int_subtract_rel = Rel (3, 3) +val nat_multiply_rel = Rel (3, 4) +val int_multiply_rel = Rel (3, 5) +val nat_divide_rel = Rel (3, 6) +val int_divide_rel = Rel (3, 7) +val nat_modulo_rel = Rel (3, 8) +val int_modulo_rel = Rel (3, 9) +val nat_less_rel = Rel (3, 10) +val int_less_rel = Rel (3, 11) +val gcd_rel = Rel (3, 12) +val lcm_rel = Rel (3, 13) +val norm_frac_rel = Rel (4, 0) + +(* int -> bool -> rel_expr *) +fun atom_for_bool j0 = Atom o Integer.add j0 o int_for_bool +(* bool -> formula *) +fun formula_for_bool b = if b then True else False + +(* int * int -> int -> int *) +fun atom_for_nat (k, j0) n = if n < 0 orelse n >= k then ~1 else n + j0 +(* int -> int *) +fun min_int_for_card k = ~k div 2 + 1 +fun max_int_for_card k = k div 2 +(* int * int -> int -> int *) +fun int_for_atom (k, j0) j = + let val j = j - j0 in if j <= max_int_for_card k then j else j - k end +fun atom_for_int (k, j0) n = + if n < min_int_for_card k orelse n > max_int_for_card k then ~1 + else if n < 0 then n + k + j0 + else n + j0 + +(* rel_expr -> bool *) +fun is_none_product (Product (r1, r2)) = + is_none_product r1 orelse is_none_product r2 + | is_none_product None = true + | is_none_product _ = false + +(* rel_expr -> bool *) +fun is_one_rel_expr (Atom _) = true + | is_one_rel_expr (AtomSeq (1, _)) = true + | is_one_rel_expr (Var _) = true + | is_one_rel_expr _ = false + +(* rel_expr -> bool *) +fun inline_rel_expr (Product (r1, r2)) = + inline_rel_expr r1 andalso inline_rel_expr r2 + | inline_rel_expr Iden = true + | inline_rel_expr Ints = true + | inline_rel_expr None = true + | inline_rel_expr Univ = true + | inline_rel_expr (Atom _) = true + | inline_rel_expr (AtomSeq _) = true + | inline_rel_expr (Rel _) = true + | inline_rel_expr (Var _) = true + | inline_rel_expr (RelReg _) = true + | inline_rel_expr _ = false + +(* rel_expr -> rel_expr -> bool option *) +fun rel_expr_equal None (Atom _) = SOME false + | rel_expr_equal None (AtomSeq (k, _)) = SOME (k = 0) + | rel_expr_equal (Atom _) None = SOME false + | rel_expr_equal (AtomSeq (k, _)) None = SOME (k = 0) + | rel_expr_equal (Atom j1) (Atom j2) = SOME (j1 = j2) + | rel_expr_equal (Atom j) (AtomSeq (k, j0)) = SOME (j = j0 andalso k = 1) + | rel_expr_equal (AtomSeq (k, j0)) (Atom j) = SOME (j = j0 andalso k = 1) + | rel_expr_equal (AtomSeq x1) (AtomSeq x2) = SOME (x1 = x2) + | rel_expr_equal r1 r2 = if r1 = r2 then SOME true else NONE + +(* rel_expr -> rel_expr -> bool option *) +fun rel_expr_intersects (Atom j1) (Atom j2) = SOME (j1 = j2) + | rel_expr_intersects (Atom j) (AtomSeq (k, j0)) = SOME (j < j0 + k) + | rel_expr_intersects (AtomSeq (k, j0)) (Atom j) = SOME (j < j0 + k) + | rel_expr_intersects (AtomSeq (k1, j01)) (AtomSeq (k2, j02)) = + SOME (k1 > 0 andalso k2 > 0 andalso j01 + k1 > j02 andalso j02 + k2 > j01) + | rel_expr_intersects r1 r2 = + if is_none_product r1 orelse is_none_product r2 then SOME false else NONE + +(* int -> rel_expr *) +fun empty_n_ary_rel 0 = raise ARG ("NitpickPeephole.empty_n_ary_rel", "0") + | empty_n_ary_rel n = funpow (n - 1) (curry Product None) None + +(* decl -> rel_expr *) +fun decl_one_set (DeclOne (_, r)) = r + | decl_one_set _ = + raise ARG ("NitpickPeephole.decl_one_set", "not \"DeclOne\"") + +(* int_expr -> bool *) +fun is_Num (Num _) = true + | is_Num _ = false +(* int_expr -> int *) +fun dest_Num (Num k) = k + | dest_Num _ = raise ARG ("NitpickPeephole.dest_Num", "not \"Num\"") +(* int -> int -> int_expr list *) +fun num_seq j0 n = map Num (index_seq j0 n) + +(* rel_expr -> rel_expr -> bool *) +fun occurs_in_union r (Union (r1, r2)) = + occurs_in_union r r1 orelse occurs_in_union r r2 + | occurs_in_union r r' = (r = r') + +(* rel_expr -> rel_expr -> rel_expr *) +fun s_and True f2 = f2 + | s_and False _ = False + | s_and f1 True = f1 + | s_and _ False = False + | s_and f1 f2 = And (f1, f2) + +type kodkod_constrs = { + kk_all: decl list -> formula -> formula, + kk_exist: decl list -> formula -> formula, + kk_formula_let: expr_assign list -> formula -> formula, + kk_formula_if: formula -> formula -> formula -> formula, + kk_or: formula -> formula -> formula, + kk_not: formula -> formula, + kk_iff: formula -> formula -> formula, + kk_implies: formula -> formula -> formula, + kk_and: formula -> formula -> formula, + kk_subset: rel_expr -> rel_expr -> formula, + kk_rel_eq: rel_expr -> rel_expr -> formula, + kk_no: rel_expr -> formula, + kk_lone: rel_expr -> formula, + kk_one: rel_expr -> formula, + kk_some: rel_expr -> formula, + kk_rel_let: expr_assign list -> rel_expr -> rel_expr, + kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr, + kk_union: rel_expr -> rel_expr -> rel_expr, + kk_difference: rel_expr -> rel_expr -> rel_expr, + kk_override: rel_expr -> rel_expr -> rel_expr, + kk_intersect: rel_expr -> rel_expr -> rel_expr, + kk_product: rel_expr -> rel_expr -> rel_expr, + kk_join: rel_expr -> rel_expr -> rel_expr, + kk_closure: rel_expr -> rel_expr, + kk_reflexive_closure: rel_expr -> rel_expr, + kk_comprehension: decl list -> formula -> rel_expr, + kk_project: rel_expr -> int_expr list -> rel_expr, + kk_project_seq: rel_expr -> int -> int -> rel_expr, + kk_not3: rel_expr -> rel_expr, + kk_nat_less: rel_expr -> rel_expr -> rel_expr, + kk_int_less: rel_expr -> rel_expr -> rel_expr +} + +(* We assume throughout that Kodkod variables have a "one" constraint. This is + always the case if Kodkod's skolemization is disabled. *) +(* bool -> int -> int -> int -> kodkod_constrs *) +fun kodkod_constrs optim nat_card int_card main_j0 = + let + val false_atom = Atom main_j0 + val true_atom = Atom (main_j0 + 1) + + (* bool -> int *) + val from_bool = atom_for_bool main_j0 + (* int -> Kodkod.rel_expr *) + fun from_nat n = Atom (n + main_j0) + val from_int = Atom o atom_for_int (int_card, main_j0) + (* int -> int *) + fun to_nat j = j - main_j0 + val to_int = int_for_atom (int_card, main_j0) + + (* decl list -> formula -> formula *) + fun s_all _ True = True + | s_all _ False = False + | s_all [] f = f + | s_all ds (All (ds', f)) = All (ds @ ds', f) + | s_all ds f = All (ds, f) + fun s_exist _ True = True + | s_exist _ False = False + | s_exist [] f = f + | s_exist ds (Exist (ds', f)) = Exist (ds @ ds', f) + | s_exist ds f = Exist (ds, f) + + (* expr_assign list -> formula -> formula *) + fun s_formula_let _ True = True + | s_formula_let _ False = False + | s_formula_let assigns f = FormulaLet (assigns, f) + + (* formula -> formula *) + fun s_not True = False + | s_not False = True + | s_not (All (ds, f)) = Exist (ds, s_not f) + | s_not (Exist (ds, f)) = All (ds, s_not f) + | s_not (Or (f1, f2)) = And (s_not f1, s_not f2) + | s_not (Implies (f1, f2)) = And (f1, s_not f2) + | s_not (And (f1, f2)) = Or (s_not f1, s_not f2) + | s_not (Not f) = f + | s_not (No r) = Some r + | s_not (Some r) = No r + | s_not f = Not f + + (* formula -> formula -> formula *) + fun s_or True _ = True + | s_or False f2 = f2 + | s_or _ True = True + | s_or f1 False = f1 + | s_or f1 f2 = if f1 = f2 then f1 else Or (f1, f2) + fun s_iff True f2 = f2 + | s_iff False f2 = s_not f2 + | s_iff f1 True = f1 + | s_iff f1 False = s_not f1 + | s_iff f1 f2 = if f1 = f2 then True else Iff (f1, f2) + fun s_implies True f2 = f2 + | s_implies False _ = True + | s_implies _ True = True + | s_implies f1 False = s_not f1 + | s_implies f1 f2 = if f1 = f2 then True else Implies (f1, f2) + + (* formula -> formula -> formula -> formula *) + fun s_formula_if True f2 _ = f2 + | s_formula_if False _ f3 = f3 + | s_formula_if f1 True f3 = s_or f1 f3 + | s_formula_if f1 False f3 = s_and (s_not f1) f3 + | s_formula_if f1 f2 True = s_implies f1 f2 + | s_formula_if f1 f2 False = s_and f1 f2 + | s_formula_if f f1 f2 = FormulaIf (f, f1, f2) + + (* rel_expr -> int_expr list -> rel_expr *) + fun s_project r is = + (case r of + Project (r1, is') => + if forall is_Num is then + s_project r1 (map (nth is' o dest_Num) is) + else + raise SAME () + | _ => raise SAME ()) + handle SAME () => + let val n = length is in + if arity_of_rel_expr r = n andalso is = num_seq 0 n then r + else Project (r, is) + end + + (* rel_expr -> formula *) + fun s_no None = True + | s_no (Product (r1, r2)) = s_or (s_no r1) (s_no r2) + | s_no (Intersect (Closure (Kodkod.Rel x), Kodkod.Iden)) = Acyclic x + | s_no r = if is_one_rel_expr r then False else No r + fun s_lone None = True + | s_lone r = if is_one_rel_expr r then True else Lone r + fun s_one None = False + | s_one r = + if is_one_rel_expr r then + True + else if inline_rel_expr r then + case arity_of_rel_expr r of + 1 => One r + | arity => foldl1 And (map (One o s_project r o single o Num) + (index_seq 0 arity)) + else + One r + fun s_some None = False + | s_some (Atom _) = True + | s_some (Product (r1, r2)) = s_and (s_some r1) (s_some r2) + | s_some r = if is_one_rel_expr r then True else Some r + + (* rel_expr -> rel_expr *) + fun s_not3 (Atom j) = Atom (if j = main_j0 then j + 1 else j - 1) + | s_not3 (r as Join (r1, r2)) = + if r2 = not3_rel then r1 else Join (r, not3_rel) + | s_not3 r = Join (r, not3_rel) + + (* rel_expr -> rel_expr -> formula *) + fun s_rel_eq r1 r2 = + (case (r1, r2) of + (Join (r11, r12), _) => + if r12 = not3_rel then s_rel_eq r11 (s_not3 r2) else raise SAME () + | (_, Join (r21, r22)) => + if r22 = not3_rel then s_rel_eq r21 (s_not3 r1) else raise SAME () + | _ => raise SAME ()) + handle SAME () => + case rel_expr_equal r1 r2 of + SOME true => True + | SOME false => False + | NONE => + case (r1, r2) of + (_, RelIf (f, r21, r22)) => + if inline_rel_expr r1 then + s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22) + else + RelEq (r1, r2) + | (RelIf (f, r11, r12), _) => + if inline_rel_expr r2 then + s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2) + else + RelEq (r1, r2) + | (_, Kodkod.None) => s_no r1 + | (Kodkod.None, _) => s_no r2 + | _ => RelEq (r1, r2) + fun s_subset (Atom j1) (Atom j2) = formula_for_bool (j1 = j2) + | s_subset (Atom j) (AtomSeq (k, j0)) = + formula_for_bool (j >= j0 andalso j < j0 + k) + | s_subset (r1 as Union (r11, r12)) r2 = + s_and (s_subset r11 r2) (s_subset r12 r2) + | s_subset r1 (r2 as Union (r21, r22)) = + if is_one_rel_expr r1 then + s_or (s_subset r1 r21) (s_subset r1 r22) + else + if s_subset r1 r21 = True orelse s_subset r1 r22 = True + orelse r1 = r2 then + True + else + Subset (r1, r2) + | s_subset r1 r2 = + if r1 = r2 orelse is_none_product r1 then True + else if is_none_product r2 then s_no r1 + else if forall is_one_rel_expr [r1, r2] then s_rel_eq r1 r2 + else Subset (r1, r2) + + (* expr_assign list -> rel_expr -> rel_expr *) + fun s_rel_let [b as AssignRelReg (x', r')] (r as RelReg x) = + if x = x' then r' else RelLet ([b], r) + | s_rel_let bs r = RelLet (bs, r) + + (* formula -> rel_expr -> rel_expr -> rel_expr *) + fun s_rel_if f r1 r2 = + (case (f, r1, r2) of + (True, _, _) => r1 + | (False, _, _) => r2 + | (No r1', None, RelIf (One r2', r3', r4')) => + if r1' = r2' andalso r2' = r3' then s_rel_if (Lone r1') r1' r4' + else raise SAME () + | _ => raise SAME ()) + handle SAME () => if r1 = r2 then r1 else RelIf (f, r1, r2) + + (* rel_expr -> rel_expr -> rel_expr *) + fun s_union r1 (Union (r21, r22)) = s_union (s_union r1 r21) r22 + | s_union r1 r2 = + if is_none_product r1 then r2 + else if is_none_product r2 then r1 + else if r1 = r2 then r1 + else if occurs_in_union r2 r1 then r1 + else Union (r1, r2) + fun s_difference r1 r2 = + if is_none_product r1 orelse is_none_product r2 then r1 + else if r1 = r2 then empty_n_ary_rel (arity_of_rel_expr r1) + else Difference (r1, r2) + fun s_override r1 r2 = + if is_none_product r2 then r1 + else if is_none_product r1 then r2 + else Override (r1, r2) + fun s_intersect r1 r2 = + case rel_expr_intersects r1 r2 of + SOME true => if r1 = r2 then r1 else Intersect (r1, r2) + | SOME false => empty_n_ary_rel (arity_of_rel_expr r1) + | NONE => if is_none_product r1 then r1 + else if is_none_product r2 then r2 + else Intersect (r1, r2) + fun s_product r1 r2 = + if is_none_product r1 then + Product (r1, empty_n_ary_rel (arity_of_rel_expr r2)) + else if is_none_product r2 then + Product (empty_n_ary_rel (arity_of_rel_expr r1), r2) + else + Product (r1, r2) + fun s_join r1 (Product (Product (r211, r212), r22)) = + Product (s_join r1 (Product (r211, r212)), r22) + | s_join (Product (r11, Product (r121, r122))) r2 = + Product (r11, s_join (Product (r121, r122)) r2) + | s_join None r = empty_n_ary_rel (arity_of_rel_expr r - 1) + | s_join r None = empty_n_ary_rel (arity_of_rel_expr r - 1) + | s_join (Product (None, None)) r = empty_n_ary_rel (arity_of_rel_expr r) + | s_join r (Product (None, None)) = empty_n_ary_rel (arity_of_rel_expr r) + | s_join Iden r2 = r2 + | s_join r1 Iden = r1 + | s_join (Product (r1, r2)) Univ = + if arity_of_rel_expr r2 = 1 then r1 + else Product (r1, s_join r2 Univ) + | s_join Univ (Product (r1, r2)) = + if arity_of_rel_expr r1 = 1 then r2 + else Product (s_join Univ r1, r2) + | s_join r1 (r2 as Product (r21, r22)) = + if arity_of_rel_expr r1 = 1 then + case rel_expr_intersects r1 r21 of + SOME true => r22 + | SOME false => empty_n_ary_rel (arity_of_rel_expr r2 - 1) + | NONE => Join (r1, r2) + else + Join (r1, r2) + | s_join (r1 as Product (r11, r12)) r2 = + if arity_of_rel_expr r2 = 1 then + case rel_expr_intersects r2 r12 of + SOME true => r11 + | SOME false => empty_n_ary_rel (arity_of_rel_expr r1 - 1) + | NONE => Join (r1, r2) + else + Join (r1, r2) + | s_join r1 (r2 as RelIf (f, r21, r22)) = + if inline_rel_expr r1 then s_rel_if f (s_join r1 r21) (s_join r1 r22) + else Join (r1, r2) + | s_join (r1 as RelIf (f, r11, r12)) r2 = + if inline_rel_expr r2 then s_rel_if f (s_join r11 r2) (s_join r12 r2) + else Join (r1, r2) + | s_join (r1 as Atom j1) (r2 as Rel (2, j2)) = + if r2 = suc_rel then + let val n = to_nat j1 + 1 in + if n < nat_card then from_nat n else None + end + else + Join (r1, r2) + | s_join r1 (r2 as Project (r21, Num k :: is)) = + if k = arity_of_rel_expr r21 - 1 andalso arity_of_rel_expr r1 = 1 then + s_project (s_join r21 r1) is + else + Join (r1, r2) + | s_join r1 (Join (r21, r22 as Rel (3, j22))) = + ((if r22 = nat_add_rel then + case (r21, r1) of + (Atom j1, Atom j2) => + let val n = to_nat j1 + to_nat j2 in + if n < nat_card then from_nat n else None + end + | (Atom j, r) => + (case to_nat j of + 0 => r + | 1 => s_join r suc_rel + | _ => raise SAME ()) + | (r, Atom j) => + (case to_nat j of + 0 => r + | 1 => s_join r suc_rel + | _ => raise SAME ()) + | _ => raise SAME () + else if r22 = nat_subtract_rel then + case (r21, r1) of + (Atom j1, Atom j2) => from_nat (to_nat j1 nat_minus to_nat j2) + | _ => raise SAME () + else if r22 = nat_multiply_rel then + case (r21, r1) of + (Atom j1, Atom j2) => + let val n = to_nat j1 * to_nat j2 in + if n < nat_card then from_nat n else None + end + | (Atom j, r) => + (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ()) + | (r, Atom j) => + (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ()) + | _ => raise SAME () + else + raise SAME ()) + handle SAME () => List.foldr Join r22 [r1, r21]) + | s_join r1 r2 = Join (r1, r2) + + (* rel_expr -> rel_expr *) + fun s_closure Iden = Iden + | s_closure r = if is_none_product r then r else Closure r + fun s_reflexive_closure Iden = Iden + | s_reflexive_closure r = + if is_none_product r then Iden else ReflexiveClosure r + + (* decl list -> formula -> rel_expr *) + fun s_comprehension ds False = empty_n_ary_rel (length ds) + | s_comprehension ds True = fold1 s_product (map decl_one_set ds) + | s_comprehension [d as DeclOne ((1, j1), r)] + (f as RelEq (Var (1, j2), Atom j)) = + if j1 = j2 andalso rel_expr_intersects (Atom j) r = SOME true then + Atom j + else + Comprehension ([d], f) + | s_comprehension ds f = Comprehension (ds, f) + + (* rel_expr -> int -> int -> rel_expr *) + fun s_project_seq r = + let + (* int -> rel_expr -> int -> int -> rel_expr *) + fun aux arity r j0 n = + if j0 = 0 andalso arity = n then + r + else case r of + RelIf (f, r1, r2) => + s_rel_if f (aux arity r1 j0 n) (aux arity r2 j0 n) + | Product (r1, r2) => + let + val arity2 = arity_of_rel_expr r2 + val arity1 = arity - arity2 + val n1 = Int.min (arity1 nat_minus j0, n) + val n2 = n - n1 + (* unit -> rel_expr *) + fun one () = aux arity1 r1 j0 n1 + fun two () = aux arity2 r2 (j0 nat_minus arity1) n2 + in + case (n1, n2) of + (0, _) => s_rel_if (s_some r1) (two ()) (empty_n_ary_rel n2) + | (_, 0) => s_rel_if (s_some r2) (one ()) (empty_n_ary_rel n1) + | _ => s_product (one ()) (two ()) + end + | _ => s_project r (num_seq j0 n) + in aux (arity_of_rel_expr r) r end + + (* rel_expr -> rel_expr -> rel_expr *) + fun s_nat_subtract r1 r2 = fold s_join [r1, r2] nat_subtract_rel + fun s_nat_less (Atom j1) (Atom j2) = from_bool (j1 < j2) + | s_nat_less r1 r2 = fold s_join [r1, r2] nat_less_rel + fun s_int_less (Atom j1) (Atom j2) = from_bool (to_int j1 < to_int j2) + | s_int_less r1 r2 = fold s_join [r1, r2] int_less_rel + + (* rel_expr -> int -> int -> rel_expr *) + fun d_project_seq r j0 n = Project (r, num_seq j0 n) + (* rel_expr -> rel_expr *) + fun d_not3 r = Join (r, not3_rel) + (* rel_expr -> rel_expr -> rel_expr *) + fun d_nat_subtract r1 r2 = List.foldl Join nat_subtract_rel [r1, r2] + fun d_nat_less r1 r2 = List.foldl Join nat_less_rel [r1, r2] + fun d_int_less r1 r2 = List.foldl Join int_less_rel [r1, r2] + in + if optim then + {kk_all = s_all, kk_exist = s_exist, kk_formula_let = s_formula_let, + kk_formula_if = s_formula_if, kk_or = s_or, kk_not = s_not, + kk_iff = s_iff, kk_implies = s_implies, kk_and = s_and, + kk_subset = s_subset, kk_rel_eq = s_rel_eq, kk_no = s_no, + kk_lone = s_lone, kk_one = s_one, kk_some = s_some, + kk_rel_let = s_rel_let, kk_rel_if = s_rel_if, kk_union = s_union, + kk_difference = s_difference, kk_override = s_override, + kk_intersect = s_intersect, kk_product = s_product, kk_join = s_join, + kk_closure = s_closure, kk_reflexive_closure = s_reflexive_closure, + kk_comprehension = s_comprehension, kk_project = s_project, + kk_project_seq = s_project_seq, kk_not3 = s_not3, + kk_nat_less = s_nat_less, kk_int_less = s_int_less} + else + {kk_all = curry All, kk_exist = curry Exist, + kk_formula_let = curry FormulaLet, kk_formula_if = curry3 FormulaIf, + kk_or = curry Or,kk_not = Not, kk_iff = curry Iff, kk_implies = curry + Implies, kk_and = curry And, kk_subset = curry Subset, kk_rel_eq = curry + RelEq, kk_no = No, kk_lone = Lone, kk_one = One, kk_some = Some, + kk_rel_let = curry RelLet, kk_rel_if = curry3 RelIf, kk_union = curry + Union, kk_difference = curry Difference, kk_override = curry Override, + kk_intersect = curry Intersect, kk_product = curry Product, + kk_join = curry Join, kk_closure = Closure, + kk_reflexive_closure = ReflexiveClosure, kk_comprehension = curry + Comprehension, kk_project = curry Project, + kk_project_seq = d_project_seq, kk_not3 = d_not3, + kk_nat_less = d_nat_less, kk_int_less = d_int_less} + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_rep.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,344 @@ +(* Title: HOL/Nitpick/Tools/nitpick_rep.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Kodkod representations of Nitpick terms. +*) + +signature NITPICK_REP = +sig + type polarity = NitpickUtil.polarity + type scope = NitpickScope.scope + + datatype rep = + Any | + Formula of polarity | + Unit | + Atom of int * int | + Struct of rep list | + Vect of int * rep | + Func of rep * rep | + Opt of rep + + exception REP of string * rep list + + val string_for_polarity : polarity -> string + val string_for_rep : rep -> string + val is_Func : rep -> bool + val is_Opt : rep -> bool + val is_opt_rep : rep -> bool + val flip_rep_polarity : rep -> rep + val card_of_rep : rep -> int + val arity_of_rep : rep -> int + val min_univ_card_of_rep : rep -> int + val is_one_rep : rep -> bool + val is_lone_rep : rep -> bool + val dest_Func : rep -> rep * rep + val smart_range_rep : int Typtab.table -> typ -> (unit -> int) -> rep -> rep + val binder_reps : rep -> rep list + val body_rep : rep -> rep + val one_rep : int Typtab.table -> typ -> rep -> rep + val optable_rep : int Typtab.table -> typ -> rep -> rep + val opt_rep : int Typtab.table -> typ -> rep -> rep + val unopt_rep : rep -> rep + val min_rep : rep -> rep -> rep + val min_reps : rep list -> rep list -> rep list + val card_of_domain_from_rep : int -> rep -> int + val rep_to_binary_rel_rep : int Typtab.table -> typ -> rep -> rep + val best_one_rep_for_type : scope -> typ -> rep + val best_opt_set_rep_for_type : scope -> typ -> rep + val best_non_opt_set_rep_for_type : scope -> typ -> rep + val best_set_rep_for_type : scope -> typ -> rep + val best_non_opt_symmetric_reps_for_fun_type : scope -> typ -> rep * rep + val atom_schema_of_rep : rep -> (int * int) list + val atom_schema_of_reps : rep list -> (int * int) list + val type_schema_of_rep : typ -> rep -> typ list + val type_schema_of_reps : typ list -> rep list -> typ list + val all_combinations_for_rep : rep -> int list list + val all_combinations_for_reps : rep list -> int list list +end; + +structure NitpickRep : NITPICK_REP = +struct + +open NitpickUtil +open NitpickHOL +open NitpickScope + +datatype rep = + Any | + Formula of polarity | + Unit | + Atom of int * int | + Struct of rep list | + Vect of int * rep | + Func of rep * rep | + Opt of rep + +exception REP of string * rep list + +(* polarity -> string *) +fun string_for_polarity Pos = "+" + | string_for_polarity Neg = "-" + | string_for_polarity Neut = "=" + +(* rep -> string *) +fun atomic_string_for_rep rep = + let val s = string_for_rep rep in + if String.isPrefix "[" s orelse not (is_substring_of " " s) then s + else "(" ^ s ^ ")" + end +(* rep -> string *) +and string_for_rep Any = "X" + | string_for_rep (Formula polar) = "F" ^ string_for_polarity polar + | string_for_rep Unit = "U" + | string_for_rep (Atom (k, j0)) = + "A" ^ string_of_int k ^ (if j0 = 0 then "" else "@" ^ string_of_int j0) + | string_for_rep (Struct rs) = "[" ^ commas (map string_for_rep rs) ^ "]" + | string_for_rep (Vect (k, R)) = + string_of_int k ^ " x " ^ atomic_string_for_rep R + | string_for_rep (Func (R1, R2)) = + atomic_string_for_rep R1 ^ " => " ^ string_for_rep R2 + | string_for_rep (Opt R) = atomic_string_for_rep R ^ "?" + +(* rep -> bool *) +fun is_Func (Func _) = true + | is_Func _ = false +fun is_Opt (Opt _) = true + | is_Opt _ = false +fun is_opt_rep (Func (_, R2)) = is_opt_rep R2 + | is_opt_rep (Opt _) = true + | is_opt_rep _ = false + +(* rep -> int *) +fun card_of_rep Any = raise REP ("NitpickRep.card_of_rep", [Any]) + | card_of_rep (Formula _) = 2 + | card_of_rep Unit = 1 + | card_of_rep (Atom (k, _)) = k + | card_of_rep (Struct rs) = Integer.prod (map card_of_rep rs) + | card_of_rep (Vect (k, R)) = reasonable_power (card_of_rep R) k + | card_of_rep (Func (R1, R2)) = + reasonable_power (card_of_rep R2) (card_of_rep R1) + | card_of_rep (Opt R) = card_of_rep R +fun arity_of_rep Any = raise REP ("NitpickRep.arity_of_rep", [Any]) + | arity_of_rep (Formula _) = 0 + | arity_of_rep Unit = 0 + | arity_of_rep (Atom _) = 1 + | arity_of_rep (Struct Rs) = Integer.sum (map arity_of_rep Rs) + | arity_of_rep (Vect (k, R)) = k * arity_of_rep R + | arity_of_rep (Func (R1, R2)) = arity_of_rep R1 + arity_of_rep R2 + | arity_of_rep (Opt R) = arity_of_rep R +fun min_univ_card_of_rep Any = + raise REP ("NitpickRep.min_univ_card_of_rep", [Any]) + | min_univ_card_of_rep (Formula _) = 0 + | min_univ_card_of_rep Unit = 0 + | min_univ_card_of_rep (Atom (k, j0)) = k + j0 + 1 + | min_univ_card_of_rep (Struct Rs) = + fold Integer.max (map min_univ_card_of_rep Rs) 0 + | min_univ_card_of_rep (Vect (_, R)) = min_univ_card_of_rep R + | min_univ_card_of_rep (Func (R1, R2)) = + Int.max (min_univ_card_of_rep R1, min_univ_card_of_rep R2) + | min_univ_card_of_rep (Opt R) = min_univ_card_of_rep R + +(* rep -> bool *) +fun is_one_rep Unit = true + | is_one_rep (Atom _) = true + | is_one_rep (Struct _) = true + | is_one_rep (Vect _) = true + | is_one_rep _ = false +fun is_lone_rep (Opt R) = is_one_rep R + | is_lone_rep R = is_one_rep R + +(* rep -> rep * rep *) +fun dest_Func (Func z) = z + | dest_Func R = raise REP ("NitpickRep.dest_Func", [R]) +(* int Typtab.table -> typ -> (unit -> int) -> rep -> rep *) +fun smart_range_rep _ _ _ Unit = Unit + | smart_range_rep _ _ _ (Vect (_, R)) = R + | smart_range_rep _ _ _ (Func (_, R2)) = R2 + | smart_range_rep ofs T ran_card (Opt R) = + Opt (smart_range_rep ofs T ran_card R) + | smart_range_rep ofs (Type ("fun", [_, T2])) _ (Atom (1, _)) = + Atom (1, offset_of_type ofs T2) + | smart_range_rep ofs (Type ("fun", [_, T2])) ran_card (Atom _) = + Atom (ran_card (), offset_of_type ofs T2) + | smart_range_rep _ _ _ R = raise REP ("NitpickRep.smart_range_rep", [R]) + +(* rep -> rep list *) +fun binder_reps (Func (R1, R2)) = R1 :: binder_reps R2 + | binder_reps R = [] +(* rep -> rep *) +fun body_rep (Func (_, R2)) = body_rep R2 + | body_rep R = R + +(* rep -> rep *) +fun flip_rep_polarity (Formula polar) = Formula (flip_polarity polar) + | flip_rep_polarity (Func (R1, R2)) = Func (R1, flip_rep_polarity R2) + | flip_rep_polarity R = R + +(* int Typtab.table -> rep -> rep *) +fun one_rep _ _ Any = raise REP ("NitpickRep.one_rep", [Any]) + | one_rep _ _ (Atom x) = Atom x + | one_rep _ _ (Struct Rs) = Struct Rs + | one_rep _ _ (Vect z) = Vect z + | one_rep ofs T (Opt R) = one_rep ofs T R + | one_rep ofs T R = Atom (card_of_rep R, offset_of_type ofs T) +fun optable_rep ofs (Type ("fun", [_, T2])) (Func (R1, R2)) = + Func (R1, optable_rep ofs T2 R2) + | optable_rep ofs T R = one_rep ofs T R +fun opt_rep ofs (Type ("fun", [_, T2])) (Func (R1, R2)) = + Func (R1, opt_rep ofs T2 R2) + | opt_rep ofs T R = Opt (optable_rep ofs T R) +(* rep -> rep *) +fun unopt_rep (Func (R1, R2)) = Func (R1, unopt_rep R2) + | unopt_rep (Opt R) = R + | unopt_rep R = R + +(* polarity -> polarity -> polarity *) +fun min_polarity polar1 polar2 = + if polar1 = polar2 then + polar1 + else if polar1 = Neut then + polar2 + else if polar2 = Neut then + polar1 + else + raise ARG ("NitpickRep.min_polarity", + commas (map (quote o string_for_polarity) [polar1, polar2])) + +(* It's important that Func is before Vect, because if the range is Opt we + could lose information by converting a Func to a Vect. *) +(* rep -> rep -> rep *) +fun min_rep (Opt R1) (Opt R2) = Opt (min_rep R1 R2) + | min_rep (Opt R) _ = Opt R + | min_rep _ (Opt R) = Opt R + | min_rep (Formula polar1) (Formula polar2) = + Formula (min_polarity polar1 polar2) + | min_rep (Formula polar) _ = Formula polar + | min_rep _ (Formula polar) = Formula polar + | min_rep Unit _ = Unit + | min_rep _ Unit = Unit + | min_rep (Atom x) _ = Atom x + | min_rep _ (Atom x) = Atom x + | min_rep (Struct Rs1) (Struct Rs2) = Struct (min_reps Rs1 Rs2) + | min_rep (Struct Rs) _ = Struct Rs + | min_rep _ (Struct Rs) = Struct Rs + | min_rep (R1 as Func (R11, R12)) (R2 as Func (R21, R22)) = + (case pairself is_opt_rep (R12, R22) of + (true, false) => R1 + | (false, true) => R2 + | _ => if R11 = R21 then Func (R11, min_rep R12 R22) + else if min_rep R11 R21 = R11 then R1 + else R2) + | min_rep (Func z) _ = Func z + | min_rep _ (Func z) = Func z + | min_rep (Vect (k1, R1)) (Vect (k2, R2)) = + if k1 < k2 then Vect (k1, R1) + else if k1 > k2 then Vect (k2, R2) + else Vect (k1, min_rep R1 R2) + | min_rep R1 R2 = raise REP ("NitpickRep.min_rep", [R1, R2]) +(* rep list -> rep list -> rep list *) +and min_reps [] _ = [] + | min_reps _ [] = [] + | min_reps (R1 :: Rs1) (R2 :: Rs2) = + if R1 = R2 then R1 :: min_reps Rs1 Rs2 + else if min_rep R1 R2 = R1 then R1 :: Rs1 + else R2 :: Rs2 + +(* int -> rep -> int *) +fun card_of_domain_from_rep ran_card R = + case R of + Unit => 1 + | Atom (k, _) => exact_log ran_card k + | Vect (k, _) => k + | Func (R1, _) => card_of_rep R1 + | Opt R => card_of_domain_from_rep ran_card R + | _ => raise REP ("NitpickRep.card_of_domain_from_rep", [R]) + +(* int Typtab.table -> typ -> rep -> rep *) +fun rep_to_binary_rel_rep ofs T R = + let + val k = exact_root 2 (card_of_domain_from_rep 2 R) + val j0 = offset_of_type ofs (fst (HOLogic.dest_prodT (domain_type T))) + in Func (Struct [Atom (k, j0), Atom (k, j0)], Formula Neut) end + +(* scope -> typ -> rep *) +fun best_one_rep_for_type (scope as {card_assigns, ...} : scope) + (Type ("fun", [T1, T2])) = + (case best_one_rep_for_type scope T2 of + Unit => Unit + | R2 => Vect (card_of_type card_assigns T1, R2)) + | best_one_rep_for_type scope (Type ("*", [T1, T2])) = + (case (best_one_rep_for_type scope T1, best_one_rep_for_type scope T2) of + (Unit, Unit) => Unit + | (R1, R2) => Struct [R1, R2]) + | best_one_rep_for_type (scope as {card_assigns, datatypes, ofs, ...}) T = + (case card_of_type card_assigns T of + 1 => if is_some (datatype_spec datatypes T) + orelse is_fp_iterator_type T then + Atom (1, offset_of_type ofs T) + else + Unit + | k => Atom (k, offset_of_type ofs T)) + +(* Datatypes are never represented by Unit, because it would confuse + "nfa_transitions_for_ctor". *) +(* scope -> typ -> rep *) +fun best_opt_set_rep_for_type scope (Type ("fun", [T1, T2])) = + Func (best_one_rep_for_type scope T1, best_opt_set_rep_for_type scope T2) + | best_opt_set_rep_for_type (scope as {ofs, ...}) T = + opt_rep ofs T (best_one_rep_for_type scope T) +fun best_non_opt_set_rep_for_type (scope as {ofs, ...}) + (Type ("fun", [T1, T2])) = + (case (best_one_rep_for_type scope T1, + best_non_opt_set_rep_for_type scope T2) of + (_, Unit) => Unit + | (Unit, Atom (2, _)) => + Func (Atom (1, offset_of_type ofs T1), Formula Neut) + | (R1, Atom (2, _)) => Func (R1, Formula Neut) + | z => Func z) + | best_non_opt_set_rep_for_type scope T = best_one_rep_for_type scope T +fun best_set_rep_for_type (scope as {datatypes, ...}) T = + (if is_precise_type datatypes T then best_non_opt_set_rep_for_type + else best_opt_set_rep_for_type) scope T +fun best_non_opt_symmetric_reps_for_fun_type (scope as {ofs, ...}) + (Type ("fun", [T1, T2])) = + (optable_rep ofs T1 (best_one_rep_for_type scope T1), + optable_rep ofs T2 (best_one_rep_for_type scope T2)) + | best_non_opt_symmetric_reps_for_fun_type _ T = + raise TYPE ("NitpickRep.best_non_opt_symmetric_reps_for_fun_type", [T], []) + +(* rep -> (int * int) list *) +fun atom_schema_of_rep Any = raise REP ("NitpickRep.atom_schema_of_rep", [Any]) + | atom_schema_of_rep (Formula _) = [] + | atom_schema_of_rep Unit = [] + | atom_schema_of_rep (Atom x) = [x] + | atom_schema_of_rep (Struct Rs) = atom_schema_of_reps Rs + | atom_schema_of_rep (Vect (k, R)) = replicate_list k (atom_schema_of_rep R) + | atom_schema_of_rep (Func (R1, R2)) = + atom_schema_of_rep R1 @ atom_schema_of_rep R2 + | atom_schema_of_rep (Opt R) = atom_schema_of_rep R +(* rep list -> (int * int) list *) +and atom_schema_of_reps Rs = maps atom_schema_of_rep Rs + +(* typ -> rep -> typ list *) +fun type_schema_of_rep _ (Formula _) = [] + | type_schema_of_rep _ Unit = [] + | type_schema_of_rep T (Atom _) = [T] + | type_schema_of_rep (Type ("*", [T1, T2])) (Struct [R1, R2]) = + type_schema_of_reps [T1, T2] [R1, R2] + | type_schema_of_rep (Type ("fun", [_, T2])) (Vect (k, R)) = + replicate_list k (type_schema_of_rep T2 R) + | type_schema_of_rep (Type ("fun", [T1, T2])) (Func (R1, R2)) = + type_schema_of_rep T1 R1 @ type_schema_of_rep T2 R2 + | type_schema_of_rep T (Opt R) = type_schema_of_rep T R + | type_schema_of_rep T R = raise REP ("NitpickRep.type_schema_of_rep", [R]) +(* typ list -> rep list -> typ list *) +and type_schema_of_reps Ts Rs = flat (map2 type_schema_of_rep Ts Rs) + +(* rep -> int list list *) +val all_combinations_for_rep = all_combinations o atom_schema_of_rep +(* rep list -> int list list *) +val all_combinations_for_reps = all_combinations o atom_schema_of_reps + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_scope.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,498 @@ +(* Title: HOL/Nitpick/Tools/nitpick_scope.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Scope enumerator for Nitpick. +*) + +signature NITPICK_SCOPE = +sig + type extended_context = NitpickHOL.extended_context + + type constr_spec = { + const: styp, + delta: int, + epsilon: int, + exclusive: bool, + explicit_max: int, + total: bool} + + type dtype_spec = { + typ: typ, + card: int, + co: bool, + precise: bool, + constrs: constr_spec list} + + type scope = { + ext_ctxt: extended_context, + card_assigns: (typ * int) list, + bisim_depth: int, + datatypes: dtype_spec list, + ofs: int Typtab.table} + + val datatype_spec : dtype_spec list -> typ -> dtype_spec option + val constr_spec : dtype_spec list -> styp -> constr_spec + val is_precise_type : dtype_spec list -> typ -> bool + val is_fully_comparable_type : dtype_spec list -> typ -> bool + val offset_of_type : int Typtab.table -> typ -> int + val spec_of_type : scope -> typ -> int * int + val pretties_for_scope : scope -> bool -> Pretty.T list + val multiline_string_for_scope : scope -> string + val scopes_equivalent : scope -> scope -> bool + val scope_less_eq : scope -> scope -> bool + val all_scopes : + extended_context -> int -> (typ option * int list) list + -> (styp option * int list) list -> (styp option * int list) list + -> int list -> typ list -> typ list -> scope list +end; + +structure NitpickScope : NITPICK_SCOPE = +struct + +open NitpickUtil +open NitpickHOL + +type constr_spec = { + const: styp, + delta: int, + epsilon: int, + exclusive: bool, + explicit_max: int, + total: bool} + +type dtype_spec = { + typ: typ, + card: int, + co: bool, + precise: bool, + constrs: constr_spec list} + +type scope = { + ext_ctxt: extended_context, + card_assigns: (typ * int) list, + bisim_depth: int, + datatypes: dtype_spec list, + ofs: int Typtab.table} + +datatype row_kind = Card of typ | Max of styp + +type row = row_kind * int list +type block = row list + +(* dtype_spec list -> typ -> dtype_spec option *) +fun datatype_spec (dtypes : dtype_spec list) T = + List.find (equal T o #typ) dtypes + +(* dtype_spec list -> styp -> constr_spec *) +fun constr_spec [] x = raise TERM ("NitpickScope.constr_spec", [Const x]) + | constr_spec ({constrs, ...} :: dtypes : dtype_spec list) (x as (s, T)) = + case List.find (equal (s, body_type T) o (apsnd body_type o #const)) + constrs of + SOME c => c + | NONE => constr_spec dtypes x + +(* dtype_spec list -> typ -> bool *) +fun is_precise_type dtypes (Type ("fun", Ts)) = + forall (is_precise_type dtypes) Ts + | is_precise_type dtypes (Type ("*", Ts)) = forall (is_precise_type dtypes) Ts + | is_precise_type dtypes T = + T <> nat_T andalso T <> int_T + andalso #precise (the (datatype_spec dtypes T)) + handle Option.Option => true +fun is_fully_comparable_type dtypes (Type ("fun", [T1, T2])) = + is_precise_type dtypes T1 andalso is_fully_comparable_type dtypes T2 + | is_fully_comparable_type dtypes (Type ("*", Ts)) = + forall (is_fully_comparable_type dtypes) Ts + | is_fully_comparable_type _ _ = true + +(* int Typtab.table -> typ -> int *) +fun offset_of_type ofs T = + case Typtab.lookup ofs T of + SOME j0 => j0 + | NONE => Typtab.lookup ofs dummyT |> the_default 0 + +(* scope -> typ -> int * int *) +fun spec_of_type ({card_assigns, ofs, ...} : scope) T = + (card_of_type card_assigns T + handle TYPE ("NitpickHOL.card_of_type", _, _) => ~1, offset_of_type ofs T) + +(* (string -> string) -> scope + -> string list * string list * string list * string list * string list *) +fun quintuple_for_scope quote ({ext_ctxt as {thy, ctxt, ...}, card_assigns, + bisim_depth, datatypes, ...} : scope) = + let + val (iter_asgns, card_asgns) = + card_assigns |> filter_out (equal @{typ bisim_iterator} o fst) + |> List.partition (is_fp_iterator_type o fst) + val (unimportant_card_asgns, important_card_asgns) = + card_asgns |> List.partition ((is_datatype thy orf is_integer_type) o fst) + val cards = + map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^ + string_of_int k) + fun maxes () = + maps (map_filter + (fn {const, explicit_max, ...} => + if explicit_max < 0 then + NONE + else + SOME (Syntax.string_of_term ctxt (Const const) ^ " = " ^ + string_of_int explicit_max)) + o #constrs) datatypes + fun iters () = + map (fn (T, k) => + quote (Syntax.string_of_term ctxt + (Const (const_for_iterator_type T))) ^ " = " ^ + string_of_int (k - 1)) iter_asgns + fun bisims () = + if bisim_depth < 0 andalso forall (not o #co) datatypes then [] + else ["bisim_depth = " ^ string_of_int bisim_depth] + in + setmp_show_all_types + (fn () => (cards important_card_asgns, cards unimportant_card_asgns, + maxes (), iters (), bisims ())) () + end + +(* scope -> bool -> Pretty.T list *) +fun pretties_for_scope scope verbose = + let + val (important_cards, unimportant_cards, maxes, iters, bisim_depths) = + quintuple_for_scope maybe_quote scope + val ss = map (prefix "card ") important_cards @ + (if verbose then + map (prefix "card ") unimportant_cards @ + map (prefix "max ") maxes @ + map (prefix "iter ") iters @ + bisim_depths + else + []) + in + if null ss then [] + else serial_commas "and" ss |> map Pretty.str |> Pretty.breaks + end + +(* scope -> string *) +fun multiline_string_for_scope scope = + let + val (important_cards, unimportant_cards, maxes, iters, bisim_depths) = + quintuple_for_scope I scope + val cards = important_cards @ unimportant_cards + in + case (if null cards then [] else ["card: " ^ commas cards]) @ + (if null maxes then [] else ["max: " ^ commas maxes]) @ + (if null iters then [] else ["iter: " ^ commas iters]) @ + bisim_depths of + [] => "empty" + | lines => space_implode "\n" lines + end + +(* scope -> scope -> bool *) +fun scopes_equivalent (s1 : scope) (s2 : scope) = + #datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2 +fun scope_less_eq (s1 : scope) (s2 : scope) = + (s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=) + +(* row -> int *) +fun rank_of_row (_, ks) = length ks +(* block -> int *) +fun rank_of_block block = fold Integer.max (map rank_of_row block) 1 +(* int -> typ * int list -> typ * int list *) +fun project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))]) +(* int -> block -> block *) +fun project_block (column, block) = map (project_row column) block + +(* (''a * ''a -> bool) -> (''a option * int list) list -> ''a -> int list *) +fun lookup_ints_assign eq asgns key = + case triple_lookup eq asgns key of + SOME ks => ks + | NONE => raise ARG ("NitpickScope.lookup_ints_assign", "") +(* theory -> (typ option * int list) list -> typ -> int list *) +fun lookup_type_ints_assign thy asgns T = + map (curry Int.max 1) (lookup_ints_assign (type_match thy) asgns T) + handle ARG ("NitpickScope.lookup_ints_assign", _) => + raise TYPE ("NitpickScope.lookup_type_ints_assign", [T], []) +(* theory -> (styp option * int list) list -> styp -> int list *) +fun lookup_const_ints_assign thy asgns x = + lookup_ints_assign (const_match thy) asgns x + handle ARG ("NitpickScope.lookup_ints_assign", _) => + raise TERM ("NitpickScope.lookup_const_ints_assign", [Const x]) + +(* theory -> (styp option * int list) list -> styp -> row option *) +fun row_for_constr thy maxes_asgns constr = + SOME (Max constr, lookup_const_ints_assign thy maxes_asgns constr) + handle TERM ("lookup_const_ints_assign", _) => NONE + +(* Proof.context -> (typ option * int list) list + -> (styp option * int list) list -> (styp option * int list) list -> int list + -> typ -> block *) +fun block_for_type ctxt cards_asgns maxes_asgns iters_asgns bisim_depths T = + let val thy = ProofContext.theory_of ctxt in + if T = @{typ bisim_iterator} then + [(Card T, map (fn k => Int.max (0, k) + 1) bisim_depths)] + else if is_fp_iterator_type T then + [(Card T, map (fn k => Int.max (0, k) + 1) + (lookup_const_ints_assign thy iters_asgns + (const_for_iterator_type T)))] + else + (Card T, lookup_type_ints_assign thy cards_asgns T) :: + (case datatype_constrs thy T of + [_] => [] + | constrs => map_filter (row_for_constr thy maxes_asgns) constrs) + end + +(* Proof.context -> (typ option * int list) list + -> (styp option * int list) list -> (styp option * int list) list -> int list + -> typ list -> typ list -> block list *) +fun blocks_for_types ctxt cards_asgns maxes_asgns iters_asgns bisim_depths + mono_Ts nonmono_Ts = + let + val thy = ProofContext.theory_of ctxt + (* typ -> block *) + val block_for = block_for_type ctxt cards_asgns maxes_asgns iters_asgns + bisim_depths + val mono_block = maps block_for mono_Ts + val nonmono_blocks = map block_for nonmono_Ts + in mono_block :: nonmono_blocks end + +val sync_threshold = 5 + +(* int list -> int list list *) +fun all_combinations_ordered_smartly ks = + let + (* int list -> int *) + fun cost_with_monos [] = 0 + | cost_with_monos (k :: ks) = + if k < sync_threshold andalso forall (equal k) ks then + k - sync_threshold + else + k * (k + 1) div 2 + Integer.sum ks + fun cost_without_monos [] = 0 + | cost_without_monos [k] = k + | cost_without_monos (_ :: k :: ks) = + if k < sync_threshold andalso forall (equal k) ks then + k - sync_threshold + else + Integer.sum (k :: ks) + in + ks |> all_combinations + |> map (`(if fst (hd ks) > 1 then cost_with_monos + else cost_without_monos)) + |> sort (int_ord o pairself fst) |> map snd + end + +(* typ -> bool *) +fun is_self_recursive_constr_type T = + exists (exists_subtype (equal (body_type T))) (binder_types T) + +(* (styp * int) list -> styp -> int *) +fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x) + +type scope_desc = (typ * int) list * (styp * int) list + +(* theory -> scope_desc -> typ * int -> bool *) +fun is_surely_inconsistent_card_assign thy (card_asgns, max_asgns) (T, k) = + case datatype_constrs thy T of + [] => false + | xs => + let + val precise_cards = + map (Integer.prod + o map (bounded_precise_card_of_type thy k 0 card_asgns) + o binder_types o snd) xs + val maxes = map (constr_max max_asgns) xs + (* int -> int -> int *) + fun effective_max 0 ~1 = k + | effective_max 0 max = max + | effective_max card ~1 = card + | effective_max card max = Int.min (card, max) + val max = map2 effective_max precise_cards maxes |> Integer.sum + (* unit -> int *) + fun doms_card () = + xs |> map (Integer.prod o map (bounded_card_of_type k ~1 card_asgns) + o binder_types o snd) + |> Integer.sum + in + max < k + orelse (forall (not_equal 0) precise_cards andalso doms_card () < k) + end + handle TYPE ("NitpickHOL.card_of_type", _, _) => false + +(* theory -> scope_desc -> bool *) +fun is_surely_inconsistent_scope_description thy (desc as (card_asgns, _)) = + exists (is_surely_inconsistent_card_assign thy desc) card_asgns + +(* theory -> scope_desc -> (typ * int) list option *) +fun repair_card_assigns thy (card_asgns, max_asgns) = + let + (* (typ * int) list -> (typ * int) list -> (typ * int) list option *) + fun aux seen [] = SOME seen + | aux seen ((T, 0) :: _) = NONE + | aux seen ((T, k) :: asgns) = + (if is_surely_inconsistent_scope_description thy + ((T, k) :: seen, max_asgns) then + raise SAME () + else + case aux ((T, k) :: seen) asgns of + SOME asgns => SOME asgns + | NONE => raise SAME ()) + handle SAME () => aux seen ((T, k - 1) :: asgns) + in aux [] (rev card_asgns) end + +(* theory -> (typ * int) list -> typ * int -> typ * int *) +fun repair_iterator_assign thy asgns (T as Type (s, Ts), k) = + (T, if T = @{typ bisim_iterator} then + let val co_cards = map snd (filter (is_codatatype thy o fst) asgns) in + Int.min (k, Integer.sum co_cards) + end + else if is_fp_iterator_type T then + case Ts of + [] => 1 + | _ => bounded_card_of_type k ~1 asgns (foldr1 HOLogic.mk_prodT Ts) + else + k) + | repair_iterator_assign _ _ asgn = asgn + +(* row -> scope_desc -> scope_desc *) +fun add_row_to_scope_descriptor (kind, ks) (card_asgns, max_asgns) = + case kind of + Card T => ((T, the_single ks) :: card_asgns, max_asgns) + | Max x => (card_asgns, (x, the_single ks) :: max_asgns) +(* block -> scope_desc *) +fun scope_descriptor_from_block block = + fold_rev add_row_to_scope_descriptor block ([], []) +(* theory -> block list -> int list -> scope_desc option *) +fun scope_descriptor_from_combination thy blocks columns = + let + val (card_asgns, max_asgns) = + maps project_block (columns ~~ blocks) |> scope_descriptor_from_block + val card_asgns = repair_card_assigns thy (card_asgns, max_asgns) |> the + in + SOME (map (repair_iterator_assign thy card_asgns) card_asgns, max_asgns) + end + handle Option.Option => NONE + +(* theory -> (typ * int) list -> dtype_spec list -> int Typtab.table *) +fun offset_table_for_card_assigns thy asgns dtypes = + let + (* int -> (int * int) list -> (typ * int) list -> int Typtab.table + -> int Typtab.table *) + fun aux next _ [] = Typtab.update_new (dummyT, next) + | aux next reusable ((T, k) :: asgns) = + if k = 1 orelse is_integer_type T then + aux next reusable asgns + else if length (these (Option.map #constrs (datatype_spec dtypes T))) + > 1 then + Typtab.update_new (T, next) #> aux (next + k) reusable asgns + else + case AList.lookup (op =) reusable k of + SOME j0 => Typtab.update_new (T, j0) #> aux next reusable asgns + | NONE => Typtab.update_new (T, next) + #> aux (next + k) ((k, next) :: reusable) asgns + in aux 0 [] asgns Typtab.empty end + +(* int -> (typ * int) list -> typ -> int *) +fun domain_card max card_asgns = + Integer.prod o map (bounded_card_of_type max max card_asgns) o binder_types + +(* scope_desc -> bool -> int -> (int -> int) -> int -> int -> bool * styp + -> constr_spec list -> constr_spec list *) +fun add_constr_spec (card_asgns, max_asgns) co card sum_dom_cards num_self_recs + num_non_self_recs (self_rec, x as (s, T)) constrs = + let + val max = constr_max max_asgns x + (* int -> int *) + fun bound k = Int.min (card, (max >= 0 ? curry Int.min max) k) + (* unit -> int *) + fun next_delta () = if null constrs then 0 else #epsilon (hd constrs) + val {delta, epsilon, exclusive, total} = + if max = 0 then + let val delta = next_delta () in + {delta = delta, epsilon = delta, exclusive = true, total = false} + end + else if not co andalso num_self_recs > 0 then + if not self_rec andalso num_non_self_recs = 1 + andalso domain_card 2 card_asgns T = 1 then + {delta = 0, epsilon = 1, exclusive = (s = @{const_name Nil}), + total = true} + else if s = @{const_name Cons} then + {delta = 1, epsilon = card, exclusive = true, total = false} + else + {delta = 0, epsilon = card, exclusive = false, total = false} + else if card = sum_dom_cards (card + 1) then + let val delta = next_delta () in + {delta = delta, epsilon = delta + domain_card card card_asgns T, + exclusive = true, total = true} + end + else + {delta = 0, epsilon = card, + exclusive = (num_self_recs + num_non_self_recs = 1), total = false} + in + {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive, + explicit_max = max, total = total} :: constrs + end + +(* extended_context -> scope_desc -> typ * int -> dtype_spec *) +fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...}) + (desc as (card_asgns, _)) (T, card) = + let + val co = is_codatatype thy T + val xs = boxed_datatype_constrs ext_ctxt T + val self_recs = map (is_self_recursive_constr_type o snd) xs + val (num_self_recs, num_non_self_recs) = + List.partition (equal true) self_recs |> pairself length + val precise = (card = bounded_precise_card_of_type thy (card + 1) 0 + card_asgns T) + (* int -> int *) + fun sum_dom_cards max = + map (domain_card max card_asgns o snd) xs |> Integer.sum + val constrs = + fold_rev (add_constr_spec desc co card sum_dom_cards num_self_recs + num_non_self_recs) (self_recs ~~ xs) [] + in {typ = T, card = card, co = co, precise = precise, constrs = constrs} end + +(* extended_context -> int -> scope_desc -> scope *) +fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break + (desc as (card_asgns, _)) = + let + val datatypes = map (datatype_spec_from_scope_descriptor ext_ctxt desc) + (filter (is_datatype thy o fst) card_asgns) + val bisim_depth = card_of_type card_asgns @{typ bisim_iterator} - 1 + in + {ext_ctxt = ext_ctxt, card_assigns = card_asgns, datatypes = datatypes, + bisim_depth = bisim_depth, + ofs = if sym_break <= 0 then Typtab.empty + else offset_table_for_card_assigns thy card_asgns datatypes} + end + +(* theory -> typ list -> (typ option * int list) list + -> (typ option * int list) list *) +fun fix_cards_assigns_wrt_boxing _ _ [] = [] + | fix_cards_assigns_wrt_boxing thy Ts ((SOME T, ks) :: cards_asgns) = + (if is_fun_type T orelse is_pair_type T then + Ts |> filter (curry (type_match thy o swap) T o unbox_type) + |> map (rpair ks o SOME) + else + [(SOME T, ks)]) @ fix_cards_assigns_wrt_boxing thy Ts cards_asgns + | fix_cards_assigns_wrt_boxing thy Ts ((NONE, ks) :: cards_asgns) = + (NONE, ks) :: fix_cards_assigns_wrt_boxing thy Ts cards_asgns + +val distinct_threshold = 512 + +(* extended_context -> int -> (typ option * int list) list + -> (styp option * int list) list -> (styp option * int list) list -> int list + -> typ list -> typ list -> scope list *) +fun all_scopes (ext_ctxt as {thy, ctxt, ...}) sym_break cards_asgns maxes_asgns + iters_asgns bisim_depths mono_Ts nonmono_Ts = + let + val cards_asgns = fix_cards_assigns_wrt_boxing thy mono_Ts cards_asgns + val blocks = blocks_for_types ctxt cards_asgns maxes_asgns iters_asgns + bisim_depths mono_Ts nonmono_Ts + val ranks = map rank_of_block blocks + val descs = all_combinations_ordered_smartly (map (rpair 0) ranks) + |> map_filter (scope_descriptor_from_combination thy blocks) + in + descs |> length descs <= distinct_threshold ? distinct (op =) + |> map (scope_from_descriptor ext_ctxt sym_break) + end + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_tests.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,334 @@ +(* Title: HOL/Nitpick/Tools/nitpick_tests.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +Unit tests for Nitpick. +*) + +signature NITPICK_TESTS = +sig + val run_all_tests : unit -> unit +end + +structure NitpickTests = +struct + +open NitpickUtil +open NitpickPeephole +open NitpickRep +open NitpickNut +open NitpickKodkod +open Nitpick + +val settings = + [("solver", "\"zChaff\""), + ("skolem_depth", "-1")] + +fun cast_to_rep R u = Op1 (Cast, type_of u, R, u) + +val unit_T = @{typ unit} +val dummy_T = @{typ 'a} + +val unity = Cst (Unity, unit_T, Unit) +val atom1_v1 = FreeName ("atom1_v1", dummy_T, Atom (1, 0)) +val atom2_v1 = FreeName ("atom2_v1", dummy_T, Atom (2, 0)) +val atom6_v1 = FreeName ("atom6_v1", dummy_T, Atom (6, 0)) +val atom16_v1 = FreeName ("atom16_v1", dummy_T, Atom (16, 0)) +val atom24_v1 = FreeName ("atom24_v1", dummy_T, Atom (24, 0)) +val atom36_v1 = FreeName ("atom36_v1", dummy_T, Atom (36, 0)) +val atom81_v1 = FreeName ("atom81_v1", dummy_T, Atom (81, 0)) +val struct_atom1_atom1_v1 = + FreeName ("struct_atom1_atom1_v1", dummy_T, Struct [Atom (1, 0), Atom (1, 0)]) +val struct_atom1_unit_v1 = + FreeName ("struct_atom1_atom1_v1", dummy_T, Struct [Atom (1, 0), Unit]) +val struct_unit_atom1_v1 = + FreeName ("struct_atom1_atom1_v1", dummy_T, Struct [Unit, Atom (1, 0)]) + +(* + Formula Unit Atom Struct Vect Func + Formula X N/A X X N/A N/A + Unit N/A N/A N/A N/A N/A N/A + Atom X N/A X X X X + Struct N/A N/A X X N/A N/A + Vect N/A N/A X N/A X X + Func N/A N/A X N/A X X +*) + +val tests = + [("rep_conversion_formula_formula", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Formula Neut) + (cast_to_rep (Formula Neut) atom2_v1), atom2_v1)), + ("rep_conversion_atom_atom", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) (cast_to_rep (Atom (16, 0)) atom16_v1), + atom16_v1)), + ("rep_conversion_struct_struct_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (4, 0), Atom (6, 0)]) + (cast_to_rep (Struct [Atom (4, 0), Atom (6, 0)]) atom24_v1), + atom24_v1)), + ("rep_conversion_struct_struct_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (4, 0), Struct [Atom (2, 0), Atom (3, 0)]]) + (cast_to_rep (Struct [Atom (4, 0), Atom (6, 0)]) atom24_v1), + atom24_v1)), + ("rep_conversion_struct_struct_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (4, 0), Atom (6, 0)]) + (cast_to_rep (Struct [Atom (4, 0), + Struct [Atom (2, 0), Atom (3, 0)]]) + atom24_v1), + atom24_v1)), + ("rep_conversion_struct_struct_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (24, 0), Unit]) + (cast_to_rep (Struct [Atom (24, 0), Atom (1, 0)]) atom24_v1), + atom24_v1)), + ("rep_conversion_struct_struct_5", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (24, 0), Atom (1, 0)]) + (cast_to_rep (Struct [Atom (24, 0), Unit]) atom24_v1), + atom24_v1)), + ("rep_conversion_struct_struct_6", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (1, 0), Atom (1, 0)]) + (cast_to_rep (Struct [Atom (1, 0), Unit]) + (cast_to_rep (Struct [Unit, Atom (1, 0)]) atom1_v1)), + atom1_v1)), + ("rep_conversion_vect_vect_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (2, Atom (4, 0))) + (cast_to_rep (Vect (2, Struct [Atom (2, 0), Atom (2, 0)])) + atom16_v1)), + atom16_v1)), + ("rep_conversion_vect_vect_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (2, Struct [Atom (2, 0), Atom (2, 0)])) + (cast_to_rep (Vect (2, Atom (4, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_vect_vect_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (2, Atom (4, 0))) + (cast_to_rep (Vect (2, Vect (2, Atom (2, 0)))) atom16_v1)), + atom16_v1)), + ("rep_conversion_vect_vect_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (2, Vect (2, Atom (2, 0)))) + (cast_to_rep (Vect (2, Atom (4, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_func_func_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Atom (2, 0), + Struct [Atom (2, 0), Atom (3, 0)])) + (cast_to_rep (Func (Atom (2, 0), Atom (6, 0))) atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Atom (2, 0), Atom (6, 0))) + (cast_to_rep (Func (Atom (2, 0), + Struct [Atom (2, 0), Atom (3, 0)])) + atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Unit, Struct [Atom (6, 0), Atom (6, 0)])) + (cast_to_rep (Func (Atom (1, 0), Atom (36, 0))) atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Atom (1, 0), Atom (36, 0))) + (cast_to_rep (Func (Unit, Struct [Atom (6, 0), Atom (6, 0)])) + atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_5", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Unit, Vect (2, Atom (6, 0)))) + (cast_to_rep (Func (Atom (1, 0), Atom (36, 0))) atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_6", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Atom (1, 0), Atom (36, 0))) + (cast_to_rep (Func (Unit, Vect (2, Atom (6, 0)))) + atom36_v1)), + atom36_v1)), + ("rep_conversion_func_func_7", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (2, 0)) + (cast_to_rep (Func (Unit, Atom (2, 0))) + (cast_to_rep (Func (Atom (1, 0), Formula Neut)) atom2_v1)), + atom2_v1)), + ("rep_conversion_func_func_8", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (2, 0)) + (cast_to_rep (Func (Atom (1, 0), Formula Neut)) + (cast_to_rep (Func (Unit, Atom (2, 0))) atom2_v1)), + atom2_v1)), + ("rep_conversion_atom_formula_atom", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (2, 0)) (cast_to_rep (Formula Neut) atom2_v1), + atom2_v1)), + ("rep_conversion_unit_atom", + Op2 (Eq, bool_T, Formula Neut, cast_to_rep (Atom (1, 0)) unity, unity)), + ("rep_conversion_atom_struct_atom1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (6, 0)) + (cast_to_rep (Struct [Atom (3, 0), Atom (2, 0)]) atom6_v1), + atom6_v1)), + ("rep_conversion_atom_struct_atom_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (24, 0)) + (cast_to_rep (Struct [Struct [Atom (3, 0), Atom (4, 0)], + Atom (2, 0)]) atom24_v1), + atom24_v1)), + ("rep_conversion_atom_struct_atom_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (6, 0)) + (cast_to_rep (Struct [Atom (6, 0), Unit]) atom6_v1), + atom6_v1)), + ("rep_conversion_atom_struct_atom_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (6, 0)) + (cast_to_rep (Struct [Struct [Atom (3, 0), Unit], Atom (2, 0)]) + atom6_v1), + atom6_v1)), + ("rep_conversion_atom_vect_func_atom_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (4, Atom (2, 0))) + (cast_to_rep (Func (Atom (4, 0), Atom (2, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_vect_func_atom_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (4, Atom (2, 0))) + (cast_to_rep (Func (Atom (4, 0), Atom (2, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_vect_func_atom_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (4, Atom (2, 0))) + (cast_to_rep (Func (Atom (4, 0), Formula Neut)) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_vect_func_atom_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (1, Atom (16, 0))) + (cast_to_rep (Func (Unit, Atom (16, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_vect_func_atom_5", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Vect (1, Atom (16, 0))) + (cast_to_rep (Func (Unit, Atom (16, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_func_vect_atom_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Func (Atom (4, 0), Atom (2, 0))) + (cast_to_rep (Vect (4, Atom (2, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_func_vect_atom_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Func (Atom (4, 0), Atom (2, 0))) + (cast_to_rep (Vect (4, Atom (2, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_func_vect_atom_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Func (Atom (4, 0), Formula Neut)) + (cast_to_rep (Vect (4, Atom (2, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_func_vect_atom_4", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Func (Unit, Atom (16, 0))) + (cast_to_rep (Vect (1, Atom (16, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_func_vect_atom_5", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (16, 0)) + (cast_to_rep (Func (Atom (1, 0), Atom (16, 0))) + (cast_to_rep (Vect (1, Atom (16, 0))) atom16_v1)), + atom16_v1)), + ("rep_conversion_atom_vect_atom", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Vect (2, Struct [Atom (2, 0), Atom (3, 0)])) + atom36_v1), + atom36_v1)), + ("rep_conversion_atom_func_atom", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Atom (36, 0)) + (cast_to_rep (Func (Atom (2, 0), + Struct [Atom (2, 0), Atom (3, 0)])) atom36_v1), + atom36_v1)), + ("rep_conversion_struct_atom1_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (1, 0), Atom (1, 0)]) atom1_v1, + struct_atom1_atom1_v1)), + ("rep_conversion_struct_atom1_2", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (1, 0), Unit]) atom1_v1, + struct_atom1_unit_v1)), + ("rep_conversion_struct_atom1_3", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Unit, Atom (1, 0)]) atom1_v1, + struct_unit_atom1_v1)) +(* + ("rep_conversion_struct_formula_struct_1", + Op2 (Eq, bool_T, Formula Neut, + cast_to_rep (Struct [Atom (2, 0), Unit]) + (cast_to_rep (Formula Neut) struct_atom_2_unit_v1), + struct_atom_2_unit_v1)) +*) + ] + +fun problem_for_nut ctxt name u = + let + val debug = false + val peephole_optim = true + val nat_card = 4 + val int_card = 9 + val j0 = 0 + val constrs = kodkod_constrs peephole_optim nat_card int_card j0 + val (free_rels, pool, table) = + rename_free_vars (fst (add_free_and_const_names u ([], []))) initial_pool + NameTable.empty + val u = Op1 (Not, type_of u, rep_of u, u) + |> rename_vars_in_nut pool table + val formula = kodkod_formula_from_nut Typtab.empty false constrs u + val bounds = map (bound_for_plain_rel ctxt debug) free_rels + val univ_card = univ_card nat_card int_card j0 bounds formula + val declarative_axioms = map (declarative_axiom_for_plain_rel constrs) + free_rels + val formula = fold_rev s_and declarative_axioms formula + in + {comment = name, settings = settings, univ_card = univ_card, + tuple_assigns = [], bounds = bounds, int_bounds = [], expr_assigns = [], + formula = formula} + end + +(* string -> unit *) +fun run_test name = + case Kodkod.solve_any_problem true NONE 0 1 + [problem_for_nut @{context} name + (the (AList.lookup (op =) tests name))] of + Kodkod.Normal ([], _) => () + | _ => warning ("Test " ^ quote name ^ " failed") + +(* unit -> unit *) +fun run_all_tests () = List.app run_test (map fst tests) + +end; diff -r 4705b7323a7d -r 79bd3fbf5d61 src/HOL/Tools/Nitpick/nitpick_util.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Mon Oct 26 14:57:49 2009 +0100 @@ -0,0 +1,307 @@ +(* Title: HOL/Nitpick/Tools/nitpick_util.ML + Author: Jasmin Blanchette, TU Muenchen + Copyright 2008, 2009 + +General-purpose functions used by the Nitpick modules. +*) + +infix 6 nat_minus +infix 7 pairf + +signature BASIC_NITPICK_UTIL = +sig + type styp = string * typ +end; + +signature NITPICK_UTIL = +sig + include BASIC_NITPICK_UTIL + + datatype polarity = Pos | Neg | Neut + + exception ARG of string * string + exception BAD of string * string + exception LIMIT of string * string + exception NOT_SUPPORTED of string + exception SAME of unit + + val nitpick_prefix : string + val curry3 : ('a * 'b * 'c -> 'd) -> 'a -> 'b -> 'c -> 'd + val pairf : ('a -> 'b) * ('a -> 'c) -> 'a -> 'b * 'c + val int_for_bool : bool -> int + val nat_minus : int * int -> int + val reasonable_power : int -> int -> int + val exact_log : int -> int -> int + val exact_root : int -> int -> int + val offset_list : int list -> int list + val index_seq : int -> int -> int list + val filter_indices : int list -> 'a list -> 'a list + val filter_out_indices : int list -> 'a list -> 'a list + val fold1 : ('a -> 'a -> 'a) -> 'a list -> 'a + val replicate_list : int -> 'a list -> 'a list + val n_fold_cartesian_product : 'a list list -> 'a list list + val all_distinct_unordered_pairs_of : ''a list -> (''a * ''a) list + val nth_combination : (int * int) list -> int -> int list + val all_combinations : (int * int) list -> int list list + val all_permutations : 'a list -> 'a list list + val batch_list : int -> 'a list -> 'a list list + val chunk_list_unevenly : int list -> 'a list -> 'a list list + val map3 : ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list + val double_lookup : + ('a * 'a -> bool) -> ('a option * 'b) list -> 'a -> 'b option + val triple_lookup : + (''a * ''a -> bool) -> (''a option * 'b) list -> ''a -> 'b option + val is_substring_of : string -> string -> bool + val serial_commas : string -> string list -> string list + val plural_s : int -> string + val plural_s_for_list : 'a list -> string + val flip_polarity : polarity -> polarity + val prop_T : typ + val bool_T : typ + val nat_T : typ + val int_T : typ + val subscript : string -> string + val nat_subscript : int -> string + val time_limit : Time.time option -> ('a -> 'b) -> 'a -> 'b + val silence : ('a -> 'b) -> 'a -> 'b + val DETERM_TIMEOUT : Time.time option -> tactic -> tactic + val setmp_show_all_types : ('a -> 'b) -> 'a -> 'b + val indent_size : int + val pstrs : string -> Pretty.T list + val plain_string_from_yxml : string -> string + val maybe_quote : string -> string +end + +structure NitpickUtil : NITPICK_UTIL = +struct + +type styp = string * typ + +datatype polarity = Pos | Neg | Neut + +exception ARG of string * string +exception BAD of string * string +exception LIMIT of string * string +exception NOT_SUPPORTED of string +exception SAME of unit + +val nitpick_prefix = "Nitpick." + +(* ('a * 'b * 'c -> 'd) -> 'a -> 'b -> 'c -> 'd *) +fun curry3 f = fn x => fn y => fn z => f (x, y, z) + +(* ('a -> 'b) * ('a -> 'c) -> 'a -> 'b * 'c *) +fun (f pairf g) = fn x => (f x, g x) + +(* bool -> int *) +fun int_for_bool b = if b then 1 else 0 +(* int * int -> int *) +fun (i nat_minus j) = if i > j then i - j else 0 + +val max_exponent = 16384 + +(* int -> int -> int *) +fun reasonable_power a 0 = 1 + | reasonable_power a 1 = a + | reasonable_power 0 _ = 0 + | reasonable_power 1 _ = 1 + | reasonable_power a b = + if b < 0 orelse b > max_exponent then + raise LIMIT ("NitpickUtil.reasonable_power", + "too large exponent (" ^ signed_string_of_int b ^ ")") + else + let + val c = reasonable_power a (b div 2) in + c * c * reasonable_power a (b mod 2) + end + +(* int -> int -> int *) +fun exact_log m n = + let + val r = Math.ln (Real.fromInt n) / Math.ln (Real.fromInt m) |> Real.round + in + if reasonable_power m r = n then + r + else + raise ARG ("NitpickUtil.exact_log", + commas (map signed_string_of_int [m, n])) + end + +(* int -> int -> int *) +fun exact_root m n = + let val r = Math.pow (Real.fromInt n, 1.0 / (Real.fromInt m)) |> Real.round in + if reasonable_power r m = n then + r + else + raise ARG ("NitpickUtil.exact_root", + commas (map signed_string_of_int [m, n])) + end + +(* ('a -> 'a -> 'a) -> 'a list -> 'a *) +fun fold1 f = foldl1 (uncurry f) + +(* int -> 'a list -> 'a list *) +fun replicate_list 0 _ = [] + | replicate_list n xs = xs @ replicate_list (n - 1) xs + +(* int list -> int list *) +fun offset_list ns = rev (tl (fold (fn x => fn xs => (x + hd xs) :: xs) ns [0])) +(* int -> int -> int list *) +fun index_seq j0 n = if j0 < 0 then j0 downto j0 - n + 1 else j0 upto j0 + n - 1 + +(* int list -> 'a list -> 'a list *) +fun filter_indices js xs = + let + (* int -> int list -> 'a list -> 'a list *) + fun aux _ [] _ = [] + | aux i (j :: js) (x :: xs) = + if i = j then x :: aux (i + 1) js xs else aux (i + 1) (j :: js) xs + | aux _ _ _ = raise ARG ("NitpickUtil.filter_indices", + "indices unordered or out of range") + in aux 0 js xs end +fun filter_out_indices js xs = + let + (* int -> int list -> 'a list -> 'a list *) + fun aux _ [] xs = xs + | aux i (j :: js) (x :: xs) = + if i = j then aux (i + 1) js xs else x :: aux (i + 1) (j :: js) xs + | aux _ _ _ = raise ARG ("NitpickUtil.filter_out_indices", + "indices unordered or out of range") + in aux 0 js xs end + +(* 'a list -> 'a list list -> 'a list list *) +fun cartesian_product [] _ = [] + | cartesian_product (x :: xs) yss = + map (cons x) yss @ cartesian_product xs yss +(* 'a list list -> 'a list list *) +fun n_fold_cartesian_product xss = fold_rev cartesian_product xss [[]] +(* ''a list -> (''a * ''a) list *) +fun all_distinct_unordered_pairs_of [] = [] + | all_distinct_unordered_pairs_of (x :: xs) = + map (pair x) xs @ all_distinct_unordered_pairs_of xs + +(* (int * int) list -> int -> int list *) +val nth_combination = + let + (* (int * int) list -> int -> int list * int *) + fun aux [] n = ([], n) + | aux ((k, j0) :: xs) n = + let val (js, n) = aux xs n in ((n mod k) + j0 :: js, n div k) end + in fst oo aux end + +(* (int * int) list -> int list list *) +val all_combinations = n_fold_cartesian_product o map (uncurry index_seq o swap) + +(* 'a list -> 'a list list *) +fun all_permutations [] = [[]] + | all_permutations xs = + maps (fn j => map (cons (nth xs j)) (all_permutations (nth_drop j xs))) + (index_seq 0 (length xs)) + +(* int -> 'a list -> 'a list list *) +fun batch_list _ [] = [] + | batch_list k xs = + if length xs <= k then [xs] + else List.take (xs, k) :: batch_list k (List.drop (xs, k)) + +(* int list -> 'a list -> 'a list list *) +fun chunk_list_unevenly _ [] = [] + | chunk_list_unevenly [] ys = map single ys + | chunk_list_unevenly (k :: ks) ys = + let val (ys1, ys2) = chop k ys in ys1 :: chunk_list_unevenly ks ys2 end + +(* ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list *) +fun map3 _ [] [] [] = [] + | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs + | map3 _ _ _ _ = raise UnequalLengths + +(* ('a * 'a -> bool) -> ('a option * 'b) list -> 'a -> 'b option *) +fun double_lookup eq ps key = + case AList.lookup (fn (SOME x, SOME y) => eq (x, y) | _ => false) ps + (SOME key) of + SOME z => SOME z + | NONE => ps |> find_first (is_none o fst) |> Option.map snd +(* (''a * ''a -> bool) -> (''a option * 'b) list -> ''a -> 'b option *) +fun triple_lookup eq ps key = + case AList.lookup (op =) ps (SOME key) of + SOME z => SOME z + | NONE => double_lookup eq ps key + +(* string -> string -> bool *) +fun is_substring_of needle stack = + not (Substring.isEmpty (snd (Substring.position needle + (Substring.full stack)))) + +(* string -> string list -> string list *) +fun serial_commas _ [] = ["??"] + | serial_commas _ [s] = [s] + | serial_commas conj [s1, s2] = [s1, conj, s2] + | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3] + | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss + +(* int -> string *) +fun plural_s n = if n = 1 then "" else "s" +(* 'a list -> string *) +fun plural_s_for_list xs = plural_s (length xs) + +(* polarity -> polarity *) +fun flip_polarity Pos = Neg + | flip_polarity Neg = Pos + | flip_polarity Neut = Neut + +val prop_T = @{typ prop} +val bool_T = @{typ bool} +val nat_T = @{typ nat} +val int_T = @{typ int} + +(* string -> string *) +val subscript = implode o map (prefix "\<^isub>") o explode +(* int -> string *) +val nat_subscript = subscript o signed_string_of_int + +(* Time.time option -> ('a -> 'b) -> 'a -> 'b *) +fun time_limit NONE f = f + | time_limit (SOME delay) f = TimeLimit.timeLimit delay f + +(* (string -> unit) Unsynchronized.ref -> ('a -> 'b) -> 'a -> 'b *) +fun silence_one out_fn = setmp_CRITICAL out_fn (K ()) + +(* ('a -> 'b) -> 'a -> 'b *) +fun silence f = + fold silence_one + [Output.writeln_fn, Output.priority_fn, Output.tracing_fn, + Output.warning_fn, Output.error_fn, Output.debug_fn] f + +(* Time.time option -> tactic -> tactic *) +fun DETERM_TIMEOUT delay tac st = + Seq.of_list (the_list (time_limit delay (fn () => SINGLE tac st) ())) + +(* ('a -> 'b) -> 'a -> 'b *) +fun setmp_show_all_types f = + setmp_CRITICAL show_all_types + (! show_types orelse ! show_sorts orelse ! show_all_types) f + +val indent_size = 2 + +(* string -> Pretty.T list *) +val pstrs = Pretty.breaks o map Pretty.str o space_explode " " + +(* XML.tree -> string *) +fun plain_string_from_xml_tree t = + Buffer.empty |> XML.add_content t |> Buffer.content +(* string -> string *) +val plain_string_from_yxml = plain_string_from_xml_tree o YXML.parse + +(* string -> bool *) +val is_long_identifier = forall Syntax.is_identifier o space_explode "." +(* string -> string *) +fun maybe_quote y = + let val s = plain_string_from_yxml y in + y |> (not (is_long_identifier (perhaps (try (unprefix "'")) s)) + orelse OuterKeyword.is_keyword s) ? quote + end + +end; + +structure BasicNitpickUtil : BASIC_NITPICK_UTIL = NitpickUtil; +open BasicNitpickUtil;